From d59b1904064942e8d375bc7091a222738e2f6539 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Thu, 21 Nov 2024 14:50:51 +0000 Subject: [PATCH 01/53] 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 02/53] 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 03/53] 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 04/53] 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 05/53] =?UTF-8?q?feat(Data/Nat/Factorization/PrimePow):=20?= =?UTF-8?q?add=20equivalence=20Primes=20=C3=97=20=E2=84=95=20=E2=89=83=20P?= =?UTF-8?q?rimePowers=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 06/53] 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 07/53] 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 08/53] 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 09/53] 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 10/53] 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 11/53] 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 12/53] 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 13/53] 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 14/53] 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 15/53] 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 16/53] 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 17/53] 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 18/53] 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 19/53] 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 20/53] 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 21/53] 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 22/53] 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 23/53] 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 24/53] 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 25/53] 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 26/53] 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 27/53] 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 28/53] 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 29/53] 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 30/53] 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 31/53] 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 32/53] 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 33/53] 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 34/53] 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 35/53] 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 36/53] 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 37/53] 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 38/53] 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 39/53] 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 40/53] 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 41/53] 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 42/53] 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 43/53] 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 44/53] 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 45/53] 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 46/53] 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 47/53] 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 48/53] 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 49/53] 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 50/53] 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 51/53] 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 52/53] 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 53/53] 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