diff --git a/Archive/Imo/Imo1962Q1.lean b/Archive/Imo/Imo1962Q1.lean index 97245b3333242..6a05b328784b8 100644 --- a/Archive/Imo/Imo1962Q1.lean +++ b/Archive/Imo/Imo1962Q1.lean @@ -117,11 +117,10 @@ theorem case_more_digits {c n : ℕ} (h1 : (digits 10 c).length ≥ 6) (h2 : Pro intro h4 have h5 : (digits 10 c).length = 0 := by simp [h4] exact case_0_digit h5 h2 - have h6 : 2 ≤ 10 := by decide calc n ≥ 10 * c := le.intro h2.left.symm - _ ≥ 10 ^ (digits 10 c).length := (base_pow_length_digits_le 10 c h6 h3) - _ ≥ 10 ^ 6 := ((pow_le_iff_le_right h6).mpr h1) + _ ≥ 10 ^ (digits 10 c).length := (base_pow_length_digits_le 10 c (by decide) h3) + _ ≥ 10 ^ 6 := pow_le_pow_right (by decide) h1 _ ≥ 153846 := by norm_num #align imo1962_q1.case_more_digits Imo1962Q1.case_more_digits diff --git a/Archive/Imo/Imo1969Q1.lean b/Archive/Imo/Imo1969Q1.lean index ce92074f30253..e9f7210443165 100644 --- a/Archive/Imo/Imo1969Q1.lean +++ b/Archive/Imo/Imo1969Q1.lean @@ -86,7 +86,7 @@ theorem aChoice_good (b : ℕ) : aChoice b ∈ goodNats := /-- `aChoice` is a strictly monotone function; this is easily proven by chaining together lemmas in the `strictMono` namespace. -/ theorem aChoice_strictMono : StrictMono aChoice := - ((strictMono_id.const_add 2).nat_pow (by decide : 0 < 4)).const_mul (by decide : 0 < 4) + ((strictMono_id.const_add 2).nat_pow (by decide)).const_mul (by decide) #align imo1969_q1.a_choice_strict_mono Imo1969Q1.aChoice_strictMono end Imo1969Q1 diff --git a/Archive/Imo/Imo2001Q2.lean b/Archive/Imo/Imo2001Q2.lean index 3e8528f7ef2ef..2276168242724 100644 --- a/Archive/Imo/Imo2001Q2.lean +++ b/Archive/Imo/Imo2001Q2.lean @@ -41,7 +41,7 @@ theorem bound (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : = a ^ 3 * (a * sqrt ((a ^ 3) ^ 2 + (8:ℝ) * b ^ 3 * c ^ 3)) := by ring _ ≤ a ^ 3 * (a ^ 4 + b ^ 4 + c ^ 4) := ?_ gcongr - apply le_of_pow_le_pow _ (by positivity) zero_lt_two + apply le_of_pow_le_pow_left two_ne_zero (by positivity) rw [mul_pow, sq_sqrt (by positivity), ← sub_nonneg] calc (a ^ 4 + b ^ 4 + c ^ 4) ^ 2 - a ^ 2 * ((a ^ 3) ^ 2 + 8 * b ^ 3 * c ^ 3) @@ -67,7 +67,7 @@ open Imo2001Q2 theorem imo2001_q2 (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : ↑1 ≤ a / sqrt (a ^ 2 + ↑8 * b * c) + b / sqrt (b ^ 2 + ↑8 * c * a) + c / sqrt (c ^ 2 + ↑8 * a * b) := have h3 : ∀ {x : ℝ}, 0 < x → (x ^ (3 : ℝ)⁻¹) ^ 3 = x := fun hx => - show ↑3 = (3 : ℝ) by norm_num ▸ rpow_nat_inv_pow_nat hx.le three_ne_zero + show ↑3 = (3 : ℝ) by norm_num ▸ rpow_inv_natCast_pow hx.le three_ne_zero calc 1 ≤ _ := imo2001_q2' (rpow_pos_of_pos ha _) (rpow_pos_of_pos hb _) (rpow_pos_of_pos hc _) _ = _ := by rw [h3 ha, h3 hb, h3 hc] diff --git a/Archive/Imo/Imo2006Q3.lean b/Archive/Imo/Imo2006Q3.lean index 914ab33ee7c8c..9a31c97393455 100644 --- a/Archive/Imo/Imo2006Q3.lean +++ b/Archive/Imo/Imo2006Q3.lean @@ -71,7 +71,7 @@ theorem subst_wlog {x y z s : ℝ} (hxy : 0 ≤ x * y) (hxyz : x + y + z = 0) : _ ≤ (2 * (x ^ 2 + y ^ 2 + (x + y) ^ 2) + 2 * s ^ 2) ^ 4 / 4 ^ 4 := by gcongr (?_ + _) ^ 4 / _ apply rhs_ineq - refine le_of_pow_le_pow 2 (by positivity) (by positivity) ?_ + refine le_of_pow_le_pow_left two_ne_zero (by positivity) ?_ calc (32 * |x * y * z * s|) ^ 2 = 32 * (2 * s ^ 2 * (16 * x ^ 2 * y ^ 2 * (x + y) ^ 2)) := by rw [mul_pow, sq_abs, hz]; ring diff --git a/Archive/Imo/Imo2008Q3.lean b/Archive/Imo/Imo2008Q3.lean index 32d6e8437a59b..f199a524867c0 100644 --- a/Archive/Imo/Imo2008Q3.lean +++ b/Archive/Imo/Imo2008Q3.lean @@ -62,10 +62,10 @@ theorem p_lemma (p : ℕ) (hpp : Nat.Prime p) (hp_mod_4_eq_1 : p ≡ 1 [MOD 4]) have hreal₂ : (p : ℝ) > 20 := by assumption_mod_cast have hreal₃ : (k : ℝ) ^ 2 + 4 ≥ p := by assumption_mod_cast have hreal₅ : (k : ℝ) > 4 := by - refine' lt_of_pow_lt_pow 2 k.cast_nonneg _ + refine' lt_of_pow_lt_pow_left 2 k.cast_nonneg _ linarith only [hreal₂, hreal₃] have hreal₆ : (k : ℝ) > sqrt (2 * n) := by - refine' lt_of_pow_lt_pow 2 k.cast_nonneg _ + refine' lt_of_pow_lt_pow_left 2 k.cast_nonneg _ rw [sq_sqrt (mul_nonneg zero_le_two n.cast_nonneg)] linarith only [hreal₁, hreal₃, hreal₅] exact ⟨n, hnat₁, by linarith only [hreal₆, hreal₁]⟩ diff --git a/Archive/Imo/Imo2008Q4.lean b/Archive/Imo/Imo2008Q4.lean index 4619f634d21ab..f1a306321dc7f 100644 --- a/Archive/Imo/Imo2008Q4.lean +++ b/Archive/Imo/Imo2008Q4.lean @@ -29,8 +29,7 @@ open Real namespace Imo2008Q4 theorem abs_eq_one_of_pow_eq_one (x : ℝ) (n : ℕ) (hn : n ≠ 0) (h : x ^ n = 1) : |x| = 1 := by - rw [← pow_left_inj (abs_nonneg x) zero_le_one (pos_iff_ne_zero.2 hn), one_pow, pow_abs, h, - abs_one] + rw [← pow_left_inj (abs_nonneg x) zero_le_one hn, one_pow, pow_abs, h, abs_one] #align imo2008_q4.abs_eq_one_of_pow_eq_one Imo2008Q4.abs_eq_one_of_pow_eq_one end Imo2008Q4 diff --git a/Archive/Wiedijk100Theorems/AbelRuffini.lean b/Archive/Wiedijk100Theorems/AbelRuffini.lean index f0b4efa515e72..0e2091001b402 100644 --- a/Archive/Wiedijk100Theorems/AbelRuffini.lean +++ b/Archive/Wiedijk100Theorems/AbelRuffini.lean @@ -126,7 +126,7 @@ theorem real_roots_Phi_ge_aux (hab : b < a) : have hfa : 0 ≤ f a := by -- Porting note: was `simp_rw` simp only [hf, ← sq] - refine' add_nonneg (sub_nonneg.mpr (pow_le_pow ha _)) _ <;> norm_num + refine' add_nonneg (sub_nonneg.mpr (pow_le_pow_right ha _)) _ <;> norm_num obtain ⟨x, ⟨-, hx1⟩, hx2⟩ := intermediate_value_Ico' hle (hc _) (Set.mem_Ioc.mpr ⟨hf1, hf0⟩) obtain ⟨y, ⟨hy1, -⟩, hy2⟩ := intermediate_value_Ioc ha (hc _) (Set.mem_Ioc.mpr ⟨hf1, hfa⟩) exact ⟨x, y, (hx1.trans hy1).ne, hx2, hy2⟩ @@ -137,7 +137,7 @@ theorem real_roots_Phi_ge_aux (hab : b < a) : f (-a) = (a : ℝ) ^ 2 - (a : ℝ) ^ 5 + b := by norm_num [hf, ← sq, sub_eq_add_neg, add_comm, Odd.neg_pow (by decide : Odd 5)] _ ≤ (a : ℝ) ^ 2 - (a : ℝ) ^ 3 + (a - 1) := by - refine' add_le_add (sub_le_sub_left (pow_le_pow ha _) _) _ <;> linarith + refine' add_le_add (sub_le_sub_left (pow_le_pow_right ha _) _) _ <;> linarith _ = -((a : ℝ) - 1) ^ 2 * (a + 1) := by ring _ ≤ 0 := by nlinarith have ha' := neg_nonpos.mpr (hle.trans ha) diff --git a/Archive/Wiedijk100Theorems/PerfectNumbers.lean b/Archive/Wiedijk100Theorems/PerfectNumbers.lean index c610c03973978..f0bd7ac012d17 100644 --- a/Archive/Wiedijk100Theorems/PerfectNumbers.lean +++ b/Archive/Wiedijk100Theorems/PerfectNumbers.lean @@ -115,7 +115,7 @@ theorem eq_two_pow_mul_prime_mersenne_of_even_perfect {n : ℕ} (ev : Even n) (p | .succ k => apply ne_of_lt _ jcon2 rw [mersenne, ← Nat.pred_eq_sub_one, Nat.lt_pred_iff, ← pow_one (Nat.succ 1)] - apply pow_lt_pow (Nat.lt_succ_self 1) (Nat.succ_lt_succ (Nat.succ_pos k)) + apply pow_lt_pow_right (Nat.lt_succ_self 1) (Nat.succ_lt_succ (Nat.succ_pos k)) contrapose! hm simp [hm] #align theorems_100.nat.eq_two_pow_mul_prime_mersenne_of_even_perfect Theorems100.Nat.eq_two_pow_mul_prime_mersenne_of_even_perfect diff --git a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean index ac3bc23dc7090..66e17dcf383cb 100644 --- a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean +++ b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean @@ -168,7 +168,7 @@ theorem card_le_two_pow {x k : ℕ} : card M₁ ≤ card (image f K) := card_le_of_subset h _ ≤ card K := card_image_le _ ≤ 2 ^ card (image Nat.succ (range k)) := by simp only [card_powerset]; rfl - _ ≤ 2 ^ card (range k) := (pow_le_pow one_le_two card_image_le) + _ ≤ 2 ^ card (range k) := (pow_le_pow_right one_le_two card_image_le) _ = 2 ^ k := by rw [card_range k] #align theorems_100.card_le_two_pow Theorems100.card_le_two_pow diff --git a/Archive/ZagierTwoSquares.lean b/Archive/ZagierTwoSquares.lean index 2b91f80fcb9e8..8b52d84d39431 100644 --- a/Archive/ZagierTwoSquares.lean +++ b/Archive/ZagierTwoSquares.lean @@ -195,5 +195,5 @@ theorem Nat.Prime.sq_add_sq' {p : ℕ} [h : Fact p.Prime] (hp : p % 4 = 1) : (Equiv.Perm.card_fixedPoints_modEq (p := 2) (n := 1) (complexInvo_sq k)) contrapose key rw [Set.not_nonempty_iff_eq_empty] at key - simp_rw [key, Fintype.card_of_isEmpty, card_fixedPoints_eq_one] + simp_rw [key, Fintype.card_eq_zero, card_fixedPoints_eq_one] decide diff --git a/Counterexamples.lean b/Counterexamples.lean index dd67bf15a550c..b637aee44f4c4 100644 --- a/Counterexamples.lean +++ b/Counterexamples.lean @@ -8,6 +8,7 @@ import Counterexamples.HomogeneousPrimeNotPrime import Counterexamples.LinearOrderWithPosMulPosEqZero import Counterexamples.MapFloor import Counterexamples.Monic_nonRegular +import Counterexamples.OrderedCancelAddCommMonoidWithBounds import Counterexamples.Phillips import Counterexamples.Pseudoelement import Counterexamples.QuadraticForm diff --git a/Counterexamples/OrderedCancelAddCommMonoidWithBounds.lean b/Counterexamples/OrderedCancelAddCommMonoidWithBounds.lean new file mode 100644 index 0000000000000..9fd427765bf90 --- /dev/null +++ b/Counterexamples/OrderedCancelAddCommMonoidWithBounds.lean @@ -0,0 +1,19 @@ +/- +Copyright (c) 2023 Martin Dvorak. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Martin Dvorak +-/ +import Mathlib.Algebra.Order.Monoid.Defs + +/-! +# Do not combine OrderedCancelAddCommMonoid with BoundedOrder + +This file shows that combining `OrderedCancelAddCommMonoid` with `BoundedOrder` is not a good idea, +as such a structure must be trivial (`⊥ = x = ⊤` for all `x`). +The same applies to any superclasses, e.g. combining `StrictOrderedSemiring` with `CompleteLattice`. +-/ + +example {α : Type*} [OrderedCancelAddCommMonoid α] [BoundedOrder α] [Nontrivial α] : False := + have top_pos := pos_of_lt_add_right (bot_le.trans_lt (add_lt_add_left bot_lt_top (⊥ : α))) + have top_add_top_lt_self := lt_add_of_le_of_pos (@le_top _ _ _ (⊤ + ⊤)) top_pos + top_add_top_lt_self.false diff --git a/Mathlib.lean b/Mathlib.lean index 9159a9eceaa8c..cf5613827a4b8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -377,6 +377,7 @@ import Mathlib.Algebra.Order.Invertible import Mathlib.Algebra.Order.Kleene import Mathlib.Algebra.Order.LatticeGroup import Mathlib.Algebra.Order.Module +import Mathlib.Algebra.Order.Module.Defs import Mathlib.Algebra.Order.Module.Synonym import Mathlib.Algebra.Order.Monoid.Basic import Mathlib.Algebra.Order.Monoid.Canonical.Defs @@ -870,6 +871,7 @@ import Mathlib.Analysis.SpecialFunctions.Log.Base import Mathlib.Analysis.SpecialFunctions.Log.Basic import Mathlib.Analysis.SpecialFunctions.Log.Deriv import Mathlib.Analysis.SpecialFunctions.Log.Monotone +import Mathlib.Analysis.SpecialFunctions.Log.NegMulLog import Mathlib.Analysis.SpecialFunctions.NonIntegrable import Mathlib.Analysis.SpecialFunctions.PolarCoord import Mathlib.Analysis.SpecialFunctions.PolynomialExp @@ -2059,6 +2061,7 @@ import Mathlib.FieldTheory.Laurent import Mathlib.FieldTheory.Minpoly.Basic import Mathlib.FieldTheory.Minpoly.Field import Mathlib.FieldTheory.Minpoly.IsIntegrallyClosed +import Mathlib.FieldTheory.Minpoly.MinpolyDiv import Mathlib.FieldTheory.MvPolynomial import Mathlib.FieldTheory.Normal import Mathlib.FieldTheory.NormalClosure @@ -2115,6 +2118,7 @@ import Mathlib.Geometry.Manifold.Diffeomorph import Mathlib.Geometry.Manifold.Instances.Real import Mathlib.Geometry.Manifold.Instances.Sphere import Mathlib.Geometry.Manifold.Instances.UnitsOfNormedAlgebra +import Mathlib.Geometry.Manifold.InteriorBoundary import Mathlib.Geometry.Manifold.LocalInvariantProperties import Mathlib.Geometry.Manifold.MFDeriv import Mathlib.Geometry.Manifold.Metrizable @@ -2355,6 +2359,7 @@ import Mathlib.LinearAlgebra.CliffordAlgebra.Even import Mathlib.LinearAlgebra.CliffordAlgebra.EvenEquiv import Mathlib.LinearAlgebra.CliffordAlgebra.Fold import Mathlib.LinearAlgebra.CliffordAlgebra.Grading +import Mathlib.LinearAlgebra.CliffordAlgebra.Inversion import Mathlib.LinearAlgebra.CliffordAlgebra.Star import Mathlib.LinearAlgebra.Coevaluation import Mathlib.LinearAlgebra.Contraction @@ -2475,6 +2480,8 @@ import Mathlib.LinearAlgebra.TensorAlgebra.Grading import Mathlib.LinearAlgebra.TensorAlgebra.ToTensorPower import Mathlib.LinearAlgebra.TensorPower import Mathlib.LinearAlgebra.TensorProduct +import Mathlib.LinearAlgebra.TensorProduct.Graded.External +import Mathlib.LinearAlgebra.TensorProduct.Graded.Internal import Mathlib.LinearAlgebra.TensorProduct.Matrix import Mathlib.LinearAlgebra.TensorProduct.Opposite import Mathlib.LinearAlgebra.TensorProduct.Prod @@ -2920,6 +2927,7 @@ import Mathlib.Order.RelSeries import Mathlib.Order.SemiconjSup import Mathlib.Order.Sublattice import Mathlib.Order.SuccPred.Basic +import Mathlib.Order.SuccPred.CompleteLinearOrder import Mathlib.Order.SuccPred.IntervalSucc import Mathlib.Order.SuccPred.Limit import Mathlib.Order.SuccPred.LinearLocallyFinite @@ -3065,6 +3073,7 @@ import Mathlib.RingTheory.Ideal.QuotientOperations import Mathlib.RingTheory.Int.Basic import Mathlib.RingTheory.IntegralClosure import Mathlib.RingTheory.IntegralDomain +import Mathlib.RingTheory.IntegralRestrict import Mathlib.RingTheory.IntegrallyClosed import Mathlib.RingTheory.IsAdjoinRoot import Mathlib.RingTheory.IsTensorProduct diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index 98887d9eb5146..65c7c472ed9ec 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -924,3 +924,27 @@ end Module example {R A} [CommSemiring R] [Semiring A] [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] : Algebra R A := Algebra.ofModule smul_mul_assoc mul_smul_comm + +section invertibility + +variable {R A B : Type*} +variable [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] + +/-- If there is a linear map `f : A →ₗ[R] B` that preserves `1`, then `algebraMap R B r` is +invertible when `algebraMap R A r` is. -/ +abbrev Invertible.algebraMapOfInvertibleAlgebraMap (f : A →ₗ[R] B) (hf : f 1 = 1) {r : R} + (h : Invertible (algebraMap R A r)) : Invertible (algebraMap R B r) where + invOf := f ⅟(algebraMap R A r) + invOf_mul_self := by rw [← Algebra.commutes, ← Algebra.smul_def, ← map_smul, Algebra.smul_def, + mul_invOf_self, hf] + mul_invOf_self := by rw [← Algebra.smul_def, ← map_smul, Algebra.smul_def, mul_invOf_self, hf] + +/-- If there is a linear map `f : A →ₗ[R] B` that preserves `1`, then `algebraMap R B r` is +a unit when `algebraMap R A r` is. -/ +lemma IsUnit.algebraMap_of_algebraMap (f : A →ₗ[R] B) (hf : f 1 = 1) {r : R} + (h : IsUnit (algebraMap R A r)) : IsUnit (algebraMap R B r) := + let ⟨i⟩ := nonempty_invertible h + letI := Invertible.algebraMapOfInvertibleAlgebraMap f hf i + isUnit_of_invertible _ + +end invertibility diff --git a/Mathlib/Algebra/CharP/Basic.lean b/Mathlib/Algebra/CharP/Basic.lean index 442eacdc5f895..5df7a76e492ee 100644 --- a/Mathlib/Algebra/CharP/Basic.lean +++ b/Mathlib/Algebra/CharP/Basic.lean @@ -145,12 +145,30 @@ theorem CharP.intCast_eq_intCast [AddGroupWithOne R] (p : ℕ) [CharP R p] {a b rw [eq_comm, ← sub_eq_zero, ← Int.cast_sub, CharP.int_cast_eq_zero_iff R p, Int.modEq_iff_dvd] #align char_p.int_cast_eq_int_cast CharP.intCast_eq_intCast -theorem CharP.natCast_eq_natCast [AddGroupWithOne R] (p : ℕ) [CharP R p] {a b : ℕ} : - (a : R) = b ↔ a ≡ b [MOD p] := by - rw [← Int.cast_ofNat, ← Int.cast_ofNat b] - exact (CharP.intCast_eq_intCast _ _).trans Int.coe_nat_modEq_iff +theorem CharP.natCast_eq_natCast' [AddMonoidWithOne R] (p : ℕ) [CharP R p] {a b : ℕ} + (h : a ≡ b [MOD p]) : (a : R) = b := by + wlog hle : a ≤ b + · exact (this R p h.symm (le_of_not_le hle)).symm + rw [Nat.modEq_iff_dvd' hle] at h + rw [← Nat.sub_add_cancel hle, Nat.cast_add, (CharP.cast_eq_zero_iff R p _).mpr h, zero_add] + +theorem CharP.natCast_eq_natCast [AddMonoidWithOne R] [IsRightCancelAdd R] (p : ℕ) [CharP R p] + {a b : ℕ} : (a : R) = b ↔ a ≡ b [MOD p] := by + wlog hle : a ≤ b + · rw [eq_comm, this R p (le_of_not_le hle), Nat.ModEq.comm] + rw [Nat.modEq_iff_dvd' hle, ← CharP.cast_eq_zero_iff R p (b - a), + ← add_right_cancel_iff (G := R) (a := a) (b := b - a), zero_add, ← Nat.cast_add, + Nat.sub_add_cancel hle, eq_comm] #align char_p.nat_cast_eq_nat_cast CharP.natCast_eq_natCast +theorem CharP.intCast_eq_intCast_mod [AddGroupWithOne R] (p : ℕ) [CharP R p] {a : ℤ} : + (a : R) = a % (p : ℤ) := + (CharP.intCast_eq_intCast R p).mpr (Int.mod_modEq a p).symm + +theorem CharP.natCast_eq_natCast_mod [AddMonoidWithOne R] (p : ℕ) [CharP R p] {a : ℕ} : + (a : R) = a % p := + CharP.natCast_eq_natCast' R p (Nat.mod_modEq a p).symm + theorem CharP.eq [AddMonoidWithOne R] {p q : ℕ} (_c1 : CharP R p) (_c2 : CharP R q) : p = q := Nat.dvd_antisymm ((CharP.cast_eq_zero_iff R p q).1 (CharP.cast_eq_zero _ _)) ((CharP.cast_eq_zero_iff R q p).1 (CharP.cast_eq_zero _ _)) diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 99d2b0e0e9b84..c5ed8501dd2b9 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -209,9 +209,9 @@ theorem sub_one_dvd_pow_sub_one [Ring α] (x : α) (n : ℕ) : theorem nat_sub_dvd_pow_sub_pow (x y n : ℕ) : x - y ∣ x ^ n - y ^ n := by cases' le_or_lt y x with h h - · have : y ^ n ≤ x ^ n := Nat.pow_le_pow_of_le_left h _ + · have : y ^ n ≤ x ^ n := Nat.pow_le_pow_left h _ exact mod_cast sub_dvd_pow_sub_pow (x : ℤ) (↑y) n - · have : x ^ n ≤ y ^ n := Nat.pow_le_pow_of_le_left h.le _ + · have : x ^ n ≤ y ^ n := Nat.pow_le_pow_left h.le _ exact (Nat.sub_eq_zero_of_le this).symm ▸ dvd_zero (x - y) #align nat_sub_dvd_pow_sub_pow nat_sub_dvd_pow_sub_pow diff --git a/Mathlib/Algebra/Group/Hom/Instances.lean b/Mathlib/Algebra/Group/Hom/Instances.lean index 372a8d01165ae..cae4b2cc6ef35 100644 --- a/Mathlib/Algebra/Group/Hom/Instances.lean +++ b/Mathlib/Algebra/Group/Hom/Instances.lean @@ -331,6 +331,14 @@ theorem AddMonoidHom.map_mul_iff (f : R →+ S) : Iff.symm AddMonoidHom.ext_iff₂ #align add_monoid_hom.map_mul_iff AddMonoidHom.map_mul_iff +lemma AddMonoidHom.mulLeft_eq_mulRight_iff_forall_commute {a : R} : + mulLeft a = mulRight a ↔ ∀ b, Commute a b := + FunLike.ext_iff + +lemma AddMonoidHom.mulRight_eq_mulLeft_iff_forall_commute {b : R} : + mulRight b = mulLeft b ↔ ∀ a, Commute a b := + FunLike.ext_iff + /-- The left multiplication map: `(a, b) ↦ a * b`. See also `AddMonoidHom.mulLeft`. -/ @[simps!] def AddMonoid.End.mulLeft : R →+ AddMonoid.End R := @@ -345,10 +353,6 @@ def AddMonoid.End.mulRight : R →+ AddMonoid.End R := #align add_monoid.End.mul_right AddMonoid.End.mulRight #align add_monoid.End.mul_right_apply_apply AddMonoid.End.mulRight_apply_apply -lemma AddMonoid.End.mulRight_eq_mulLeft_of_commute (a : R) (h : ∀ (b : R), Commute a b) : - mulRight a = mulLeft a := - AddMonoidHom.ext fun _ ↦ (h _).eq.symm - end Semiring section CommSemiring @@ -357,9 +361,9 @@ variable {R S : Type*} [NonUnitalNonAssocCommSemiring R] namespace AddMonoid.End -lemma comm_mulRight_eq_mulLeft : mulRight = (mulLeft : R →+ AddMonoid.End R) := by - ext a - exact mulRight_eq_mulLeft_of_commute _ (Commute.all _) +lemma mulRight_eq_mulLeft : mulRight = (mulLeft : R →+ AddMonoid.End R) := + AddMonoidHom.ext fun _ => + Eq.symm <| AddMonoidHom.mulLeft_eq_mulRight_iff_forall_commute.2 (.all _) end AddMonoid.End diff --git a/Mathlib/Algebra/GroupPower/CovariantClass.lean b/Mathlib/Algebra/GroupPower/CovariantClass.lean index fde8ae6af4676..a89c22c72e81f 100644 --- a/Mathlib/Algebra/GroupPower/CovariantClass.lean +++ b/Mathlib/Algebra/GroupPower/CovariantClass.lean @@ -31,15 +31,15 @@ section Left variable [CovariantClass M M (· * ·) (· ≤ ·)] {x : M} -@[to_additive (attr := mono) nsmul_le_nsmul_of_le_right] -theorem pow_le_pow_of_le_left' [CovariantClass M M (swap (· * ·)) (· ≤ ·)] {a b : M} (hab : a ≤ b) : +@[to_additive (attr := mono) nsmul_le_nsmul_right] +theorem pow_le_pow_left' [CovariantClass M M (swap (· * ·)) (· ≤ ·)] {a b : M} (hab : a ≤ b) : ∀ i : ℕ, a ^ i ≤ b ^ i | 0 => by simp | k + 1 => by rw [pow_succ, pow_succ] - exact mul_le_mul' hab (pow_le_pow_of_le_left' hab k) -#align pow_le_pow_of_le_left' pow_le_pow_of_le_left' -#align nsmul_le_nsmul_of_le_right nsmul_le_nsmul_of_le_right + exact mul_le_mul' hab (pow_le_pow_left' hab k) +#align pow_le_pow_of_le_left' pow_le_pow_left' +#align nsmul_le_nsmul_of_le_right nsmul_le_nsmul_right @[to_additive nsmul_nonneg] theorem one_le_pow_of_one_le' {a : M} (H : 1 ≤ a) : ∀ n : ℕ, 1 ≤ a ^ n @@ -56,20 +56,20 @@ theorem pow_le_one' {a : M} (H : a ≤ 1) (n : ℕ) : a ^ n ≤ 1 := #align pow_le_one' pow_le_one' #align nsmul_nonpos nsmul_nonpos -@[to_additive nsmul_le_nsmul] -theorem pow_le_pow' {a : M} {n m : ℕ} (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := +@[to_additive nsmul_le_nsmul_left] +theorem pow_le_pow_right' {a : M} {n m : ℕ} (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := let ⟨k, hk⟩ := Nat.le.dest h calc a ^ n ≤ a ^ n * a ^ k := le_mul_of_one_le_right' (one_le_pow_of_one_le' ha _) _ = a ^ m := by rw [← hk, pow_add] -#align pow_le_pow' pow_le_pow' -#align nsmul_le_nsmul nsmul_le_nsmul +#align pow_le_pow' pow_le_pow_right' +#align nsmul_le_nsmul nsmul_le_nsmul_left -@[to_additive nsmul_le_nsmul_of_nonpos] -theorem pow_le_pow_of_le_one' {a : M} {n m : ℕ} (ha : a ≤ 1) (h : n ≤ m) : a ^ m ≤ a ^ n := - @pow_le_pow' Mᵒᵈ _ _ _ _ _ _ ha h -#align pow_le_pow_of_le_one' pow_le_pow_of_le_one' -#align nsmul_le_nsmul_of_nonpos nsmul_le_nsmul_of_nonpos +@[to_additive nsmul_le_nsmul_left_of_nonpos] +theorem pow_le_pow_right_of_le_one' {a : M} {n m : ℕ} (ha : a ≤ 1) (h : n ≤ m) : a ^ m ≤ a ^ n := + pow_le_pow_right' (M := Mᵒᵈ) ha h +#align pow_le_pow_of_le_one' pow_le_pow_right_of_le_one' +#align nsmul_le_nsmul_of_nonpos nsmul_le_nsmul_left_of_nonpos @[to_additive nsmul_pos] theorem one_lt_pow' {a : M} (ha : 1 < a) {k : ℕ} (hk : k ≠ 0) : 1 < a ^ k := by @@ -88,20 +88,20 @@ theorem pow_lt_one' {a : M} (ha : a < 1) {k : ℕ} (hk : k ≠ 0) : a ^ k < 1 := #align pow_lt_one' pow_lt_one' #align nsmul_neg nsmul_neg -@[to_additive nsmul_lt_nsmul] -theorem pow_lt_pow' [CovariantClass M M (· * ·) (· < ·)] {a : M} {n m : ℕ} (ha : 1 < a) +@[to_additive nsmul_lt_nsmul_left] +theorem pow_lt_pow_right' [CovariantClass M M (· * ·) (· < ·)] {a : M} {n m : ℕ} (ha : 1 < a) (h : n < m) : a ^ n < a ^ m := by rcases Nat.le.dest h with ⟨k, rfl⟩; clear h rw [pow_add, pow_succ', mul_assoc, ← pow_succ] exact lt_mul_of_one_lt_right' _ (one_lt_pow' ha k.succ_ne_zero) -#align pow_lt_pow' pow_lt_pow' -#align nsmul_lt_nsmul nsmul_lt_nsmul +#align pow_lt_pow_right' pow_lt_pow_right' +#align nsmul_lt_nsmul_left nsmul_lt_nsmul_left -@[to_additive nsmul_strictMono_right] -theorem pow_strictMono_left [CovariantClass M M (· * ·) (· < ·)] {a : M} (ha : 1 < a) : - StrictMono ((a ^ ·) : ℕ → M) := fun _ _ => pow_lt_pow' ha -#align pow_strict_mono_left pow_strictMono_left -#align nsmul_strict_mono_right nsmul_strictMono_right +@[to_additive nsmul_left_strictMono] +theorem pow_right_strictMono' [CovariantClass M M (· * ·) (· < ·)] {a : M} (ha : 1 < a) : + StrictMono ((a ^ ·) : ℕ → M) := fun _ _ => pow_lt_pow_right' ha +#align pow_strict_mono_left pow_right_strictMono' +#align nsmul_strict_mono_right nsmul_left_strictMono @[to_additive Left.pow_nonneg] theorem Left.one_le_pow_of_le (hx : 1 ≤ x) : ∀ {n : ℕ}, 1 ≤ x ^ n @@ -150,24 +150,22 @@ end Right section CovariantLTSwap variable [Preorder β] [CovariantClass M M (· * ·) (· < ·)] - [CovariantClass M M (swap (· * ·)) (· < ·)] {f : β → M} + [CovariantClass M M (swap (· * ·)) (· < ·)] {f : β → M} {n : ℕ} -@[to_additive StrictMono.nsmul_left] -theorem StrictMono.pow_right' (hf : StrictMono f) : ∀ {n : ℕ}, n ≠ 0 → StrictMono fun a => f a ^ n +@[to_additive StrictMono.const_nsmul] +theorem StrictMono.pow_const (hf : StrictMono f) : ∀ {n : ℕ}, n ≠ 0 → StrictMono (f · ^ n) | 0, hn => (hn rfl).elim | 1, _ => by simpa | Nat.succ <| Nat.succ n, _ => by - simp_rw [pow_succ _ (n + 1)] - exact hf.mul' (StrictMono.pow_right' hf n.succ_ne_zero) -#align strict_mono.pow_right' StrictMono.pow_right' -#align strict_mono.nsmul_left StrictMono.nsmul_left - -/-- See also `pow_strictMono_right` -/ -@[to_additive nsmul_strictMono_left] -- Porting note: nolint to_additive_doc -theorem pow_strictMono_right' {n : ℕ} (hn : n ≠ 0) : StrictMono fun a : M => a ^ n := - strictMono_id.pow_right' hn -#align pow_strict_mono_right' pow_strictMono_right' -#align nsmul_strict_mono_left nsmul_strictMono_left + simpa only [pow_succ] using hf.mul' (hf.pow_const n.succ_ne_zero) +#align strict_mono.pow_const StrictMono.pow_const +#align strict_mono.const_nsmul StrictMono.const_nsmul + +/-- See also `pow_left_strictMonoOn`. -/ +@[to_additive nsmul_right_strictMono] -- Porting note: nolint to_additive_doc +theorem pow_left_strictMono (hn : n ≠ 0) : StrictMono (· ^ n : M → M) := strictMono_id.pow_const hn +#align pow_strict_mono_right' pow_left_strictMono +#align nsmul_strict_mono_left nsmul_right_strictMono end CovariantLTSwap @@ -176,14 +174,14 @@ section CovariantLESwap variable [Preorder β] [CovariantClass M M (· * ·) (· ≤ ·)] [CovariantClass M M (swap (· * ·)) (· ≤ ·)] -@[to_additive Monotone.nsmul_left] +@[to_additive Monotone.const_nsmul] theorem Monotone.pow_right {f : β → M} (hf : Monotone f) : ∀ n : ℕ, Monotone fun a => f a ^ n | 0 => by simpa using monotone_const | n + 1 => by simp_rw [pow_succ] exact hf.mul' (Monotone.pow_right hf _) #align monotone.pow_right Monotone.pow_right -#align monotone.nsmul_left Monotone.nsmul_left +#align monotone.const_nsmul Monotone.const_nsmul @[to_additive nsmul_mono_left] theorem pow_mono_right (n : ℕ) : Monotone fun a : M => a ^ n := @@ -258,17 +256,17 @@ theorem pow_eq_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n = 1 ↔ x = 1 := variable [CovariantClass M M (· * ·) (· < ·)] {a : M} {m n : ℕ} -@[to_additive nsmul_le_nsmul_iff] -theorem pow_le_pow_iff' (ha : 1 < a) : a ^ m ≤ a ^ n ↔ m ≤ n := - (pow_strictMono_left ha).le_iff_le -#align pow_le_pow_iff' pow_le_pow_iff' -#align nsmul_le_nsmul_iff nsmul_le_nsmul_iff +@[to_additive nsmul_le_nsmul_iff_left] +theorem pow_le_pow_iff_right' (ha : 1 < a) : a ^ m ≤ a ^ n ↔ m ≤ n := + (pow_right_strictMono' ha).le_iff_le +#align pow_le_pow_iff' pow_le_pow_iff_right' +#align nsmul_le_nsmul_iff nsmul_le_nsmul_iff_left -@[to_additive nsmul_lt_nsmul_iff] -theorem pow_lt_pow_iff' (ha : 1 < a) : a ^ m < a ^ n ↔ m < n := - (pow_strictMono_left ha).lt_iff_lt -#align pow_lt_pow_iff' pow_lt_pow_iff' -#align nsmul_lt_nsmul_iff nsmul_lt_nsmul_iff +@[to_additive nsmul_lt_nsmul_iff_left] +theorem pow_lt_pow_iff_right' (ha : 1 < a) : a ^ m < a ^ n ↔ m < n := + (pow_right_strictMono' ha).lt_iff_lt +#align pow_lt_pow_iff' pow_lt_pow_iff_right' +#align nsmul_lt_nsmul_iff nsmul_lt_nsmul_iff_left end CovariantLE @@ -276,11 +274,11 @@ section CovariantLESwap variable [CovariantClass M M (· * ·) (· ≤ ·)] [CovariantClass M M (swap (· * ·)) (· ≤ ·)] -@[to_additive lt_of_nsmul_lt_nsmul] -theorem lt_of_pow_lt_pow' {a b : M} (n : ℕ) : a ^ n < b ^ n → a < b := +@[to_additive lt_of_nsmul_lt_nsmul_right] +theorem lt_of_pow_lt_pow_left' {a b : M} (n : ℕ) : a ^ n < b ^ n → a < b := (pow_mono_right _).reflect_lt -#align lt_of_pow_lt_pow' lt_of_pow_lt_pow' -#align lt_of_nsmul_lt_nsmul lt_of_nsmul_lt_nsmul +#align lt_of_pow_lt_pow' lt_of_pow_lt_pow_left' +#align lt_of_nsmul_lt_nsmul lt_of_nsmul_lt_nsmul_right @[to_additive min_lt_of_add_lt_two_nsmul] theorem min_lt_of_mul_lt_sq {a b c : M} (h : a * b < c ^ 2) : min a b < c := by @@ -300,11 +298,11 @@ section CovariantLTSwap variable [CovariantClass M M (· * ·) (· < ·)] [CovariantClass M M (swap (· * ·)) (· < ·)] -@[to_additive le_of_nsmul_le_nsmul] -theorem le_of_pow_le_pow' {a b : M} {n : ℕ} (hn : n ≠ 0) : a ^ n ≤ b ^ n → a ≤ b := - (pow_strictMono_right' hn).le_iff_le.1 -#align le_of_pow_le_pow' le_of_pow_le_pow' -#align le_of_nsmul_le_nsmul le_of_nsmul_le_nsmul +@[to_additive le_of_nsmul_le_nsmul_right'] +theorem le_of_pow_le_pow_left' {a b : M} {n : ℕ} (hn : n ≠ 0) : a ^ n ≤ b ^ n → a ≤ b := + (pow_left_strictMono hn).le_iff_le.1 +#align le_of_pow_le_pow' le_of_pow_le_pow_left' +#align le_of_nsmul_le_nsmul le_of_nsmul_le_nsmul_right' @[to_additive min_le_of_add_le_two_nsmul] theorem min_le_of_mul_le_sq {a b c : M} (h : a * b ≤ c ^ 2) : min a b ≤ c := by diff --git a/Mathlib/Algebra/GroupPower/Lemmas.lean b/Mathlib/Algebra/GroupPower/Lemmas.lean index 5f081f7272ad8..9662035c32d90 100644 --- a/Mathlib/Algebra/GroupPower/Lemmas.lean +++ b/Mathlib/Algebra/GroupPower/Lemmas.lean @@ -326,12 +326,12 @@ theorem one_lt_zpow' (ha : 1 < a) {k : ℤ} (hk : (0 : ℤ) < k) : 1 < a ^ k := #align zsmul_pos zsmul_pos @[to_additive zsmul_strictMono_left] -theorem zpow_strictMono_right (ha : 1 < a) : StrictMono fun n : ℤ => a ^ n := fun m n h => +theorem zpow_right_strictMono (ha : 1 < a) : StrictMono fun n : ℤ => a ^ n := fun m n h => calc a ^ m = a ^ m * 1 := (mul_one _).symm _ < a ^ m * a ^ (n - m) := mul_lt_mul_left' (one_lt_zpow' ha <| sub_pos_of_lt h) _ _ = a ^ n := by rw [← zpow_add]; simp -#align zpow_strict_mono_right zpow_strictMono_right +#align zpow_strict_mono_right zpow_right_strictMono #align zsmul_strict_mono_left zsmul_strictMono_left @[to_additive zsmul_mono_left] @@ -351,19 +351,19 @@ theorem zpow_le_zpow (ha : 1 ≤ a) (h : m ≤ n) : a ^ m ≤ a ^ n := @[to_additive] theorem zpow_lt_zpow (ha : 1 < a) (h : m < n) : a ^ m < a ^ n := - zpow_strictMono_right ha h + zpow_right_strictMono ha h #align zpow_lt_zpow zpow_lt_zpow #align zsmul_lt_zsmul zsmul_lt_zsmul @[to_additive] theorem zpow_le_zpow_iff (ha : 1 < a) : a ^ m ≤ a ^ n ↔ m ≤ n := - (zpow_strictMono_right ha).le_iff_le + (zpow_right_strictMono ha).le_iff_le #align zpow_le_zpow_iff zpow_le_zpow_iff #align zsmul_le_zsmul_iff zsmul_le_zsmul_iff @[to_additive] theorem zpow_lt_zpow_iff (ha : 1 < a) : a ^ m < a ^ n ↔ m < n := - (zpow_strictMono_right ha).lt_iff_lt + (zpow_right_strictMono ha).lt_iff_lt #align zpow_lt_zpow_iff zpow_lt_zpow_iff #align zsmul_lt_zsmul_iff zsmul_lt_zsmul_iff @@ -746,9 +746,9 @@ theorem strictMono_pow_bit1 (n : ℕ) : StrictMono fun a : R => a ^ bit1 n := by cases' le_total a 0 with ha ha · cases' le_or_lt b 0 with hb hb · rw [← neg_lt_neg_iff, ← neg_pow_bit1, ← neg_pow_bit1] - exact pow_lt_pow_of_lt_left (neg_lt_neg hab) (neg_nonneg.2 hb) (bit1_pos (zero_le n)) + exact pow_lt_pow_left (neg_lt_neg hab) (neg_nonneg.2 hb) n.bit1_ne_zero · exact (pow_bit1_nonpos_iff.2 ha).trans_lt (pow_bit1_pos_iff.2 hb) - · exact pow_lt_pow_of_lt_left hab ha (bit1_pos (zero_le n)) + · exact pow_lt_pow_left hab ha n.bit1_ne_zero #align strict_mono_pow_bit1 strictMono_pow_bit1 end bit1 diff --git a/Mathlib/Algebra/GroupPower/Order.lean b/Mathlib/Algebra/GroupPower/Order.lean index d24b7c00f7990..b75018c35ff98 100644 --- a/Mathlib/Algebra/GroupPower/Order.lean +++ b/Mathlib/Algebra/GroupPower/Order.lean @@ -19,6 +19,8 @@ Note that some lemmas are in `Algebra/GroupPower/Lemmas.lean` as they import fil depend on this file. -/ +open Function + variable {M R : Type*} namespace CanonicallyOrderedCommSemiring @@ -83,32 +85,26 @@ theorem one_le_pow_of_one_le (H : 1 ≤ a) : ∀ n : ℕ, 1 ≤ a ^ n mul_le_mul H (one_le_pow_of_one_le H n) zero_le_one (le_trans zero_le_one H) #align one_le_pow_of_one_le one_le_pow_of_one_le -theorem pow_mono (h : 1 ≤ a) : Monotone fun n : ℕ => a ^ n := +theorem pow_right_mono (h : 1 ≤ a) : Monotone (a ^ ·) := monotone_nat_of_le_succ fun n => by rw [pow_succ] exact le_mul_of_one_le_left (pow_nonneg (zero_le_one.trans h) _) h -#align pow_mono pow_mono +#align pow_mono pow_right_mono -theorem pow_le_pow (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := - pow_mono ha h -#align pow_le_pow pow_le_pow +theorem pow_le_pow_right (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := pow_right_mono ha h +#align pow_le_pow pow_le_pow_right -theorem le_self_pow (ha : 1 ≤ a) (h : m ≠ 0) : a ≤ a ^ m := - (pow_one a).symm.trans_le (pow_le_pow ha <| pos_iff_ne_zero.mpr h) +theorem le_self_pow (ha : 1 ≤ a) (h : m ≠ 0) : a ≤ a ^ m := by + simpa only [pow_one] using pow_le_pow_right ha <| pos_iff_ne_zero.2 h +#align self_le_pow le_self_pow #align le_self_pow le_self_pow @[mono] -theorem pow_le_pow_of_le_left {a b : R} (ha : 0 ≤ a) (hab : a ≤ b) : ∀ i : ℕ, a ^ i ≤ b ^ i := by - intro i - induction i with - | zero => simp - | succ k ih => - rw [pow_succ, pow_succ] - apply mul_le_mul hab - apply ih - apply pow_nonneg ha - apply le_trans ha hab -#align pow_le_pow_of_le_left pow_le_pow_of_le_left +theorem pow_le_pow_left {a b : R} (ha : 0 ≤ a) (hab : a ≤ b) : ∀ n, a ^ n ≤ b ^ n + | 0 => by simp + | n + 1 => by simpa only [pow_succ] + using mul_le_mul hab (pow_le_pow_left ha hab _) (pow_nonneg ha _) (ha.trans hab) +#align pow_le_pow_of_le_left pow_le_pow_left theorem one_lt_pow (ha : 1 < a) : ∀ {n : ℕ} (_ : n ≠ 0), 1 < a ^ n | 0, h => (h rfl).elim @@ -123,60 +119,52 @@ section StrictOrderedSemiring variable [StrictOrderedSemiring R] {a x y : R} {n m : ℕ} -theorem pow_lt_pow_of_lt_left (h : x < y) (hx : 0 ≤ x) : ∀ {n : ℕ}, 0 < n → x ^ n < y ^ n +theorem pow_lt_pow_left (h : x < y) (hx : 0 ≤ x) : ∀ {n : ℕ}, n ≠ 0 → x ^ n < y ^ n | 0, hn => by contradiction | n + 1, _ => by simpa only [pow_succ'] using - mul_lt_mul_of_le_of_le' (pow_le_pow_of_le_left hx h.le _) h (pow_pos (hx.trans_lt h) _) hx -#align pow_lt_pow_of_lt_left pow_lt_pow_of_lt_left + mul_lt_mul_of_le_of_le' (pow_le_pow_left hx h.le _) h (pow_pos (hx.trans_lt h) _) hx +#align pow_lt_pow_of_lt_left pow_lt_pow_left -theorem strictMonoOn_pow (hn : 0 < n) : StrictMonoOn (fun x : R => x ^ n) (Set.Ici 0) := - fun _ hx _ _ h => pow_lt_pow_of_lt_left h hx hn -#align strict_mono_on_pow strictMonoOn_pow +/-- See also `pow_left_strictMonoOn`. -/ +lemma pow_left_strictMonoOn (hn : n ≠ 0) : StrictMonoOn (· ^ n : R → R) (Set.Ici 0) := + fun _a ha _b _ hab ↦ pow_lt_pow_left hab ha hn +#align strict_mono_on_pow pow_left_strictMonoOn -theorem pow_strictMono_right (h : 1 < a) : StrictMono fun n : ℕ => a ^ n := +/-- See also `pow_right_strictMono'`. -/ +lemma pow_right_strictMono (h : 1 < a) : StrictMono (a ^ ·) := have : 0 < a := zero_le_one.trans_lt h strictMono_nat_of_lt_succ fun n => by simpa only [one_mul, pow_succ] using mul_lt_mul h (le_refl (a ^ n)) (pow_pos this _) this.le -#align pow_strict_mono_right pow_strictMono_right - -theorem pow_lt_pow (h : 1 < a) (h2 : n < m) : a ^ n < a ^ m := - pow_strictMono_right h h2 -#align pow_lt_pow pow_lt_pow +#align pow_strict_mono_right pow_right_strictMono -theorem pow_lt_pow_iff (h : 1 < a) : a ^ n < a ^ m ↔ n < m := - (pow_strictMono_right h).lt_iff_lt -#align pow_lt_pow_iff pow_lt_pow_iff +theorem pow_lt_pow_right (h : 1 < a) (hmn : m < n) : a ^ m < a ^ n := pow_right_strictMono h hmn +#align pow_lt_pow_right pow_lt_pow_right -theorem pow_le_pow_iff (h : 1 < a) : a ^ n ≤ a ^ m ↔ n ≤ m := - (pow_strictMono_right h).le_iff_le -#align pow_le_pow_iff pow_le_pow_iff +lemma pow_lt_pow_iff_right (h : 1 < a) : a ^ n < a ^ m ↔ n < m := (pow_right_strictMono h).lt_iff_lt +#align pow_lt_pow_iff_ pow_lt_pow_iff_right -theorem self_lt_pow (h : 1 < a) (h2 : 1 < m) : a < a ^ m := by - calc - a = a ^ 1 := (pow_one _).symm - _ < a ^ m := pow_lt_pow h h2 +lemma pow_le_pow_iff_right (h : 1 < a) : a ^ n ≤ a ^ m ↔ n ≤ m := (pow_right_strictMono h).le_iff_le +#align pow_le_pow_iff pow_le_pow_iff_right -theorem self_le_pow (h : 1 ≤ a) (h2 : 1 ≤ m) : a ≤ a ^ m := - le_self_pow h <| Nat.one_le_iff_ne_zero.mp h2 +theorem lt_self_pow (h : 1 < a) (hm : 1 < m) : a < a ^ m := by + simpa only [pow_one] using pow_lt_pow_right h hm -theorem strictAnti_pow (h₀ : 0 < a) (h₁ : a < 1) : StrictAnti fun n : ℕ => a ^ n := +theorem pow_right_strictAnti (h₀ : 0 < a) (h₁ : a < 1) : StrictAnti (a ^ ·) := strictAnti_nat_of_succ_lt fun n => by simpa only [pow_succ, one_mul] using mul_lt_mul h₁ le_rfl (pow_pos h₀ n) zero_le_one -#align strict_anti_pow strictAnti_pow +#align strict_anti_pow pow_right_strictAnti -theorem pow_lt_pow_iff_of_lt_one (h₀ : 0 < a) (h₁ : a < 1) : a ^ m < a ^ n ↔ n < m := - (strictAnti_pow h₀ h₁).lt_iff_lt -#align pow_lt_pow_iff_of_lt_one pow_lt_pow_iff_of_lt_one +theorem pow_lt_pow_iff_right_of_lt_one (h₀ : 0 < a) (h₁ : a < 1) : a ^ m < a ^ n ↔ n < m := + (pow_right_strictAnti h₀ h₁).lt_iff_lt +#align pow_lt_pow_iff_of_lt_one pow_lt_pow_iff_right_of_lt_one -theorem pow_lt_pow_of_lt_one (h : 0 < a) (ha : a < 1) {i j : ℕ} (hij : i < j) : a ^ j < a ^ i := - (pow_lt_pow_iff_of_lt_one h ha).2 hij -#align pow_lt_pow_of_lt_one pow_lt_pow_of_lt_one +theorem pow_lt_pow_right_of_lt_one (h₀ : 0 < a) (h₁ : a < 1) (hmn : m < n) : a ^ n < a ^ m := + (pow_lt_pow_iff_right_of_lt_one h₀ h₁).2 hmn +#align pow_lt_pow_of_lt_one pow_lt_pow_right_of_lt_one -theorem pow_lt_self_of_lt_one (h₀ : 0 < a) (h₁ : a < 1) (hn : 1 < n) : a ^ n < a := - calc - a ^ n < a ^ 1 := pow_lt_pow_of_lt_one h₀ h₁ hn - _ = a := pow_one _ +theorem pow_lt_self_of_lt_one (h₀ : 0 < a) (h₁ : a < 1) (hn : 1 < n) : a ^ n < a := by + simpa only [pow_one] using pow_lt_pow_right_of_lt_one h₀ h₁ hn #align pow_lt_self_of_lt_one pow_lt_self_of_lt_one theorem sq_pos_of_pos (ha : 0 < a) : 0 < a ^ 2 := by @@ -208,34 +196,46 @@ theorem sq_pos_of_neg (ha : a < 0) : 0 < a ^ 2 := end StrictOrderedRing section LinearOrderedSemiring +variable [LinearOrderedSemiring R] {a b : R} {m n : ℕ} + +lemma pow_le_pow_iff_left (ha : 0 ≤ a) (hb : 0 ≤ b) (hn : n ≠ 0) : a ^ n ≤ b ^ n ↔ a ≤ b := + (pow_left_strictMonoOn hn).le_iff_le ha hb + +lemma pow_lt_pow_iff_left (ha : 0 ≤ a) (hb : 0 ≤ b) (hn : n ≠ 0) : a ^ n < b ^ n ↔ a < b := + (pow_left_strictMonoOn hn).lt_iff_lt ha hb + +@[simp] +lemma pow_left_inj (ha : 0 ≤ a) (hb : 0 ≤ b) (hn : n ≠ 0) : a ^ n = b ^ n ↔ a = b := + (pow_left_strictMonoOn hn).eq_iff_eq ha hb +#align pow_left_inj pow_left_inj + +lemma pow_right_injective (ha₀ : 0 < a) (ha₁ : a ≠ 1) : Injective (a ^ ·) := by + obtain ha₁ | ha₁ := ha₁.lt_or_lt + · exact (pow_right_strictAnti ha₀ ha₁).injective + · exact (pow_right_strictMono ha₁).injective -variable [LinearOrderedSemiring R] {a b : R} +@[simp] +lemma pow_right_inj (ha₀ : 0 < a) (ha₁ : a ≠ 1) : a ^ m = a ^ n ↔ m = n := + (pow_right_injective ha₀ ha₁).eq_iff -theorem pow_le_one_iff_of_nonneg {a : R} (ha : 0 ≤ a) {n : ℕ} (hn : n ≠ 0) : a ^ n ≤ 1 ↔ a ≤ 1 := by - refine' ⟨_, pow_le_one n ha⟩ - rw [← not_lt, ← not_lt] - exact mt fun h => one_lt_pow h hn +theorem pow_le_one_iff_of_nonneg (ha : 0 ≤ a) (hn : n ≠ 0) : a ^ n ≤ 1 ↔ a ≤ 1 := by + simpa only [one_pow] using pow_le_pow_iff_left ha zero_le_one hn #align pow_le_one_iff_of_nonneg pow_le_one_iff_of_nonneg -theorem one_le_pow_iff_of_nonneg {a : R} (ha : 0 ≤ a) {n : ℕ} (hn : n ≠ 0) : 1 ≤ a ^ n ↔ 1 ≤ a := by - refine' ⟨_, fun h => one_le_pow_of_one_le h n⟩ - rw [← not_lt, ← not_lt] - exact mt fun h => pow_lt_one ha h hn +theorem one_le_pow_iff_of_nonneg (ha : 0 ≤ a) (hn : n ≠ 0) : 1 ≤ a ^ n ↔ 1 ≤ a := by + simpa only [one_pow] using pow_le_pow_iff_left (zero_le_one' R) ha hn #align one_le_pow_iff_of_nonneg one_le_pow_iff_of_nonneg -lemma pow_eq_one_iff_of_nonneg {a : R} (ha : 0 ≤ a) - {n : ℕ} (hn : n ≠ 0) : a ^ n = 1 ↔ a = 1 := - ⟨fun h ↦ le_antisymm ((pow_le_one_iff_of_nonneg ha hn).mp h.le) - ((one_le_pow_iff_of_nonneg ha hn).mp h.ge), - fun h ↦ by rw [h]; exact one_pow _⟩ +theorem pow_lt_one_iff_of_nonneg (ha : 0 ≤ a) (hn : n ≠ 0) : a ^ n < 1 ↔ a < 1 := + lt_iff_lt_of_le_iff_le (one_le_pow_iff_of_nonneg ha hn) +#align pow_lt_one_iff_of_nonneg pow_lt_one_iff_of_nonneg -theorem one_lt_pow_iff_of_nonneg {a : R} (ha : 0 ≤ a) {n : ℕ} (hn : n ≠ 0) : 1 < a ^ n ↔ 1 < a := - lt_iff_lt_of_le_iff_le (pow_le_one_iff_of_nonneg ha hn) +theorem one_lt_pow_iff_of_nonneg (ha : 0 ≤ a) (hn : n ≠ 0) : 1 < a ^ n ↔ 1 < a := by + simpa only [one_pow] using pow_lt_pow_iff_left (zero_le_one' R) ha hn #align one_lt_pow_iff_of_nonneg one_lt_pow_iff_of_nonneg -theorem pow_lt_one_iff_of_nonneg {a : R} (ha : 0 ≤ a) {n : ℕ} (hn : n ≠ 0) : a ^ n < 1 ↔ a < 1 := - lt_iff_lt_of_le_iff_le (one_le_pow_iff_of_nonneg ha hn) -#align pow_lt_one_iff_of_nonneg pow_lt_one_iff_of_nonneg +lemma pow_eq_one_iff_of_nonneg (ha : 0 ≤ a) (hn : n ≠ 0) : a ^ n = 1 ↔ a = 1 := by + simpa only [one_pow] using pow_left_inj ha zero_le_one hn theorem sq_le_one_iff {a : R} (ha : 0 ≤ a) : a ^ 2 ≤ 1 ↔ a ≤ 1 := pow_le_one_iff_of_nonneg ha (Nat.succ_ne_zero _) @@ -253,19 +253,13 @@ theorem one_lt_sq_iff {a : R} (ha : 0 ≤ a) : 1 < a ^ 2 ↔ 1 < a := one_lt_pow_iff_of_nonneg ha (Nat.succ_ne_zero _) #align one_lt_sq_iff one_lt_sq_iff -@[simp] -theorem pow_left_inj {x y : R} {n : ℕ} (Hxpos : 0 ≤ x) (Hypos : 0 ≤ y) (Hnpos : 0 < n) : - x ^ n = y ^ n ↔ x = y := - (@strictMonoOn_pow R _ _ Hnpos).eq_iff_eq Hxpos Hypos -#align pow_left_inj pow_left_inj - -theorem lt_of_pow_lt_pow {a b : R} (n : ℕ) (hb : 0 ≤ b) (h : a ^ n < b ^ n) : a < b := - lt_of_not_ge fun hn => not_lt_of_ge (pow_le_pow_of_le_left hb hn _) h -#align lt_of_pow_lt_pow lt_of_pow_lt_pow +theorem lt_of_pow_lt_pow_left (n : ℕ) (hb : 0 ≤ b) (h : a ^ n < b ^ n) : a < b := + lt_of_not_ge fun hn => not_lt_of_ge (pow_le_pow_left hb hn _) h +#align lt_of_pow_lt_pow lt_of_pow_lt_pow_left -theorem le_of_pow_le_pow {a b : R} (n : ℕ) (hb : 0 ≤ b) (hn : 0 < n) (h : a ^ n ≤ b ^ n) : a ≤ b := - le_of_not_lt fun h1 => not_le_of_lt (pow_lt_pow_of_lt_left h1 hb hn) h -#align le_of_pow_le_pow le_of_pow_le_pow +theorem le_of_pow_le_pow_left (hn : n ≠ 0) (hb : 0 ≤ b) (h : a ^ n ≤ b ^ n) : a ≤ b := + le_of_not_lt fun h1 => not_le_of_lt (pow_lt_pow_left h1 hb hn) h +#align le_of_pow_le_pow le_of_pow_le_pow_left @[simp] theorem sq_eq_sq {a b : R} (ha : 0 ≤ a) (hb : 0 ≤ b) : a ^ 2 = b ^ 2 ↔ a = b := @@ -274,7 +268,7 @@ theorem sq_eq_sq {a b : R} (ha : 0 ≤ a) (hb : 0 ≤ b) : a ^ 2 = b ^ 2 ↔ a = theorem lt_of_mul_self_lt_mul_self (hb : 0 ≤ b) : a * a < b * b → a < b := by simp_rw [← sq] - exact lt_of_pow_lt_pow _ hb + exact lt_of_pow_lt_pow_left _ hb #align lt_of_mul_self_lt_mul_self lt_of_mul_self_lt_mul_self end LinearOrderedSemiring @@ -290,7 +284,7 @@ theorem pow_abs (a : R) (n : ℕ) : |a| ^ n = |a ^ n| := theorem abs_neg_one_pow (n : ℕ) : |(-1 : R) ^ n| = 1 := by rw [← pow_abs, abs_neg, abs_one, one_pow] #align abs_neg_one_pow abs_neg_one_pow -theorem abs_pow_eq_one (a : R) {n : ℕ} (h : 0 < n) : |a ^ n| = 1 ↔ |a| = 1 := by +theorem abs_pow_eq_one (a : R) {n : ℕ} (h : n ≠ 0) : |a ^ n| = 1 ↔ |a| = 1 := by convert pow_left_inj (abs_nonneg a) zero_le_one h exacts [(pow_abs _ _).symm, (one_pow _).symm] #align abs_pow_eq_one abs_pow_eq_one @@ -346,7 +340,7 @@ theorem abs_sq (x : R) : |x ^ 2| = x ^ 2 := by simpa only [sq] using abs_mul_sel theorem sq_lt_sq : x ^ 2 < y ^ 2 ↔ |x| < |y| := by simpa only [sq_abs] using - (@strictMonoOn_pow R _ _ two_pos).lt_iff_lt (abs_nonneg x) (abs_nonneg y) + (pow_left_strictMonoOn two_ne_zero).lt_iff_lt (abs_nonneg x) (abs_nonneg y) #align sq_lt_sq sq_lt_sq theorem sq_lt_sq' (h1 : -y < x) (h2 : x < y) : x ^ 2 < y ^ 2 := @@ -355,7 +349,7 @@ theorem sq_lt_sq' (h1 : -y < x) (h2 : x < y) : x ^ 2 < y ^ 2 := theorem sq_le_sq : x ^ 2 ≤ y ^ 2 ↔ |x| ≤ |y| := by simpa only [sq_abs] using - (@strictMonoOn_pow R _ _ two_pos).le_iff_le (abs_nonneg x) (abs_nonneg y) + (pow_left_strictMonoOn two_ne_zero).le_iff_le (abs_nonneg x) (abs_nonneg y) #align sq_le_sq sq_le_sq theorem sq_le_sq' (h1 : -y ≤ x) (h2 : x ≤ y) : x ^ 2 ≤ y ^ 2 := @@ -403,7 +397,7 @@ theorem one_lt_sq_iff_one_lt_abs (x : R) : 1 < x ^ 2 ↔ 1 < |x| := by #align one_lt_sq_iff_one_lt_abs one_lt_sq_iff_one_lt_abs theorem pow_four_le_pow_two_of_pow_two_le {x y : R} (h : x ^ 2 ≤ y) : x ^ 4 ≤ y ^ 2 := - (pow_mul x 2 2).symm ▸ pow_le_pow_of_le_left (sq_nonneg x) h 2 + (pow_mul x 2 2).symm ▸ pow_le_pow_left (sq_nonneg x) h 2 #align pow_four_le_pow_two_of_pow_two_le pow_four_le_pow_two_of_pow_two_le end LinearOrderedRing @@ -441,10 +435,10 @@ theorem pow_lt_pow_succ (ha : 1 < a) : a ^ n < a ^ n.succ := by exact mul_lt_right₀ _ ha (pow_ne_zero _ (zero_lt_one.trans ha).ne') #align pow_lt_pow_succ pow_lt_pow_succ -theorem pow_lt_pow₀ (ha : 1 < a) (hmn : m < n) : a ^ m < a ^ n := by +theorem pow_lt_pow_right₀ (ha : 1 < a) (hmn : m < n) : a ^ m < a ^ n := by induction' hmn with n _ ih exacts [pow_lt_pow_succ ha, lt_trans ih (pow_lt_pow_succ ha)] -#align pow_lt_pow₀ pow_lt_pow₀ +#align pow_lt_pow₀ pow_lt_pow_right₀ end LinearOrderedCommGroupWithZero diff --git a/Mathlib/Algebra/Homology/LocalCohomology.lean b/Mathlib/Algebra/Homology/LocalCohomology.lean index 48486eaa8f1de..196d87baa329b 100644 --- a/Mathlib/Algebra/Homology/LocalCohomology.lean +++ b/Mathlib/Algebra/Homology/LocalCohomology.lean @@ -145,7 +145,7 @@ variable {R : Type u} [CommRing R] /-- The functor sending a natural number `i` to the `i`-th power of the ideal `J` -/ def idealPowersDiagram (J : Ideal R) : ℕᵒᵖ ⥤ Ideal R where obj t := J ^ unop t - map w := ⟨⟨Ideal.pow_le_pow w.unop.down.down⟩⟩ + map w := ⟨⟨Ideal.pow_le_pow_right w.unop.down.down⟩⟩ #align local_cohomology.ideal_powers_diagram localCohomology.idealPowersDiagram /-- The full subcategory of all ideals with radical containing `J` -/ @@ -234,7 +234,7 @@ theorem Ideal.exists_pow_le_of_le_radical_of_fG (hIJ : I ≤ J.radical) (hJ : J. obtain ⟨k, hk⟩ := J.exists_radical_pow_le_of_fg hJ use k calc - I ^ k ≤ J.radical ^ k := Ideal.pow_mono hIJ _ + I ^ k ≤ J.radical ^ k := Ideal.pow_right_mono hIJ _ _ ≤ J := hk #align local_cohomology.ideal.exists_pow_le_of_le_radical_of_fg localCohomology.Ideal.exists_pow_le_of_le_radical_of_fG diff --git a/Mathlib/Algebra/IsPrimePow.lean b/Mathlib/Algebra/IsPrimePow.lean index e3f8d507fdcf8..4b342e70dd7aa 100644 --- a/Mathlib/Algebra/IsPrimePow.lean +++ b/Mathlib/Algebra/IsPrimePow.lean @@ -88,7 +88,7 @@ theorem isPrimePow_nat_iff_bounded (n : ℕ) : rintro ⟨p, k, hp, hk, rfl⟩ refine' ⟨p, _, k, (Nat.lt_pow_self hp.one_lt _).le, hp, hk, rfl⟩ conv => { lhs; rw [← (pow_one p)] } - exact (Nat.pow_le_iff_le_right hp.two_le).mpr hk + exact pow_le_pow_right hp.one_lt.le hk #align is_prime_pow_nat_iff_bounded isPrimePow_nat_iff_bounded instance {n : ℕ} : Decidable (IsPrimePow n) := diff --git a/Mathlib/Algebra/Module/Zlattice.lean b/Mathlib/Algebra/Module/Zlattice.lean index 172a7b74fe320..add3590eb674c 100644 --- a/Mathlib/Algebra/Module/Zlattice.lean +++ b/Mathlib/Algebra/Module/Zlattice.lean @@ -388,7 +388,7 @@ theorem Zlattice.FG : AddSubgroup.FG L := by rw [← hs, ← h_span] exact span_mono (by simp only [Subtype.range_coe_subtype, Set.setOf_mem_eq, subset_rfl])) rw [show span ℤ s = span ℤ (Set.range b) by simp [Basis.coe_mk, Subtype.range_coe_subtype]] - have : Fintype s := Set.Finite.fintype h_lind.finite + have : Fintype s := h_lind.setFinite.fintype refine Set.Finite.of_finite_image (f := ((↑) : _ → E) ∘ Zspan.quotientEquiv b) ?_ (Function.Injective.injOn (Subtype.coe_injective.comp (Zspan.quotientEquiv b).injective) _) have : Set.Finite ((Zspan.fundamentalDomain b) ∩ L) := @@ -406,7 +406,7 @@ theorem Zlattice.FG : AddSubgroup.FG L := by exact span_le.mpr h_incl · -- `span ℤ s` is finitely generated because `s` is finite rw [ker_mkQ, inf_of_le_right (span_le.mpr h_incl)] - exact fg_span (LinearIndependent.finite h_lind) + exact fg_span (LinearIndependent.setFinite h_lind) theorem Zlattice.module_finite : Module.Finite ℤ L := Module.Finite.iff_addGroup_fg.mpr ((AddGroup.fg_iff_addSubgroup_fg L).mpr (FG K hs)) diff --git a/Mathlib/Algebra/MonoidAlgebra/Basic.lean b/Mathlib/Algebra/MonoidAlgebra/Basic.lean index e8f00eb595151..4d7bb6fdf2b7e 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Basic.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Basic.lean @@ -50,8 +50,6 @@ Similarly, I attempted to just define `Multiplicative G = G` leaks through everywhere, and seems impossible to use. -/ -set_option autoImplicit true - noncomputable section @@ -118,7 +116,8 @@ theorem single_add (a : G) (b₁ b₂ : k) : single a (b₁ + b₂) = single a b Finsupp.single_add a b₁ b₂ @[simp] -theorem sum_single_index [AddCommMonoid N] {a : G} {b : k} {h : G → k → N} (h_zero : h a 0 = 0) : +theorem sum_single_index {N} [AddCommMonoid N] {a : G} {b : k} {h : G → k → N} + (h_zero : h a 0 = 0) : (single a b).sum h = h a b := Finsupp.sum_single_index h_zero @[simp] @@ -452,6 +451,17 @@ theorem single_mul_single [Mul G] {a₁ a₂ : G} {b₁ b₂ : k} : (sum_single_index (by rw [mul_zero, single_zero])) #align monoid_algebra.single_mul_single MonoidAlgebra.single_mul_single +theorem single_commute_single [Mul G] {a₁ a₂ : G} {b₁ b₂ : k} + (ha : Commute a₁ a₂) (hb : Commute b₁ b₂) : + Commute (single a₁ b₁) (single a₂ b₂) := + single_mul_single.trans <| congr_arg₂ single ha hb |>.trans single_mul_single.symm + +theorem single_commute [Mul G] {a : G} {b : k} (ha : ∀ a', Commute a a') (hb : ∀ b', Commute b b') : + ∀ f : MonoidAlgebra k G, Commute (single a b) f := + suffices AddMonoidHom.mulLeft (single a b) = AddMonoidHom.mulRight (single a b) from + FunLike.congr_fun this + addHom_ext' fun a' => AddMonoidHom.ext fun b' => single_commute_single (ha a') (hb b') + @[simp] theorem single_pow [Monoid G] {a : G} {b : k} : ∀ n : ℕ, single a b ^ n = single (a ^ n) (b ^ n) | 0 => by @@ -516,6 +526,10 @@ theorem of_injective [MulOneClass G] [Nontrivial k] : simpa using (single_eq_single_iff _ _ _ _).mp h #align monoid_algebra.of_injective MonoidAlgebra.of_injective +theorem of_commute [MulOneClass G] {a : G} (h : ∀ a', Commute a a') (f : MonoidAlgebra k G) : + Commute (of k G a) f := + single_commute h Commute.one_left f + /-- `Finsupp.single` as a `MonoidHom` from the product type into the monoid algebra. Note the order of the elements of the product are reversed compared to the arguments of @@ -729,10 +743,8 @@ section Algebra -- attribute [local reducible] MonoidAlgebra -- Porting note: `reducible` cannot be `local`. theorem single_one_comm [CommSemiring k] [MulOneClass G] (r : k) (f : MonoidAlgebra k G) : - single (1 : G) r * f = f * single (1 : G) r := by - -- Porting note: `ext` → `refine Finsupp.ext fun _ => ?_` - refine Finsupp.ext fun _ => ?_ - rw [single_one_mul_apply, mul_single_one_apply, mul_comm] + single (1 : G) r * f = f * single (1 : G) r := + single_commute Commute.one_left (Commute.all _) f #align monoid_algebra.single_one_comm MonoidAlgebra.single_one_comm /-- `Finsupp.single 1` as a `RingHom` -/ @@ -1235,7 +1247,8 @@ theorem single_add (a : G) (b₁ b₂ : k) : single a (b₁ + b₂) = single a b Finsupp.single_add a b₁ b₂ @[simp] -theorem sum_single_index [AddCommMonoid N] {a : G} {b : k} {h : G → k → N} (h_zero : h a 0 = 0) : +theorem sum_single_index {N} [AddCommMonoid N] {a : G} {b : k} {h : G → k → N} + (h_zero : h a 0 = 0) : (single a b).sum h = h a b := Finsupp.sum_single_index h_zero @[simp] @@ -1554,6 +1567,11 @@ theorem single_mul_single [Add G] {a₁ a₂ : G} {b₁ b₂ : k} : @MonoidAlgebra.single_mul_single k (Multiplicative G) _ _ _ _ _ _ #align add_monoid_algebra.single_mul_single AddMonoidAlgebra.single_mul_single +theorem single_commute_single [Add G] {a₁ a₂ : G} {b₁ b₂ : k} + (ha : AddCommute a₁ a₂) (hb : Commute b₁ b₂) : + Commute (single a₁ b₁) (single a₂ b₂) := + @MonoidAlgebra.single_commute_single k (Multiplicative G) _ _ _ _ _ _ ha hb + -- This should be a `@[simp]` lemma, but the simp_nf linter times out if we add this. -- Probably the correct fix is to make a `[Add]MonoidAlgebra.single` with the correct type, -- instead of relying on `Finsupp.single`. @@ -1632,6 +1650,11 @@ theorem of_injective [Nontrivial k] [AddZeroClass G] : Function.Injective (of k MonoidAlgebra.of_injective #align add_monoid_algebra.of_injective AddMonoidAlgebra.of_injective +theorem of'_commute [Semiring k] [AddZeroClass G] {a : G} (h : ∀ a', AddCommute a a') + (f : AddMonoidAlgebra k G) : + Commute (of' k G a) f := + MonoidAlgebra.of_commute (G := Multiplicative G) h f + /-- `Finsupp.single` as a `MonoidHom` from the product type into the additive monoid algebra. Note the order of the elements of the product are reversed compared to the arguments of diff --git a/Mathlib/Algebra/Order/Archimedean.lean b/Mathlib/Algebra/Order/Archimedean.lean index 52931a45ca0b0..002d784a47374 100644 --- a/Mathlib/Algebra/Order/Archimedean.lean +++ b/Mathlib/Algebra/Order/Archimedean.lean @@ -51,7 +51,7 @@ instance OrderDual.archimedean [OrderedAddCommGroup α] [Archimedean α] : Archi theorem exists_lt_nsmul [OrderedAddCommMonoid M] [Archimedean M] [CovariantClass M M (· + ·) (· < ·)] {a : M} (ha : 0 < a) (b : M) : ∃ n : ℕ, b < n • a := - let ⟨k, hk⟩ := Archimedean.arch b ha; ⟨k + 1, hk.trans_lt <| nsmul_lt_nsmul ha k.lt_succ_self⟩ + let ⟨k, hk⟩ := Archimedean.arch b ha; ⟨k + 1, hk.trans_lt $ nsmul_lt_nsmul_left ha k.lt_succ_self⟩ section LinearOrderedAddCommGroup diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 22490b49341c4..078231becfd77 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -613,13 +613,14 @@ theorem one_div_strictAntiOn : StrictAntiOn (fun x : α => 1 / x) (Set.Ioi 0) := theorem one_div_pow_le_one_div_pow_of_le (a1 : 1 ≤ a) {m n : ℕ} (mn : m ≤ n) : 1 / a ^ n ≤ 1 / a ^ m := by - refine' (one_div_le_one_div _ _).mpr (pow_le_pow a1 mn) <;> + refine' (one_div_le_one_div _ _).mpr (pow_le_pow_right a1 mn) <;> exact pow_pos (zero_lt_one.trans_le a1) _ #align one_div_pow_le_one_div_pow_of_le one_div_pow_le_one_div_pow_of_le theorem one_div_pow_lt_one_div_pow_of_lt (a1 : 1 < a) {m n : ℕ} (mn : m < n) : 1 / a ^ n < 1 / a ^ m := by - refine' (one_div_lt_one_div _ _).mpr (pow_lt_pow a1 mn) <;> exact pow_pos (zero_lt_one.trans a1) _ + refine (one_div_lt_one_div ?_ ?_).2 (pow_lt_pow_right a1 mn) <;> + exact pow_pos (zero_lt_one.trans a1) _ #align one_div_pow_lt_one_div_pow_of_lt one_div_pow_lt_one_div_pow_of_lt theorem one_div_pow_anti (a1 : 1 ≤ a) : Antitone fun n : ℕ => 1 / a ^ n := fun _ _ => diff --git a/Mathlib/Algebra/Order/Group/Defs.lean b/Mathlib/Algebra/Order/Group/Defs.lean index 0a29356297ca9..50d8582b69997 100644 --- a/Mathlib/Algebra/Order/Group/Defs.lean +++ b/Mathlib/Algebra/Order/Group/Defs.lean @@ -705,7 +705,7 @@ section Right variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d : α} -@[to_additive (attr := simp)] +@[to_additive] theorem div_le_div_iff_right (c : α) : a / c ≤ b / c ↔ a ≤ b := by simpa only [div_eq_mul_inv] using mul_le_mul_iff_right _ #align div_le_div_iff_right div_le_div_iff_right @@ -727,7 +727,7 @@ alias ⟨le_of_sub_nonneg, sub_nonneg_of_le⟩ := sub_nonneg #align sub_nonneg_of_le sub_nonneg_of_le #align le_of_sub_nonneg le_of_sub_nonneg -@[to_additive (attr := simp) sub_nonpos] +@[to_additive sub_nonpos] theorem div_le_one' : a / b ≤ 1 ↔ a ≤ b := by rw [← mul_le_mul_iff_right b, one_mul, div_eq_mul_inv, inv_mul_cancel_right] #align div_le_one' div_le_one' @@ -753,6 +753,10 @@ theorem div_le_iff_le_mul : a / c ≤ b ↔ a ≤ b * c := by #align div_le_iff_le_mul div_le_iff_le_mul #align sub_le_iff_le_add sub_le_iff_le_add +-- Note: we intentionally don't have `@[simp]` for the additive version, +-- since the LHS simplifies with `tsub_le_iff_right` +attribute [simp] div_le_iff_le_mul + -- TODO: Should we get rid of `sub_le_iff_le_add` in favor of -- (a renamed version of) `tsub_le_iff_right`? -- see Note [lower instance priority] @@ -769,7 +773,7 @@ variable [CovariantClass α α (· * ·) (· ≤ ·)] variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} -@[to_additive (attr := simp)] +@[to_additive] theorem div_le_div_iff_left (a : α) : a / b ≤ a / c ↔ c ≤ b := by rw [div_eq_mul_inv, div_eq_mul_inv, ← mul_le_mul_iff_left a⁻¹, inv_mul_cancel_left, inv_mul_cancel_left, inv_le_inv_iff] @@ -1070,7 +1074,7 @@ theorem div_le_inv_mul_iff [CovariantClass α α (swap (· * ·)) (· ≤ ·)] : -- What is the point of this lemma? See comment about `div_le_inv_mul_iff` above. -- Note: we intentionally don't have `@[simp]` for the additive version, -- since the LHS simplifies with `tsub_le_iff_right` -@[to_additive, simp] +@[to_additive] theorem div_le_div_flip {α : Type*} [CommGroup α] [LinearOrder α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} : a / b ≤ b / a ↔ a ≤ b := by rw [div_eq_mul_inv b, mul_comm] diff --git a/Mathlib/Algebra/Order/Interval.lean b/Mathlib/Algebra/Order/Interval.lean index c10dda5864a46..f314bb0db2ab0 100644 --- a/Mathlib/Algebra/Order/Interval.lean +++ b/Mathlib/Algebra/Order/Interval.lean @@ -230,7 +230,7 @@ end Mul -- TODO: if `to_additive` gets improved sufficiently, derive this from `hasPow` instance NonemptyInterval.hasNSMul [AddMonoid α] [Preorder α] [CovariantClass α α (· + ·) (· ≤ ·)] [CovariantClass α α (swap (· + ·)) (· ≤ ·)] : SMul ℕ (NonemptyInterval α) := - ⟨fun n s => ⟨(n • s.fst, n • s.snd), nsmul_le_nsmul_of_le_right s.fst_le_snd _⟩⟩ + ⟨fun n s => ⟨(n • s.fst, n • s.snd), nsmul_le_nsmul_right s.fst_le_snd _⟩⟩ #align nonempty_interval.has_nsmul NonemptyInterval.hasNSMul section Pow @@ -240,7 +240,7 @@ variable [Monoid α] [Preorder α] [CovariantClass α α (· * ·) (· ≤ ·)] @[to_additive existing] instance NonemptyInterval.hasPow : Pow (NonemptyInterval α) ℕ := - ⟨fun s n => ⟨s.toProd ^ n, pow_le_pow_of_le_left' s.fst_le_snd _⟩⟩ + ⟨fun s n => ⟨s.toProd ^ n, pow_le_pow_left' s.fst_le_snd _⟩⟩ #align nonempty_interval.has_pow NonemptyInterval.hasPow namespace NonemptyInterval diff --git a/Mathlib/Algebra/Order/Module.lean b/Mathlib/Algebra/Order/Module.lean index 9293e492bc803..f8c47a176922a 100644 --- a/Mathlib/Algebra/Order/Module.lean +++ b/Mathlib/Algebra/Order/Module.lean @@ -12,6 +12,10 @@ import Mathlib.Algebra.Order.SMul In this file we provide lemmas about `OrderedSMul` that hold once a module structure is present. +## TODO + +Generalise lemmas to the framework from `Mathlib.Algebra.Order.Module.Defs`. + ## References * https://en.wikipedia.org/wiki/Ordered_vector_space @@ -65,12 +69,6 @@ theorem lt_of_smul_lt_smul_of_nonpos (h : c • a < c • b) (hc : c ≤ 0) : b exact lt_of_smul_lt_smul_of_nonneg h (neg_nonneg_of_nonpos hc) #align lt_of_smul_lt_smul_of_nonpos lt_of_smul_lt_smul_of_nonpos -lemma smul_le_smul_of_nonneg_right (h : c ≤ d) (hb : 0 ≤ b) : c • b ≤ d • b := by - rw [← sub_nonneg, ← sub_smul]; exact smul_nonneg (sub_nonneg.2 h) hb - -lemma smul_le_smul (hcd : c ≤ d) (hab : a ≤ b) (hc : 0 ≤ c) (hb : 0 ≤ b) : c • a ≤ d • b := - (smul_le_smul_of_nonneg_left hab hc).trans $ smul_le_smul_of_nonneg_right hcd hb - theorem smul_lt_smul_iff_of_neg (hc : c < 0) : c • a < c • b ↔ b < a := by rw [← neg_neg c, neg_smul, neg_smul (-c), neg_lt_neg_iff] exact smul_lt_smul_iff_of_pos (neg_pos_of_neg hc) @@ -86,23 +84,16 @@ theorem smul_pos_iff_of_neg (hc : c < 0) : 0 < c • a ↔ a < 0 := by exact smul_neg_iff_of_pos (neg_pos_of_neg hc) #align smul_pos_iff_of_neg smul_pos_iff_of_neg -theorem smul_nonpos_of_nonpos_of_nonneg (hc : c ≤ 0) (ha : 0 ≤ a) : c • a ≤ 0 := - calc - c • a ≤ c • (0 : M) := smul_le_smul_of_nonpos ha hc - _ = 0 := smul_zero c #align smul_nonpos_of_nonpos_of_nonneg smul_nonpos_of_nonpos_of_nonneg theorem smul_nonneg_of_nonpos_of_nonpos (hc : c ≤ 0) (ha : a ≤ 0) : 0 ≤ c • a := - @smul_nonpos_of_nonpos_of_nonneg k Mᵒᵈ _ _ _ _ _ _ hc ha + smul_nonpos_of_nonpos_of_nonneg (β := Mᵒᵈ) hc ha #align smul_nonneg_of_nonpos_of_nonpos smul_nonneg_of_nonpos_of_nonpos alias ⟨_, smul_pos_of_neg_of_neg⟩ := smul_pos_iff_of_neg #align smul_pos_of_neg_of_neg smul_pos_of_neg_of_neg -alias ⟨_, smul_neg_of_pos_of_neg⟩ := smul_neg_iff_of_pos #align smul_neg_of_pos_of_neg smul_neg_of_pos_of_neg - -alias ⟨_, smul_neg_of_neg_of_pos⟩ := smul_neg_iff_of_neg #align smul_neg_of_neg_of_pos smul_neg_of_neg_of_pos theorem antitone_smul_left (hc : c ≤ 0) : Antitone (SMul.smul c : M → M) := fun _ _ h => diff --git a/Mathlib/Algebra/Order/Module/Defs.lean b/Mathlib/Algebra/Order/Module/Defs.lean new file mode 100644 index 0000000000000..540341d9ebf36 --- /dev/null +++ b/Mathlib/Algebra/Order/Module/Defs.lean @@ -0,0 +1,891 @@ +/- +Copyright (c) 2023 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Algebra.Order.Module.Synonym +import Mathlib.Algebra.Order.Ring.Lemmas + +/-! +# Monotonicity of scalar multiplication by positive elements + +This file defines typeclasses to reason about monotonicity of the operations +* `b ↦ a • b`, "left scalar multiplication" +* `a ↦ a • b`, "right scalar multiplication" + +We use eight typeclasses to encode the various properties we care about for those two operations. +These typeclasses are meant to be mostly internal to this file, to set up each lemma in the +appropriate generality. + +Less granular typeclasses like `OrderedAddCommMonoid`, `LinearOrderedField`, `OrderedSMul` should be +enough for most purposes, and the system is set up so that they imply the correct granular +typeclasses here. If those are enough for you, you may stop reading here! Else, beware that what +follows is a bit technical. + +## Definitions + +In all that follows, `α` and `β` are orders which have a `0` and such that `α` acts on `β` by scalar +multiplication. Note however that we do not use lawfulness of this action in most of the file. Hence +`•` should be considered here as a mostly arbitrary function `α → β → β`. + +We use the following four typeclasses to reason about left scalar multiplication (`b ↦ a • b`): +* `PosSMulMono`: If `a ≥ 0`, then `b₁ ≤ b₂` implies `a • b₁ ≤ a • b₂`. +* `PosSMulStrictMono`: If `a > 0`, then `b₁ < b₂` implies `a • b₁ < a • b₂`. +* `PosSMulReflectLT`: If `a ≥ 0`, then `a • b₁ < a • b₂` implies `b₁ < b₂`. +* `PosSMulReflectLE`: If `a > 0`, then `a • b₁ ≤ a • b₂` implies `b₁ ≤ b₂`. + +We use the following four typeclasses to reason about right scalar multiplication (`a ↦ a • b`): +* `SMulPosMono`: If `b ≥ 0`, then `a₁ ≤ a₂` implies `a₁ • b ≤ a₂ • b`. +* `SMulPosStrictMono`: If `b > 0`, then `a₁ < a₂` implies `a₁ • b < a₂ • b`. +* `SMulPosReflectLT`: If `b ≥ 0`, then `a₁ • b < a₂ • b` implies `a₁ < a₂`. +* `SMulPosReflectLE`: If `b > 0`, then `a₁ • b ≤ a₂ • b` implies `a₁ ≤ a₂`. + +## Constructors + +The four typeclasses about nonnegativity can usually be checked only on positive inputs due to their +condition becoming trivial when `a = 0` or `b = 0`. We therefore make the following constructors +available: `PosSMulMono.of_pos`, `PosSMulReflectLT.of_pos`, `SMulPosMono.of_pos`, +`SMulPosReflectLT.of_pos` + +## Implications + +As `α` and `β` get more and more structure, those typeclasses end up being equivalent. The commonly +used implications are: +* When `α`, `β` are partial orders: + * `PosSMulStrictMono → PosSMulMono` + * `SMulPosStrictMono → SMulPosMono` + * `PosSMulReflectLE → PosSMulReflectLT` + * `SMulPosReflectLE → SMulPosReflectLT` +* When `β` is a linear order: `PosSMulStrictMono → PosSMulReflectLE` +* When `α` is a linear order: `SMulPosStrictMono → SMulPosReflectLE` +* When `α` is an ordered ring, `β` an ordered group and also an `α`-module: + * `PosSMulMono → SMulPosMono` + * `PosSMulStrictMono → SMulPosStrictMono` +* When `α` is an ordered semifield, `β` is an `α`-module: + * `PosSMulStrictMono → PosSMulReflectLT` + * `PosSMulMono → PosSMulReflectLE` + +Further, the bundled non-granular typeclasses imply the granular ones like so: +* `OrderedSMul → PosSMulStrictMono` +* `OrderedSMul → PosSMulReflectLT` + +All these are registered as instances, which means that in practice you should not worry about these +implications. However, if you encounter a case where you think a statement is true but not covered +by the current implications, please bring it up on Zulip! + +## Implementation notes + +This file uses custom typeclasses instead of abbreviations of `CovariantClass`/`ContravariantClass` +because: +* They get displayed as classes in the docs. In particular, one can see their list of instances, + instead of their instances being invariably dumped to the `CovariantClass`/`ContravariantClass` + list. +* They don't pollute other typeclass searches. Having many abbreviations of the same typeclass for + different purposes always felt like a performance issue (more instances with the same key, for no + added benefit), and indeed making the classes here abbreviation previous creates timeouts due to + the higher number of `CovariantClass`/`ContravariantClass` instances. +* `SMulPosReflectLT`/`SMulPosReflectLE` do not fit in the framework since they relate `≤` on two + different types. So we would have to generalise `CovariantClass`/`ContravariantClass` to three + types and two relations. +* Very minor, but the constructors let you work with `a : α`, `h : 0 ≤ a` instead of + `a : {a : α // 0 ≤ a}`. This actually makes some instances surprisingly cleaner to prove. +* The `CovariantClass`/`ContravariantClass` framework is only useful to automate very simple logic + anyway. It is easily copied over. + +In the future, it would be good to make the corresponding typeclasses in +`Mathlib.Algebra.Order.Ring.Lemmas` custom typeclasses too. + +## TODO + +This file acts as a substitute for `Mathlib.Algebra.Order.SMul`. We now need to +* finish the transition by deleting the duplicate lemmas +* rearrange the non-duplicate lemmas into new files +* generalise (most of) the lemmas from `Mathlib.Algebra.Order.Module` to here +* rethink `OrderedSMul` +-/ + +variable (α β : Type*) + +section Defs +variable [SMul α β] [Preorder α] [Preorder β] + +section Left +variable [Zero α] + +/-- Typeclass for monotonicity of scalar multiplication by nonnegative elements on the left, +namely `b₁ ≤ b₂ → a • b₁ ≤ a • b₂` if `0 ≤ a`. + +You should usually not use this very granular typeclass directly, but rather a typeclass like +`OrderedSMul`. -/ +class PosSMulMono : Prop where + /-- Do not use this. Use `smul_le_smul_of_nonneg_left` instead. -/ + protected elim ⦃a : α⦄ (ha : 0 ≤ a) ⦃b₁ b₂ : β⦄ (hb : b₁ ≤ b₂) : a • b₁ ≤ a • b₂ + +/-- Typeclass for strict monotonicity of scalar multiplication by positive elements on the left, +namely `b₁ < b₂ → a • b₁ < a • b₂` if `0 < a`. + +You should usually not use this very granular typeclass directly, but rather a typeclass like +`OrderedSMul`. -/ +class PosSMulStrictMono : Prop where + /-- Do not use this. Use `smul_lt_smul_of_pos_left` instead. -/ + protected elim ⦃a : α⦄ (ha : 0 < a) ⦃b₁ b₂ : β⦄ (hb : b₁ < b₂) : a • b₁ < a • b₂ + +/-- Typeclass for strict reverse monotonicity of scalar multiplication by nonnegative elements on +the left, namely `a • b₁ < a • b₂ → b₁ < b₂` if `0 ≤ a`. + +You should usually not use this very granular typeclass directly, but rather a typeclass like +`OrderedSMul`. -/ +class PosSMulReflectLT : Prop where + /-- Do not use this. Use `lt_of_smul_lt_smul_left` instead. -/ + protected elim ⦃a : α⦄ (ha : 0 ≤ a) ⦃b₁ b₂ : β⦄ (hb : a • b₁ < a • b₂) : b₁ < b₂ + +/-- Typeclass for reverse monotonicity of scalar multiplication by positive elements on the left, +namely `a • b₁ ≤ a • b₂ → b₁ ≤ b₂` if `0 < a`. + +You should usually not use this very granular typeclass directly, but rather a typeclass like +`OrderedSMul`. -/ +class PosSMulReflectLE : Prop where + /-- Do not use this. Use `le_of_smul_lt_smul_left` instead. -/ + protected elim ⦃a : α⦄ (ha : 0 < a) ⦃b₁ b₂ : β⦄ (hb : a • b₁ ≤ a • b₂) : b₁ ≤ b₂ + +end Left + +section Right +variable [Zero β] + +/-- Typeclass for monotonicity of scalar multiplication by nonnegative elements on the left, +namely `a₁ ≤ a₂ → a₁ • b ≤ a₂ • b` if `0 ≤ b`. + +You should usually not use this very granular typeclass directly, but rather a typeclass like +`OrderedSMul`. -/ +class SMulPosMono : Prop where + /-- Do not use this. Use `smul_le_smul_of_nonneg_right` instead. -/ + protected elim ⦃b : β⦄ (hb : 0 ≤ b) ⦃a₁ a₂ : α⦄ (ha : a₁ ≤ a₂) : a₁ • b ≤ a₂ • b + +/-- Typeclass for strict monotonicity of scalar multiplication by positive elements on the left, +namely `a₁ < a₂ → a₁ • b < a₂ • b` if `0 < b`. + +You should usually not use this very granular typeclass directly, but rather a typeclass like +`OrderedSMul`. -/ +class SMulPosStrictMono : Prop where + /-- Do not use this. Use `smul_lt_smul_of_pos_right` instead. -/ + protected elim ⦃b : β⦄ (hb : 0 < b) ⦃a₁ a₂ : α⦄ (ha : a₁ < a₂) : a₁ • b < a₂ • b + +/-- Typeclass for strict reverse monotonicity of scalar multiplication by nonnegative elements on +the left, namely `a₁ • b < a₂ • b → a₁ < a₂` if `0 ≤ b`. + +You should usually not use this very granular typeclass directly, but rather a typeclass like +`OrderedSMul`. -/ +class SMulPosReflectLT : Prop where + /-- Do not use this. Use `lt_of_smul_lt_smul_right` instead. -/ + protected elim ⦃b : β⦄ (hb : 0 ≤ b) ⦃a₁ a₂ : α⦄ (hb : a₁ • b < a₂ • b) : a₁ < a₂ + +/-- Typeclass for reverse monotonicity of scalar multiplication by positive elements on the left, +namely `a₁ • b ≤ a₂ • b → a₁ ≤ a₂` if `0 < b`. + +You should usually not use this very granular typeclass directly, but rather a typeclass like +`OrderedSMul`. -/ +class SMulPosReflectLE : Prop where + /-- Do not use this. Use `le_of_smul_lt_smul_right` instead. -/ + protected elim ⦃b : β⦄ (hb : 0 < b) ⦃a₁ a₂ : α⦄ (hb : a₁ • b ≤ a₂ • b) : a₁ ≤ a₂ + +end Right +end Defs + +variable {α β} {a a₁ a₂ : α} {b b₁ b₂ : β} + +section Mul +variable [Zero α] [Mul α] [Preorder α] + +-- See note [lower instance priority] +instance (priority := 100) PosMulMono.toPosSMulMono [PosMulMono α] : PosSMulMono α α where + elim _a ha _b₁ _b₂ hb := mul_le_mul_of_nonneg_left hb ha + +-- See note [lower instance priority] +instance (priority := 100) PosMulStrictMono.toPosSMulStrictMono [PosMulStrictMono α] : + PosSMulStrictMono α α where + elim _a ha _b₁ _b₂ hb := mul_lt_mul_of_pos_left hb ha + +-- See note [lower instance priority] +instance (priority := 100) PosMulReflectLT.toPosSMulReflectLT [PosMulReflectLT α] : + PosSMulReflectLT α α where + elim _a ha _b₁ _b₂ h := lt_of_mul_lt_mul_left h ha + +-- See note [lower instance priority] +instance (priority := 100) PosMulReflectLE.toPosSMulReflectLE [PosMulReflectLE α] : + PosSMulReflectLE α α where + elim _a ha _b₁ _b₂ h := le_of_mul_le_mul_left h ha + +-- See note [lower instance priority] +instance (priority := 100) MulPosMono.toSMulPosMono [MulPosMono α] : SMulPosMono α α where + elim _b hb _a₁ _a₂ ha := mul_le_mul_of_nonneg_right ha hb + +-- See note [lower instance priority] +instance (priority := 100) MulPosStrictMono.toSMulPosStrictMono [MulPosStrictMono α] : + SMulPosStrictMono α α where + elim _b hb _a₁ _a₂ ha := mul_lt_mul_of_pos_right ha hb + +-- See note [lower instance priority] +instance (priority := 100) MulPosReflectLT.toSMulPosReflectLT [MulPosReflectLT α] : + SMulPosReflectLT α α where + elim _b hb _a₁ _a₂ h := lt_of_mul_lt_mul_right h hb + +-- See note [lower instance priority] +instance (priority := 100) MulPosReflectLE.toSMulPosReflectLE [MulPosReflectLE α] : + SMulPosReflectLE α α where + elim _b hb _a₁ _a₂ h := le_of_mul_le_mul_right h hb + +end Mul + +section SMul +variable [SMul α β] + +section Preorder +variable [Preorder α] [Preorder β] + +section Left +variable [Zero α] + +lemma monotone_smul_left_of_nonneg [PosSMulMono α β] (ha : 0 ≤ a) : Monotone ((a • ·) : β → β) := + PosSMulMono.elim ha + +lemma strictMono_smul_left_of_pos [PosSMulStrictMono α β] (ha : 0 < a) : + StrictMono ((a • ·) : β → β) := PosSMulStrictMono.elim ha + +lemma smul_le_smul_of_nonneg_left [PosSMulMono α β] (hb : b₁ ≤ b₂) (ha : 0 ≤ a) : a • b₁ ≤ a • b₂ := + monotone_smul_left_of_nonneg ha hb + +lemma smul_lt_smul_of_pos_left [PosSMulStrictMono α β] (hb : b₁ < b₂) (ha : 0 < a) : + a • b₁ < a • b₂ := strictMono_smul_left_of_pos ha hb + +lemma lt_of_smul_lt_smul_left [PosSMulReflectLT α β] (h : a • b₁ < a • b₂) (ha : 0 ≤ a) : b₁ < b₂ := + PosSMulReflectLT.elim ha h + +lemma le_of_smul_le_smul_left [PosSMulReflectLE α β] (h : a • b₁ ≤ a • b₂) (ha : 0 < a) : b₁ ≤ b₂ := + PosSMulReflectLE.elim ha h + +alias lt_of_smul_lt_smul_of_nonneg_left := lt_of_smul_lt_smul_left +alias le_of_smul_le_smul_of_pos_left := le_of_smul_le_smul_left + +@[simp] +lemma smul_le_smul_iff_of_pos_left [PosSMulMono α β] [PosSMulReflectLE α β] (ha : 0 < a) : + a • b₁ ≤ a • b₂ ↔ b₁ ≤ b₂ := + ⟨fun h ↦ le_of_smul_le_smul_left h ha, fun h ↦ smul_le_smul_of_nonneg_left h ha.le⟩ + +@[simp] +lemma smul_lt_smul_iff_of_pos_left [PosSMulStrictMono α β] [PosSMulReflectLT α β] (ha : 0 < a) : + a • b₁ < a • b₂ ↔ b₁ < b₂ := + ⟨fun h ↦ lt_of_smul_lt_smul_left h ha.le, fun hb ↦ smul_lt_smul_of_pos_left hb ha⟩ + +end Left + +section Right +variable [Zero β] + +lemma monotone_smul_right_of_nonneg [SMulPosMono α β] (hb : 0 ≤ b) : Monotone ((· • b) : α → β) := + SMulPosMono.elim hb + +lemma strictMono_smul_right_of_pos [SMulPosStrictMono α β] (hb : 0 < b) : + StrictMono ((· • b) : α → β) := SMulPosStrictMono.elim hb + +lemma smul_le_smul_of_nonneg_right [SMulPosMono α β] (ha : a₁ ≤ a₂) (hb : 0 ≤ b) : + a₁ • b ≤ a₂ • b := monotone_smul_right_of_nonneg hb ha + +lemma smul_lt_smul_of_pos_right [SMulPosStrictMono α β] (ha : a₁ < a₂) (hb : 0 < b) : + a₁ • b < a₂ • b := strictMono_smul_right_of_pos hb ha + +lemma lt_of_smul_lt_smul_right [SMulPosReflectLT α β] (h : a₁ • b < a₂ • b) (hb : 0 ≤ b) : + a₁ < a₂ := SMulPosReflectLT.elim hb h + +lemma le_of_smul_le_smul_right [SMulPosReflectLE α β] (h : a₁ • b ≤ a₂ • b) (hb : 0 < b) : + a₁ ≤ a₂ := SMulPosReflectLE.elim hb h + +alias lt_of_smul_lt_smul_of_nonneg_right := lt_of_smul_lt_smul_right +alias le_of_smul_le_smul_of_pos_right := le_of_smul_le_smul_right + +@[simp] +lemma smul_le_smul_iff_of_pos_right [SMulPosMono α β] [SMulPosReflectLE α β] (hb : 0 < b) : + a₁ • b ≤ a₂ • b ↔ a₁ ≤ a₂ := + ⟨fun h ↦ le_of_smul_le_smul_right h hb, fun ha ↦ smul_le_smul_of_nonneg_right ha hb.le⟩ + +@[simp] +lemma smul_lt_smul_iff_of_pos_right [SMulPosStrictMono α β] [SMulPosReflectLT α β] (hb : 0 < b) : + a₁ • b < a₂ • b ↔ a₁ < a₂ := + ⟨fun h ↦ lt_of_smul_lt_smul_right h hb.le, fun ha ↦ smul_lt_smul_of_pos_right ha hb⟩ + +end Right + +section LeftRight +variable [Zero α] [Zero β] + +lemma smul_lt_smul_of_le_of_lt [PosSMulStrictMono α β] [SMulPosMono α β] (ha : a₁ ≤ a₂) + (hb : b₁ < b₂) (h₁ : 0 < a₁) (h₂ : 0 ≤ b₂) : a₁ • b₁ < a₂ • b₂ := + (smul_lt_smul_of_pos_left hb h₁).trans_le (smul_le_smul_of_nonneg_right ha h₂) + +lemma smul_lt_smul_of_le_of_lt' [PosSMulStrictMono α β] [SMulPosMono α β] (ha : a₁ ≤ a₂) + (hb : b₁ < b₂) (h₂ : 0 < a₂) (h₁ : 0 ≤ b₁) : a₁ • b₁ < a₂ • b₂ := + (smul_le_smul_of_nonneg_right ha h₁).trans_lt (smul_lt_smul_of_pos_left hb h₂) + +lemma smul_lt_smul_of_lt_of_le [PosSMulMono α β] [SMulPosStrictMono α β] (ha : a₁ < a₂) + (hb : b₁ ≤ b₂) (h₁ : 0 ≤ a₁) (h₂ : 0 < b₂) : a₁ • b₁ < a₂ • b₂ := + (smul_le_smul_of_nonneg_left hb h₁).trans_lt (smul_lt_smul_of_pos_right ha h₂) + +lemma smul_lt_smul_of_lt_of_le' [PosSMulMono α β] [SMulPosStrictMono α β] (ha : a₁ < a₂) + (hb : b₁ ≤ b₂) (h₂ : 0 ≤ a₂) (h₁ : 0 < b₁) : a₁ • b₁ < a₂ • b₂ := + (smul_lt_smul_of_pos_right ha h₁).trans_le (smul_le_smul_of_nonneg_left hb h₂) + +lemma smul_lt_smul [PosSMulStrictMono α β] [SMulPosStrictMono α β] (ha : a₁ < a₂) (hb : b₁ < b₂) + (h₁ : 0 < a₁) (h₂ : 0 < b₂) : a₁ • b₁ < a₂ • b₂ := + (smul_lt_smul_of_pos_left hb h₁).trans (smul_lt_smul_of_pos_right ha h₂) + +lemma smul_lt_smul' [PosSMulStrictMono α β] [SMulPosStrictMono α β] (ha : a₁ < a₂) (hb : b₁ < b₂) + (h₂ : 0 < a₂) (h₁ : 0 < b₁) : a₁ • b₁ < a₂ • b₂ := + (smul_lt_smul_of_pos_right ha h₁).trans (smul_lt_smul_of_pos_left hb h₂) + +lemma smul_le_smul [PosSMulMono α β] [SMulPosMono α β] (ha : a₁ ≤ a₂) (hb : b₁ ≤ b₂) + (h₁ : 0 ≤ a₁) (h₂ : 0 ≤ b₂) : a₁ • b₁ ≤ a₂ • b₂ := + (smul_le_smul_of_nonneg_left hb h₁).trans (smul_le_smul_of_nonneg_right ha h₂) + +lemma smul_le_smul' [PosSMulMono α β] [SMulPosMono α β] (ha : a₁ ≤ a₂) (hb : b₁ ≤ b₂) (h₂ : 0 ≤ a₂) + (h₁ : 0 ≤ b₁) : a₁ • b₁ ≤ a₂ • b₂ := + (smul_le_smul_of_nonneg_right ha h₁).trans (smul_le_smul_of_nonneg_left hb h₂) + +end LeftRight +end Preorder + +section LinearOrder +variable [Preorder α] [LinearOrder β] + +section Left +variable [Zero α] + +-- See note [lower instance priority] +instance (priority := 100) PosSMulStrictMono.toPosSMulReflectLE [PosSMulStrictMono α β] : + PosSMulReflectLE α β where + elim _a ha _b₁ _b₂ := (strictMono_smul_left_of_pos ha).le_iff_le.1 + +lemma PosSMulReflectLE.toPosSMulStrictMono [PosSMulReflectLE α β] : PosSMulStrictMono α β where + elim _a ha _b₁ _b₂ hb := not_le.1 fun h ↦ hb.not_le $ le_of_smul_le_smul_left h ha + +lemma posSMulStrictMono_iff_PosSMulReflectLE : PosSMulStrictMono α β ↔ PosSMulReflectLE α β := + ⟨fun _ ↦ inferInstance, fun _ ↦ PosSMulReflectLE.toPosSMulStrictMono⟩ + +instance PosSMulMono.toPosSMulReflectLT [PosSMulMono α β] : PosSMulReflectLT α β where + elim _a ha _b₁ _b₂ := (monotone_smul_left_of_nonneg ha).reflect_lt + +lemma PosSMulReflectLT.toPosSMulMono [PosSMulReflectLT α β] : PosSMulMono α β where + elim _a ha _b₁ _b₂ hb := not_lt.1 fun h ↦ hb.not_lt $ lt_of_smul_lt_smul_left h ha + +lemma posSMulMono_iff_posSMulReflectLT : PosSMulMono α β ↔ PosSMulReflectLT α β := + ⟨fun _ ↦ PosSMulMono.toPosSMulReflectLT, fun _ ↦ PosSMulReflectLT.toPosSMulMono⟩ + +lemma smul_max_of_nonneg [PosSMulMono α β] (ha : 0 ≤ a) (b₁ b₂ : β) : + a • max b₁ b₂ = max (a • b₁) (a • b₂) := (monotone_smul_left_of_nonneg ha).map_max + +lemma smul_min_of_nonneg [PosSMulMono α β] (ha : 0 ≤ a) (b₁ b₂ : β) : + a • min b₁ b₂ = min (a • b₁) (a • b₂) := (monotone_smul_left_of_nonneg ha).map_min + +end Left + +section Right +variable [Zero β] + +lemma SMulPosReflectLE.toSMulPosStrictMono [SMulPosReflectLE α β] : SMulPosStrictMono α β where + elim _b hb _a₁ _a₂ ha := not_le.1 fun h ↦ ha.not_le $ le_of_smul_le_smul_of_pos_right h hb + +lemma SMulPosReflectLT.toSMulPosMono [SMulPosReflectLT α β] : SMulPosMono α β where + elim _b hb _a₁ _a₂ ha := not_lt.1 fun h ↦ ha.not_lt $ lt_of_smul_lt_smul_right h hb + +end Right +end LinearOrder + +section LinearOrder +variable [LinearOrder α] [Preorder β] + +section Right +variable [Zero β] + +-- See note [lower instance priority] +instance (priority := 100) SMulPosStrictMono.toSMulPosReflectLE [SMulPosStrictMono α β] : + SMulPosReflectLE α β where + elim _b hb _a₁ _a₂ h := not_lt.1 fun ha ↦ h.not_lt $ smul_lt_smul_of_pos_right ha hb + +lemma SMulPosMono.toSMulPosReflectLT [SMulPosMono α β] : SMulPosReflectLT α β where + elim _b hb _a₁ _a₂ h := not_le.1 fun ha ↦ h.not_le $ smul_le_smul_of_nonneg_right ha hb + +end Right +end LinearOrder + +section LinearOrder +variable [LinearOrder α] [LinearOrder β] + +section Right +variable [Zero β] + +lemma smulPosStrictMono_iff_SMulPosReflectLE : SMulPosStrictMono α β ↔ SMulPosReflectLE α β := + ⟨fun _ ↦ SMulPosStrictMono.toSMulPosReflectLE, fun _ ↦ SMulPosReflectLE.toSMulPosStrictMono⟩ + +lemma smulPosMono_iff_smulPosReflectLT : SMulPosMono α β ↔ SMulPosReflectLT α β := + ⟨fun _ ↦ SMulPosMono.toSMulPosReflectLT, fun _ ↦ SMulPosReflectLT.toSMulPosMono⟩ + +end Right +end LinearOrder +end SMul + +section SMulZeroClass +variable [Zero α] [Zero β] [SMulZeroClass α β] + +section Preorder +variable [Preorder α] [Preorder β] + +lemma smul_pos [PosSMulStrictMono α β] (ha : 0 < a) (hb : 0 < b) : 0 < a • b := by + simpa only [smul_zero] using smul_lt_smul_of_pos_left hb ha + +lemma smul_neg_of_pos_of_neg [PosSMulStrictMono α β] (ha : 0 < a) (hb : b < 0) : a • b < 0 := by + simpa only [smul_zero] using smul_lt_smul_of_pos_left hb ha + +@[simp] +lemma smul_pos_iff_of_pos_left [PosSMulStrictMono α β] [PosSMulReflectLT α β] (ha : 0 < a) : + 0 < a • b ↔ 0 < b := by + simpa only [smul_zero] using smul_lt_smul_iff_of_pos_left ha (b₁ := 0) (b₂ := b) + +lemma smul_nonneg [PosSMulMono α β] (ha : 0 ≤ a) (hb : 0 ≤ b₁) : 0 ≤ a • b₁ := by + simpa only [smul_zero] using smul_le_smul_of_nonneg_left hb ha + +lemma smul_nonpos_of_nonneg_of_nonpos [PosSMulMono α β] (ha : 0 ≤ a) (hb : b ≤ 0) : a • b ≤ 0 := by + simpa only [smul_zero] using smul_le_smul_of_nonneg_left hb ha + +lemma pos_of_smul_pos_right [PosSMulReflectLT α β] (h : 0 < a • b) (ha : 0 ≤ a) : 0 < b := + lt_of_smul_lt_smul_left (by rwa [smul_zero]) ha + +end Preorder +end SMulZeroClass + +section SMulWithZero +variable [Zero α] [Zero β] [SMulWithZero α β] + +section Preorder +variable [Preorder α] [Preorder β] + +lemma smul_pos' [SMulPosStrictMono α β] (ha : 0 < a) (hb : 0 < b) : 0 < a • b := by + simpa only [zero_smul] using smul_lt_smul_of_pos_right ha hb + +lemma smul_neg_of_neg_of_pos [SMulPosStrictMono α β] (ha : a < 0) (hb : 0 < b) : a • b < 0 := by + simpa only [zero_smul] using smul_lt_smul_of_pos_right ha hb + +@[simp] +lemma smul_pos_iff_of_pos_right [SMulPosStrictMono α β] [SMulPosReflectLT α β] (hb : 0 < b) : + 0 < a • b ↔ 0 < a := by + simpa only [zero_smul] using smul_lt_smul_iff_of_pos_right hb (a₁ := 0) (a₂ := a) + +lemma smul_nonneg' [SMulPosMono α β] (ha : 0 ≤ a) (hb : 0 ≤ b₁) : 0 ≤ a • b₁ := by + simpa only [zero_smul] using smul_le_smul_of_nonneg_right ha hb + +lemma smul_nonpos_of_nonpos_of_nonneg [SMulPosMono α β] (ha : a ≤ 0) (hb : 0 ≤ b) : a • b ≤ 0 := by + simpa only [zero_smul] using smul_le_smul_of_nonneg_right ha hb + +lemma pos_of_smul_pos_left [SMulPosReflectLT α β] (h : 0 < a • b) (hb : 0 ≤ b) : 0 < a := + lt_of_smul_lt_smul_right (by rwa [zero_smul]) hb + +lemma pos_iff_pos_of_smul_pos [PosSMulReflectLT α β] [SMulPosReflectLT α β] (hab : 0 < a • b) : + 0 < a ↔ 0 < b := + ⟨pos_of_smul_pos_right hab ∘ le_of_lt, pos_of_smul_pos_left hab ∘ le_of_lt⟩ + +end Preorder + +section PartialOrder +variable [PartialOrder α] [Preorder β] + +/-- A constructor for `PosSMulMono` requiring you to prove `b₁ ≤ b₂ → a • b₁ ≤ a • b₂` only when +`0 < a`-/ +lemma PosSMulMono.of_pos (h₀ : ∀ a : α, 0 < a → ∀ b₁ b₂ : β, b₁ ≤ b₂ → a • b₁ ≤ a • b₂) : + PosSMulMono α β where + elim a ha b₁ b₂ h := by + obtain ha | ha := ha.eq_or_lt + · simp [← ha] + · exact h₀ _ ha _ _ h + +/-- A constructor for `PosSMulReflectLT` requiring you to prove `a • b₁ < a • b₂ → b₁ < b₂` only +when `0 < a`-/ +lemma PosSMulReflectLT.of_pos (h₀ : ∀ a : α, 0 < a → ∀ b₁ b₂ : β, a • b₁ < a • b₂ → b₁ < b₂) : + PosSMulReflectLT α β where + elim a ha b₁ b₂ h := by + obtain ha | ha := ha.eq_or_lt + · simp [← ha] at h + · exact h₀ _ ha _ _ h + +end PartialOrder + +section PartialOrder +variable [Preorder α] [PartialOrder β] + +/-- A constructor for `SMulPosMono` requiring you to prove `a₁ ≤ a₂ → a₁ • b ≤ a₂ • b` only when +`0 < b`-/ +lemma SMulPosMono.of_pos (h₀ : ∀ b : β, 0 < b → ∀ a₁ a₂ : α, a₁ ≤ a₂ → a₁ • b ≤ a₂ • b) : + SMulPosMono α β where + elim b hb a₁ a₂ h := by + obtain hb | hb := hb.eq_or_lt + · simp [← hb] + · exact h₀ _ hb _ _ h + +/-- A constructor for `SMulPosReflectLT` requiring you to prove `a₁ • b < a₂ • b → a₁ < a₂` only +when `0 < b`-/ +lemma SMulPosReflectLT.of_pos (h₀ : ∀ b : β, 0 < b → ∀ a₁ a₂ : α, a₁ • b < a₂ • b → a₁ < a₂) : + SMulPosReflectLT α β where + elim b hb a₁ a₂ h := by + obtain hb | hb := hb.eq_or_lt + · simp [← hb] at h + · exact h₀ _ hb _ _ h + +end PartialOrder + +section PartialOrder +variable [PartialOrder α] [PartialOrder β] + +-- See note [lower instance priority] +instance (priority := 100) PosSMulStrictMono.toPosSMulMono [PosSMulStrictMono α β] : + PosSMulMono α β := + PosSMulMono.of_pos fun _a ha ↦ (strictMono_smul_left_of_pos ha).monotone + +-- See note [lower instance priority] +instance (priority := 100) SMulPosStrictMono.toSMulPosMono [SMulPosStrictMono α β] : + SMulPosMono α β := + SMulPosMono.of_pos fun _b hb ↦ (strictMono_smul_right_of_pos hb).monotone + +-- See note [lower instance priority] +instance (priority := 100) PosSMulReflectLE.toPosSMulReflectLT [PosSMulReflectLE α β] : + PosSMulReflectLT α β := + PosSMulReflectLT.of_pos fun a ha b₁ b₂ h ↦ + (le_of_smul_le_smul_of_pos_left h.le ha).lt_of_ne $ by rintro rfl; simp at h + +-- See note [lower instance priority] +instance (priority := 100) SMulPosReflectLE.toSMulPosReflectLT [SMulPosReflectLE α β] : + SMulPosReflectLT α β := + SMulPosReflectLT.of_pos fun b hb a₁ a₂ h ↦ + (le_of_smul_le_smul_of_pos_right h.le hb).lt_of_ne $ by rintro rfl; simp at h + +lemma smul_eq_smul_iff_eq_and_eq_of_pos [PosSMulStrictMono α β] [SMulPosStrictMono α β] + (ha : a₁ ≤ a₂) (hb : b₁ ≤ b₂) (h₁ : 0 < a₁) (h₂ : 0 < b₂) : + a₁ • b₁ = a₂ • b₂ ↔ a₁ = a₂ ∧ b₁ = b₂ := by + refine ⟨fun h ↦ ?_, by rintro ⟨rfl, rfl⟩; rfl⟩ + simp only [eq_iff_le_not_lt, ha, hb, true_and] + refine ⟨fun ha ↦ h.not_lt ?_, fun hb ↦ h.not_lt ?_⟩ + · exact (smul_le_smul_of_nonneg_left hb h₁.le).trans_lt (smul_lt_smul_of_pos_right ha h₂) + · exact (smul_lt_smul_of_pos_left hb h₁).trans_le (smul_le_smul_of_nonneg_right ha h₂.le) + +lemma smul_eq_smul_iff_eq_and_eq_of_pos' [PosSMulStrictMono α β] [SMulPosStrictMono α β] + (ha : a₁ ≤ a₂) (hb : b₁ ≤ b₂) (h₂ : 0 < a₂) (h₁ : 0 < b₁) : + a₁ • b₁ = a₂ • b₂ ↔ a₁ = a₂ ∧ b₁ = b₂ := by + refine ⟨fun h ↦ ?_, by rintro ⟨rfl, rfl⟩; rfl⟩ + simp only [eq_iff_le_not_lt, ha, hb, true_and] + refine ⟨fun ha ↦ h.not_lt ?_, fun hb ↦ h.not_lt ?_⟩ + · exact (smul_lt_smul_of_pos_right ha h₁).trans_le (smul_le_smul_of_nonneg_left hb h₂.le) + · exact (smul_le_smul_of_nonneg_right ha h₁.le).trans_lt (smul_lt_smul_of_pos_left hb h₂) + +end PartialOrder + +section LinearOrder +variable [LinearOrder α] [LinearOrder β] + +lemma pos_and_pos_or_neg_and_neg_of_smul_pos [PosSMulMono α β] [SMulPosMono α β] (hab : 0 < a • b) : + 0 < a ∧ 0 < b ∨ a < 0 ∧ b < 0 := by + obtain ha | rfl | ha := lt_trichotomy a 0 + · refine' Or.inr ⟨ha, lt_imp_lt_of_le_imp_le (fun hb ↦ _) hab⟩ + exact smul_nonpos_of_nonpos_of_nonneg ha.le hb + · rw [zero_smul] at hab + exact hab.false.elim + · refine' Or.inl ⟨ha, lt_imp_lt_of_le_imp_le (fun hb ↦ _) hab⟩ + exact smul_nonpos_of_nonneg_of_nonpos ha.le hb + +lemma neg_of_smul_pos_right [PosSMulMono α β] [SMulPosMono α β] (h : 0 < a • b) (ha : a ≤ 0) : + b < 0 := ((pos_and_pos_or_neg_and_neg_of_smul_pos h).resolve_left fun h ↦ h.1.not_le ha).2 + +lemma neg_of_smul_pos_left [PosSMulMono α β] [SMulPosMono α β] (h : 0 < a • b) (ha : b ≤ 0) : + a < 0 := ((pos_and_pos_or_neg_and_neg_of_smul_pos h).resolve_left fun h ↦ h.2.not_le ha).1 + +lemma neg_iff_neg_of_smul_pos [PosSMulMono α β] [SMulPosMono α β] (hab : 0 < a • b) : + a < 0 ↔ b < 0 := + ⟨neg_of_smul_pos_right hab ∘ le_of_lt, neg_of_smul_pos_left hab ∘ le_of_lt⟩ + +lemma neg_of_smul_neg_left [PosSMulMono α β] (h : a • b < 0) (ha : 0 ≤ a) : b < 0 := + lt_of_not_ge fun hb ↦ (smul_nonneg ha hb).not_lt h + +lemma neg_of_smul_neg_left' [SMulPosMono α β] (h : a • b < 0) (ha : 0 ≤ a) : b < 0 := + lt_of_not_ge fun hb ↦ (smul_nonneg' ha hb).not_lt h + +lemma neg_of_smul_neg_right [PosSMulMono α β] (h : a • b < 0) (hb : 0 ≤ b) : a < 0 := + lt_of_not_ge fun ha ↦ (smul_nonneg ha hb).not_lt h + +lemma neg_of_smul_neg_right' [SMulPosMono α β] (h : a • b < 0) (hb : 0 ≤ b) : a < 0 := + lt_of_not_ge fun ha ↦ (smul_nonneg' ha hb).not_lt h + +end LinearOrder +end SMulWithZero + +section MulAction +variable [Monoid α] [Zero α] [Zero β] [MulAction α β] + +section Preorder +variable [Preorder α] [Preorder β] + +@[simp] +lemma le_smul_iff_one_le_left [SMulPosMono α β] [SMulPosReflectLE α β] (hb : 0 < b) : + b ≤ a • b ↔ 1 ≤ a := Iff.trans (by rw [one_smul]) (smul_le_smul_iff_of_pos_right hb) + +@[simp] +lemma lt_smul_iff_one_lt_left [SMulPosStrictMono α β] [SMulPosReflectLT α β] (hb : 0 < b) : + b < a • b ↔ 1 < a := Iff.trans (by rw [one_smul]) (smul_lt_smul_iff_of_pos_right hb) + +@[simp] +lemma smul_le_iff_le_one_left [SMulPosMono α β] [SMulPosReflectLE α β] (hb : 0 < b) : + a • b ≤ b ↔ a ≤ 1 := Iff.trans (by rw [one_smul]) (smul_le_smul_iff_of_pos_right hb) + +@[simp] +lemma smul_lt_iff_lt_one_left [SMulPosStrictMono α β] [SMulPosReflectLT α β] (hb : 0 < b) : + a • b < b ↔ a < 1 := Iff.trans (by rw [one_smul]) (smul_lt_smul_iff_of_pos_right hb) + +lemma smul_le_of_le_one_left [SMulPosMono α β] (hb : 0 ≤ b) (h : a ≤ 1) : a • b ≤ b := by + simpa only [one_smul] using smul_le_smul_of_nonneg_right h hb + +lemma le_smul_of_one_le_left [SMulPosMono α β] (hb : 0 ≤ b) (h : 1 ≤ a) : b ≤ a • b := by + simpa only [one_smul] using smul_le_smul_of_nonneg_right h hb + +lemma smul_lt_of_lt_one_left [SMulPosStrictMono α β] (hb : 0 < b) (h : a < 1) : a • b < b := by + simpa only [one_smul] using smul_lt_smul_of_pos_right h hb + +lemma lt_smul_of_one_lt_left [SMulPosStrictMono α β] (hb : 0 < b) (h : 1 < a) : b < a • b := by + simpa only [one_smul] using smul_lt_smul_of_pos_right h hb + +end Preorder +end MulAction + +section Semiring +variable [Semiring α] [AddCommGroup β] [Module α β] [NoZeroSMulDivisors α β] + +section PartialOrder +variable [Preorder α] [PartialOrder β] + +lemma PosSMulMono.toPosSMulStrictMono [PosSMulMono α β] : PosSMulStrictMono α β := + ⟨fun _a ha _b₁ _b₂ hb ↦ (smul_le_smul_of_nonneg_left hb.le ha.le).lt_of_ne $ + (smul_right_injective _ ha.ne').ne hb.ne⟩ + +instance PosSMulReflectLT.toPosSMulReflectLE [PosSMulReflectLT α β] : PosSMulReflectLE α β := + ⟨fun _a ha _b₁ _b₂ h ↦ h.eq_or_lt.elim (fun h ↦ (smul_right_injective _ ha.ne' h).le) fun h' ↦ + (lt_of_smul_lt_smul_left h' ha.le).le⟩ + +end PartialOrder + +section PartialOrder +variable [PartialOrder α] [PartialOrder β] + +lemma posSMulMono_iff_posSMulStrictMono : PosSMulMono α β ↔ PosSMulStrictMono α β := + ⟨fun _ ↦ PosSMulMono.toPosSMulStrictMono, fun _ ↦ inferInstance⟩ + +lemma PosSMulReflectLE_iff_posSMulReflectLT : PosSMulReflectLE α β ↔ PosSMulReflectLT α β := + ⟨fun _ ↦ inferInstance, fun _ ↦ PosSMulReflectLT.toPosSMulReflectLE⟩ + +end PartialOrder +end Semiring + +section Ring +variable [Ring α] [AddCommGroup β] [Module α β] [NoZeroSMulDivisors α β] + +section PartialOrder +variable [PartialOrder α] [PartialOrder β] + +lemma SMulPosMono.toSMulPosStrictMono [SMulPosMono α β] : SMulPosStrictMono α β := + ⟨fun _b hb _a₁ _a₂ ha ↦ (smul_le_smul_of_nonneg_right ha.le hb.le).lt_of_ne $ + (smul_left_injective _ hb.ne').ne ha.ne⟩ + +lemma smulPosMono_iff_smulPosStrictMono : SMulPosMono α β ↔ SMulPosStrictMono α β := + ⟨fun _ ↦ SMulPosMono.toSMulPosStrictMono, fun _ ↦ inferInstance⟩ + +lemma SMulPosReflectLT.toSMulPosReflectLE [SMulPosReflectLT α β] : SMulPosReflectLE α β := + ⟨fun _b hb _a₁ _a₂ h ↦ h.eq_or_lt.elim (fun h ↦ (smul_left_injective _ hb.ne' h).le) fun h' ↦ + (lt_of_smul_lt_smul_right h' hb.le).le⟩ + +lemma SMulPosReflectLE_iff_smulPosReflectLT : SMulPosReflectLE α β ↔ SMulPosReflectLT α β := + ⟨fun _ ↦ inferInstance, fun _ ↦ SMulPosReflectLT.toSMulPosReflectLE⟩ + +end PartialOrder +end Ring + +section OrderedRing +variable [OrderedRing α] [OrderedAddCommGroup β] [Module α β] + +instance PosSMulMono.toSMulPosMono [PosSMulMono α β] : SMulPosMono α β where + elim _b hb a₁ a₂ ha := by rw [← sub_nonneg, ← sub_smul]; exact smul_nonneg (sub_nonneg.2 ha) hb + +instance PosSMulStrictMono.toSMulPosStrictMono [PosSMulStrictMono α β] : SMulPosStrictMono α β where + elim _b hb a₁ a₂ ha := by rw [← sub_pos, ← sub_smul]; exact smul_pos (sub_pos.2 ha) hb + +end OrderedRing + +section Field +variable [GroupWithZero α] [Preorder α] [Preorder β] [MulAction α β] + +lemma inv_smul_le_iff_of_pos [PosSMulMono α β] [PosSMulReflectLE α β] (ha : 0 < a) : + a⁻¹ • b₁ ≤ b₂ ↔ b₁ ≤ a • b₂ := by rw [← smul_le_smul_iff_of_pos_left ha, smul_inv_smul₀ ha.ne'] + +lemma le_inv_smul_iff_of_pos [PosSMulMono α β] [PosSMulReflectLE α β] (ha : 0 < a) : + b₁ ≤ a⁻¹ • b₂ ↔ a • b₁ ≤ b₂ := by rw [← smul_le_smul_iff_of_pos_left ha, smul_inv_smul₀ ha.ne'] + +lemma inv_smul_lt_iff_of_pos [PosSMulStrictMono α β] [PosSMulReflectLT α β] (ha : 0 < a) : + a⁻¹ • b₁ < b₂ ↔ b₁ < a • b₂ := by rw [← smul_lt_smul_iff_of_pos_left ha, smul_inv_smul₀ ha.ne'] + +lemma lt_inv_smul_iff_of_pos [PosSMulStrictMono α β] [PosSMulReflectLT α β] (ha : 0 < a) : + b₁ < a⁻¹ • b₂ ↔ a • b₁ < b₂ := by rw [← smul_lt_smul_iff_of_pos_left ha, smul_inv_smul₀ ha.ne'] + +end Field + +section LinearOrderedSemifield +variable [LinearOrderedSemifield α] [AddCommGroup β] [PartialOrder β] + +-- See note [lower instance priority] +instance (priority := 100) PosSMulMono.toPosSMulReflectLE [MulAction α β] [PosSMulMono α β] : + PosSMulReflectLE α β where + elim _a ha b₁ b₂ h := by simpa [ha.ne'] using smul_le_smul_of_nonneg_left h $ inv_nonneg.2 ha.le + +-- See note [lower instance priority] +instance (priority := 100) PosSMulStrictMono.toPosSMulReflectLT [MulActionWithZero α β] + [PosSMulStrictMono α β] : PosSMulReflectLT α β := + PosSMulReflectLT.of_pos fun a ha b₁ b₂ h ↦ by + simpa [ha.ne'] using smul_lt_smul_of_pos_left h $ inv_pos.2 ha + +end LinearOrderedSemifield + +namespace OrderDual + +section Left +variable [Preorder α] [Preorder β] [SMul α β] [Zero α] + +instance instPosSMulMono [PosSMulMono α β] : PosSMulMono α βᵒᵈ where + elim _a ha _b₁ _b₂ hb := smul_le_smul_of_nonneg_left (β := β) hb ha +instance instPosSMulStrictMono [PosSMulStrictMono α β] : PosSMulStrictMono α βᵒᵈ where + elim _a ha _b₁ _b₂ hb := smul_lt_smul_of_pos_left (β := β) hb ha +instance instPosSMulReflectLT [PosSMulReflectLT α β] : PosSMulReflectLT α βᵒᵈ where + elim _a ha _b₁ _b₂ h := lt_of_smul_lt_smul_of_nonneg_left (β := β) h ha +instance instPosSMulReflectLE [PosSMulReflectLE α β] : PosSMulReflectLE α βᵒᵈ where + elim _a ha _b₁ _b₂ h := le_of_smul_le_smul_of_pos_left (β := β) h ha + +end Left + +section Right +variable [Preorder α] [Ring α] [OrderedAddCommGroup β] [Module α β] + +instance instSMulPosMono [SMulPosMono α β] : SMulPosMono α βᵒᵈ where + elim _b hb a₁ a₂ ha := by + rw [← neg_le_neg_iff, ← smul_neg, ← smul_neg] + exact smul_le_smul_of_nonneg_right (β := β) ha $ neg_nonneg.2 hb + +instance instSMulPosStrictMono [SMulPosStrictMono α β] : SMulPosStrictMono α βᵒᵈ where + elim _b hb a₁ a₂ ha := by + rw [← neg_lt_neg_iff, ← smul_neg, ← smul_neg] + exact smul_lt_smul_of_pos_right (β := β) ha $ neg_pos.2 hb + +instance instSMulPosReflectLT [SMulPosReflectLT α β] : SMulPosReflectLT α βᵒᵈ where + elim _b hb a₁ a₂ h := by + rw [← neg_lt_neg_iff, ← smul_neg, ← smul_neg] at h + exact lt_of_smul_lt_smul_right (β := β) h $ neg_nonneg.2 hb + +instance instSMulPosReflectLE [SMulPosReflectLE α β] : SMulPosReflectLE α βᵒᵈ where + elim _b hb a₁ a₂ h := by + rw [← neg_le_neg_iff, ← smul_neg, ← smul_neg] at h + exact le_of_smul_le_smul_right (β := β) h $ neg_pos.2 hb + +end Right +end OrderDual + +namespace Prod + +end Prod + +namespace Pi +variable {ι : Type*} {β : ι → Type*} [Zero α] [∀ i, Zero (β i)] + +section SMulZeroClass +variable [Preorder α] [∀ i, Preorder (β i)] [∀ i, SMulZeroClass α (β i)] + +instance instPosSMulMono [∀ i, PosSMulMono α (β i)] : PosSMulMono α (∀ i, β i) where + elim _a ha _b₁ _b₂ hb i := smul_le_smul_of_nonneg_left (hb i) ha + +instance instSMulPosMono [∀ i, SMulPosMono α (β i)] : SMulPosMono α (∀ i, β i) where + elim _b hb _a₁ _a₂ ha i := smul_le_smul_of_nonneg_right ha (hb i) + +instance instPosSMulReflectLE [∀ i, PosSMulReflectLE α (β i)] : PosSMulReflectLE α (∀ i, β i) where + elim _a ha _b₁ _b₂ h i := le_of_smul_le_smul_left (h i) ha + +instance instSMulPosReflectLE [∀ i, SMulPosReflectLE α (β i)] : SMulPosReflectLE α (∀ i, β i) where + elim _b hb _a₁ _a₂ h := by + obtain ⟨-, i, hi⟩ := lt_def.1 hb; exact le_of_smul_le_smul_right (h _) hi + +end SMulZeroClass + +section SMulWithZero +variable [PartialOrder α] [∀ i, PartialOrder (β i)] [∀ i, SMulWithZero α (β i)] + +instance instPosSMulStrictMono [∀ i, PosSMulStrictMono α (β i)] : + PosSMulStrictMono α (∀ i, β i) where + elim := by + simp_rw [lt_def] + rintro _a ha _b₁ _b₂ ⟨hb, i, hi⟩ + exact ⟨smul_le_smul_of_nonneg_left hb ha.le, i, smul_lt_smul_of_pos_left hi ha⟩ + +instance instSMulPosStrictMono [∀ i, SMulPosStrictMono α (β i)] : + SMulPosStrictMono α (∀ i, β i) where + elim := by + simp_rw [lt_def] + rintro a ⟨ha, i, hi⟩ _b₁ _b₂ hb + exact ⟨smul_le_smul_of_nonneg_right hb.le ha, i, smul_lt_smul_of_pos_right hb hi⟩ + +-- Note: There is no interesting instance for `PosSMulReflectLT α (∀ i, β i)` that's not already +-- implied by the other instances + +instance instSMulPosReflectLT [∀ i, SMulPosReflectLT α (β i)] : SMulPosReflectLT α (∀ i, β i) where + elim := by + simp_rw [lt_def] + rintro b hb _a₁ _a₂ ⟨-, i, hi⟩ + exact lt_of_smul_lt_smul_right hi $ hb _ + +end SMulWithZero +end Pi + +section Lift +variable {γ : Type*} [Zero α] [Preorder α] [Zero β] [Preorder β] [Zero γ] [Preorder γ] + [SMul α β] [SMul α γ] (f : β → γ) (hf : ∀ {b₁ b₂}, f b₁ ≤ f b₂ ↔ b₁ ≤ b₂) + (smul : ∀ (a : α) b, f (a • b) = a • f b) (zero : f 0 = 0) + +lemma PosSMulMono.lift [PosSMulMono α γ] : PosSMulMono α β where + elim a ha b₁ b₂ hb := by simp only [← hf, smul] at *; exact smul_le_smul_of_nonneg_left hb ha + +lemma PosSMulStrictMono.lift [PosSMulStrictMono α γ] : PosSMulStrictMono α β where + elim a ha b₁ b₂ hb := by + simp only [← lt_iff_lt_of_le_iff_le' hf hf, smul] at *; exact smul_lt_smul_of_pos_left hb ha + +lemma PosSMulReflectLE.lift [PosSMulReflectLE α γ] : PosSMulReflectLE α β where + elim a ha b₁ b₂ h := hf.1 $ le_of_smul_le_smul_left (by simpa only [smul] using hf.2 h) ha + +lemma PosSMulReflectLT.lift [PosSMulReflectLT α γ] : PosSMulReflectLT α β where + elim a ha b₁ b₂ h := by + simp only [← lt_iff_lt_of_le_iff_le' hf hf, smul] at *; exact lt_of_smul_lt_smul_left h ha + +lemma SMulPosMono.lift [SMulPosMono α γ] : SMulPosMono α β where + elim b hb a₁ a₂ ha := by + simp only [← hf, zero, smul] at *; exact smul_le_smul_of_nonneg_right ha hb + +lemma SMulPosStrictMono.lift [SMulPosStrictMono α γ] : SMulPosStrictMono α β where + elim b hb a₁ a₂ ha := by + simp only [← lt_iff_lt_of_le_iff_le' hf hf, zero, smul] at * + exact smul_lt_smul_of_pos_right ha hb + +lemma SMulPosReflectLE.lift [SMulPosReflectLE α γ] : SMulPosReflectLE α β where + elim b hb a₁ a₂ h := by + simp only [← hf, ← lt_iff_lt_of_le_iff_le' hf hf, zero, smul] at * + exact le_of_smul_le_smul_right h hb + +lemma SMulPosReflectLT.lift [SMulPosReflectLT α γ] : SMulPosReflectLT α β where + elim b hb a₁ a₂ h := by + simp only [← hf, ← lt_iff_lt_of_le_iff_le' hf hf, zero, smul] at * + exact lt_of_smul_lt_smul_right h hb + +end Lift diff --git a/Mathlib/Algebra/Order/Monovary.lean b/Mathlib/Algebra/Order/Monovary.lean index 30fe1693595b8..dec508e80301d 100644 --- a/Mathlib/Algebra/Order/Monovary.lean +++ b/Mathlib/Algebra/Order/Monovary.lean @@ -86,10 +86,10 @@ lemma antivaryOn_inv_right : AntivaryOn f g⁻¹ s ↔ MonovaryOn f g s := by AntivaryOn (f₁ / f₂) g s := fun _i hi _j hj hij ↦ div_le_div'' (h₁ hi hj hij) (h₂ hi hj hij) @[to_additive] lemma MonovaryOn.pow_left (hfg : MonovaryOn f g s) (n : ℕ) : - MonovaryOn (f ^ n) g s := fun _i hi _j hj hij ↦ pow_le_pow_of_le_left' (hfg hi hj hij) _ + MonovaryOn (f ^ n) g s := fun _i hi _j hj hij ↦ pow_le_pow_left' (hfg hi hj hij) _ @[to_additive] lemma AntivaryOn.pow_left (hfg : AntivaryOn f g s) (n : ℕ) : - AntivaryOn (f ^ n) g s := fun _i hi _j hj hij ↦ pow_le_pow_of_le_left' (hfg hi hj hij) _ + AntivaryOn (f ^ n) g s := fun _i hi _j hj hij ↦ pow_le_pow_left' (hfg hi hj hij) _ @[to_additive] lemma Monovary.mul_left (h₁ : Monovary f₁ g) (h₂ : Monovary f₂ g) : Monovary (f₁ * f₂) g := @@ -108,10 +108,10 @@ lemma Antivary.div_left (h₁ : Antivary f₁ g) (h₂ : Monovary f₂ g) : Anti fun _i _j hij ↦ div_le_div'' (h₁ hij) (h₂ hij) @[to_additive] lemma Monovary.pow_left (hfg : Monovary f g) (n : ℕ) : Monovary (f ^ n) g := - fun _i _j hij ↦ pow_le_pow_of_le_left' (hfg hij) _ + fun _i _j hij ↦ pow_le_pow_left' (hfg hij) _ @[to_additive] lemma Antivary.pow_left (hfg : Antivary f g) (n : ℕ) : Antivary (f ^ n) g := - fun _i _j hij ↦ pow_le_pow_of_le_left' (hfg hij) _ + fun _i _j hij ↦ pow_le_pow_left' (hfg hij) _ end OrderedCommGroup @@ -136,10 +136,10 @@ variable [OrderedCommGroup α] [LinearOrderedCommGroup β] {s : Set ι} {f f₁ fun _i hi _j hj hij ↦ (lt_or_lt_of_div_lt_div hij).elim (h₁ hi hj) $ h₂ hj hi @[to_additive] lemma MonovaryOn.pow_right (hfg : MonovaryOn f g s) (n : ℕ) : - MonovaryOn f (g ^ n) s := fun _i hi _j hj hij ↦ hfg hi hj $ lt_of_pow_lt_pow' _ hij + MonovaryOn f (g ^ n) s := fun _i hi _j hj hij ↦ hfg hi hj $ lt_of_pow_lt_pow_left' _ hij @[to_additive] lemma AntivaryOn.pow_right (hfg : AntivaryOn f g s) (n : ℕ) : - AntivaryOn f (g ^ n) s := fun _i hi _j hj hij ↦ hfg hi hj $ lt_of_pow_lt_pow' _ hij + AntivaryOn f (g ^ n) s := fun _i hi _j hj hij ↦ hfg hi hj $ lt_of_pow_lt_pow_left' _ hij @[to_additive] lemma Monovary.mul_right (h₁ : Monovary f g₁) (h₂ : Monovary f g₂) : Monovary f (g₁ * g₂) := @@ -158,10 +158,10 @@ variable [OrderedCommGroup α] [LinearOrderedCommGroup β] {s : Set ι} {f f₁ fun _i _j hij ↦ (lt_or_lt_of_div_lt_div hij).elim (fun h ↦ h₁ h) $ fun h ↦ h₂ h @[to_additive] lemma Monovary.pow_right (hfg : Monovary f g) (n : ℕ) : Monovary f (g ^ n) := - fun _i _j hij ↦ hfg $ lt_of_pow_lt_pow' _ hij + fun _i _j hij ↦ hfg $ lt_of_pow_lt_pow_left' _ hij @[to_additive] lemma Antivary.pow_right (hfg : Antivary f g) (n : ℕ) : Antivary f (g ^ n) := - fun _i _j hij ↦ hfg $ lt_of_pow_lt_pow' _ hij + fun _i _j hij ↦ hfg $ lt_of_pow_lt_pow_left' _ hij end LinearOrderedCommGroup @@ -178,11 +178,11 @@ lemma AntivaryOn.mul_left₀ (hf₁ : ∀ i ∈ s, 0 ≤ f₁ i) (hf₂ : ∀ i lemma MonovaryOn.pow_left₀ (hf : ∀ i ∈ s, 0 ≤ f i) (hfg : MonovaryOn f g s) (n : ℕ) : MonovaryOn (f ^ n) g s := - fun _i hi _j hj hij ↦ pow_le_pow_of_le_left (hf _ hi) (hfg hi hj hij) _ + fun _i hi _j hj hij ↦ pow_le_pow_left (hf _ hi) (hfg hi hj hij) _ lemma AntivaryOn.pow_left₀ (hf : ∀ i ∈ s, 0 ≤ f i) (hfg : AntivaryOn f g s) (n : ℕ) : AntivaryOn (f ^ n) g s := - fun _i hi _j hj hij ↦ pow_le_pow_of_le_left (hf _ hj) (hfg hi hj hij) _ + fun _i hi _j hj hij ↦ pow_le_pow_left (hf _ hj) (hfg hi hj hij) _ lemma Monovary.mul_left₀ (hf₁ : 0 ≤ f₁) (hf₂ : 0 ≤ f₂) (h₁ : Monovary f₁ g) (h₂ : Monovary f₂ g) : Monovary (f₁ * f₂) g := fun _i _j hij ↦ mul_le_mul (h₁ hij) (h₂ hij) (hf₂ _) (hf₁ _) @@ -191,10 +191,10 @@ lemma Antivary.mul_left₀ (hf₁ : 0 ≤ f₁) (hf₂ : 0 ≤ f₂) (h₁ : Ant Antivary (f₁ * f₂) g := fun _i _j hij ↦ mul_le_mul (h₁ hij) (h₂ hij) (hf₂ _) (hf₁ _) lemma Monovary.pow_left₀ (hf : 0 ≤ f) (hfg : Monovary f g) (n : ℕ) : Monovary (f ^ n) g := - fun _i _j hij ↦ pow_le_pow_of_le_left (hf _) (hfg hij) _ + fun _i _j hij ↦ pow_le_pow_left (hf _) (hfg hij) _ lemma Antivary.pow_left₀ (hf : 0 ≤ f) (hfg : Antivary f g) (n : ℕ) : Antivary (f ^ n) g := - fun _i _j hij ↦ pow_le_pow_of_le_left (hf _) (hfg hij) _ + fun _i _j hij ↦ pow_le_pow_left (hf _) (hfg hij) _ end OrderedSemiring diff --git a/Mathlib/Algebra/Order/Ring/Star.lean b/Mathlib/Algebra/Order/Ring/Star.lean index 13b4b35f6981f..3d5b271af22e8 100644 --- a/Mathlib/Algebra/Order/Ring/Star.lean +++ b/Mathlib/Algebra/Order/Ring/Star.lean @@ -27,17 +27,8 @@ type class assumptions to be satisfied without a `CommSemiring` intance already it is impossible, only that it shouldn't occur in practice. -/ example {R : Type*} [OrderedSemiring R] [StarOrderedRing R] {x y : R} (hx : 0 ≤ x) (hy : 0 ≤ y) : x * y = y * x := by - -- nonnegative elements are self-adjoint; we prove it by hand to avoid adding imports - have : ∀ z : R, 0 ≤ z → star z = z := by - intros z hz - rw [nonneg_iff] at hz - induction hz using AddSubmonoid.closure_induction' with - | Hs _ h => obtain ⟨x, rfl⟩ := h; simp - | H1 => simp - | Hmul x hx y hy => simp only [← nonneg_iff] at hx hy; aesop - -- `0 ≤ y * x`, and hence `y * x` is self-adjoint - have := this _ <| mul_nonneg hy hx - aesop + rw [← IsSelfAdjoint.of_nonneg (mul_nonneg hy hx), star_mul, IsSelfAdjoint.of_nonneg hx, + IsSelfAdjoint.of_nonneg hy] /- This will be implied by the instance below, we only prove it to avoid duplicating the argument in the instance below for `mul_le_mul_of_nonneg_right`. -/ diff --git a/Mathlib/Algebra/Order/SMul.lean b/Mathlib/Algebra/Order/SMul.lean index 099d97214c37e..3938b6e825f39 100644 --- a/Mathlib/Algebra/Order/SMul.lean +++ b/Mathlib/Algebra/Order/SMul.lean @@ -5,7 +5,7 @@ Authors: Frédéric Dupuis -/ import Mathlib.Algebra.Module.Pi import Mathlib.Algebra.Module.Prod -import Mathlib.Algebra.Order.Module.Synonym +import Mathlib.Algebra.Order.Module.Defs import Mathlib.Algebra.Order.Monoid.Prod import Mathlib.Algebra.Order.Pi import Mathlib.Data.Set.Pointwise.SMul @@ -31,6 +31,10 @@ In this file we define * To get ordered modules and ordered vector spaces, it suffices to replace the `OrderedAddCommMonoid` and the `OrderedSemiring` as desired. +## TODO + +Delete the lemmas that have been generalised by `PosSMulMono` and friends. + ## References * https://en.wikipedia.org/wiki/Ordered_vector_space @@ -66,29 +70,21 @@ section OrderedSMul variable [OrderedSemiring R] [OrderedAddCommMonoid M] [SMulWithZero R M] [OrderedSMul R M] {s : Set M} {a b : M} {c : R} -@[gcongr] theorem smul_lt_smul_of_pos : a < b → 0 < c → c • a < c • b := - OrderedSMul.smul_lt_smul_of_pos -#align smul_lt_smul_of_pos smul_lt_smul_of_pos +instance OrderedSMul.toPosSMulStrictMono : PosSMulStrictMono R M where + elim _a ha _b₁ _b₂ hb := OrderedSMul.smul_lt_smul_of_pos hb ha -@[gcongr] theorem smul_le_smul_of_nonneg (h₁ : a ≤ b) (h₂ : 0 ≤ c) : c • a ≤ c • b := by - rcases h₁.eq_or_lt with (rfl | hab) - · rfl - · rcases h₂.eq_or_lt with (rfl | hc) - · rw [zero_smul, zero_smul] - · exact (smul_lt_smul_of_pos hab hc).le -#align smul_le_smul_of_nonneg smul_le_smul_of_nonneg +instance OrderedSMul.toPosSMulReflectLT : PosSMulReflectLT R M := + PosSMulReflectLT.of_pos $ fun _a ha _b₁ _b₂ h ↦ OrderedSMul.lt_of_smul_lt_smul_of_pos h ha + +@[gcongr] theorem smul_lt_smul_of_pos : a < b → 0 < c → c • a < c • b := smul_lt_smul_of_pos_left +#align smul_lt_smul_of_pos smul_lt_smul_of_pos -- TODO: Remove `smul_le_smul_of_nonneg` completely -alias smul_le_smul_of_nonneg_left := smul_le_smul_of_nonneg +@[gcongr] theorem smul_le_smul_of_nonneg (h₁ : a ≤ b) (h₂ : 0 ≤ c) : c • a ≤ c • b := + smul_le_smul_of_nonneg_left h₁ h₂ +#align smul_le_smul_of_nonneg smul_le_smul_of_nonneg -theorem smul_nonneg (hc : 0 ≤ c) (ha : 0 ≤ a) : 0 ≤ c • a := - calc - (0 : M) = c • (0 : M) := (smul_zero c).symm - _ ≤ c • a := smul_le_smul_of_nonneg ha hc #align smul_nonneg smul_nonneg - -theorem smul_nonpos_of_nonneg_of_nonpos (hc : 0 ≤ c) (ha : a ≤ 0) : c • a ≤ 0 := - @smul_nonneg R Mᵒᵈ _ _ _ _ _ _ hc ha #align smul_nonpos_of_nonneg_of_nonpos smul_nonpos_of_nonneg_of_nonpos theorem eq_of_smul_eq_smul_of_pos_of_le (h₁ : c • a = c • b) (hc : 0 < c) (hle : a ≤ b) : a = b := @@ -96,30 +92,24 @@ theorem eq_of_smul_eq_smul_of_pos_of_le (h₁ : c • a = c • b) (hc : 0 < c) #align eq_of_smul_eq_smul_of_pos_of_le eq_of_smul_eq_smul_of_pos_of_le theorem lt_of_smul_lt_smul_of_nonneg (h : c • a < c • b) (hc : 0 ≤ c) : a < b := - hc.eq_or_lt.elim - (fun hc => False.elim <| lt_irrefl (0 : M) <| by rwa [← hc, zero_smul, zero_smul] at h) - (OrderedSMul.lt_of_smul_lt_smul_of_pos h) + lt_of_smul_lt_smul_of_nonneg_left h hc #align lt_of_smul_lt_smul_of_nonneg lt_of_smul_lt_smul_of_nonneg theorem smul_lt_smul_iff_of_pos (hc : 0 < c) : c • a < c • b ↔ a < b := - ⟨fun h => lt_of_smul_lt_smul_of_nonneg h hc.le, fun h => smul_lt_smul_of_pos h hc⟩ + smul_lt_smul_iff_of_pos_left hc #align smul_lt_smul_iff_of_pos smul_lt_smul_iff_of_pos -theorem smul_pos_iff_of_pos (hc : 0 < c) : 0 < c • a ↔ 0 < a := - calc - 0 < c • a ↔ c • (0 : M) < c • a := by rw [smul_zero] - _ ↔ 0 < a := smul_lt_smul_iff_of_pos hc +theorem smul_pos_iff_of_pos (hc : 0 < c) : 0 < c • a ↔ 0 < a := smul_pos_iff_of_pos_left hc #align smul_pos_iff_of_pos smul_pos_iff_of_pos -alias ⟨_, smul_pos⟩ := smul_pos_iff_of_pos #align smul_pos smul_pos -theorem monotone_smul_left (hc : 0 ≤ c) : Monotone (SMul.smul c : M → M) := fun _ _ h => - smul_le_smul_of_nonneg h hc +theorem monotone_smul_left (hc : 0 ≤ c) : Monotone (SMul.smul c : M → M) := + monotone_smul_left_of_nonneg hc #align monotone_smul_left monotone_smul_left -theorem strictMono_smul_left (hc : 0 < c) : StrictMono (SMul.smul c : M → M) := fun _ _ h => - smul_lt_smul_of_pos h hc +theorem strictMono_smul_left (hc : 0 < c) : StrictMono (SMul.smul c : M → M) := + strictMono_smul_left_of_pos hc #align strict_mono_smul_left strictMono_smul_left theorem smul_lowerBounds_subset_lowerBounds_smul (hc : 0 ≤ c) : @@ -178,11 +168,11 @@ instance LinearOrderedSemiring.toOrderedSMul : OrderedSMul R R := #align linear_ordered_semiring.to_ordered_smul LinearOrderedSemiring.toOrderedSMul theorem smul_max (ha : 0 ≤ a) (b₁ b₂ : M) : a • max b₁ b₂ = max (a • b₁) (a • b₂) := - (monotone_smul_left ha : Monotone (_ : M → M)).map_max + smul_max_of_nonneg ha _ _ #align smul_max smul_max theorem smul_min (ha : 0 ≤ a) (b₁ b₂ : M) : a • min b₁ b₂ = min (a • b₁) (a • b₂) := - (monotone_smul_left ha : Monotone (_ : M → M)).map_min + smul_min_of_nonneg ha _ _ #align smul_min smul_min end LinearOrderedSemiring @@ -231,26 +221,19 @@ instance Pi.orderedSMul'' : OrderedSMul 𝕜 (ι → 𝕜) := variable [OrderedSMul 𝕜 M] {s : Set M} {a b : M} {c : 𝕜} theorem smul_le_smul_iff_of_pos (hc : 0 < c) : c • a ≤ c • b ↔ a ≤ b := - ⟨fun h => - inv_smul_smul₀ hc.ne' a ▸ - inv_smul_smul₀ hc.ne' b ▸ smul_le_smul_of_nonneg h (inv_nonneg.2 hc.le), - fun h => smul_le_smul_of_nonneg h hc.le⟩ + smul_le_smul_iff_of_pos_left hc #align smul_le_smul_iff_of_pos smul_le_smul_iff_of_pos -theorem inv_smul_le_iff (h : 0 < c) : c⁻¹ • a ≤ b ↔ a ≤ c • b := by - rw [← smul_le_smul_iff_of_pos h, smul_inv_smul₀ h.ne'] +theorem inv_smul_le_iff (h : 0 < c) : c⁻¹ • a ≤ b ↔ a ≤ c • b := inv_smul_le_iff_of_pos h #align inv_smul_le_iff inv_smul_le_iff -theorem inv_smul_lt_iff (h : 0 < c) : c⁻¹ • a < b ↔ a < c • b := by - rw [← smul_lt_smul_iff_of_pos h, smul_inv_smul₀ h.ne'] +theorem inv_smul_lt_iff (h : 0 < c) : c⁻¹ • a < b ↔ a < c • b := inv_smul_lt_iff_of_pos h #align inv_smul_lt_iff inv_smul_lt_iff -theorem le_inv_smul_iff (h : 0 < c) : a ≤ c⁻¹ • b ↔ c • a ≤ b := by - rw [← smul_le_smul_iff_of_pos h, smul_inv_smul₀ h.ne'] +theorem le_inv_smul_iff (h : 0 < c) : a ≤ c⁻¹ • b ↔ c • a ≤ b := le_inv_smul_iff_of_pos h #align le_inv_smul_iff le_inv_smul_iff -theorem lt_inv_smul_iff (h : 0 < c) : a < c⁻¹ • b ↔ c • a < b := by - rw [← smul_lt_smul_iff_of_pos h, smul_inv_smul₀ h.ne'] +theorem lt_inv_smul_iff (h : 0 < c) : a < c⁻¹ • b ↔ c • a < b := lt_inv_smul_iff_of_pos h #align lt_inv_smul_iff lt_inv_smul_iff variable (M) diff --git a/Mathlib/Algebra/Order/Sub/Defs.lean b/Mathlib/Algebra/Order/Sub/Defs.lean index 964d88879b82a..cbe68060983a5 100644 --- a/Mathlib/Algebra/Order/Sub/Defs.lean +++ b/Mathlib/Algebra/Order/Sub/Defs.lean @@ -59,13 +59,14 @@ class OrderedSub (α : Type*) [LE α] [Add α] [Sub α] : Prop where section Add -variable [Preorder α] [Add α] [Sub α] [OrderedSub α] {a b c d : α} - @[simp] -theorem tsub_le_iff_right : a - b ≤ c ↔ a ≤ c + b := +theorem tsub_le_iff_right [LE α] [Add α] [Sub α] [OrderedSub α] {a b c : α} : + a - b ≤ c ↔ a ≤ c + b := OrderedSub.tsub_le_iff_right a b c #align tsub_le_iff_right tsub_le_iff_right +variable [Preorder α] [Add α] [Sub α] [OrderedSub α] {a b c d : α} + /-- See `add_tsub_cancel_right` for the equality if `ContravariantClass α α (+) (≤)`. -/ theorem add_tsub_le_right : a + b - b ≤ a := tsub_le_iff_right.mpr le_rfl diff --git a/Mathlib/Algebra/Star/Order.lean b/Mathlib/Algebra/Star/Order.lean index 715476260bc80..fe8c3756b67a2 100644 --- a/Mathlib/Algebra/Star/Order.lean +++ b/Mathlib/Algebra/Star/Order.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Mathlib.Algebra.Star.Basic +import Mathlib.Algebra.Star.SelfAdjoint import Mathlib.GroupTheory.Submonoid.Basic #align_import algebra.star.order from "leanprover-community/mathlib"@"31c24aa72e7b3e5ed97a8412470e904f82b81004" @@ -185,4 +185,75 @@ theorem conjugate_le_conjugate' {a b : R} (hab : a ≤ b) (c : R) : c * a * star by simpa only [star_star] using conjugate_le_conjugate hab (star c) #align conjugate_le_conjugate' conjugate_le_conjugate' +@[simp] +lemma star_le_star_iff {x y : R} : star x ≤ star y ↔ x ≤ y := by + suffices ∀ x y, x ≤ y → star x ≤ star y from + ⟨by simpa only [star_star] using this (star x) (star y), this x y⟩ + intro x y h + rw [StarOrderedRing.le_iff] at h ⊢ + obtain ⟨d, hd, rfl⟩ := h + refine ⟨starAddEquiv d, ?_, star_add _ _⟩ + refine AddMonoidHom.mclosure_preimage_le _ _ <| AddSubmonoid.closure_mono ?_ hd + rintro - ⟨s, rfl⟩ + exact ⟨s, by simp⟩ + +@[simp] +lemma star_lt_star_iff {x y : R} : star x < star y ↔ x < y := by + by_cases h : x = y + · simp [h] + · simpa [le_iff_lt_or_eq, h] using star_le_star_iff (x := x) (y := y) + +lemma star_le_iff {x y : R} : star x ≤ y ↔ x ≤ star y := by rw [← star_le_star_iff, star_star] + +lemma star_lt_iff {x y : R} : star x < y ↔ x < star y := by rw [← star_lt_star_iff, star_star] + +@[simp] +lemma star_nonneg_iff {x : R} : 0 ≤ star x ↔ 0 ≤ x := by + simpa using star_le_star_iff (x := 0) (y := x) + +@[simp] +lemma star_nonpos_iff {x : R} : star x ≤ 0 ↔ x ≤ 0 := by + simpa using star_le_star_iff (x := x) (y := 0) + +@[simp] +lemma star_pos_iff {x : R} : 0 < star x ↔ 0 < x := by + simpa using star_lt_star_iff (x := 0) (y := x) + +@[simp] +lemma star_neg_iff {x : R} : star x < 0 ↔ x < 0 := by + simpa using star_lt_star_iff (x := x) (y := 0) + +lemma IsSelfAdjoint.mono {x y : R} (h : x ≤ y) (hx : IsSelfAdjoint x) : IsSelfAdjoint y := by + rw [StarOrderedRing.le_iff] at h + obtain ⟨d, hd, rfl⟩ := h + rw [IsSelfAdjoint, star_add, hx.star_eq] + congr + refine AddMonoidHom.eqOn_closureM (f := starAddEquiv (R := R)) (g := .id R) ?_ hd + rintro - ⟨s, rfl⟩ + simp + +lemma IsSelfAdjoint.of_nonneg {x : R} (hx : 0 ≤ x) : IsSelfAdjoint x := + (isSelfAdjoint_zero R).mono hx + end NonUnitalSemiring + +section Semiring +variable [Semiring R] [PartialOrder R] [StarOrderedRing R] + +@[simp] +lemma one_le_star_iff {x : R} : 1 ≤ star x ↔ 1 ≤ x := by + simpa using star_le_star_iff (x := 1) (y := x) + +@[simp] +lemma star_le_one_iff {x : R} : star x ≤ 1 ↔ x ≤ 1 := by + simpa using star_le_star_iff (x := x) (y := 1) + +@[simp] +lemma one_lt_star_iff {x : R} : 1 < star x ↔ 1 < x := by + simpa using star_lt_star_iff (x := 1) (y := x) + +@[simp] +lemma star_lt_one_iff {x : R} : star x < 1 ↔ x < 1 := by + simpa using star_lt_star_iff (x := x) (y := 1) + +end Semiring diff --git a/Mathlib/Algebra/Tropical/Basic.lean b/Mathlib/Algebra/Tropical/Basic.lean index de073e688ab86..c595d65b5d981 100644 --- a/Mathlib/Algebra/Tropical/Basic.lean +++ b/Mathlib/Algebra/Tropical/Basic.lean @@ -544,8 +544,8 @@ theorem add_pow [LinearOrder R] [AddMonoid R] [CovariantClass R R (· + ·) (· [CovariantClass R R (Function.swap (· + ·)) (· ≤ ·)] (x y : Tropical R) (n : ℕ) : (x + y) ^ n = x ^ n + y ^ n := by cases' le_total x y with h h - · rw [add_eq_left h, add_eq_left (pow_le_pow_of_le_left' h _)] - · rw [add_eq_right h, add_eq_right (pow_le_pow_of_le_left' h _)] + · rw [add_eq_left h, add_eq_left (pow_le_pow_left' h _)] + · rw [add_eq_right h, add_eq_right (pow_le_pow_left' h _)] #align tropical.add_pow Tropical.add_pow end Distrib diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean index fdb14960f8315..b4d70283c0093 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean @@ -142,7 +142,7 @@ many of the paths do not have defeq starting/ending points, so we end up needing /-- Interpret a homotopy `H : C(I × X, Y)` as a map `C(ULift I × X, Y)` -/ def uliftMap : C(TopCat.of (ULift.{u} I × X), Y) := ⟨fun x => H (x.1.down, x.2), - H.continuous.comp ((continuous_induced_dom.comp continuous_fst).prod_mk continuous_snd)⟩ + H.continuous.comp ((continuous_uLift_down.comp continuous_fst).prod_mk continuous_snd)⟩ #align continuous_map.homotopy.ulift_map ContinuousMap.Homotopy.uliftMap -- This lemma has always been bad, but the linter only noticed after lean4#2644. diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 6a070f62ee43a..8a8945c8e29a6 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -499,7 +499,7 @@ theorem comp_summable_nnreal (q : FormalMultilinearSeries 𝕜 F G) (p : FormalM simp only [Finset.prod_mul_distrib, Finset.prod_pow_eq_pow_sum, c.sum_blocksFun] _ ≤ ∏ _i : Fin c.length, Cp := (Finset.prod_le_prod' fun i _ => hCp _) _ = Cp ^ c.length := by simp - _ ≤ Cp ^ n := pow_le_pow hCp1 c.length_le + _ ≤ Cp ^ n := pow_le_pow_right hCp1 c.length_le calc ‖q.compAlongComposition p c‖₊ * r ^ n ≤ (‖q c.length‖₊ * ∏ i, ‖p (c.blocksFun i)‖₊) * r ^ n := @@ -842,7 +842,7 @@ theorem HasFPowerSeriesAt.comp {g : F → G} {f : E → F} {q : FormalMultilinea _ ≤ ‖compAlongComposition q p c‖ * (r : ℝ) ^ n := by apply mul_le_mul_of_nonneg_left _ (norm_nonneg _) rw [Finset.prod_const, Finset.card_fin] - apply pow_le_pow_of_le_left (norm_nonneg _) + apply pow_le_pow_left (norm_nonneg _) rw [EMetric.mem_ball, edist_eq_coe_nnnorm] at hy have := le_trans (le_of_lt hy) (min_le_right _ _) rwa [ENNReal.coe_le_coe, ← NNReal.coe_le_coe, coe_nnnorm] at this diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index 4a6ecf09d8d13..4f8636d9e4510 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -1664,14 +1664,13 @@ theorem IsBigOWith.pow [NormOneClass R] {f : α → R} {g : α → 𝕜} (h : Is theorem IsBigOWith.of_pow {n : ℕ} {f : α → 𝕜} {g : α → R} (h : IsBigOWith c l (f ^ n) (g ^ n)) (hn : n ≠ 0) (hc : c ≤ c' ^ n) (hc' : 0 ≤ c') : IsBigOWith c' l f g := - IsBigOWith.of_bound <| - (h.weaken hc).bound.mono fun x hx => - le_of_pow_le_pow n (mul_nonneg hc' <| norm_nonneg _) hn.bot_lt <| - calc - ‖f x‖ ^ n = ‖f x ^ n‖ := (norm_pow _ _).symm - _ ≤ c' ^ n * ‖g x ^ n‖ := hx - _ ≤ c' ^ n * ‖g x‖ ^ n := by gcongr; exact norm_pow_le' _ hn.bot_lt - _ = (c' * ‖g x‖) ^ n := (mul_pow _ _ _).symm + IsBigOWith.of_bound <| (h.weaken hc).bound.mono fun x hx ↦ + le_of_pow_le_pow_left hn (by positivity) <| + calc + ‖f x‖ ^ n = ‖f x ^ n‖ := (norm_pow _ _).symm + _ ≤ c' ^ n * ‖g x ^ n‖ := hx + _ ≤ c' ^ n * ‖g x‖ ^ n := by gcongr; exact norm_pow_le' _ hn.bot_lt + _ = (c' * ‖g x‖) ^ n := (mul_pow _ _ _).symm #align asymptotics.is_O_with.of_pow Asymptotics.IsBigOWith.of_pow theorem IsBigO.pow {f : α → R} {g : α → 𝕜} (h : f =O[l] g) (n : ℕ) : diff --git a/Mathlib/Analysis/BoxIntegral/Basic.lean b/Mathlib/Analysis/BoxIntegral/Basic.lean index f3c8fe8c90b1a..5f127649461bd 100644 --- a/Mathlib/Analysis/BoxIntegral/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Basic.lean @@ -766,7 +766,7 @@ theorem HasIntegral.of_bRiemann_eq_false_of_forall_isLittleO (hl : l.bRiemann = refine' (sum_le_sum _).trans (hεs _ _) · rintro b - rw [← Nat.cast_two, ← Nat.cast_pow, ← nsmul_eq_mul] - refine' nsmul_le_nsmul (hεs0 _).le _ + refine' nsmul_le_nsmul_left (hεs0 _).le _ refine' (Finset.card_le_of_subset _).trans ((hπδ.isHenstock hlH).card_filter_tag_eq_le b) exact filter_subset_filter _ (filter_subset _ _) · rw [Finset.coe_image, Set.image_subset_iff] diff --git a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean index 21b0746eb5faf..1e45579362f25 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean @@ -328,7 +328,7 @@ theorem D_subset_differentiable_set {K : Set (E →L[𝕜] F)} (hK : IsComplete -- the scale is large enough (as `y` is small enough) have k_gt : n e < k := by have : ((1 : ℝ) / 2) ^ (k + 1) < (1 / 2) ^ (n e + 1) := lt_trans hk y_lt - rw [pow_lt_pow_iff_of_lt_one (by norm_num : (0 : ℝ) < 1 / 2) (by norm_num)] at this + rw [pow_lt_pow_iff_right_of_lt_one (by norm_num : (0 : ℝ) < 1 / 2) (by norm_num)] at this linarith set m := k - 1 have m_ge : n e ≤ m := Nat.le_sub_one_of_lt k_gt @@ -691,7 +691,7 @@ theorem D_subset_differentiable_set {K : Set F} (hK : IsComplete K) : -- the scale is large enough (as `y - x` is small enough) have k_gt : n e < k := by have : ((1 : ℝ) / 2) ^ (k + 1) < (1 / 2) ^ (n e + 1) := lt_of_lt_of_le hk y_le - rw [pow_lt_pow_iff_of_lt_one (by norm_num : (0 : ℝ) < 1 / 2) (by norm_num)] at this + rw [pow_lt_pow_iff_right_of_lt_one (by norm_num : (0 : ℝ) < 1 / 2) (by norm_num)] at this linarith set m := k - 1 have m_ge : n e ≤ m := Nat.le_sub_one_of_lt k_gt diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index 56292525977b9..9e1c2f1a459ea 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -194,9 +194,8 @@ theorem nnnorm_int (n : ℤ) : ‖(n : ℂ)‖₊ = ‖n‖₊ := Subtype.ext norm_int #align complex.nnnorm_int Complex.nnnorm_int -theorem nnnorm_eq_one_of_pow_eq_one {ζ : ℂ} {n : ℕ} (h : ζ ^ n = 1) (hn : n ≠ 0) : ‖ζ‖₊ = 1 := by - refine' (@pow_left_inj NNReal _ _ _ _ zero_le' zero_le' hn.bot_lt).mp _ - rw [← nnnorm_pow, h, nnnorm_one, one_pow] +theorem nnnorm_eq_one_of_pow_eq_one {ζ : ℂ} {n : ℕ} (h : ζ ^ n = 1) (hn : n ≠ 0) : ‖ζ‖₊ = 1 := + (pow_left_inj zero_le' zero_le' hn).1 $ by rw [← nnnorm_pow, h, nnnorm_one, one_pow] #align complex.nnnorm_eq_one_of_pow_eq_one Complex.nnnorm_eq_one_of_pow_eq_one theorem norm_eq_one_of_pow_eq_one {ζ : ℂ} {n : ℕ} (h : ζ ^ n = 1) (hn : n ≠ 0) : ‖ζ‖ = 1 := diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index e5010c8add85b..7120d654c35e3 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -190,7 +190,7 @@ theorem cmp_dist_eq_cmp_dist_coe_center (z w : ℍ) (r : ℝ) : have hr₀' : 0 ≤ w.im * Real.sinh r := mul_nonneg w.im_pos.le (sinh_nonneg_iff.2 hr₀) have hzw₀ : 0 < 2 * z.im * w.im := mul_pos (mul_pos two_pos z.im_pos) w.im_pos simp only [← cosh_strictMonoOn.cmp_map_eq dist_nonneg hr₀, ← - (@strictMonoOn_pow ℝ _ _ two_pos).cmp_map_eq dist_nonneg hr₀', dist_coe_center_sq] + (pow_left_strictMonoOn two_ne_zero).cmp_map_eq dist_nonneg hr₀', dist_coe_center_sq] rw [← cmp_mul_pos_left hzw₀, ← cmp_sub_zero, ← mul_sub, ← cmp_add_right, zero_add] #align upper_half_plane.cmp_dist_eq_cmp_dist_coe_center UpperHalfPlane.cmp_dist_eq_cmp_dist_coe_center diff --git a/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean b/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean index f78ba29edff39..662229221fc00 100644 --- a/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean +++ b/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean @@ -52,14 +52,14 @@ theorem strictConvexOn_exp : StrictConvexOn ℝ univ exp := by calc exp y - exp x = exp y - exp y * exp (x - y) := by rw [← exp_add]; ring_nf _ = exp y * (1 - exp (x - y)) := by ring - _ < exp y * -(x - y) := by gcongr; linarith [add_one_lt_exp_of_nonzero h2.ne] + _ < exp y * -(x - y) := by gcongr; linarith [add_one_lt_exp h2.ne] _ = exp y * (y - x) := by ring · have h1 : 0 < z - y := by linarith rw [lt_div_iff h1] calc exp y * (z - y) < exp y * (exp (z - y) - 1) := by gcongr _ * ?_ - linarith [add_one_lt_exp_of_nonzero h1.ne'] + linarith [add_one_lt_exp h1.ne'] _ = exp (z - y) * exp y - exp y := by ring _ ≤ exp z - exp y := by rw [← exp_add]; ring_nf; rfl #align strict_convex_on_exp strictConvexOn_exp @@ -216,3 +216,16 @@ theorem strictConcaveOn_log_Iio : StrictConcaveOn ℝ (Iio 0) log := by _ = _ := by rw [log_neg_eq_log] #align strict_concave_on_log_Iio strictConcaveOn_log_Iio + +namespace Real + +lemma exp_mul_le_cosh_add_mul_sinh {t : ℝ} (ht : |t| ≤ 1) (x : ℝ) : + exp (t * x) ≤ cosh x + t * sinh x := by + rw [abs_le] at ht + calc + _ = exp ((1 + t) / 2 * x + (1 - t) / 2 * (-x)) := by ring_nf + _ ≤ (1 + t) / 2 * exp x + (1 - t) / 2 * exp (-x) := + convexOn_exp.2 (Set.mem_univ _) (Set.mem_univ _) (by linarith) (by linarith) $ by ring + _ = _ := by rw [cosh_eq, sinh_eq]; ring + +end Real diff --git a/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean b/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean index 95872b239c3f4..101bc4ca6a098 100644 --- a/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean +++ b/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean @@ -38,9 +38,8 @@ open scoped BigOperators NNReal theorem strictConvexOn_pow {n : ℕ} (hn : 2 ≤ n) : StrictConvexOn ℝ (Ici 0) fun x : ℝ => x ^ n := by apply StrictMonoOn.strictConvexOn_of_deriv (convex_Ici _) (continuousOn_pow _) rw [deriv_pow', interior_Ici] - exact fun x (hx : 0 < x) y hy hxy => - mul_lt_mul_of_pos_left (pow_lt_pow_of_lt_left hxy hx.le <| Nat.sub_pos_of_lt hn) - (Nat.cast_pos.2 <| zero_lt_two.trans_le hn) + exact fun x (hx : 0 < x) y _ hxy => mul_lt_mul_of_pos_left + (pow_lt_pow_left hxy hx.le <| Nat.sub_ne_zero_of_lt hn) (by positivity) #align strict_convex_on_pow strictConvexOn_pow /-- `x^n`, `n : ℕ` is strictly convex on the whole real line whenever `n ≠ 0` is even. -/ diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 45dab0b1f2d27..87a3d62095514 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -629,7 +629,7 @@ theorem _root_.Function.HasTemperateGrowth.norm_iteratedFDeriv_le_uniform_aux {f · simp only [Finset.le_sup'_iff, le_max_iff] right exact ⟨N, hN, rfl.le⟩ - refine' pow_le_pow (by simp only [le_add_iff_nonneg_right, norm_nonneg]) _ + refine' pow_le_pow_right (by simp only [le_add_iff_nonneg_right, norm_nonneg]) _ exact Finset.le_sup hN #align function.has_temperate_growth.norm_iterated_fderiv_le_uniform_aux Function.HasTemperateGrowth.norm_iteratedFDeriv_le_uniform_aux @@ -777,7 +777,7 @@ def bilinLeftCLM (B : E →L[ℝ] F →L[ℝ] G) {g : D → F} (hg : g.HasTemper refine' mul_le_mul_of_nonneg_right _ (norm_nonneg _) rw [pow_add] refine' mul_le_mul_of_nonneg_left _ (by positivity) - refine' pow_le_pow_of_le_left (norm_nonneg _) _ _ + refine' pow_le_pow_left (norm_nonneg _) _ _ simp only [zero_le_one, le_add_iff_nonneg_left]) #align schwartz_map.bilin_left_clm SchwartzMap.bilinLeftCLM @@ -819,7 +819,7 @@ def compCLM {g : D → E} (hg : g.HasTemperateGrowth) 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] - refine' pow_le_pow_of_le_left (by positivity) _ _ + refine' pow_le_pow_left (by positivity) _ _ rw [add_mul] refine' add_le_add _ (hg_upper' x) nth_rw 1 [← one_mul (1 : ℝ)] @@ -844,7 +844,7 @@ def compCLM {g : D → E} (hg : g.HasTemperateGrowth) simp only [le_add_iff_nonneg_right, norm_nonneg] have := norm_iteratedFDeriv_comp_le f.smooth' hg.1 le_top x hbound hgrowth' have hxk : ‖x‖ ^ k ≤ (1 + ‖x‖) ^ k := - pow_le_pow_of_le_left (norm_nonneg _) (by simp only [zero_le_one, le_add_iff_nonneg_left]) _ + pow_le_pow_left (norm_nonneg _) (by simp only [zero_le_one, le_add_iff_nonneg_left]) _ refine' le_trans (mul_le_mul hxk this (by positivity) (by positivity)) _ have rearrange : (1 + ‖x‖) ^ k * diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 9d0e64cc20201..4b9c1bf5d02a9 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -1200,7 +1200,7 @@ instance (priority := 100) InnerProductSpace.toUniformConvexSpace : UniformConve refine' le_sqrt_of_sq_le _ rw [sq, eq_sub_iff_add_eq.2 (parallelogram_law_with_norm ℝ x y), ← sq ‖x - y‖, hx, hy] ring_nf - exact sub_le_sub_left (pow_le_pow_of_le_left hε.le hxy _) 4⟩ + exact sub_le_sub_left (pow_le_pow_left hε.le hxy _) 4⟩ #align inner_product_space.to_uniform_convex_space InnerProductSpace.toUniformConvexSpace section Complex diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index a7663356774b9..a2ad682841a9e 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -805,7 +805,7 @@ theorem Orthonormal.exists_orthonormalBasis_extension (hv : Orthonormal 𝕜 (( ∃ (u : Finset E) (b : OrthonormalBasis u 𝕜 E), v ⊆ u ∧ ⇑b = ((↑) : u → E) := by obtain ⟨u₀, hu₀s, hu₀, hu₀_max⟩ := exists_maximal_orthonormal hv rw [maximal_orthonormal_iff_orthogonalComplement_eq_bot hu₀] at hu₀_max - have hu₀_finite : u₀.Finite := hu₀.linearIndependent.finite + have hu₀_finite : u₀.Finite := hu₀.linearIndependent.setFinite let u : Finset E := hu₀_finite.toFinset let fu : ↥u ≃ ↥u₀ := hu₀_finite.subtypeEquivToFinset.symm have hu : Orthonormal 𝕜 ((↑) : u → E) := by simpa using hu₀.comp _ fu.injective diff --git a/Mathlib/Analysis/InnerProductSpace/Projection.lean b/Mathlib/Analysis/InnerProductSpace/Projection.lean index 25c01f65fd1e6..fae9a45834782 100644 --- a/Mathlib/Analysis/InnerProductSpace/Projection.lean +++ b/Mathlib/Analysis/InnerProductSpace/Projection.lean @@ -492,7 +492,7 @@ def orthogonalProjection : E →L[𝕜] K := simp [eq_orthogonalProjectionFn_of_mem_of_inner_eq_zero hm ho] } 1 fun x => by simp only [one_mul, LinearMap.coe_mk] - refine' le_of_pow_le_pow 2 (norm_nonneg _) (by norm_num) _ + refine' le_of_pow_le_pow_left two_ne_zero (norm_nonneg _) _ change ‖orthogonalProjectionFn K x‖ ^ 2 ≤ ‖x‖ ^ 2 nlinarith [orthogonalProjectionFn_norm_sq K x] #align orthogonal_projection orthogonalProjection @@ -1190,7 +1190,8 @@ theorem LinearIsometryEquiv.reflections_generate_dim_aux [FiniteDimensional ℝ · -- Base case: `n = 0`, the fixed subspace is the whole space, so `φ = id` refine' ⟨[], rfl.le, show φ = 1 from _⟩ have : ker (ContinuousLinearMap.id ℝ F - φ) = ⊤ := by - rwa [Nat.zero_eq, le_zero_iff, finrank_eq_zero, Submodule.orthogonal_eq_bot_iff] at hn + rwa [Nat.zero_eq, le_zero_iff, Submodule.finrank_eq_zero, + Submodule.orthogonal_eq_bot_iff] at hn symm ext x have := LinearMap.congr_fun (LinearMap.ker_eq_top.mp this) x diff --git a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean index d1a8ee6492fe2..32a3c36a44303 100644 --- a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean +++ b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean @@ -448,7 +448,7 @@ theorem nonneg_inner_and_areaForm_eq_zero_iff_sameRay (x y : E) : suffices ↑0 ≤ a * ‖x‖ ^ 2 ∧ b * ‖x‖ ^ 2 = 0 → SameRay ℝ x (a • x + b • J x) by rw [← (o.basisRightAngleRotation x hx).sum_repr y] simp only [Fin.sum_univ_succ, coe_basisRightAngleRotation, Matrix.cons_val_zero, - Fin.succ_zero_eq_one', Fintype.univ_of_isEmpty, Finset.sum_empty, areaForm_apply_self, + Fin.succ_zero_eq_one', Finset.univ_eq_empty, Finset.sum_empty, areaForm_apply_self, map_smul, map_add, real_inner_smul_right, inner_add_right, Matrix.cons_val_one, Matrix.head_cons, Algebra.id.smul_eq_mul, areaForm_rightAngleRotation_right, mul_zero, add_zero, zero_add, neg_zero, inner_rightAngleRotation_right, diff --git a/Mathlib/Analysis/InnerProductSpace/l2Space.lean b/Mathlib/Analysis/InnerProductSpace/l2Space.lean index 32bae3a00162f..2b4a47b2e7dd9 100644 --- a/Mathlib/Analysis/InnerProductSpace/l2Space.lean +++ b/Mathlib/Analysis/InnerProductSpace/l2Space.lean @@ -584,7 +584,8 @@ theorem _root_.Orthonormal.exists_hilbertBasis_extension {s : Set E} ∃ (w : Set E) (b : HilbertBasis w 𝕜 E), s ⊆ w ∧ ⇑b = ((↑) : w → E) := let ⟨w, hws, hw_ortho, hw_max⟩ := exists_maximal_orthonormal hs ⟨w, HilbertBasis.mkOfOrthogonalEqBot hw_ortho - (by simpa [maximal_orthonormal_iff_orthogonalComplement_eq_bot hw_ortho] using hw_max), + (by simpa only [Subtype.range_coe_subtype, Set.setOf_mem_eq, + maximal_orthonormal_iff_orthogonalComplement_eq_bot hw_ortho] using hw_max), hws, HilbertBasis.coe_mkOfOrthogonalEqBot _ _⟩ #align orthonormal.exists_hilbert_basis_extension Orthonormal.exists_hilbertBasis_extension diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index e442b2c0557bc..7310ec13f143f 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -200,7 +200,7 @@ for two `NNReal` numbers. -/ theorem geom_mean_le_arith_mean2_weighted (w₁ w₂ p₁ p₂ : ℝ≥0) : w₁ + w₂ = 1 → p₁ ^ (w₁ : ℝ) * p₂ ^ (w₂ : ℝ) ≤ w₁ * p₁ + w₂ * p₂ := by simpa only [Fin.prod_univ_succ, Fin.sum_univ_succ, Finset.prod_empty, Finset.sum_empty, - Fintype.univ_of_isEmpty, Fin.cons_succ, Fin.cons_zero, add_zero, mul_one] using + Finset.univ_eq_empty, Fin.cons_succ, Fin.cons_zero, add_zero, mul_one] using geom_mean_le_arith_mean_weighted univ ![w₁, w₂] ![p₁, p₂] #align nnreal.geom_mean_le_arith_mean2_weighted NNReal.geom_mean_le_arith_mean2_weighted @@ -208,7 +208,7 @@ theorem geom_mean_le_arith_mean3_weighted (w₁ w₂ w₃ p₁ p₂ p₃ : ℝ w₁ + w₂ + w₃ = 1 → p₁ ^ (w₁ : ℝ) * p₂ ^ (w₂ : ℝ) * p₃ ^ (w₃ : ℝ) ≤ w₁ * p₁ + w₂ * p₂ + w₃ * p₃ := by simpa only [Fin.prod_univ_succ, Fin.sum_univ_succ, Finset.prod_empty, Finset.sum_empty, - Fintype.univ_of_isEmpty, Fin.cons_succ, Fin.cons_zero, add_zero, mul_one, ← add_assoc, + Finset.univ_eq_empty, Fin.cons_succ, Fin.cons_zero, add_zero, mul_one, ← add_assoc, mul_assoc] using geom_mean_le_arith_mean_weighted univ ![w₁, w₂, w₃] ![p₁, p₂, p₃] #align nnreal.geom_mean_le_arith_mean3_weighted NNReal.geom_mean_le_arith_mean3_weighted @@ -217,7 +217,7 @@ theorem geom_mean_le_arith_mean4_weighted (w₁ w₂ w₃ w₄ p₁ p₂ p₃ p p₁ ^ (w₁ : ℝ) * p₂ ^ (w₂ : ℝ) * p₃ ^ (w₃ : ℝ) * p₄ ^ (w₄ : ℝ) ≤ w₁ * p₁ + w₂ * p₂ + w₃ * p₃ + w₄ * p₄ := by simpa only [Fin.prod_univ_succ, Fin.sum_univ_succ, Finset.prod_empty, Finset.sum_empty, - Fintype.univ_of_isEmpty, Fin.cons_succ, Fin.cons_zero, add_zero, mul_one, ← add_assoc, + Finset.univ_eq_empty, Fin.cons_succ, Fin.cons_zero, add_zero, mul_one, ← add_assoc, mul_assoc] using geom_mean_le_arith_mean_weighted univ ![w₁, w₂, w₃, w₄] ![p₁, p₂, p₃, p₄] #align nnreal.geom_mean_le_arith_mean4_weighted NNReal.geom_mean_le_arith_mean4_weighted diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index 426ecac3abf1a..ee66ae3aeab78 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -101,6 +101,24 @@ instance (priority := 100) NormedRing.toNonUnitalNormedRing [β : NormedRing α] { β with } #align normed_ring.to_non_unital_normed_ring NormedRing.toNonUnitalNormedRing +/-- A non-unital seminormed commutative ring is a non-unital commutative ring endowed with a +seminorm which satisfies the inequality `‖x y‖ ≤ ‖x‖ ‖y‖`. -/ +class NonUnitalSeminormedCommRing (α : Type*) extends NonUnitalSeminormedRing α where + /-- Multiplication is commutative. -/ + mul_comm : ∀ x y : α, x * y = y * x + +/-- A non-unital normed commutative ring is a non-unital commutative ring endowed with a +norm which satisfies the inequality `‖x y‖ ≤ ‖x‖ ‖y‖`. -/ +class NonUnitalNormedCommRing (α : Type*) extends NonUnitalNormedRing α where + /-- Multiplication is commutative. -/ + mul_comm : ∀ x y : α, x * y = y * x + +-- see Note [lower instance priority] +/-- A non-unital normed commutative ring is a non-unital seminormed commutative ring. -/ +instance (priority := 100) NonUnitalNormedCommRing.toNonUnitalSeminormedCommRing + [β : NonUnitalNormedCommRing α] : NonUnitalSeminormedCommRing α := + { β with } + /-- A seminormed commutative ring is a commutative ring endowed with a seminorm which satisfies the inequality `‖x y‖ ≤ ‖x‖ ‖y‖`. -/ class SeminormedCommRing (α : Type*) extends SeminormedRing α where @@ -115,6 +133,18 @@ class NormedCommRing (α : Type*) extends NormedRing α where mul_comm : ∀ x y : α, x * y = y * x #align normed_comm_ring NormedCommRing +-- see Note [lower instance priority] +/-- A seminormed commutative ring is a non-unital seminormed commutative ring. -/ +instance (priority := 100) SeminormedCommRing.toNonUnitalSeminormedCommRing + [β : SeminormedCommRing α] : NonUnitalSeminormedCommRing α := + { β with } + +-- see Note [lower instance priority] +/-- A normed commutative ring is a non-unital normed commutative ring. -/ +instance (priority := 100) NormedCommRing.toNonUnitalNormedCommRing + [β : NormedCommRing α] : NonUnitalNormedCommRing α := + { β with } + -- see Note [lower instance priority] /-- A normed commutative ring is a seminormed commutative ring. -/ instance (priority := 100) NormedCommRing.toSeminormedCommRing [β : NormedCommRing α] : @@ -147,6 +177,11 @@ theorem NormOneClass.nontrivial (α : Type*) [SeminormedAddCommGroup α] [One α nontrivial_of_ne 0 1 <| ne_of_apply_ne norm <| by simp #align norm_one_class.nontrivial NormOneClass.nontrivial +-- see Note [lower instance priority] +instance (priority := 100) NonUnitalSeminormedCommRing.toNonUnitalCommRing + [β : NonUnitalSeminormedCommRing α] : NonUnitalCommRing α := + { β with } + -- see Note [lower instance priority] instance (priority := 100) SeminormedCommRing.toCommRing [β : SeminormedCommRing α] : CommRing α := { β with } @@ -463,6 +498,53 @@ instance MulOpposite.normedRing : NormedRing αᵐᵒᵖ := end NormedRing +section NonUnitalSeminormedCommRing + +variable [NonUnitalSeminormedCommRing α] + +instance ULift.nonUnitalSeminormedCommRing : NonUnitalSeminormedCommRing (ULift α) := + { ULift.nonUnitalSeminormedRing, ULift.nonUnitalCommRing with } + +/-- Non-unital seminormed commutative ring structure on the product of two non-unital seminormed +commutative rings, using the sup norm. -/ +instance Prod.nonUnitalSeminormedCommRing [NonUnitalSeminormedCommRing β] : + NonUnitalSeminormedCommRing (α × β) := + { nonUnitalSeminormedRing, instNonUnitalCommRing with } + +/-- Non-unital seminormed commutative ring structure on the product of finitely many non-unital +seminormed commutative rings, using the sup norm. -/ +instance Pi.nonUnitalSeminormedCommRing {π : ι → Type*} [Fintype ι] + [∀ i, NonUnitalSeminormedCommRing (π i)] : NonUnitalSeminormedCommRing (∀ i, π i) := + { Pi.nonUnitalSeminormedRing, Pi.nonUnitalCommRing with } + +instance MulOpposite.nonUnitalSeminormedCommRing : NonUnitalSeminormedCommRing αᵐᵒᵖ := + { MulOpposite.nonUnitalSeminormedRing, MulOpposite.nonUnitalCommRing α with } + +end NonUnitalSeminormedCommRing +section NonUnitalNormedCommRing + +variable [NonUnitalNormedCommRing α] + +instance ULift.nonUnitalNormedCommRing : NonUnitalNormedCommRing (ULift α) := + { ULift.nonUnitalSeminormedCommRing, ULift.normedAddCommGroup with } + +/-- Non-unital normed commutative ring structure on the product of two non-unital normed +commutative rings, using the sup norm. -/ +instance Prod.nonUnitalNormedCommRing [NonUnitalNormedCommRing β] : + NonUnitalNormedCommRing (α × β) := + { Prod.nonUnitalSeminormedCommRing, Prod.normedAddCommGroup with } + +/-- Normed commutative ring structure on the product of finitely many non-unital normed +commutative rings, using the sup norm. -/ +instance Pi.nonUnitalNormedCommRing {π : ι → Type*} [Fintype ι] + [∀ i, NonUnitalNormedCommRing (π i)] : NonUnitalNormedCommRing (∀ i, π i) := + { Pi.nonUnitalSeminormedCommRing, Pi.normedAddCommGroup with } + +instance MulOpposite.nonUnitalNormedCommRing : NonUnitalNormedCommRing αᵐᵒᵖ := + { MulOpposite.nonUnitalSeminormedCommRing, MulOpposite.normedAddCommGroup with } + +end NonUnitalNormedCommRing + -- see Note [lower instance priority] instance (priority := 100) semi_normed_ring_top_monoid [NonUnitalSeminormedRing α] : ContinuousMul α := @@ -707,7 +789,7 @@ instance (priority := 100) NormedDivisionRing.to_topologicalDivisionRing : Topol theorem norm_map_one_of_pow_eq_one [Monoid β] (φ : β →* α) {x : β} {k : ℕ+} (h : x ^ (k : ℕ) = 1) : ‖φ x‖ = 1 := by rw [← pow_left_inj, ← norm_pow, ← map_pow, h, map_one, norm_one, one_pow] - exacts [norm_nonneg _, zero_le_one, k.pos] + exacts [norm_nonneg _, zero_le_one, k.ne_zero] #align norm_map_one_of_pow_eq_one norm_map_one_of_pow_eq_one theorem norm_one_of_pow_eq_one {x : α} {k : ℕ+} (h : x ^ (k : ℕ) = 1) : ‖x‖ = 1 := @@ -1021,6 +1103,23 @@ def NormedRing.induced [Ring R] [NormedRing S] [NonUnitalRingHomClass F R S] (f { NonUnitalSeminormedRing.induced R S f, NormedAddCommGroup.induced R S f hf, ‹Ring R› with } #align normed_ring.induced NormedRing.induced +/-- A non-unital ring homomorphism from a `NonUnitalCommRing` to a `NonUnitalSeminormedCommRing` +induces a `NonUnitalSeminormedCommRing` structure on the domain. + +See note [reducible non-instances] -/ +@[reducible] +def NonUnitalSeminormedCommRing.induced [NonUnitalCommRing R] [NonUnitalSeminormedCommRing S] + [NonUnitalRingHomClass F R S] (f : F) : NonUnitalSeminormedCommRing R := + { NonUnitalSeminormedRing.induced R S f, ‹NonUnitalCommRing R› with } + +/-- An injective non-unital ring homomorphism from a `NonUnitalCommRing` to a +`NonUnitalNormedCommRing` induces a `NonUnitalNormedCommRing` structure on the domain. + +See note [reducible non-instances] -/ +@[reducible] +def NonUnitalNormedCommRing.induced [NonUnitalCommRing R] [NonUnitalNormedCommRing S] + [NonUnitalRingHomClass F R S] (f : F) (hf : Function.Injective f) : NonUnitalNormedCommRing R := + { NonUnitalNormedRing.induced R S f hf, ‹NonUnitalCommRing R› with } /-- A non-unital ring homomorphism from a `CommRing` to a `SeminormedRing` induces a `SeminormedCommRing` structure on the domain. diff --git a/Mathlib/Analysis/NormedSpace/FiniteDimension.lean b/Mathlib/Analysis/NormedSpace/FiniteDimension.lean index de87532d58585..a904d1090e685 100644 --- a/Mathlib/Analysis/NormedSpace/FiniteDimension.lean +++ b/Mathlib/Analysis/NormedSpace/FiniteDimension.lean @@ -283,7 +283,7 @@ theorem Basis.op_nnnorm_le {ι : Type*} [Fintype ι] (v : Basis ι 𝕜 E) {u : gcongr calc ∑ i, ‖v.equivFun e i‖₊ ≤ Fintype.card ι • ‖φ e‖₊ := Pi.sum_nnnorm_apply_le_nnnorm _ - _ ≤ Fintype.card ι • (‖φ‖₊ * ‖e‖₊) := nsmul_le_nsmul_of_le_right (φ.le_op_nnnorm e) _ + _ ≤ Fintype.card ι • (‖φ‖₊ * ‖e‖₊) := nsmul_le_nsmul_right (φ.le_op_nnnorm e) _ _ = Fintype.card ι • ‖φ‖₊ * M * ‖e‖₊ := by simp only [smul_mul_assoc, mul_right_comm] #align basis.op_nnnorm_le Basis.op_nnnorm_le diff --git a/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean b/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean index fb07426e93fcc..0062fb9613dbe 100644 --- a/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean +++ b/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean @@ -5,6 +5,7 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Analysis.NormedSpace.HahnBanach.Extension import Mathlib.Analysis.NormedSpace.HahnBanach.Separation +import Mathlib.LinearAlgebra.Dual /-! # Spaces with separating dual @@ -22,6 +23,7 @@ equivalences acts transitively on the set of nonzero vectors. /-- When `E` is a topological module over a topological ring `R`, the class `SeparatingDual R E` registers that continuous linear forms on `E` separate points of `E`. -/ +@[mk_iff separatingDual_def] class SeparatingDual (R V : Type*) [Ring R] [AddCommGroup V] [TopologicalSpace V] [TopologicalSpace R] [Module R V] : Prop := /-- Any nonzero vector can be mapped by a continuous linear map to a nonzero scalar. -/ @@ -73,6 +75,26 @@ section Field variable {R V : Type*} [Field R] [AddCommGroup V] [TopologicalSpace R] [TopologicalSpace V] [TopologicalRing R] [TopologicalAddGroup V] [Module R V] [SeparatingDual R V] +-- TODO (@alreadydone): this could generalize to CommRing R if we were to add a section +theorem _root_.separatingDual_iff_injective : SeparatingDual R V ↔ + Function.Injective (ContinuousLinearMap.coeLM (R := R) R (M := V) (N₃ := R)).flip := by + simp_rw [separatingDual_def, Ne, injective_iff_map_eq_zero] + congrm ∀ v, ?_ + rw [not_imp_comm, LinearMap.ext_iff] + push_neg; rfl + +open Function in +/-- Given a finite-dimensional subspace `W` of a space `V` with separating dual, any + linear functional on `W` extends to a continuous linear functional on `V`. + This is stated more generally for an injective linear map from `W` to `V`. -/ +theorem dualMap_surjective_iff {W} [AddCommGroup W] [Module R W] [FiniteDimensional R W] + {f : W →ₗ[R] V} : Surjective (f.dualMap ∘ ContinuousLinearMap.toLinearMap) ↔ Injective f := by + constructor <;> intro hf + · exact LinearMap.dualMap_surjective_iff.mp hf.of_comp + have := (separatingDual_iff_injective.mp ‹_›).comp hf + rw [← LinearMap.coe_comp] at this + exact LinearMap.flip_surjective_iff₁.mpr this + lemma exists_eq_one {x : V} (hx : x ≠ 0) : ∃ f : V →L[R] R, f x = 1 := by rcases exists_ne_zero (R := R) hx with ⟨f, hf⟩ diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 703113e545b05..193888790ef6c 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -663,7 +663,7 @@ theorem continuous_eval : Continuous _ ≤ (‖p‖ + 1) * Fintype.card ι * (‖p‖ + 1) ^ (Fintype.card ι - 1) * ‖q - p‖ + ‖q - p‖ * ∏ i, ‖p.2 i‖ := by apply_rules [add_le_add, mul_le_mul, le_refl, le_trans (norm_fst_le q) A, Nat.cast_nonneg, - mul_nonneg, pow_le_pow_of_le_left, pow_nonneg, norm_snd_le (q - p), norm_nonneg, + mul_nonneg, pow_le_pow_left, pow_nonneg, norm_snd_le (q - p), norm_nonneg, norm_fst_le (q - p), prod_nonneg] _ = ((‖p‖ + 1) * Fintype.card ι * (‖p‖ + 1) ^ (Fintype.card ι - 1) + ∏ i, ‖p.2 i‖) * dist q p := by diff --git a/Mathlib/Analysis/NormedSpace/Star/Multiplier.lean b/Mathlib/Analysis/NormedSpace/Star/Multiplier.lean index a43fb4eda6c5e..7a790033364cb 100644 --- a/Mathlib/Analysis/NormedSpace/Star/Multiplier.lean +++ b/Mathlib/Analysis/NormedSpace/Star/Multiplier.lean @@ -709,7 +709,7 @@ instance instCstarRing : CstarRing 𝓜(𝕜, A) where rintro - ⟨y, hy, rfl⟩ exact key (star x) y ((nnnorm_star x).trans_le hx') (mem_closedBall_zero_iff.1 hy) · simpa only [a.central, star_star, CstarRing.nnnorm_star_mul_self, NNReal.sq_sqrt, ← sq] - using pow_lt_pow_of_lt_left hxr zero_le' two_pos + using pow_lt_pow_left hxr zero_le' two_ne_zero end DenselyNormed diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index 6a4e0d1481561..3739a9197021f 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -86,7 +86,7 @@ theorem sum_condensed_le' (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → f n ≤ f m theorem sum_condensed_le (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → f n ≤ f m) (n : ℕ) : (∑ k in range (n + 1), 2 ^ k • f (2 ^ k)) ≤ f 1 + 2 • ∑ k in Ico 2 (2 ^ n + 1), f k := by - convert add_le_add_left (nsmul_le_nsmul_of_le_right (sum_condensed_le' hf n) 2) (f 1) + convert add_le_add_left (nsmul_le_nsmul_right (sum_condensed_le' hf n) 2) (f 1) simp [sum_range_succ', add_comm, pow_succ, mul_nsmul', sum_nsmul] #align finset.sum_condensed_le Finset.sum_condensed_le @@ -111,7 +111,7 @@ theorem tsum_condensed_le (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → f n ≤ f m iSup_le fun n => le_trans _ (add_le_add_left - (nsmul_le_nsmul_of_le_right (ENNReal.sum_le_tsum <| Finset.Ico 2 (2 ^ n + 1)) _) _) + (nsmul_le_nsmul_right (ENNReal.sum_le_tsum <| Finset.Ico 2 (2 ^ n + 1)) _) _) simpa using Finset.sum_condensed_le hf n #align ennreal.tsum_condensed_le ENNReal.tsum_condensed_le @@ -332,7 +332,7 @@ theorem sum_Ioo_inv_sq_le (k n : ℕ) : (∑ i in Ioo k n, (i ^ 2 : α)⁻¹) have A : (1 : α) ≤ k + 1 := by simp only [le_add_iff_nonneg_left, Nat.cast_nonneg] simp_rw [← one_div] gcongr - simpa using pow_le_pow A one_le_two + simpa using pow_le_pow_right A one_le_two _ = 2 / (k + 1) := by ring #align sum_Ioo_inv_sq_le sum_Ioo_inv_sq_le diff --git a/Mathlib/Analysis/SpecialFunctions/Bernstein.lean b/Mathlib/Analysis/SpecialFunctions/Bernstein.lean index 11b25b4bee846..29acf116dd222 100644 --- a/Mathlib/Analysis/SpecialFunctions/Bernstein.lean +++ b/Mathlib/Analysis/SpecialFunctions/Bernstein.lean @@ -236,7 +236,6 @@ theorem bernsteinApproximation_uniform (f : C(I, ℝ)) : filter_upwards [nhds_zero.eventually (gt_mem_nhds (half_pos h)), eventually_gt_atTop 0] with n nh npos' have npos : 0 < (n : ℝ) := by positivity - have w₂ : 0 ≤ δ ^ (-2:ℤ) := zpow_neg_two_nonneg _ -- TODO: need a positivity extension for `zpow` -- As `[0,1]` is compact, it suffices to check the inequality pointwise. rw [ContinuousMap.norm_lt_iff _ h] intro x diff --git a/Mathlib/Analysis/SpecialFunctions/CompareExp.lean b/Mathlib/Analysis/SpecialFunctions/CompareExp.lean index 09612131c4d31..7e6bae15b2529 100644 --- a/Mathlib/Analysis/SpecialFunctions/CompareExp.lean +++ b/Mathlib/Analysis/SpecialFunctions/CompareExp.lean @@ -136,7 +136,7 @@ theorem isLittleO_log_abs_re (hl : IsExpCmpFilter l) : (fun z => Real.log (abs z have hm₀ : 0 < max z.re |z.im| := lt_max_iff.2 (Or.inl <| one_pos.trans_le hz) rw [one_mul, Real.norm_eq_abs, _root_.abs_of_nonneg (Real.log_nonneg hz')] refine' le_trans _ (le_abs_self _) - rw [← Real.log_mul, Real.log_le_log, ← _root_.abs_of_nonneg (le_trans zero_le_one hz)] + rw [← Real.log_mul, Real.log_le_log_iff, ← _root_.abs_of_nonneg (le_trans zero_le_one hz)] exacts [abs_le_sqrt_two_mul_max z, one_pos.trans_le hz', mul_pos h2 hm₀, h2.ne', hm₀.ne'] _ =o[l] re := IsLittleO.add (isLittleO_const_left.2 <| Or.inr <| hl.tendsto_abs_re) <| diff --git a/Mathlib/Analysis/SpecialFunctions/Exp.lean b/Mathlib/Analysis/SpecialFunctions/Exp.lean index 7aacf5b2df0a8..4dcce95fbf30f 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exp.lean @@ -225,7 +225,7 @@ theorem tendsto_exp_div_pow_atTop (n : ℕ) : Tendsto (fun x => exp x / x ^ n) a have hx₀ : 0 < x := (Nat.cast_nonneg N).trans_lt hx rw [Set.mem_Ici, le_div_iff (pow_pos hx₀ _), ← le_div_iff' hC₀] calc - x ^ n ≤ ⌈x⌉₊ ^ n := mod_cast pow_le_pow_of_le_left hx₀.le (Nat.le_ceil _) _ + x ^ n ≤ ⌈x⌉₊ ^ n := mod_cast pow_le_pow_left hx₀.le (Nat.le_ceil _) _ _ ≤ exp ⌈x⌉₊ / (exp 1 * C) := mod_cast (hN _ (Nat.lt_ceil.2 hx).le).le _ ≤ exp (x + 1) / (exp 1 * C) := (div_le_div_of_le (mul_pos (exp_pos _) hC₀).le diff --git a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean index b73d3edc39161..ac0728f081189 100644 --- a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean +++ b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean @@ -83,7 +83,7 @@ theorem finite_integral_rpow_sub_one_pow_aux {r : ℝ} (n : ℕ) (hnr : (n : ℝ ENNReal.ofReal ((x ^ (-r⁻¹) - 1) ^ n) ≤ ENNReal.ofReal (x ^ (-(r⁻¹ * n))) := fun x hx ↦ by apply ENNReal.ofReal_le_ofReal rw [← neg_mul, rpow_mul hx.1.le, rpow_nat_cast] - refine' pow_le_pow_of_le_left _ (by simp only [sub_le_self_iff, zero_le_one]) n + refine' pow_le_pow_left _ (by simp only [sub_le_self_iff, zero_le_one]) n rw [le_sub_iff_add_le', add_zero] refine' Real.one_le_rpow_of_pos_of_le_one_of_nonpos hx.1 hx.2 _ rw [Right.neg_nonpos_iff, inv_nonneg] diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean index c5a2600b595db..e51f3d6dfe94e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean @@ -182,7 +182,7 @@ private theorem b_ne_one' : b ≠ 1 := by linarith @[simp] theorem logb_le_logb (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ x ≤ y := by - rw [logb, logb, div_le_div_right (log_pos hb), log_le_log h h₁] + rw [logb, logb, div_le_div_right (log_pos hb), log_le_log_iff h h₁] #align real.logb_le_logb Real.logb_le_logb @[gcongr] @@ -295,7 +295,7 @@ private theorem b_ne_one : b ≠ 1 := by linarith @[simp] theorem logb_le_logb_of_base_lt_one (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ y ≤ x := by - rw [logb, logb, div_le_div_right_of_neg (log_neg b_pos b_lt_one), log_le_log h₁ h] + rw [logb, logb, div_le_div_right_of_neg (log_neg b_pos b_lt_one), log_le_log_iff h₁ h] #align real.logb_le_logb_of_base_lt_one Real.logb_le_logb_of_base_lt_one theorem logb_lt_logb_of_base_lt_one (hx : 0 < x) (hxy : x < y) : logb b y < logb b x := by diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean index 487700b077ee6..24ba679fa2b05 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean @@ -141,22 +141,18 @@ theorem log_inv (x : ℝ) : log x⁻¹ = -log x := by rw [← exp_eq_exp, exp_log_eq_abs (inv_ne_zero hx), exp_neg, exp_log_eq_abs hx, abs_inv] #align real.log_inv Real.log_inv -theorem log_le_log (h : 0 < x) (h₁ : 0 < y) : log x ≤ log y ↔ x ≤ y := by +theorem log_le_log_iff (h : 0 < x) (h₁ : 0 < y) : log x ≤ log y ↔ x ≤ y := by rw [← exp_le_exp, exp_log h, exp_log h₁] -#align real.log_le_log Real.log_le_log +#align real.log_le_log Real.log_le_log_iff @[gcongr] -theorem log_lt_log (hx : 0 < x) : x < y → log x < log y := by - intro h - rwa [← exp_lt_exp, exp_log hx, exp_log (lt_trans hx h)] -#align real.log_lt_log Real.log_lt_log +lemma log_le_log (hx : 0 < x) (hxy : x ≤ y) : log x ≤ log y := + (log_le_log_iff hx (hx.trans_le hxy)).2 hxy @[gcongr] -theorem log_le_log' (hx : 0 < x) : x ≤ y → log x ≤ log y := by - intro hxy - cases hxy.eq_or_lt with - | inl h_eq => simp [h_eq] - | inr hlt => exact le_of_lt <| log_lt_log hx hlt +theorem log_lt_log (hx : 0 < x) (h : x < y) : log x < log y := by + rwa [← exp_lt_exp, exp_log hx, exp_log (lt_trans hx h)] +#align real.log_lt_log Real.log_lt_log theorem log_lt_log_iff (hx : 0 < x) (hy : 0 < y) : log x < log y ↔ x < y := by rw [← exp_lt_exp, exp_log hx, exp_log hy] @@ -265,7 +261,7 @@ theorem log_lt_sub_one_of_pos (hx1 : 0 < x) (hx2 : x ≠ 1) : log x < x - 1 := b have h : log x ≠ 0 · rwa [← log_one, log_injOn_pos.ne_iff hx1] exact mem_Ioi.mpr zero_lt_one - linarith [add_one_lt_exp_of_nonzero h, exp_log hx1] + linarith [add_one_lt_exp h, exp_log hx1] #align real.log_lt_sub_one_of_pos Real.log_lt_sub_one_of_pos theorem eq_one_of_pos_of_log_eq_zero {x : ℝ} (h₁ : 0 < x) (h₂ : log x = 0) : x = 1 := diff --git a/Mathlib/Analysis/SpecialFunctions/Log/NegMulLog.lean b/Mathlib/Analysis/SpecialFunctions/Log/NegMulLog.lean new file mode 100644 index 0000000000000..066f81e187631 --- /dev/null +++ b/Mathlib/Analysis/SpecialFunctions/Log/NegMulLog.lean @@ -0,0 +1,133 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Analysis.SpecialFunctions.Log.Deriv +import Mathlib.Analysis.SpecialFunctions.Pow.Asymptotics + +/-! +# The function `x ↦ - x * log x` + +The purpose of this file is to record basic analytic properties of the function `x ↦ - x * log x`, +which is notably used in the theory of Shannon entropy. + +## Main definitions + +* `negMulLog`: the function `x ↦ - x * log x` from `ℝ` to `ℝ`. + +-/ + +open scoped Topology + +namespace Real + +lemma continuous_mul_log : Continuous fun x ↦ x * log x := by + rw [continuous_iff_continuousAt] + intro x + obtain hx | rfl := ne_or_eq x 0 + · exact (continuous_id'.continuousAt).mul (continuousAt_log hx) + rw [ContinuousAt, zero_mul] + simp_rw [mul_comm _ (log _)] + nth_rewrite 1 [← nhdsWithin_univ] + have : (Set.univ : Set ℝ) = Set.Iio 0 ∪ Set.Ioi 0 ∪ {0} := by ext; simp [em] + rw [this, nhdsWithin_union, nhdsWithin_union] + simp only [nhdsWithin_singleton, sup_le_iff, Filter.nonpos_iff, Filter.tendsto_sup] + refine ⟨⟨tendsto_log_mul_self_nhds_zero_left, ?_⟩, ?_⟩ + · simpa only [rpow_one] using tendsto_log_mul_rpow_nhds_zero zero_lt_one + · convert tendsto_pure_nhds (fun x ↦ log x * x) 0 + simp + +lemma differentiableOn_mul_log : DifferentiableOn ℝ (fun x ↦ x * log x) {0}ᶜ := + differentiable_id'.differentiableOn.mul differentiableOn_log + +lemma deriv_mul_log {x : ℝ} (hx : x ≠ 0) : deriv (fun x ↦ x * log x) x = log x + 1 := by + rw [deriv_mul differentiableAt_id' (differentiableAt_log hx)] + simp only [deriv_id'', one_mul, deriv_log', ne_eq, add_right_inj] + exact mul_inv_cancel hx + +lemma hasDerivAt_mul_log {x : ℝ} (hx : x ≠ 0) : HasDerivAt (fun x ↦ x * log x) (log x + 1) x := by + rw [← deriv_mul_log hx, hasDerivAt_deriv_iff] + refine DifferentiableOn.differentiableAt differentiableOn_mul_log ?_ + simp [hx] + +lemma deriv2_mul_log {x : ℝ} (hx : x ≠ 0) : deriv^[2] (fun x ↦ x * log x) x = x⁻¹ := by + simp only [Function.iterate_succ, Function.iterate_zero, Function.comp.left_id, + Function.comp_apply] + suffices ∀ᶠ y in (𝓝 x), deriv (fun x ↦ x * log x) y = log y + 1 by + refine (Filter.EventuallyEq.deriv_eq this).trans ?_ + rw [deriv_add_const, deriv_log x] + filter_upwards [eventually_ne_nhds hx] with y hy using deriv_mul_log hy + +lemma strictConvexOn_mul_log : StrictConvexOn ℝ (Set.Ici (0 : ℝ)) (fun x ↦ x * log x) := by + refine strictConvexOn_of_deriv2_pos (convex_Ici 0) (continuous_mul_log.continuousOn) ?_ + intro x hx + simp only [Set.nonempty_Iio, interior_Ici', Set.mem_Ioi] at hx + rw [deriv2_mul_log hx.ne'] + positivity + +lemma convexOn_mul_log : ConvexOn ℝ (Set.Ici (0 : ℝ)) (fun x ↦ x * log x) := + strictConvexOn_mul_log.convexOn + +lemma mul_log_nonneg {x : ℝ} (hx : 1 ≤ x) : 0 ≤ x * log x := + mul_nonneg (zero_le_one.trans hx) (log_nonneg hx) + +lemma mul_log_nonpos {x : ℝ} (hx₀ : 0 ≤ x) (hx₁ : x ≤ 1) : x * log x ≤ 0 := + mul_nonpos_of_nonneg_of_nonpos hx₀ (log_nonpos hx₀ hx₁) + +section negMulLog + +/-- The function `x ↦ - x * log x` from `ℝ` to `ℝ`. -/ +noncomputable def negMulLog (x : ℝ) : ℝ := - x * log x + +lemma negMulLog_def : negMulLog = fun x ↦ - x * log x := rfl + +lemma negMulLog_eq_neg : negMulLog = fun x ↦ - (x * log x) := by simp [negMulLog_def] + +@[simp] lemma negMulLog_zero : negMulLog (0 : ℝ) = 0 := by simp [negMulLog] + +@[simp] lemma negMulLog_one : negMulLog (1 : ℝ) = 0 := by simp [negMulLog] + +lemma negMulLog_nonneg {x : ℝ} (h1 : 0 ≤ x) (h2 : x ≤ 1) : 0 ≤ negMulLog x := by + simpa only [negMulLog_eq_neg, neg_nonneg] using mul_log_nonpos h1 h2 + +lemma negMulLog_mul (x y : ℝ) : negMulLog (x * y) = y * negMulLog x + x * negMulLog y := by + simp only [negMulLog, neg_mul, neg_add_rev] + by_cases hx : x = 0 + · simp [hx] + by_cases hy : y = 0 + · simp [hy] + rw [log_mul hx hy] + ring + +lemma continuous_negMulLog : Continuous negMulLog := by + simpa only [negMulLog_eq_neg] using continuous_mul_log.neg + +lemma differentiableOn_negMulLog : DifferentiableOn ℝ negMulLog {0}ᶜ := by + simpa only [negMulLog_eq_neg] using differentiableOn_mul_log.neg + +lemma deriv_negMulLog {x : ℝ} (hx : x ≠ 0) : deriv negMulLog x = - log x - 1 := by + rw [negMulLog_eq_neg, deriv.neg, deriv_mul_log hx] + ring + +lemma hasDerivAt_negMulLog {x : ℝ} (hx : x ≠ 0) : HasDerivAt negMulLog (- log x - 1) x := by + rw [← deriv_negMulLog hx, hasDerivAt_deriv_iff] + refine DifferentiableOn.differentiableAt differentiableOn_negMulLog ?_ + simp [hx] + +lemma deriv2_negMulLog {x : ℝ} (hx : x ≠ 0) : deriv^[2] negMulLog x = - x⁻¹ := by + rw [negMulLog_eq_neg] + have h := deriv2_mul_log hx + simp only [Function.iterate_succ, Function.iterate_zero, Function.comp.left_id, + Function.comp_apply, deriv.neg', differentiableAt_id', differentiableAt_log_iff, ne_eq] at h ⊢ + rw [h] + +lemma strictConcaveOn_negMulLog : StrictConcaveOn ℝ (Set.Ici (0 : ℝ)) negMulLog := by + simpa only [negMulLog_eq_neg] using strictConvexOn_mul_log.neg + +lemma concaveOn_negMulLog : ConcaveOn ℝ (Set.Ici (0 : ℝ)) negMulLog := + strictConcaveOn_negMulLog.concaveOn + +end negMulLog + +end Real diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean index 39f42c5b2e481..515844e618e49 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean @@ -380,3 +380,17 @@ theorem tendsto_log_mul_rpow_nhds_zero {r : ℝ} (hr : 0 < r) : (tendsto_log_div_rpow_nhds_zero <| neg_lt_zero.2 hr).congr' <| eventually_mem_nhdsWithin.mono fun x hx => by rw [rpow_neg hx.out.le, div_inv_eq_mul] #align tendsto_log_mul_rpow_nhds_zero tendsto_log_mul_rpow_nhds_zero + +lemma tendsto_log_mul_self_nhds_zero_left : Filter.Tendsto (fun x ↦ log x * x) (𝓝[<] 0) (𝓝 0) := by + have h := tendsto_log_mul_rpow_nhds_zero zero_lt_one + simp only [Real.rpow_one] at h + have h_eq : ∀ x ∈ Set.Iio 0, (- (fun x ↦ log x * x) ∘ (fun x ↦ |x|)) x = log x * x := by + simp only [Set.mem_Iio, Pi.neg_apply, Function.comp_apply, log_abs] + intro x hx + simp only [abs_of_nonpos hx.le, mul_neg, neg_neg] + refine tendsto_nhdsWithin_congr h_eq ?_ + nth_rewrite 3 [← neg_zero] + refine (h.comp (tendsto_abs_nhdsWithin_zero.mono_left ?_)).neg + refine nhdsWithin_mono 0 (fun x hx ↦ ?_) + simp only [Set.mem_Iio] at hx + simp only [Set.mem_compl_iff, Set.mem_singleton_iff, hx.ne, not_false_eq_true] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean index 7dce3327aaad1..cc599f1328a88 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean @@ -451,7 +451,7 @@ theorem eventually_pow_one_div_le (x : ℝ≥0) {y : ℝ≥0} (hy : 1 < y) : rw [tsub_add_cancel_of_le hy.le] at hm refine' eventually_atTop.2 ⟨m + 1, fun n hn => _⟩ simpa only [NNReal.rpow_one_div_le_iff (Nat.cast_pos.2 <| m.succ_pos.trans_le hn), - NNReal.rpow_nat_cast] using hm.le.trans (pow_le_pow hy.le (m.le_succ.trans hn)) + NNReal.rpow_nat_cast] using hm.le.trans (pow_le_pow_right hy.le (m.le_succ.trans hn)) #align nnreal.eventually_pow_one_div_le NNReal.eventually_pow_one_div_le end NNReal diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean index 920e9ce3a48f8..9a4f0606447eb 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean @@ -320,15 +320,15 @@ theorem rpow_one_div_eq_iff {x y : ℝ≥0} {z : ℝ} (hz : z ≠ 0) : x ^ (1 / rw [← rpow_eq_rpow_iff hz, rpow_self_rpow_inv hz] #align nnreal.rpow_one_div_eq_iff NNReal.rpow_one_div_eq_iff -theorem pow_nat_rpow_nat_inv (x : ℝ≥0) {n : ℕ} (hn : n ≠ 0) : (x ^ n) ^ (n⁻¹ : ℝ) = x := by +theorem pow_rpow_inv_natCast (x : ℝ≥0) {n : ℕ} (hn : n ≠ 0) : (x ^ n) ^ (n⁻¹ : ℝ) = x := by rw [← NNReal.coe_eq, coe_rpow, NNReal.coe_pow] - exact Real.pow_nat_rpow_nat_inv x.2 hn -#align nnreal.pow_nat_rpow_nat_inv NNReal.pow_nat_rpow_nat_inv + exact Real.pow_rpow_inv_natCast x.2 hn +#align nnreal.pow_nat_rpow_nat_inv NNReal.pow_rpow_inv_natCast -theorem rpow_nat_inv_pow_nat (x : ℝ≥0) {n : ℕ} (hn : n ≠ 0) : (x ^ (n⁻¹ : ℝ)) ^ n = x := by +theorem rpow_inv_natCast_pow (x : ℝ≥0) {n : ℕ} (hn : n ≠ 0) : (x ^ (n⁻¹ : ℝ)) ^ n = x := by rw [← NNReal.coe_eq, NNReal.coe_pow, coe_rpow] - exact Real.rpow_nat_inv_pow_nat x.2 hn -#align nnreal.rpow_nat_inv_pow_nat NNReal.rpow_nat_inv_pow_nat + exact Real.rpow_inv_natCast_pow x.2 hn +#align nnreal.rpow_nat_inv_pow_nat NNReal.rpow_inv_natCast_pow theorem _root_.Real.toNNReal_rpow_of_nonneg {x y : ℝ} (hx : 0 ≤ x) : Real.toNNReal (x ^ y) = Real.toNNReal x ^ y := by diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index b4dff8e01d1ef..58b55feae704e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -26,6 +26,7 @@ open Finset Set ## Definitions -/ namespace Real +variable {x y z : ℝ} /-- The real power function `x ^ y`, defined as the real part of the complex power function. For `x > 0`, it is equal to `exp (y log x)`. For `x = 0`, one sets `0 ^ 0=1` and `0 ^ y=0` for @@ -58,15 +59,31 @@ theorem rpow_def_of_pos {x : ℝ} (hx : 0 < x) (y : ℝ) : x ^ y = exp (log x * theorem exp_mul (x y : ℝ) : exp (x * y) = exp x ^ y := by rw [rpow_def_of_pos (exp_pos _), log_exp] #align real.exp_mul Real.exp_mul +@[simp, norm_cast] +theorem rpow_int_cast (x : ℝ) (n : ℤ) : x ^ (n : ℝ) = x ^ n := by + simp only [rpow_def, ← Complex.ofReal_zpow, Complex.cpow_int_cast, Complex.ofReal_int_cast, + Complex.ofReal_re] +#align real.rpow_int_cast Real.rpow_int_cast + +@[simp, norm_cast] +theorem rpow_nat_cast (x : ℝ) (n : ℕ) : x ^ (n : ℝ) = x ^ n := by simpa using rpow_int_cast x n +#align real.rpow_nat_cast Real.rpow_nat_cast + @[simp] theorem exp_one_rpow (x : ℝ) : exp 1 ^ x = exp x := by rw [← exp_mul, one_mul] #align real.exp_one_rpow Real.exp_one_rpow -theorem rpow_eq_zero_iff_of_nonneg {x y : ℝ} (hx : 0 ≤ x) : x ^ y = 0 ↔ x = 0 ∧ y ≠ 0 := by +@[simp] lemma exp_one_pow (n : ℕ) : exp 1 ^ n = exp n := by rw [← rpow_nat_cast, exp_one_rpow] + +theorem rpow_eq_zero_iff_of_nonneg (hx : 0 ≤ x) : x ^ y = 0 ↔ x = 0 ∧ y ≠ 0 := by simp only [rpow_def_of_nonneg hx] split_ifs <;> simp [*, exp_ne_zero] #align real.rpow_eq_zero_iff_of_nonneg Real.rpow_eq_zero_iff_of_nonneg +@[simp] +lemma rpow_eq_zero (hx : 0 ≤ x) (hy : y ≠ 0) : x ^ y = 0 ↔ x = 0 := by + simp [rpow_eq_zero_iff_of_nonneg, *] + open Real theorem rpow_def_of_neg {x : ℝ} (hx : x < 0) (y : ℝ) : x ^ y = exp (log x * y) * cos (y * π) := by @@ -315,6 +332,39 @@ theorem abs_cpow_eq_rpow_re_of_nonneg {x : ℝ} (hx : 0 ≤ x) {y : ℂ} (hy : r end Complex +/-! ### Positivity extension -/ + +namespace Mathlib.Meta.Positivity +open Lean Meta Qq + +/-- Extension for the `positivity` tactic: exponentiation by a real number is positive (namely 1) +when the exponent is zero. The other cases are done in `evalRpow`. -/ +@[positivity (_ : ℝ) ^ (0 : ℝ), Pow.pow (_ : ℝ) (0 : ℝ), Real.rpow (_ : ℝ) (0 : ℝ)] +def evalRpowZero : Mathlib.Meta.Positivity.PositivityExt where eval {_ _} _ _ e := do + let .app (.app (f : Q(ℝ → ℝ → ℝ)) (a : Q(ℝ))) (_ : Q(ℝ)) ← withReducible (whnf e) + | throwError "not Real.rpow" + guard <|← withDefault <| withNewMCtxDepth <| isDefEq f q(Real.rpow) + pure (.positive (q(Real.rpow_zero_pos $a) : Expr)) + +/-- Extension for the `positivity` tactic: exponentiation by a real number is nonnegative when +the base is nonnegative and positive when the base is positive. -/ +@[positivity (_ : ℝ) ^ (_ : ℝ), Pow.pow (_ : ℝ) (_ : ℝ), Real.rpow (_ : ℝ) (_ : ℝ)] +def evalRpow : Mathlib.Meta.Positivity.PositivityExt where eval {_ _} zα pα e := do + let .app (.app (f : Q(ℝ → ℝ → ℝ)) (a : Q(ℝ))) (b : Q(ℝ)) ← withReducible (whnf e) + | throwError "not Real.rpow" + guard <| ← withDefault <| withNewMCtxDepth <| isDefEq f q(Real.rpow) + let ra ← core zα pα a + match ra with + | .positive pa => + have pa' : Q(0 < $a) := pa + pure (.positive (q(Real.rpow_pos_of_pos $pa' $b) : Expr)) + | .nonnegative pa => + have pa' : Q(0 ≤ $a) := pa + pure (.nonnegative (q(Real.rpow_nonneg_of_nonneg $pa' $b) : Expr)) + | _ => pure .none + +end Mathlib.Meta.Positivity + /-! ## Further algebraic properties of `rpow` -/ @@ -322,7 +372,7 @@ end Complex namespace Real -variable {x y z : ℝ} +variable {x y z : ℝ} {n : ℕ} theorem rpow_mul {x : ℝ} (hx : 0 ≤ x) (y z : ℝ) : x ^ (y * z) = (x ^ y) ^ z := by rw [← Complex.ofReal_inj, Complex.ofReal_cpow (rpow_nonneg_of_nonneg hx _), @@ -349,6 +399,18 @@ theorem rpow_sub_nat {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℕ) : x ^ (y - n) by simpa using rpow_sub_int hx y n #align real.rpow_sub_nat Real.rpow_sub_nat +lemma rpow_add_int' (hx : 0 ≤ x) {n : ℤ} (h : y + n ≠ 0) : x ^ (y + n) = x ^ y * x ^ n := by + rw [rpow_add' hx h, rpow_int_cast] + +lemma rpow_add_nat' (hx : 0 ≤ x) (h : y + n ≠ 0) : x ^ (y + n) = x ^ y * x ^ n := by + rw [rpow_add' hx h, rpow_nat_cast] + +lemma rpow_sub_int' (hx : 0 ≤ x) {n : ℤ} (h : y - n ≠ 0) : x ^ (y - n) = x ^ y / x ^ n := by + rw [rpow_sub' hx h, rpow_int_cast] + +lemma rpow_sub_nat' (hx : 0 ≤ x) (h : y - n ≠ 0) : x ^ (y - n) = x ^ y / x ^ n := by + rw [rpow_sub' hx h, rpow_nat_cast] + theorem rpow_add_one {x : ℝ} (hx : x ≠ 0) (y : ℝ) : x ^ (y + 1) = x ^ y * x := by simpa using rpow_add_nat hx y 1 #align real.rpow_add_one Real.rpow_add_one @@ -357,16 +419,17 @@ theorem rpow_sub_one {x : ℝ} (hx : x ≠ 0) (y : ℝ) : x ^ (y - 1) = x ^ y / simpa using rpow_sub_nat hx y 1 #align real.rpow_sub_one Real.rpow_sub_one -@[simp, norm_cast] -theorem rpow_int_cast (x : ℝ) (n : ℤ) : x ^ (n : ℝ) = x ^ n := by - simp only [rpow_def, ← Complex.ofReal_zpow, Complex.cpow_int_cast, Complex.ofReal_int_cast, - Complex.ofReal_re] -#align real.rpow_int_cast Real.rpow_int_cast +lemma rpow_add_one' (hx : 0 ≤ x) (h : y + 1 ≠ 0) : x ^ (y + 1) = x ^ y * x := by + rw [rpow_add' hx h, rpow_one] -@[simp, norm_cast] -theorem rpow_nat_cast (x : ℝ) (n : ℕ) : x ^ (n : ℝ) = x ^ n := - by simpa using rpow_int_cast x n -#align real.rpow_nat_cast Real.rpow_nat_cast +lemma rpow_one_add' (hx : 0 ≤ x) (h : 1 + y ≠ 0) : x ^ (1 + y) = x * x ^ y := by + rw [rpow_add' hx h, rpow_one] + +lemma rpow_sub_one' (hx : 0 ≤ x) (h : y - 1 ≠ 0) : x ^ (y - 1) = x ^ y / x := by + rw [rpow_sub' hx h, rpow_one] + +lemma rpow_one_sub' (hx : 0 ≤ x) (h : 1 - y ≠ 0) : x ^ (1 - y) = x / x ^ y := by + rw [rpow_sub' hx h, rpow_one] @[simp] theorem rpow_two (x : ℝ) : x ^ (2 : ℝ) = x ^ 2 := by @@ -379,25 +442,10 @@ theorem rpow_neg_one (x : ℝ) : x ^ (-1 : ℝ) = x⁻¹ := by simp only [rpow_int_cast, zpow_one, zpow_neg] #align real.rpow_neg_one Real.rpow_neg_one -theorem mul_rpow {x y z : ℝ} (h : 0 ≤ x) (h₁ : 0 ≤ y) : (x * y) ^ z = x ^ z * y ^ z := by +theorem mul_rpow (hx : 0 ≤ x) (hy : 0 ≤ y) : (x * y) ^ z = x ^ z * y ^ z := by iterate 2 rw [Real.rpow_def_of_nonneg]; split_ifs with h_ifs <;> simp_all - · exact h - · rw [not_or] at h_ifs - have hx : 0 < x := by - cases' lt_or_eq_of_le h with h₂ h₂ - · exact h₂ - exfalso - apply h_ifs.1 - exact Eq.symm h₂ - have hy : 0 < y := by - cases' lt_or_eq_of_le h₁ with h₂ h₂ - · exact h₂ - exfalso - apply h_ifs.2 - exact Eq.symm h₂ - rw [log_mul (ne_of_gt hx) (ne_of_gt hy), add_mul, exp_add, rpow_def_of_pos hx, - rpow_def_of_pos hy] - · exact mul_nonneg h h₁ + · rw [log_mul ‹_› ‹_›, add_mul, exp_add, rpow_def_of_pos (hy.lt_of_ne' ‹_›)] + all_goals positivity #align real.mul_rpow Real.mul_rpow theorem inv_rpow (hx : 0 ≤ x) (y : ℝ) : x⁻¹ ^ y = (x ^ y)⁻¹ := by @@ -418,6 +466,34 @@ theorem mul_log_eq_log_iff {x y z : ℝ} (hx : 0 < x) (hz : 0 < z) : ⟨fun h ↦ log_injOn_pos (rpow_pos_of_pos hx _) hz <| log_rpow hx _ |>.trans h, by rintro rfl; rw [log_rpow hx]⟩ +@[simp] lemma rpow_rpow_inv (hx : 0 ≤ x) (hy : y ≠ 0) : (x ^ y) ^ y⁻¹ = x := by + rw [← rpow_mul hx, mul_inv_cancel hy, rpow_one] + +@[simp] lemma rpow_inv_rpow (hx : 0 ≤ x) (hy : y ≠ 0) : (x ^ y⁻¹) ^ y = x := by + rw [← rpow_mul hx, inv_mul_cancel hy, rpow_one] + +theorem pow_rpow_inv_natCast (hx : 0 ≤ x) (hn : n ≠ 0) : (x ^ n) ^ (n⁻¹ : ℝ) = x := by + have hn0 : (n : ℝ) ≠ 0 := Nat.cast_ne_zero.2 hn + rw [← rpow_nat_cast, ← rpow_mul hx, mul_inv_cancel hn0, rpow_one] +#align real.pow_nat_rpow_nat_inv Real.pow_rpow_inv_natCast + +theorem rpow_inv_natCast_pow (hx : 0 ≤ x) (hn : n ≠ 0) : (x ^ (n⁻¹ : ℝ)) ^ n = x := by + have hn0 : (n : ℝ) ≠ 0 := Nat.cast_ne_zero.2 hn + rw [← rpow_nat_cast, ← rpow_mul hx, inv_mul_cancel hn0, rpow_one] +#align real.rpow_nat_inv_pow_nat Real.rpow_inv_natCast_pow + +lemma rpow_natCast_mul (hx : 0 ≤ x) (n : ℕ) (z : ℝ) : x ^ (n * z) = (x ^ n) ^ z := by + rw [rpow_mul hx, rpow_nat_cast] + +lemma rpow_mul_natCast (hx : 0 ≤ x) (y : ℝ) (n : ℕ) : x ^ (y * n) = (x ^ y) ^ n := by + rw [rpow_mul hx, rpow_nat_cast] + +lemma rpow_intCast_mul (hx : 0 ≤ x) (n : ℤ) (z : ℝ) : x ^ (n * z) = (x ^ n) ^ z := by + rw [rpow_mul hx, rpow_int_cast] + +lemma rpow_mul_intCast (hx : 0 ≤ x) (y : ℝ) (n : ℤ) : x ^ (y * n) = (x ^ y) ^ n := by + rw [rpow_mul hx, rpow_int_cast] + /-! Note: lemmas about `(∏ i in s, f i ^ r)` such as `Real.finset_prod_rpow` are proved in `Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean` instead. -/ @@ -450,6 +526,18 @@ theorem monotoneOn_rpow_Ici_of_exponent_nonneg {r : ℝ} (hr : 0 ≤ r) : MonotoneOn (fun (x : ℝ) => x ^ r) (Set.Ici 0) := fun _ ha _ _ hab => rpow_le_rpow ha hab hr +lemma rpow_lt_rpow_of_neg (hx : 0 < x) (hxy : x < y) (hz : z < 0) : y ^ z < x ^ z := by + have := hx.trans hxy + rw [← inv_lt_inv, ← rpow_neg, ← rpow_neg] + refine rpow_lt_rpow ?_ hxy (neg_pos.2 hz) + all_goals positivity + +lemma rpow_le_rpow_of_nonpos (hx : 0 < x) (hxy : x ≤ y) (hz : z ≤ 0) : y ^ z ≤ x ^ z := by + have := hx.trans_le hxy + rw [← inv_le_inv, ← rpow_neg, ← rpow_neg] + refine rpow_le_rpow ?_ hxy (neg_nonneg.2 hz) + all_goals positivity + theorem rpow_lt_rpow_iff (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 < z) : x ^ z < y ^ z ↔ x < y := ⟨lt_imp_lt_of_le_imp_le fun h => rpow_le_rpow hy h (le_of_lt hz), fun h => rpow_lt_rpow hx h hz⟩ #align real.rpow_lt_rpow_iff Real.rpow_lt_rpow_iff @@ -458,38 +546,43 @@ theorem rpow_le_rpow_iff (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 < z) : x ^ z ≤ le_iff_le_iff_lt_iff_lt.2 <| rpow_lt_rpow_iff hy hx hz #align real.rpow_le_rpow_iff Real.rpow_le_rpow_iff +lemma rpow_lt_rpow_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : x ^ z < y ^ z ↔ y < x := + ⟨lt_imp_lt_of_le_imp_le fun h ↦ rpow_le_rpow_of_nonpos hx h hz.le, + fun h ↦ rpow_lt_rpow_of_neg hy h hz⟩ + +lemma rpow_le_rpow_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : x ^ z ≤ y ^ z ↔ y ≤ x := + le_iff_le_iff_lt_iff_lt.2 $ rpow_lt_rpow_iff_of_neg hy hx hz + +lemma le_rpow_inv_iff_of_pos (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 < z) : x ≤ y ^ z⁻¹ ↔ x ^ z ≤ y := by + rw [← rpow_le_rpow_iff hx _ hz, rpow_inv_rpow] <;> positivity + +lemma rpow_inv_le_iff_of_pos (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 < z) : x ^ z⁻¹ ≤ y ↔ x ≤ y ^ z := by + rw [← rpow_le_rpow_iff _ hy hz, rpow_inv_rpow] <;> positivity + +lemma lt_rpow_inv_iff_of_pos (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 < z) : x < y ^ z⁻¹ ↔ x ^ z < y := + lt_iff_lt_of_le_iff_le $ rpow_inv_le_iff_of_pos hy hx hz + +lemma rpow_inv_lt_iff_of_pos (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 < z) : x ^ z⁻¹ < y ↔ x < y ^ z := + lt_iff_lt_of_le_iff_le $ le_rpow_inv_iff_of_pos hy hx hz + theorem le_rpow_inv_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : x ≤ y ^ z⁻¹ ↔ y ≤ x ^ z := by - have hz' : 0 < -z := by rwa [lt_neg, neg_zero] - have hxz : 0 < x ^ (-z) := Real.rpow_pos_of_pos hx _ - have hyz : 0 < y ^ z⁻¹ := Real.rpow_pos_of_pos hy _ - rw [← Real.rpow_le_rpow_iff hx.le hyz.le hz', ← Real.rpow_mul hy.le] - simp only [ne_of_lt hz, Real.rpow_neg_one, mul_neg, inv_mul_cancel, Ne.def, not_false_iff] - rw [le_inv hxz hy, ← Real.rpow_neg_one, ← Real.rpow_mul hx.le] - simp + rw [← rpow_le_rpow_iff_of_neg _ hx hz, rpow_inv_rpow _ hz.ne] <;> positivity #align real.le_rpow_inv_iff_of_neg Real.le_rpow_inv_iff_of_neg theorem lt_rpow_inv_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : x < y ^ z⁻¹ ↔ y < x ^ z := by - have hz' : 0 < -z := by rwa [lt_neg, neg_zero] - have hxz : 0 < x ^ (-z) := Real.rpow_pos_of_pos hx _ - have hyz : 0 < y ^ z⁻¹ := Real.rpow_pos_of_pos hy _ - rw [← Real.rpow_lt_rpow_iff hx.le hyz.le hz', ← Real.rpow_mul hy.le] - simp only [ne_of_lt hz, Real.rpow_neg_one, mul_neg, inv_mul_cancel, Ne.def, not_false_iff] - rw [lt_inv hxz hy, ← Real.rpow_neg_one, ← Real.rpow_mul hx.le] - simp + rw [← rpow_lt_rpow_iff_of_neg _ hx hz, rpow_inv_rpow _ hz.ne] <;> positivity #align real.lt_rpow_inv_iff_of_neg Real.lt_rpow_inv_iff_of_neg theorem rpow_inv_lt_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : x ^ z⁻¹ < y ↔ y ^ z < x := by - convert lt_rpow_inv_iff_of_neg (Real.rpow_pos_of_pos hx z⁻¹) (Real.rpow_pos_of_pos hy z) hz <;> - simp [← Real.rpow_mul hx.le, ← Real.rpow_mul hy.le, ne_of_lt hz] + rw [← rpow_lt_rpow_iff_of_neg hy _ hz, rpow_inv_rpow _ hz.ne] <;> positivity #align real.rpow_inv_lt_iff_of_neg Real.rpow_inv_lt_iff_of_neg theorem rpow_inv_le_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : x ^ z⁻¹ ≤ y ↔ y ^ z ≤ x := by - convert le_rpow_inv_iff_of_neg (Real.rpow_pos_of_pos hx z⁻¹) (Real.rpow_pos_of_pos hy z) hz <;> - simp [← Real.rpow_mul hx.le, ← Real.rpow_mul hy.le, ne_of_lt hz] + rw [← rpow_le_rpow_iff_of_neg hy _ hz, rpow_inv_rpow _ hz.ne] <;> positivity #align real.rpow_inv_le_iff_of_neg Real.rpow_inv_le_iff_of_neg theorem rpow_lt_rpow_of_exponent_lt (hx : 1 < x) (hyz : y < z) : x ^ y < x ^ z := by @@ -532,7 +625,7 @@ theorem antitoneOn_rpow_Ioi_of_exponent_nonpos {r : ℝ} (hr : r ≤ 0) : @[simp] theorem rpow_le_rpow_left_iff (hx : 1 < x) : x ^ y ≤ x ^ z ↔ y ≤ z := by have x_pos : 0 < x := lt_trans zero_lt_one hx - rw [← log_le_log (rpow_pos_of_pos x_pos y) (rpow_pos_of_pos x_pos z), log_rpow x_pos, + rw [← log_le_log_iff (rpow_pos_of_pos x_pos y) (rpow_pos_of_pos x_pos z), log_rpow x_pos, log_rpow x_pos, mul_le_mul_right (log_pos hx)] #align real.rpow_le_rpow_left_iff Real.rpow_le_rpow_left_iff @@ -554,7 +647,7 @@ theorem rpow_le_rpow_of_exponent_ge (hx0 : 0 < x) (hx1 : x ≤ 1) (hyz : z ≤ y @[simp] theorem rpow_le_rpow_left_iff_of_base_lt_one (hx0 : 0 < x) (hx1 : x < 1) : x ^ y ≤ x ^ z ↔ z ≤ y := by - rw [← log_le_log (rpow_pos_of_pos hx0 y) (rpow_pos_of_pos hx0 z), log_rpow hx0, log_rpow hx0, + rw [← log_le_log_iff (rpow_pos_of_pos hx0 y) (rpow_pos_of_pos hx0 z), log_rpow hx0, log_rpow hx0, mul_le_mul_right_of_neg (log_neg hx0 hx1)] #align real.rpow_le_rpow_left_iff_of_base_lt_one Real.rpow_le_rpow_left_iff_of_base_lt_one @@ -644,8 +737,17 @@ theorem rpow_left_injOn {x : ℝ} (hx : x ≠ 0) : InjOn (fun y : ℝ => y ^ x) rw [← rpow_one y, ← rpow_one z, ← _root_.mul_inv_cancel hx, rpow_mul hy, rpow_mul hz, hyz] #align real.rpow_left_inj_on Real.rpow_left_injOn +lemma rpow_left_inj (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : z ≠ 0) : x ^ z = y ^ z ↔ x = y := + (rpow_left_injOn hz).eq_iff hx hy + +lemma rpow_inv_eq (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : z ≠ 0) : x ^ z⁻¹ = y ↔ x = y ^ z := by + rw [← rpow_left_inj _ hy hz, rpow_inv_rpow hx hz]; positivity + +lemma eq_rpow_inv (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : z ≠ 0) : x = y ^ z⁻¹ ↔ x ^ z = y := by + rw [← rpow_left_inj hx _ hz, rpow_inv_rpow hy hz]; positivity + theorem le_rpow_iff_log_le (hx : 0 < x) (hy : 0 < y) : x ≤ y ^ z ↔ Real.log x ≤ z * Real.log y := by - rw [← Real.log_le_log hx (Real.rpow_pos_of_pos hy z), Real.log_rpow hy] + rw [← Real.log_le_log_iff hx (Real.rpow_pos_of_pos hy z), Real.log_rpow hy] #align real.le_rpow_iff_log_le Real.le_rpow_iff_log_le theorem le_rpow_of_log_le (hx : 0 ≤ x) (hy : 0 < y) (h : Real.log x ≤ z * Real.log y) : @@ -678,18 +780,6 @@ theorem abs_log_mul_self_rpow_lt (x t : ℝ) (h1 : 0 < x) (h2 : x ≤ 1) (ht : 0 rwa [log_rpow h1, mul_assoc, abs_mul, abs_of_pos ht, mul_comm] at this #align real.abs_log_mul_self_rpow_lt Real.abs_log_mul_self_rpow_lt -theorem pow_nat_rpow_nat_inv {x : ℝ} (hx : 0 ≤ x) {n : ℕ} (hn : n ≠ 0) : - (x ^ n) ^ (n⁻¹ : ℝ) = x := by - have hn0 : (n : ℝ) ≠ 0 := Nat.cast_ne_zero.2 hn - rw [← rpow_nat_cast, ← rpow_mul hx, mul_inv_cancel hn0, rpow_one] -#align real.pow_nat_rpow_nat_inv Real.pow_nat_rpow_nat_inv - -theorem rpow_nat_inv_pow_nat {x : ℝ} (hx : 0 ≤ x) {n : ℕ} (hn : n ≠ 0) : - (x ^ (n⁻¹ : ℝ)) ^ n = x := by - have hn0 : (n : ℝ) ≠ 0 := Nat.cast_ne_zero.2 hn - rw [← rpow_nat_cast, ← rpow_mul hx, inv_mul_cancel hn0, rpow_one] -#align real.rpow_nat_inv_pow_nat Real.rpow_nat_inv_pow_nat - lemma strictMono_rpow_of_base_gt_one {b : ℝ} (hb : 1 < b) : StrictMono (rpow b) := by show StrictMono (fun (x:ℝ) => b ^ x) @@ -756,7 +846,7 @@ theorem exists_rat_pow_btwn_rat_aux (hn : n ≠ 0) (x y : ℝ) (h : x < y) (hy : have hq := this.trans_lt hxq replace hxq := rpow_lt_rpow this hxq hn' replace hqy := rpow_lt_rpow hq.le hqy hn' - rw [rpow_nat_cast, rpow_nat_cast, rpow_nat_inv_pow_nat _ hn] at hxq hqy + rw [rpow_nat_cast, rpow_nat_cast, rpow_inv_natCast_pow _ hn] at hxq hqy · exact ⟨q, mod_cast hq, (le_max_right _ _).trans_lt hxq, hqy⟩ · exact hy.le · exact le_max_left _ _ @@ -866,36 +956,4 @@ section Tactics -- end NormNum -namespace Mathlib.Meta.Positivity - -open Lean Meta Qq - -/-- Extension for the `positivity` tactic: exponentiation by a real number is positive (namely 1) -when the exponent is zero. The other cases are done in `evalRpow`. -/ -@[positivity (_ : ℝ) ^ (0 : ℝ), Pow.pow (_ : ℝ) (0 : ℝ), Real.rpow (_ : ℝ) (0 : ℝ)] -def evalRpowZero : Mathlib.Meta.Positivity.PositivityExt where eval {_ _} _ _ e := do - let .app (.app (f : Q(ℝ → ℝ → ℝ)) (a : Q(ℝ))) (_ : Q(ℝ)) ← withReducible (whnf e) - | throwError "not Real.rpow" - guard <|← withDefault <| withNewMCtxDepth <| isDefEq f q(Real.rpow) - pure (.positive (q(Real.rpow_zero_pos $a) : Expr)) - -/-- Extension for the `positivity` tactic: exponentiation by a real number is nonnegative when -the base is nonnegative and positive when the base is positive. -/ -@[positivity (_ : ℝ) ^ (_ : ℝ), Pow.pow (_ : ℝ) (_ : ℝ), Real.rpow (_ : ℝ) (_ : ℝ)] -def evalRpow : Mathlib.Meta.Positivity.PositivityExt where eval {_ _} zα pα e := do - let .app (.app (f : Q(ℝ → ℝ → ℝ)) (a : Q(ℝ))) (b : Q(ℝ)) ← withReducible (whnf e) - | throwError "not Real.rpow" - guard <|← withDefault <| withNewMCtxDepth <| isDefEq f q(Real.rpow) - let ra ← core zα pα a - match ra with - | .positive pa => - have pa' : Q(0 < $a) := pa - pure (.positive (q(Real.rpow_pos_of_pos $pa' $b) : Expr)) - | .nonnegative pa => - have pa' : Q(0 ≤ $a) := pa - pure (.nonnegative (q(Real.rpow_nonneg_of_nonneg $pa' $b) : Expr)) - | _ => pure .none - -end Mathlib.Meta.Positivity - end Tactics diff --git a/Mathlib/Analysis/SpecialFunctions/Stirling.lean b/Mathlib/Analysis/SpecialFunctions/Stirling.lean index 73b4b04f4e010..695e5114d1e7c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Stirling.lean +++ b/Mathlib/Analysis/SpecialFunctions/Stirling.lean @@ -188,7 +188,7 @@ theorem stirlingSeq'_bounded_by_pos_constant : ∃ a, 0 < a ∧ ∀ n : ℕ, a /-- The sequence `stirlingSeq ∘ succ` is monotone decreasing -/ theorem stirlingSeq'_antitone : Antitone (stirlingSeq ∘ succ) := fun n m h => - (log_le_log (stirlingSeq'_pos m) (stirlingSeq'_pos n)).mp (log_stirlingSeq'_antitone h) + (log_le_log_iff (stirlingSeq'_pos m) (stirlingSeq'_pos n)).mp (log_stirlingSeq'_antitone h) #align stirling.stirling_seq'_antitone Stirling.stirlingSeq'_antitone /-- The limit `a` of the sequence `stirlingSeq` satisfies `0 < a` -/ @@ -229,7 +229,8 @@ theorem stirlingSeq_pow_four_div_stirlingSeq_pow_two_eq (n : ℕ) (hn : n ≠ 0) simp_rw [div_pow, mul_pow] rw [sq_sqrt, sq_sqrt] any_goals positivity - field_simp; ring + field_simp [← exp_nsmul] + ring_nf #align stirling.stirling_seq_pow_four_div_stirling_seq_pow_two_eq Stirling.stirlingSeq_pow_four_div_stirlingSeq_pow_two_eq /-- Suppose the sequence `stirlingSeq` (defined above) has the limit `a ≠ 0`. diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index 97b565219fce2..20e45259933f5 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -96,7 +96,7 @@ theorem tendsto_coe_nat_div_add_atTop {𝕜 : Type*} [DivisionRing 𝕜] [Topolo theorem tendsto_add_one_pow_atTop_atTop_of_pos [LinearOrderedSemiring α] [Archimedean α] {r : α} (h : 0 < r) : Tendsto (fun n : ℕ ↦ (r + 1) ^ n) atTop atTop := - (tendsto_atTop_atTop_of_monotone' fun _ _ ↦ pow_le_pow (le_add_of_nonneg_left (le_of_lt h))) <| + tendsto_atTop_atTop_of_monotone' (fun _ _ ↦ pow_le_pow_right $ le_add_of_nonneg_left h.le) <| not_bddAbove_iff.2 fun _ ↦ Set.exists_range_iff.2 <| add_one_pow_unbounded_of_pos _ h #align tendsto_add_one_pow_at_top_at_top_of_pos tendsto_add_one_pow_atTop_atTop_of_pos @@ -131,7 +131,7 @@ theorem tendsto_pow_atTop_nhds_0_of_lt_1 {𝕜 : Type*} [LinearOrderedField 𝕜 simp only [hr.symm, one_pow] at h exact zero_ne_one <| tendsto_nhds_unique h tendsto_const_nhds · apply @not_tendsto_nhds_of_tendsto_atTop 𝕜 ℕ _ _ _ _ atTop _ (fun n ↦ |r| ^ n) _ 0 _ - refine (pow_strictMono_right $ lt_of_le_of_ne (le_of_not_lt hr_le) + refine (pow_right_strictMono $ lt_of_le_of_ne (le_of_not_lt hr_le) hr).monotone.tendsto_atTop_atTop (fun b ↦ ?_) obtain ⟨n, hn⟩ := (pow_unbounded_of_one_lt b (lt_of_le_of_ne (le_of_not_lt hr_le) hr)) exacts [⟨n, le_of_lt hn⟩, by simpa only [← abs_pow]] @@ -465,7 +465,7 @@ theorem summable_one_div_pow_of_le {m : ℝ} {f : ℕ → ℕ} (hm : 1 < m) (fi (summable_geometric_of_lt_1 (one_div_nonneg.mpr (zero_le_one.trans hm.le)) ((one_div_lt (zero_lt_one.trans hm) zero_lt_one).mpr (one_div_one.le.trans_lt hm))) rw [div_pow, one_pow] - refine' (one_div_le_one_div _ _).mpr (pow_le_pow hm.le (fi a)) <;> + refine' (one_div_le_one_div _ _).mpr (pow_le_pow_right hm.le (fi a)) <;> exact pow_pos (zero_lt_one.trans hm) _ #align summable_one_div_pow_of_le summable_one_div_pow_of_le diff --git a/Mathlib/Analysis/SpecificLimits/FloorPow.lean b/Mathlib/Analysis/SpecificLimits/FloorPow.lean index 9be939fc3a6d3..a9a3408c43b1e 100644 --- a/Mathlib/Analysis/SpecificLimits/FloorPow.lean +++ b/Mathlib/Analysis/SpecificLimits/FloorPow.lean @@ -250,7 +250,7 @@ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc have : c ^ 3 = c ^ 2 * c := by ring simp only [mul_sub, this, mul_one, inv_pow, sub_le_sub_iff_left] rw [mul_assoc, mul_comm c, ← mul_assoc, mul_inv_cancel (sq_pos_of_pos cpos).ne', one_mul] - simpa using pow_le_pow hc.le one_le_two + simpa using pow_le_pow_right hc.le one_le_two calc (∑ i in (range N).filter fun i => j < c ^ i, (1 : ℝ) / (c ^ i) ^ 2) ≤ ∑ i in Ico ⌊Real.log j / Real.log c⌋₊ N, (1 : ℝ) / (c ^ i) ^ 2 := by @@ -307,7 +307,7 @@ theorem mul_pow_le_nat_floor_pow {c : ℝ} (hc : 1 < c) (i : ℕ) : (1 - c⁻¹) (1 - c⁻¹) * c ^ i = c ^ i - c ^ i * c⁻¹ := by ring _ ≤ c ^ i - 1 := by simpa only [← div_eq_mul_inv, sub_le_sub_iff_left, one_le_div cpos, pow_one] using - pow_le_pow hc.le hident + pow_le_pow_right hc.le hident _ ≤ ⌊c ^ i⌋₊ := (Nat.sub_one_lt_floor _).le #align mul_pow_le_nat_floor_pow mul_pow_le_nat_floor_pow @@ -337,7 +337,7 @@ theorem sum_div_nat_floor_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : simp only [Nat.le_floor, one_le_pow_of_one_le, hc.le, Nat.one_le_cast, Nat.cast_one] · exact sq_pos_of_pos (pow_pos cpos _) rw [one_mul, ← mul_pow] - apply pow_le_pow_of_le_left (pow_nonneg cpos.le _) + apply pow_le_pow_left (pow_nonneg cpos.le _) rw [← div_eq_inv_mul, le_div_iff A, mul_comm] exact mul_pow_le_nat_floor_pow hc i _ ≤ (1 - c⁻¹)⁻¹ ^ 2 * (c ^ 3 * (c - 1)⁻¹) / j ^ 2 := by diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index 9a93c7c549190..4240f7b44c35e 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -319,7 +319,7 @@ theorem summable_geometric_iff_norm_lt_1 : (Summable fun n : ℕ ↦ ξ ^ n) ↔ (h.tendsto_cofinite_zero.eventually (ball_mem_nhds _ zero_lt_one)).exists simp only [norm_pow, dist_zero_right] at hk rw [← one_pow k] at hk - exact lt_of_pow_lt_pow _ zero_le_one hk + exact lt_of_pow_lt_pow_left _ zero_le_one hk #align summable_geometric_iff_norm_lt_1 summable_geometric_iff_norm_lt_1 end Geometric diff --git a/Mathlib/Combinatorics/Additive/Behrend.lean b/Mathlib/Combinatorics/Additive/Behrend.lean index 0b20c91e57703..247a9d984680f 100644 --- a/Mathlib/Combinatorics/Additive/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/Behrend.lean @@ -196,7 +196,7 @@ theorem addSalemSpencer_image_sphere : theorem sum_sq_le_of_mem_box (hx : x ∈ box n d) : ∑ i : Fin n, x i ^ 2 ≤ n * (d - 1) ^ 2 := by rw [mem_box] at hx have : ∀ i, x i ^ 2 ≤ (d - 1) ^ 2 := fun i => - Nat.pow_le_pow_of_le_left (Nat.le_sub_one_of_lt (hx i)) _ + Nat.pow_le_pow_left (Nat.le_sub_one_of_lt (hx i)) _ exact (sum_le_card_nsmul univ _ _ fun i _ => this i).trans (by rw [card_fin, smul_eq_mul]) #align behrend.sum_sq_le_of_mem_box Behrend.sum_sq_le_of_mem_box @@ -305,15 +305,11 @@ theorem two_div_one_sub_two_div_e_le_eight : 2 / (1 - 2 / exp 1) ≤ 8 := by #align behrend.two_div_one_sub_two_div_e_le_eight Behrend.two_div_one_sub_two_div_e_le_eight theorem le_sqrt_log (hN : 4096 ≤ N) : log (2 / (1 - 2 / exp 1)) * (69 / 50) ≤ sqrt (log ↑N) := by - have : ((12 : ℕ) : ℝ) * log 2 ≤ log N := by - rw [← log_rpow zero_lt_two, log_le_log, rpow_nat_cast] - · norm_num1 - exact mod_cast hN - · exact rpow_pos_of_pos zero_lt_two _ - rw [cast_pos] - exact hN.trans_lt' (by norm_num1) - refine' (mul_le_mul_of_nonneg_right ((log_le_log _ <| by norm_num1).2 - two_div_one_sub_two_div_e_le_eight) <| by norm_num1).trans _ + have : (12 : ℕ) * log 2 ≤ log N := by + rw [← log_rpow zero_lt_two, rpow_nat_cast] + exact log_le_log (by positivity) (mod_cast hN) + refine (mul_le_mul_of_nonneg_right (log_le_log ?_ two_div_one_sub_two_div_e_le_eight) <| by + norm_num1).trans ?_ · refine' div_pos zero_lt_two _ rw [sub_pos, div_lt_one (exp_pos _)] exact exp_one_gt_d9.trans_le' (by norm_num1) @@ -340,8 +336,7 @@ theorem exp_neg_two_mul_le {x : ℝ} (hx : 0 < x) : exp (-2 * x) < exp (2 - ⌈x refine' le_trans _ (div_le_div_of_le_of_nonneg (exp_le_exp.2 h₂) <| add_nonneg hx.le zero_le_one) rw [le_div_iff (add_pos hx zero_lt_one), ← le_div_iff' (exp_pos _), ← exp_sub, neg_mul, sub_neg_eq_add, two_mul, sub_add_add_cancel, add_comm _ x] - refine' le_trans _ (add_one_le_exp_of_nonneg <| add_nonneg hx.le zero_le_one) - exact le_add_of_nonneg_right zero_le_one + exact le_trans (le_add_of_nonneg_right zero_le_one) (add_one_le_exp _) #align behrend.exp_neg_two_mul_le Behrend.exp_neg_two_mul_le theorem div_lt_floor {x : ℝ} (hx : 2 / (1 - 2 / exp 1) ≤ x) : x / exp 1 < (⌊x / 2⌋₊ : ℝ) := by @@ -372,22 +367,14 @@ noncomputable def nValue (N : ℕ) : ℕ := #align behrend.n_value Behrend.nValue /-- The (almost) optimal value of `d` in `Behrend.bound_aux`. -/ -noncomputable def dValue (N : ℕ) : ℕ := - ⌊(N : ℝ) ^ (1 / nValue N : ℝ) / 2⌋₊ +noncomputable def dValue (N : ℕ) : ℕ := ⌊(N : ℝ) ^ (nValue N : ℝ)⁻¹ / 2⌋₊ #align behrend.d_value Behrend.dValue theorem nValue_pos (hN : 2 ≤ N) : 0 < nValue N := ceil_pos.2 <| Real.sqrt_pos.2 <| log_pos <| one_lt_cast.2 <| hN #align behrend.n_value_pos Behrend.nValue_pos -theorem two_le_nValue (hN : 3 ≤ N) : 2 ≤ nValue N := by - refine' succ_le_of_lt (lt_ceil.2 <| lt_sqrt_of_sq_lt _) - rw [cast_one, one_pow, lt_log_iff_exp_lt] - refine' lt_of_lt_of_le _ (cast_le.2 hN) - · exact exp_one_lt_d9.trans_le (by norm_num) - rw [cast_pos] - exact (zero_lt_succ _).trans_le hN -#align behrend.two_le_n_value Behrend.two_le_nValue +#noalign behrend.two_le_n_value theorem three_le_nValue (hN : 64 ≤ N) : 3 ≤ nValue N := by rw [nValue, ← lt_iff_add_one_le, lt_ceil, cast_two] @@ -395,18 +382,16 @@ theorem three_le_nValue (hN : 64 ≤ N) : 3 ≤ nValue N := by have : (2 : ℝ) ^ ((6 : ℕ) : ℝ) ≤ N := by rw [rpow_nat_cast] exact (cast_le.2 hN).trans' (by norm_num1) - apply lt_of_lt_of_le _ ((log_le_log (rpow_pos_of_pos zero_lt_two _) _).2 this) + apply lt_of_lt_of_le _ (log_le_log (rpow_pos_of_pos zero_lt_two _) this) rw [log_rpow zero_lt_two, ← div_lt_iff'] · exact log_two_gt_d9.trans_le' (by norm_num1) · norm_num1 - rw [cast_pos] - exact hN.trans_lt' (by norm_num1) #align behrend.three_le_n_value Behrend.three_le_nValue theorem dValue_pos (hN₃ : 8 ≤ N) : 0 < dValue N := by have hN₀ : 0 < (N : ℝ) := cast_pos.2 (succ_pos'.trans_le hN₃) - rw [dValue, floor_pos, ← log_le_log zero_lt_one, log_one, log_div _ two_ne_zero, log_rpow hN₀, - div_mul_eq_mul_div, one_mul, sub_nonneg, le_div_iff] + rw [dValue, floor_pos, ← log_le_log_iff zero_lt_one, log_one, log_div _ two_ne_zero, log_rpow hN₀, + inv_mul_eq_div, sub_nonneg, le_div_iff] · have : (nValue N : ℝ) ≤ 2 * sqrt (log N) := by apply (ceil_lt_add_one <| sqrt_nonneg _).le.trans rw [two_mul, add_le_add_iff_left] @@ -417,9 +402,7 @@ theorem dValue_pos (hN₃ : 8 ≤ N) : 0 < dValue N := by rw [← mul_assoc, ← le_div_iff (Real.sqrt_pos.2 <| log_pos <| one_lt_cast.2 _), div_sqrt] · apply log_two_mul_two_le_sqrt_log_eight.trans apply Real.sqrt_le_sqrt - rw [log_le_log _ hN₀] - · exact mod_cast hN₃ - · norm_num + exact log_le_log (by norm_num) (mod_cast hN₃) exact hN₃.trans_lt' (by norm_num) · exact cast_pos.2 (nValue_pos <| hN₃.trans' <| by norm_num) · exact (rpow_pos_of_pos hN₀ _).ne' @@ -428,13 +411,13 @@ theorem dValue_pos (hN₃ : 8 ≤ N) : 0 < dValue N := by theorem le_N (hN : 2 ≤ N) : (2 * dValue N - 1) ^ nValue N ≤ N := by have : (2 * dValue N - 1) ^ nValue N ≤ (2 * dValue N) ^ nValue N := - Nat.pow_le_pow_of_le_left (Nat.sub_le _ _) _ + Nat.pow_le_pow_left (Nat.sub_le _ _) _ apply this.trans suffices ((2 * dValue N) ^ nValue N : ℝ) ≤ N from mod_cast this - suffices i : (2 * dValue N : ℝ) ≤ (N : ℝ) ^ (1 / nValue N : ℝ) + suffices i : (2 * dValue N : ℝ) ≤ (N : ℝ) ^ (nValue N : ℝ)⁻¹ · rw [← rpow_nat_cast] apply (rpow_le_rpow (mul_nonneg zero_le_two (cast_nonneg _)) i (cast_nonneg _)).trans - rw [← rpow_mul (cast_nonneg _), one_div_mul_cancel, rpow_one] + rw [← rpow_mul (cast_nonneg _), inv_mul_cancel, rpow_one] rw [cast_ne_zero] apply (nValue_pos hN).ne' rw [← le_div_iff'] @@ -443,9 +426,9 @@ theorem le_N (hN : 2 ≤ N) : (2 * dValue N - 1) ^ nValue N ≤ N := by set_option linter.uppercaseLean3 false in #align behrend.le_N Behrend.le_N -theorem bound (hN : 4096 ≤ N) : (N : ℝ) ^ (1 / nValue N : ℝ) / exp 1 < dValue N := by +theorem bound (hN : 4096 ≤ N) : (N : ℝ) ^ (nValue N : ℝ)⁻¹ / exp 1 < dValue N := by apply div_lt_floor _ - rw [← log_le_log, log_rpow, mul_comm, ← div_eq_mul_one_div] + rw [← log_le_log_iff, log_rpow, mul_comm, ← div_eq_mul_inv] · apply le_trans _ (div_le_div_of_le_left _ _ (ceil_lt_mul _).le) rw [mul_comm, ← div_div, div_sqrt, le_div_iff] · norm_num; exact le_sqrt_log hN @@ -458,13 +441,9 @@ theorem bound (hN : 4096 ≤ N) : (N : ℝ) ^ (1 / nValue N : ℝ) / exp 1 < dVa rw [one_lt_cast] exact hN.trans_lt' (by norm_num1) apply le_sqrt_of_sq_le - have : ((12 : ℕ) : ℝ) * log 2 ≤ log N := by - rw [← log_rpow zero_lt_two, log_le_log, rpow_nat_cast] - · norm_num1 - exact mod_cast hN - · exact rpow_pos_of_pos zero_lt_two _ - rw [cast_pos] - exact hN.trans_lt' (by norm_num1) + have : (12 : ℕ) * log 2 ≤ log N := by + rw [← log_rpow zero_lt_two, rpow_nat_cast] + exact log_le_log (by positivity) (mod_cast hN) refine' le_trans _ this rw [← div_le_iff'] · exact log_two_gt_d9.le.trans' (by norm_num1) @@ -474,9 +453,7 @@ theorem bound (hN : 4096 ≤ N) : (N : ℝ) ^ (1 / nValue N : ℝ) / exp 1 < dVa · refine' div_pos zero_lt_two _ rw [sub_pos, div_lt_one (exp_pos _)] exact lt_of_le_of_lt (by norm_num1) exp_one_gt_d9 - apply rpow_pos_of_pos - rw [cast_pos] - exact hN.trans_lt' (by norm_num1) + positivity #align behrend.bound Behrend.bound theorem roth_lower_bound_explicit (hN : 4096 ≤ N) : @@ -485,14 +462,15 @@ theorem roth_lower_bound_explicit (hN : 4096 ≤ N) : have hn : 0 < (n : ℝ) := cast_pos.2 (nValue_pos <| hN.trans' <| by norm_num1) have hd : 0 < dValue N := dValue_pos (hN.trans' <| by norm_num1) have hN₀ : 0 < (N : ℝ) := cast_pos.2 (hN.trans' <| by norm_num1) - have hn₂ : 2 ≤ n := two_le_nValue (hN.trans' <| by norm_num1) + have hn₂ : 2 < n := three_le_nValue $ hN.trans' $ by norm_num1 have : (2 * dValue N - 1) ^ n ≤ N := le_N (hN.trans' <| by norm_num1) - refine' ((bound_aux hd.ne' hn₂).trans <| cast_le.2 <| rothNumberNat.mono this).trans_lt' _ - refine' (div_lt_div_of_lt hn <| pow_lt_pow_of_lt_left (bound hN) _ _).trans_le' _ - · exact div_nonneg (rpow_nonneg_of_nonneg (cast_nonneg _) _) (exp_pos _).le - · exact tsub_pos_of_lt (three_le_nValue <| hN.trans' <| by norm_num1) + calc + _ ≤ (N ^ (nValue N : ℝ)⁻¹ / rexp 1 : ℝ) ^ (n - 2) / n := ?_ + _ < _ := by gcongr; exacts [(tsub_pos_of_lt hn₂).ne', bound hN] + _ ≤ rothNumberNat ((2 * dValue N - 1) ^ n) := bound_aux hd.ne' hn₂.le + _ ≤ rothNumberNat N := mod_cast rothNumberNat.mono this rw [← rpow_nat_cast, div_rpow (rpow_nonneg_of_nonneg hN₀.le _) (exp_pos _).le, ← rpow_mul hN₀.le, - mul_comm (_ / _), mul_one_div, cast_sub hn₂, cast_two, same_sub_div hn.ne', exp_one_rpow, + inv_mul_eq_div, cast_sub hn₂.le, cast_two, same_sub_div hn.ne', exp_one_rpow, div_div, rpow_sub hN₀, rpow_one, div_div, div_eq_mul_inv] refine' mul_le_mul_of_nonneg_left _ (cast_nonneg _) rw [mul_inv, mul_inv, ← exp_neg, ← rpow_neg (cast_nonneg _), neg_sub, ← div_eq_mul_inv] @@ -528,7 +506,7 @@ theorem four_zero_nine_six_lt_exp_sixteen : 4096 < exp 16 := by theorem lower_bound_le_one' (hN : 2 ≤ N) (hN' : N ≤ 4096) : (N : ℝ) * exp (-4 * sqrt (log N)) ≤ 1 := by - rw [← log_le_log (mul_pos (cast_pos.2 (zero_lt_two.trans_le hN)) (exp_pos _)) zero_lt_one, + rw [← log_le_log_iff (mul_pos (cast_pos.2 (zero_lt_two.trans_le hN)) (exp_pos _)) zero_lt_one, log_one, log_mul (cast_pos.2 (zero_lt_two.trans_le hN)).ne' (exp_pos _).ne', log_exp, neg_mul, ← sub_eq_add_neg, sub_nonpos, ← div_le_iff (Real.sqrt_pos.2 <| log_pos <| one_lt_cast.2 <| one_lt_two.trans_le hN), div_sqrt, diff --git a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean index eb813b4bb3210..b0551f599bb24 100644 --- a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean +++ b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Data.Finset.Pointwise +import Mathlib.SetTheory.Cardinal.Finite #align_import combinatorics.additive.ruzsa_covering from "leanprover-community/mathlib"@"b363547b3113d350d053abdf2884e9850a56b205" @@ -54,3 +55,23 @@ theorem exists_subset_mul_div (ht : t.Nonempty) : #align finset.exists_subset_add_sub Finset.exists_subset_add_sub end Finset + +namespace Set +variable {α : Type*} [CommGroup α] {s t : Set α} + +/-- **Ruzsa's covering lemma** for sets. See also `Finset.exists_subset_mul_div`. -/ +@[to_additive "**Ruzsa's covering lemma**. Version for sets. For finsets, +see `Finset.exists_subset_add_sub`."] +lemma exists_subset_mul_div (hs : s.Finite) (ht' : t.Finite) (ht : t.Nonempty) : + ∃ u : Set α, Nat.card u * Nat.card t ≤ Nat.card (s * t) ∧ s ⊆ u * t / t ∧ u.Finite := by + lift s to Finset α using hs + lift t to Finset α using ht' + classical + obtain ⟨u, hu, hsut⟩ := Finset.exists_subset_mul_div s ht + refine ⟨u, ?_⟩ + -- `norm_cast` would find these automatically, but breaks `to_additive` when it does so + rw [← Finset.coe_mul, ← Finset.coe_mul, ← Finset.coe_div] + norm_cast + simp [*] + +end Set diff --git a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean index 12049feaf31d6..a75d667611083 100644 --- a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean +++ b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean @@ -358,4 +358,5 @@ lemma Finset.le_card_diffs_mul_card_diffs (s t : Finset α) : /-- The **Marica-Schönheim Inequality**. -/ lemma Finset.card_le_card_diffs (s : Finset α) : s.card ≤ (s \\ s).card := - le_of_pow_le_pow 2 (zero_le _) two_pos $ by simpa [← sq] using s.le_card_diffs_mul_card_diffs s + le_of_pow_le_pow_left two_ne_zero (zero_le _) $ by + simpa [← sq] using s.le_card_diffs_mul_card_diffs s diff --git a/Mathlib/Combinatorics/SetFamily/Kleitman.lean b/Mathlib/Combinatorics/SetFamily/Kleitman.lean index 21c5d6084a869..559b18e16d228 100644 --- a/Mathlib/Combinatorics/SetFamily/Kleitman.lean +++ b/Mathlib/Combinatorics/SetFamily/Kleitman.lean @@ -81,6 +81,6 @@ theorem Finset.card_biUnion_le_of_intersecting (s : Finset ι) (f : ι → Finse (ih _ (fun i hi ↦ (hf₁ _ <| subset_cons _ hi).2.2) ((card_le_of_subset <| subset_cons _).trans hs)) _).trans _ rw [mul_tsub, two_mul, ← pow_succ, - ← add_tsub_assoc_of_le (pow_le_pow' (one_le_two : (1 : ℕ) ≤ 2) tsub_le_self), + ← add_tsub_assoc_of_le (pow_le_pow_right' (one_le_two : (1 : ℕ) ≤ 2) tsub_le_self), tsub_add_eq_add_tsub hs, card_cons, add_tsub_add_eq_tsub_right] #align finset.card_bUnion_le_of_intersecting Finset.card_biUnion_le_of_intersecting diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean index 866a22b248f8c..87b8f92135ef7 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean @@ -74,7 +74,7 @@ private theorem eps_pos {ε : ℝ} {n : ℕ} (h : 100 ≤ (4 : ℝ) ^ n * ε ^ 5 (pos_of_mul_pos_right ((show 0 < (100 : ℝ) by norm_num).trans_le h) (by positivity)) private theorem m_pos [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) : 0 < m := - Nat.div_pos ((Nat.mul_le_mul_left _ <| Nat.pow_le_pow_of_le_left (by norm_num) _).trans hPα) <| + Nat.div_pos ((Nat.mul_le_mul_left _ <| Nat.pow_le_pow_left (by norm_num) _).trans hPα) <| stepBound_pos (P.parts_nonempty <| univ_nonempty.ne_empty).card_pos /-- Local extension for the `positivity` tactic: A few facts that are needed many times for the diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean index 4d389e9f879b4..0b50659e47bda 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean @@ -120,7 +120,7 @@ private theorem card_nonuniformWitness_sdiff_biUnion_star (hV : V ∈ P.parts) ( rw [sum_const] refine' mul_le_mul_right' _ _ have t := card_filter_atomise_le_two_pow (s := U) hX - refine' t.trans (pow_le_pow (by norm_num) <| tsub_le_tsub_right _ _) + refine' t.trans (pow_le_pow_right (by norm_num) <| tsub_le_tsub_right _ _) exact card_image_le.trans (card_le_of_subset <| filter_subset _ _) private theorem one_sub_eps_mul_card_nonuniformWitness_le_card_star (hV : V ∈ P.parts) @@ -245,7 +245,7 @@ private theorem m_add_one_div_m_le_one_add [Nonempty α] have : ↑1 + ↑1 / (m : ℝ) ≤ ↑1 + ε ^ 5 / 100 := by rw [add_le_add_iff_left, ← one_div_div (100 : ℝ)] exact one_div_le_one_div_of_le (by sz_positivity) (hundred_div_ε_pow_five_le_m hPα hPε) - refine' (pow_le_pow_of_le_left _ this 2).trans _ + refine' (pow_le_pow_left _ this 2).trans _ · positivity rw [add_sq, one_pow, add_assoc, add_le_add_iff_left, mul_one, ← le_sub_iff_add_le', div_eq_mul_one_div _ (49 : ℝ), mul_div_left_comm (2 : ℝ), ← mul_sub_left_distrib, div_pow, @@ -367,7 +367,7 @@ private theorem edgeDensity_chunk_aux [Nonempty α] · rw [biUnion_parts, biUnion_parts] · rw [card_chunk (m_pos hPα).ne', card_chunk (m_pos hPα).ne', ← cast_mul, ← mul_pow, cast_pow] norm_cast - refine' le_trans _ (pow_le_pow_of_le_left hGε this 2) + refine' le_trans _ (pow_le_pow_left hGε this 2) rw [sub_sq, sub_add, sub_le_sub_iff_left] refine' (sub_le_self _ <| sq_nonneg <| ε ^ 5 / 50).trans _ rw [mul_right_comm, mul_div_left_comm, div_eq_mul_inv (ε ^ 5), diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean index 2e0c36d7492bf..0be6b52ec2371 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean @@ -65,7 +65,7 @@ variable {hP G ε} theorem card_increment (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPG : ¬P.IsUniform G ε) : (increment hP G ε).parts.card = stepBound P.parts.card := by have hPα' : stepBound P.parts.card ≤ card α := - (mul_le_mul_left' (pow_le_pow_of_le_left' (by norm_num) _) _).trans hPα + (mul_le_mul_left' (pow_le_pow_left' (by norm_num) _) _).trans hPα have hPpos : 0 < stepBound P.parts.card := stepBound_pos (nonempty_of_not_uniform hPG).card_pos rw [increment, card_bind] simp_rw [chunk, apply_dite Finpartition.parts, apply_dite card, sum_dite] diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean index ab1ac9b9e648f..caad9b8ea7bf5 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean @@ -104,7 +104,7 @@ theorem szemeredi_regularity (hε : 0 < ε) (hl : l ≤ card α) : refine' ⟨P, hP₁, (le_initialBound _ _).trans hP₂, hP₃.trans _, hP₄.resolve_right fun hPenergy => lt_irrefl (1 : ℝ) _⟩ · rw [iterate_succ_apply'] - exact mul_le_mul_left' (pow_le_pow_of_le_left (by norm_num) (by norm_num) _) _ + exact mul_le_mul_left' (pow_le_pow_left (by norm_num) (by norm_num) _) _ calc (1 : ℝ) = ε ^ 5 / ↑4 * (↑4 / ε ^ 5) := by rw [mul_comm, div_mul_div_cancel 4 (pow_pos hε 5).ne']; norm_num @@ -131,7 +131,7 @@ theorem szemeredi_regularity (hε : 0 < ε) (hl : l ≤ card α) : -- We gather a few numerical facts. have hεl' : ↑100 < ↑4 ^ P.parts.card * ε ^ 5 := (hundred_lt_pow_initialBound_mul hε l).trans_le - (mul_le_mul_of_nonneg_right (pow_le_pow (by norm_num) hP₂) <| by positivity) + (mul_le_mul_of_nonneg_right (pow_le_pow_right (by norm_num) hP₂) <| by positivity) have hi : (i : ℝ) ≤ 4 / ε ^ 5 := by have hi : ε ^ 5 / 4 * ↑i ≤ 1 := hP₄.trans (mod_cast P.energy_le_one G) rw [div_mul_eq_mul_div, div_le_iff (show (0 : ℝ) < 4 by norm_num)] at hi diff --git a/Mathlib/Computability/Ackermann.lean b/Mathlib/Computability/Ackermann.lean index 2a4625465eeec..436e1d581b3a1 100644 --- a/Mathlib/Computability/Ackermann.lean +++ b/Mathlib/Computability/Ackermann.lean @@ -101,7 +101,7 @@ theorem ack_three (n : ℕ) : ack 3 n = 2 ^ (n + 3) - 3 := by Nat.mul_sub_left_distrib, ← Nat.sub_add_comm, two_mul 3, Nat.add_sub_add_right] have H : 2 * 3 ≤ 2 * 2 ^ 3 := by norm_num apply H.trans - simp [pow_le_pow (show 1 ≤ 2 by norm_num)] + simp [pow_le_pow_right (show 1 ≤ 2 by norm_num)] #align ack_three ack_three theorem ack_pos : ∀ m n, 0 < ack m n @@ -297,7 +297,7 @@ theorem ack_ack_lt_ack_max_add_two (m n k : ℕ) : ack m (ack n k) < ack (max m theorem ack_add_one_sq_lt_ack_add_four (m n : ℕ) : ack m ((n + 1) ^ 2) < ack (m + 4) n := calc ack m ((n + 1) ^ 2) < ack m ((ack m n + 1) ^ 2) := - ack_strictMono_right m <| pow_lt_pow_of_lt_left (succ_lt_succ <| lt_ack_right m n) zero_lt_two + ack_strictMono_right m <| Nat.pow_lt_pow_left (succ_lt_succ <| lt_ack_right m n) two_ne_zero _ ≤ ack m (ack (m + 3) n) := ack_mono_right m <| ack_add_one_sq_lt_ack_add_three m n _ ≤ ack (m + 2) (ack (m + 3) n) := ack_mono_left _ <| by linarith _ = ack (m + 3) (n + 1) := (ack_succ_succ _ n).symm @@ -332,7 +332,7 @@ theorem exists_lt_ack_of_nat_primrec {f : ℕ → ℕ} (hf : Nat.Primrec f) : · refine' ⟨max a b + 3, fun n => (pair_lt_max_add_one_sq _ _).trans_le <| - (pow_le_pow_of_le_left (add_le_add_right _ _) 2).trans <| + (Nat.pow_le_pow_left (add_le_add_right _ _) 2).trans <| ack_add_one_sq_lt_ack_add_three _ _⟩ rw [max_ack_left] exact max_le_max (ha n).le (hb n).le diff --git a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean index 191629749c7b8..8a595d5a9522b 100644 --- a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean +++ b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean @@ -302,7 +302,7 @@ lemma one_add_smoothingFn_le_two {x : ℝ} (hx : exp 1 ≤ x) : 1 + ε x ≤ 2 : _ ≤ x := hx rw [div_le_one (log_pos this)] calc 1 = log (exp 1) := by simp - _ ≤ log x := log_le_log' (exp_pos _) hx + _ ≤ log x := log_le_log (exp_pos _) hx lemma isLittleO_smoothingFn_one : ε =o[atTop] (fun _ => (1:ℝ)) := by unfold smoothingFn diff --git a/Mathlib/Control/Random.lean b/Mathlib/Control/Random.lean index 1ebd61c4d1d9e..e904a25a471b4 100644 --- a/Mathlib/Control/Random.lean +++ b/Mathlib/Control/Random.lean @@ -132,7 +132,7 @@ instance : BoundedRandom m Int where Int.le_add_of_nonneg_left (Int.ofNat_zero_le z), Int.add_le_of_le_sub_right $ Int.le_trans (Int.ofNat_le.mpr h2) - (le_of_eq $ Int.ofNat_natAbs_eq_of_nonneg _ $ Int.sub_nonneg_of_le h)⟩ + (le_of_eq $ Int.natAbs_of_nonneg $ Int.sub_nonneg_of_le h)⟩ instance {n : Nat} : BoundedRandom m (Fin n) where randomR lo hi h _ := do diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index eb88a2a8a4df1..dd362c906dc7a 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -541,6 +541,9 @@ theorem exp_sum {α : Type*} (s : Finset α) (f : α → ℂ) : @map_prod (Multiplicative ℂ) α ℂ _ _ _ _ expMonoidHom f s #align complex.exp_sum Complex.exp_sum +lemma exp_nsmul (x : ℂ) (n : ℕ) : exp (n • x) = exp x ^ n := + @MonoidHom.map_pow (Multiplicative ℂ) ℂ _ _ expMonoidHom _ _ + theorem exp_nat_mul (x : ℂ) : ∀ n : ℕ, exp (n * x) = exp x ^ n | 0 => by rw [Nat.cast_zero, zero_mul, exp_zero, pow_zero] | Nat.succ n => by rw [pow_succ', Nat.cast_add_one, add_mul, exp_add, ← exp_nat_mul _ n, one_mul] @@ -1159,6 +1162,9 @@ theorem exp_sum {α : Type*} (s : Finset α) (f : α → ℝ) : @map_prod (Multiplicative ℝ) α ℝ _ _ _ _ expMonoidHom f s #align real.exp_sum Real.exp_sum +lemma exp_nsmul (x : ℝ) (n : ℕ) : exp (n • x) = exp x ^ n := + @MonoidHom.map_pow (Multiplicative ℝ) ℝ _ _ expMonoidHom _ _ + nonrec theorem exp_nat_mul (x : ℝ) (n : ℕ) : exp (n * x) = exp x ^ n := ofReal_injective (by simp [exp_nat_mul]) #align real.exp_nat_mul Real.exp_nat_mul @@ -1475,6 +1481,12 @@ theorem sum_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) (n : ℕ) : ∑ i in range _ = exp x := by rw [exp, Complex.exp, ← cauSeqRe, lim_re] #align real.sum_le_exp_of_nonneg Real.sum_le_exp_of_nonneg +lemma pow_div_factorial_le_exp (hx : 0 ≤ x) (n : ℕ) : x ^ n / n ! ≤ exp x := + calc + x ^ n / n ! ≤ ∑ k in range (n + 1), x ^ k / k ! := + single_le_sum (f := fun k ↦ x ^ k / k !) (fun k _ ↦ by positivity) (self_mem_range_succ n) + _ ≤ exp x := sum_le_exp_of_nonneg hx _ + theorem quadratic_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) : 1 + x + x ^ 2 / 2 ≤ exp x := calc 1 + x + x ^ 2 / 2 = ∑ i in range 3, x ^ i / i ! := by @@ -1485,17 +1497,13 @@ theorem quadratic_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) : 1 + x + x ^ 2 / 2 _ ≤ exp x := sum_le_exp_of_nonneg hx 3 #align real.quadratic_le_exp_of_nonneg Real.quadratic_le_exp_of_nonneg -theorem add_one_lt_exp_of_pos {x : ℝ} (hx : 0 < x) : x + 1 < exp x := +private theorem add_one_lt_exp_of_pos {x : ℝ} (hx : 0 < x) : x + 1 < exp x := (by nlinarith : x + 1 < 1 + x + x ^ 2 / 2).trans_le (quadratic_le_exp_of_nonneg hx.le) -#align real.add_one_lt_exp_of_pos Real.add_one_lt_exp_of_pos -/-- This is an intermediate result that is later replaced by `Real.add_one_le_exp`; use that lemma -instead. -/ -theorem add_one_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) : x + 1 ≤ exp x := by +private theorem add_one_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) : x + 1 ≤ exp x := by rcases eq_or_lt_of_le hx with (rfl | h) · simp exact (add_one_lt_exp_of_pos h).le -#align real.add_one_le_exp_of_nonneg Real.add_one_le_exp_of_nonneg theorem one_le_exp {x : ℝ} (hx : 0 ≤ x) : 1 ≤ exp x := by linarith [add_one_le_exp_of_nonneg hx] #align real.one_le_exp Real.one_le_exp @@ -1506,11 +1514,16 @@ theorem exp_pos (x : ℝ) : 0 < exp x := exact inv_pos.2 (lt_of_lt_of_le zero_lt_one (one_le_exp (neg_nonneg.2 h))) #align real.exp_pos Real.exp_pos +lemma exp_nonneg (x : ℝ) : 0 ≤ exp x := x.exp_pos.le + @[simp] theorem abs_exp (x : ℝ) : |exp x| = exp x := abs_of_pos (exp_pos _) #align real.abs_exp Real.abs_exp +lemma exp_abs_le (x : ℝ) : exp |x| ≤ exp x + exp (-x) := by + cases le_total x 0 <;> simp [abs_of_nonpos, _root_.abs_of_nonneg, exp_nonneg, *] + @[mono] theorem exp_strictMono : StrictMono exp := fun x y h => by rw [← sub_add_cancel y x, Real.exp_add] @@ -1577,7 +1590,7 @@ theorem cosh_pos (x : ℝ) : 0 < Real.cosh x := #align real.cosh_pos Real.cosh_pos theorem sinh_lt_cosh : sinh x < cosh x := - lt_of_pow_lt_pow 2 (cosh_pos _).le <| (cosh_sq x).symm ▸ lt_add_one _ + lt_of_pow_lt_pow_left 2 (cosh_pos _).le <| (cosh_sq x).symm ▸ lt_add_one _ #align real.sinh_lt_cosh Real.sinh_lt_cosh end Real @@ -1970,48 +1983,42 @@ theorem exp_bound_div_one_sub_of_interval {x : ℝ} (h1 : 0 ≤ x) (h2 : x < 1) · exact (exp_bound_div_one_sub_of_interval' h1 h2).le #align real.exp_bound_div_one_sub_of_interval Real.exp_bound_div_one_sub_of_interval -theorem one_sub_lt_exp_minus_of_pos {y : ℝ} (h : 0 < y) : 1 - y < Real.exp (-y) := by - cases' le_or_lt 1 y with h' h' - · linarith [(-y).exp_pos] - rw [exp_neg, lt_inv _ y.exp_pos, inv_eq_one_div] - · exact exp_bound_div_one_sub_of_interval' h h' - · linarith -#align real.one_sub_le_exp_minus_of_pos Real.one_sub_lt_exp_minus_of_pos +theorem add_one_lt_exp {x : ℝ} (hx : x ≠ 0) : x + 1 < Real.exp x := by + obtain hx | hx := hx.symm.lt_or_lt + · exact add_one_lt_exp_of_pos hx + obtain h' | h' := le_or_lt 1 (-x) + · linarith [x.exp_pos] + have hx' : 0 < x + 1 := by linarith + simpa [add_comm, exp_neg, inv_lt_inv (exp_pos _) hx'] + using exp_bound_div_one_sub_of_interval' (neg_pos.2 hx) h' +#align real.add_one_lt_exp_of_nonzero Real.add_one_lt_exp +#align real.add_one_lt_exp_of_pos Real.add_one_lt_exp -theorem one_sub_le_exp_minus_of_nonneg {y : ℝ} (h : 0 ≤ y) : 1 - y ≤ Real.exp (-y) := by - rcases eq_or_lt_of_le h with (rfl | h) +theorem add_one_le_exp (x : ℝ) : x + 1 ≤ Real.exp x := by + obtain rfl | hx := eq_or_ne x 0 · simp - · exact (one_sub_lt_exp_minus_of_pos h).le -#align real.one_sub_le_exp_minus_of_nonneg Real.one_sub_le_exp_minus_of_nonneg - -theorem add_one_lt_exp_of_neg {x : ℝ} (h : x < 0) : x + 1 < Real.exp x := by - have h1 : 0 < -x := by linarith - simpa [add_comm] using one_sub_lt_exp_minus_of_pos h1 -#align real.add_one_lt_exp_of_neg Real.add_one_lt_exp_of_neg + · exact (add_one_lt_exp hx).le +#align real.add_one_le_exp Real.add_one_le_exp +#align real.add_one_le_exp_of_nonneg Real.add_one_le_exp -theorem add_one_lt_exp_of_nonzero {x : ℝ} (hx : x ≠ 0) : x + 1 < Real.exp x := by - cases' lt_or_gt_of_ne hx with h h - · exact add_one_lt_exp_of_neg h - exact add_one_lt_exp_of_pos h -#align real.add_one_lt_exp_of_nonzero Real.add_one_lt_exp_of_nonzero +lemma one_sub_lt_exp_neg {x : ℝ} (hx : x ≠ 0) : 1 - x < exp (-x) := + (sub_eq_neg_add _ _).trans_lt $ add_one_lt_exp $ neg_ne_zero.2 hx -theorem add_one_le_exp (x : ℝ) : x + 1 ≤ Real.exp x := by - cases' le_or_lt 0 x with h h - · exact Real.add_one_le_exp_of_nonneg h - exact (add_one_lt_exp_of_neg h).le -#align real.add_one_le_exp Real.add_one_le_exp +lemma one_sub_le_exp_neg (x : ℝ) : 1 - x ≤ exp (-x) := + (sub_eq_neg_add _ _).trans_le $ add_one_le_exp _ +#align real.one_sub_le_exp_minus_of_pos Real.one_sub_le_exp_neg +#align real.one_sub_le_exp_minus_of_nonneg Real.one_sub_le_exp_neg theorem one_sub_div_pow_le_exp_neg {n : ℕ} {t : ℝ} (ht' : t ≤ n) : (1 - t / n) ^ n ≤ exp (-t) := by rcases eq_or_ne n 0 with (rfl | hn) · simp rwa [Nat.cast_zero] at ht' - convert pow_le_pow_of_le_left ?_ (add_one_le_exp (-(t / n))) n using 2 - · abel + convert pow_le_pow_left ?_ (one_sub_le_exp_neg (t / n)) n using 2 · rw [← Real.exp_nat_mul] congr 1 field_simp ring_nf - · rwa [add_comm, ← sub_eq_add_neg, sub_nonneg, div_le_one] + · rwa [sub_nonneg, div_le_one] positivity #align real.one_sub_div_pow_le_exp_neg Real.one_sub_div_pow_le_exp_neg diff --git a/Mathlib/Data/DFinsupp/Order.lean b/Mathlib/Data/DFinsupp/Order.lean index fa58b58413fd1..f0c4c51158c53 100644 --- a/Mathlib/Data/DFinsupp/Order.lean +++ b/Mathlib/Data/DFinsupp/Order.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import Mathlib.Algebra.Order.Module.Defs import Mathlib.Data.DFinsupp.Basic #align_import data.dfinsupp.order from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" @@ -150,6 +151,42 @@ instance [∀ i, OrderedAddCommMonoid (α i)] [∀ i, ContravariantClass (α i) ContravariantClass (Π₀ i, α i) (Π₀ i, α i) (· + ·) (· ≤ ·) := ⟨fun _ _ _ H i ↦ le_of_add_le_add_left (H i)⟩ +section Module +variable {α : Type*} {β : ι → Type*} [Semiring α] [Preorder α] [∀ i, AddCommMonoid (β i)] + [∀ i, Preorder (β i)] [∀ i, Module α (β i)] + +instance instPosSMulMono [∀ i, PosSMulMono α (β i)] : PosSMulMono α (Π₀ i, β i) := + PosSMulMono.lift _ coe_le_coe coe_smul + +instance instSMulPosMono [∀ i, SMulPosMono α (β i)] : SMulPosMono α (Π₀ i, β i) := + SMulPosMono.lift _ coe_le_coe coe_smul coe_zero + +instance instPosSMulReflectLE [∀ i, PosSMulReflectLE α (β i)] : PosSMulReflectLE α (Π₀ i, β i) := + PosSMulReflectLE.lift _ coe_le_coe coe_smul + +instance instSMulPosReflectLE [∀ i, SMulPosReflectLE α (β i)] : SMulPosReflectLE α (Π₀ i, β i) := + SMulPosReflectLE.lift _ coe_le_coe coe_smul coe_zero + +end Module + +section Module +variable {α : Type*} {β : ι → Type*} [Semiring α] [PartialOrder α] [∀ i, AddCommMonoid (β i)] + [∀ i, PartialOrder (β i)] [∀ i, Module α (β i)] + +instance instPosSMulStrictMono [∀ i, PosSMulStrictMono α (β i)] : PosSMulStrictMono α (Π₀ i, β i) := + PosSMulStrictMono.lift _ coe_le_coe coe_smul + +instance instSMulPosStrictMono [∀ i, SMulPosStrictMono α (β i)] : SMulPosStrictMono α (Π₀ i, β i) := + SMulPosStrictMono.lift _ coe_le_coe coe_smul coe_zero + +-- Note: There is no interesting instance for `PosSMulReflectLT α (Π₀ i, β i)` that's not already +-- implied by the other instances + +instance instSMulPosReflectLT [∀ i, SMulPosReflectLT α (β i)] : SMulPosReflectLT α (Π₀ i, β i) := + SMulPosReflectLT.lift _ coe_le_coe coe_smul coe_zero + +end Module + section CanonicallyOrderedAddCommMonoid -- porting note: Split into 2 lines to satisfy the unusedVariables linter. diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index eae91f1518a91..cf1d121c7c88d 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -267,6 +267,11 @@ theorem fiber_card_ne_zero_iff_mem_image (s : Finset α) (f : α → β) [Decida rw [← pos_iff_ne_zero, card_pos, fiber_nonempty_iff_mem_image] #align finset.fiber_card_ne_zero_iff_mem_image Finset.fiber_card_ne_zero_iff_mem_image +lemma card_filter_le_iff (s : Finset α) (P : α → Prop) [DecidablePred P] (n : ℕ) : + (s.filter P).card ≤ n ↔ ∀ s' ⊆ s, n < s'.card → ∃ a ∈ s', ¬ P a := + (s.1.card_filter_le_iff P n).trans ⟨fun H s' hs' h ↦ H s'.1 (by aesop) h, + fun H s' hs' h ↦ H ⟨s', nodup_of_le hs' s.2⟩ (fun x hx ↦ subset_of_le hs' hx) h⟩ + @[simp] theorem card_map (f : α ↪ β) : (s.map f).card = s.card := Multiset.card_map _ _ diff --git a/Mathlib/Data/Finset/Pointwise.lean b/Mathlib/Data/Finset/Pointwise.lean index 9d3ab9242823f..7feea71720368 100644 --- a/Mathlib/Data/Finset/Pointwise.lean +++ b/Mathlib/Data/Finset/Pointwise.lean @@ -8,6 +8,7 @@ import Mathlib.Data.Finset.Preimage import Mathlib.Data.Set.Pointwise.Finite import Mathlib.Data.Set.Pointwise.SMul import Mathlib.Data.Set.Pointwise.ListOfFn +import Mathlib.SetTheory.Cardinal.Finite #align_import data.finset.pointwise from "leanprover-community/mathlib"@"eba7871095e834365616b5e43c8c7bb0b37058d0" @@ -2309,4 +2310,47 @@ theorem Finite.toFinset_vsub (hs : s.Finite) (ht : t.Finite) (hf := hs.vsub ht) end VSub +section MulAction +variable [Group α] [MulAction α β] + +@[to_additive (attr := simp)] +lemma card_smul_set (a : α) (s : Set β) : Nat.card ↥(a • s) = Nat.card s := + Nat.card_image_of_injective (MulAction.injective a) _ + +end MulAction + +section IsCancelMul +variable [Mul α] [IsCancelMul α] {s t : Set α} + +@[to_additive] +lemma card_mul_le : Nat.card (s * t) ≤ Nat.card s * Nat.card t := by + classical + obtain h | h := (s * t).infinite_or_finite + · simp [Set.Infinite.card_eq_zero h] + obtain ⟨hs, ht⟩ | rfl | rfl := finite_mul.1 h + · lift s to Finset α using hs + lift t to Finset α using ht + rw [← Finset.coe_mul] + simpa [-Finset.coe_mul] using Finset.card_mul_le + all_goals simp + +end IsCancelMul + +section InvolutiveInv +variable [InvolutiveInv α] {s t : Set α} + +@[to_additive (attr := simp)] +lemma card_inv (s : Set α) : Nat.card ↥(s⁻¹) = Nat.card s := by + rw [← image_inv, Nat.card_image_of_injective inv_injective] + +end InvolutiveInv + +section Group +variable [Group α] {s t : Set α} + +@[to_additive] +lemma card_div_le : Nat.card (s / t) ≤ Nat.card s * Nat.card t := by + rw [div_eq_mul_inv, ← card_inv t]; exact card_mul_le + +end Group end Set diff --git a/Mathlib/Data/Finset/Powerset.lean b/Mathlib/Data/Finset/Powerset.lean index 37696b7f88740..f7bf0c5bd810b 100644 --- a/Mathlib/Data/Finset/Powerset.lean +++ b/Mathlib/Data/Finset/Powerset.lean @@ -218,7 +218,7 @@ theorem card_powersetCard (n : ℕ) (s : Finset α) : #align finset.card_powerset_len Finset.card_powersetCard @[simp] -theorem powersetCard_zero (s : Finset α) : Finset.powersetCard 0 s = {∅} := by +theorem powersetCard_zero (s : Finset α) : s.powersetCard 0 = {∅} := by ext; rw [mem_powersetCard, mem_singleton, card_eq_zero] refine' ⟨fun h => h.2, fun h => by @@ -226,6 +226,16 @@ theorem powersetCard_zero (s : Finset α) : Finset.powersetCard 0 s = {∅} := b exact ⟨empty_subset s, rfl⟩⟩ #align finset.powerset_len_zero Finset.powersetCard_zero +@[simp] +theorem map_val_val_powersetCard (s : Finset α) (i : ℕ) : + (s.powersetCard i).val.map Finset.val = s.1.powersetCard i := by + simp [Finset.powersetCard, map_pmap, pmap_eq_map, map_id'] +#align finset.map_val_val_powerset_len Finset.map_val_val_powersetCard + +theorem powersetCard_one (s : Finset α) : + s.powersetCard 1 = s.map ⟨_, Finset.singleton_injective⟩ := + eq_of_veq <| Multiset.map_injective val_injective <| by simp [Multiset.powersetCard_one] + @[simp] theorem powersetCard_empty (n : ℕ) {s : Finset α} (h : s.card < n) : powersetCard n s = ∅ := Finset.card_eq_zero.mp (by rw [card_powersetCard, Nat.choose_eq_zero_of_lt h]) @@ -324,12 +334,6 @@ theorem powersetCard_card_add (s : Finset α) {i : ℕ} (hi : 0 < i) : Finset.powersetCard_empty _ (lt_add_of_pos_right (Finset.card s) hi) #align finset.powerset_len_card_add Finset.powersetCard_card_add -@[simp] -theorem map_val_val_powersetCard (s : Finset α) (i : ℕ) : - (s.powersetCard i).val.map Finset.val = s.1.powersetCard i := by - simp [Finset.powersetCard, map_pmap, pmap_eq_map, map_id'] -#align finset.map_val_val_powerset_len Finset.map_val_val_powersetCard - theorem powersetCard_map {β : Type*} (f : α ↪ β) (n : ℕ) (s : Finset α) : powersetCard n (s.map f) = (powersetCard n s).map (mapEmbedding f).toEmbedding := ext <| fun t => by diff --git a/Mathlib/Data/Finsupp/Order.lean b/Mathlib/Data/Finsupp/Order.lean index 9bed6cd3e4157..baa09ec25ec16 100644 --- a/Mathlib/Data/Finsupp/Order.lean +++ b/Mathlib/Data/Finsupp/Order.lean @@ -3,7 +3,8 @@ Copyright (c) 2021 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Aaron Anderson -/ -import Mathlib.Data.Finsupp.Defs +import Mathlib.Algebra.Order.Module.Defs +import Mathlib.Data.Finsupp.Basic #align_import data.finsupp.order from "leanprover-community/mathlib"@"1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29" @@ -30,7 +31,7 @@ open BigOperators open Finset -variable {ι α : Type*} +variable {ι α β : Type*} namespace Finsupp @@ -149,6 +150,39 @@ instance contravariantClass [OrderedAddCommMonoid α] [ContravariantClass α α ContravariantClass (ι →₀ α) (ι →₀ α) (· + ·) (· ≤ ·) := ⟨fun _f _g _h H x => le_of_add_le_add_left <| H x⟩ +section SMulZeroClass +variable [Zero α] [Preorder α] [Zero β] [Preorder β] [SMulZeroClass α β] + +instance instPosSMulMono [PosSMulMono α β] : PosSMulMono α (ι →₀ β) := + PosSMulMono.lift _ coe_le_coe coe_smul + +instance instSMulPosMono [SMulPosMono α β] : SMulPosMono α (ι →₀ β) := + SMulPosMono.lift _ coe_le_coe coe_smul coe_zero + +instance instPosSMulReflectLE [PosSMulReflectLE α β] : PosSMulReflectLE α (ι →₀ β) := + PosSMulReflectLE.lift _ coe_le_coe coe_smul + +instance instSMulPosReflectLE [SMulPosReflectLE α β] : SMulPosReflectLE α (ι →₀ β) := + SMulPosReflectLE.lift _ coe_le_coe coe_smul coe_zero + +end SMulZeroClass + +section SMulWithZero +variable [Zero α] [PartialOrder α] [Zero β] [PartialOrder β] [SMulWithZero α β] + +instance instPosSMulStrictMono [PosSMulStrictMono α β] : PosSMulStrictMono α (ι →₀ β) := + PosSMulStrictMono.lift _ coe_le_coe coe_smul + +instance instSMulPosStrictMono [SMulPosStrictMono α β] : SMulPosStrictMono α (ι →₀ β) := + SMulPosStrictMono.lift _ coe_le_coe coe_smul coe_zero + +-- `PosSMulReflectLT α (ι →₀ β)` already follows from the other instances + +instance instSMulPosReflectLT [SMulPosReflectLT α β] : SMulPosReflectLT α (ι →₀ β) := + SMulPosReflectLT.lift _ coe_le_coe coe_smul coe_zero + +end SMulWithZero + section CanonicallyOrderedAddCommMonoid variable [CanonicallyOrderedAddCommMonoid α] {f g : ι →₀ α} diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index 9c075e7a1c66c..e2b0a9d38c4d7 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -596,19 +596,22 @@ theorem univ_ofSubsingleton (a : α) [Subsingleton α] : @univ _ (ofSubsingleton rfl #align fintype.univ_of_subsingleton Fintype.univ_ofSubsingleton --- see Note [lower instance priority] -instance (priority := 100) ofIsEmpty [IsEmpty α] : Fintype α := +/-- An empty type is a fintype. Not registered as an instance, to make sure that there aren't two +conflicting `Fintype ι` instances around when casing over whether a fintype `ι` is empty or not. -/ +def ofIsEmpty [IsEmpty α] : Fintype α := ⟨∅, isEmptyElim⟩ #align fintype.of_is_empty Fintype.ofIsEmpty --- no-lint since while `Finset.univ_eq_empty` can prove this, it isn't applicable for `dsimp`. /-- Note: this lemma is specifically about `Fintype.of_isEmpty`. For a statement about arbitrary `Fintype` instances, use `Finset.univ_eq_empty`. -/ -@[simp, nolint simpNF] -theorem univ_of_isEmpty [IsEmpty α] : @univ α _ = ∅ := +@[simp] +theorem univ_of_isEmpty [IsEmpty α] : @univ α Fintype.ofIsEmpty = ∅ := rfl #align fintype.univ_of_is_empty Fintype.univ_of_isEmpty +instance : Fintype Empty := Fintype.ofIsEmpty +instance : Fintype PEmpty := Fintype.ofIsEmpty + end Fintype namespace Set @@ -757,6 +760,7 @@ theorem toFinset_univ [Fintype α] [Fintype (Set.univ : Set α)] : @[simp] theorem toFinset_eq_empty [Fintype s] : s.toFinset = ∅ ↔ s = ∅ := by + let A : Fintype (∅ : Set α) := Fintype.ofIsEmpty rw [← toFinset_empty, toFinset_inj] #align set.to_finset_eq_empty Set.toFinset_eq_empty diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index 312325eea6653..69e9fecfd4d40 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -222,10 +222,10 @@ theorem card_unique [Unique α] [h : Fintype α] : Fintype.card α = 1 := Subsingleton.elim (ofSubsingleton default) h ▸ card_ofSubsingleton _ #align fintype.card_unique Fintype.card_unique -/-- Note: this lemma is specifically about `Fintype.of_is_empty`. For a statement about -arbitrary `Fintype` instances, use `Fintype.card_eq_zero_iff`. -/ +/-- Note: this lemma is specifically about `Fintype.ofIsEmpty`. For a statement about +arbitrary `Fintype` instances, use `Fintype.card_eq_zero`. -/ @[simp] -theorem card_of_isEmpty [IsEmpty α] : Fintype.card α = 0 := +theorem card_of_isEmpty [IsEmpty α] : @Fintype.card α Fintype.ofIsEmpty = 0 := rfl #align fintype.card_of_is_empty Fintype.card_of_isEmpty @@ -356,12 +356,10 @@ theorem Fintype.card_subtype_eq' (y : α) [Fintype { x // y = x }] : Fintype.card_unique #align fintype.card_subtype_eq' Fintype.card_subtype_eq' -@[simp] theorem Fintype.card_empty : Fintype.card Empty = 0 := rfl #align fintype.card_empty Fintype.card_empty -@[simp] theorem Fintype.card_pempty : Fintype.card PEmpty = 0 := rfl #align fintype.card_pempty Fintype.card_pempty @@ -534,7 +532,7 @@ theorem card_eq_zero_iff : card α = 0 ↔ IsEmpty α := by rw [card, Finset.card_eq_zero, univ_eq_empty_iff] #align fintype.card_eq_zero_iff Fintype.card_eq_zero_iff -theorem card_eq_zero [IsEmpty α] : card α = 0 := +@[simp] theorem card_eq_zero [IsEmpty α] : card α = 0 := card_eq_zero_iff.2 ‹_› #align fintype.card_eq_zero Fintype.card_eq_zero @@ -1227,7 +1225,7 @@ theorem List.exists_pw_disjoint_with_card {α : Type*} [Fintype α] intro a ha conv_rhs => rw [← List.map_id a] rw [List.map_pmap] - simp only [Fin.valEmbedding_apply, Fin.val_mk, List.pmap_eq_map, List.map_id'', List.map_id] + simp only [Fin.valEmbedding_apply, Fin.val_mk, List.pmap_eq_map, List.map_id', List.map_id] use l.map (List.map (Fintype.equivFin α).symm) constructor · -- length diff --git a/Mathlib/Data/Fintype/CardEmbedding.lean b/Mathlib/Data/Fintype/CardEmbedding.lean index a2fb358165c70..5ddebd9f5afae 100644 --- a/Mathlib/Data/Fintype/CardEmbedding.lean +++ b/Mathlib/Data/Fintype/CardEmbedding.lean @@ -53,7 +53,6 @@ theorem card_embedding_eq {α β : Type*} [Fintype α] [Fintype β] [emb : Finty /- The cardinality of embeddings from an infinite type to a finite type is zero. This is a re-statement of the pigeonhole principle. -/ -@[simp] theorem card_embedding_eq_of_infinite {α β : Type*} [Infinite α] [Fintype β] [Fintype (α ↪ β)] : ‖α ↪ β‖ = 0 := card_eq_zero diff --git a/Mathlib/Data/Int/ModEq.lean b/Mathlib/Data/Int/ModEq.lean index 870b4111ef194..fcdf45a50b175 100644 --- a/Mathlib/Data/Int/ModEq.lean +++ b/Mathlib/Data/Int/ModEq.lean @@ -315,7 +315,7 @@ theorem exists_unique_equiv (a : ℤ) {b : ℤ} (hb : 0 < b) : theorem exists_unique_equiv_nat (a : ℤ) {b : ℤ} (hb : 0 < b) : ∃ z : ℕ, ↑z < b ∧ ↑z ≡ a [ZMOD b] := let ⟨z, hz1, hz2, hz3⟩ := exists_unique_equiv a hb ⟨z.natAbs, by - constructor <;> rw [ofNat_natAbs_eq_of_nonneg z hz1] <;> assumption⟩ + constructor <;> rw [natAbs_of_nonneg hz1] <;> assumption⟩ #align int.exists_unique_equiv_nat Int.exists_unique_equiv_nat theorem mod_mul_right_mod (a b c : ℤ) : a % (b * c) % b = a % b := diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index c2161c8f7cbf8..df16f00de16cd 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -1726,14 +1726,11 @@ theorem map_concat (f : α → β) (a : α) (l : List α) : induction l <;> [rfl; simp only [*, concat_eq_append, cons_append, map, map_append]] #align list.map_concat List.map_concat -@[simp] -theorem map_id'' (l : List α) : map (fun x => x) l = l := - map_id _ -#align list.map_id'' List.map_id'' +#align list.map_id'' List.map_id' -theorem map_id' {f : α → α} (h : ∀ x, f x = x) (l : List α) : map f l = l := by +theorem map_id'' {f : α → α} (h : ∀ x, f x = x) (l : List α) : map f l = l := by simp [show f = id from funext h] -#align list.map_id' List.map_id' +#align list.map_id' List.map_id'' theorem eq_nil_of_map_eq_nil {f : α → β} {l : List α} (h : map f l = nil) : l = nil := eq_nil_of_length_eq_zero <| by rw [← length_map l f, h]; rfl diff --git a/Mathlib/Data/List/BigOperators/Basic.lean b/Mathlib/Data/List/BigOperators/Basic.lean index 3f9dd02a8f11d..b9e9540de0234 100644 --- a/Mathlib/Data/List/BigOperators/Basic.lean +++ b/Mathlib/Data/List/BigOperators/Basic.lean @@ -348,7 +348,7 @@ theorem prod_lt_prod_of_ne_nil [Preorder M] [CovariantClass M M (· * ·) (· < theorem prod_le_pow_card [Preorder M] [CovariantClass M M (Function.swap (· * ·)) (· ≤ ·)] [CovariantClass M M (· * ·) (· ≤ ·)] (l : List M) (n : M) (h : ∀ x ∈ l, x ≤ n) : l.prod ≤ n ^ l.length := by - simpa only [map_id'', map_const', prod_replicate] using prod_le_prod' h + simpa only [map_id', map_const', prod_replicate] using prod_le_prod' h #align list.prod_le_pow_card List.prod_le_pow_card #align list.sum_le_card_nsmul List.sum_le_card_nsmul diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index e54d2e02fbe24..4d9ab07fff437 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -175,7 +175,7 @@ theorem sublists_cons (a : α) (l : List α) : theorem sublists_concat (l : List α) (a : α) : sublists (l ++ [a]) = sublists l ++ map (fun x => x ++ [a]) (sublists l) := by rw [sublists_append, sublists_singleton, bind_eq_bind, cons_bind, cons_bind, nil_bind, - map_id' append_nil, append_nil] + map_id'' append_nil, append_nil] #align list.sublists_concat List.sublists_concat theorem sublists_reverse (l : List α) : sublists (reverse l) = map reverse (sublists' l) := by @@ -189,7 +189,7 @@ theorem sublists_eq_sublists' (l : List α) : sublists l = map reverse (sublists #align list.sublists_eq_sublists' List.sublists_eq_sublists' theorem sublists'_reverse (l : List α) : sublists' (reverse l) = map reverse (sublists l) := by - simp only [sublists_eq_sublists', map_map, map_id' reverse_reverse, Function.comp] + simp only [sublists_eq_sublists', map_map, map_id'' reverse_reverse, Function.comp] #align list.sublists'_reverse List.sublists'_reverse theorem sublists'_eq_sublists (l : List α) : sublists' l = map reverse (sublists (reverse l)) := by @@ -274,6 +274,10 @@ theorem sublistsLen_succ_cons {α : Type*} (n) (a : α) (l) : append_nil]; rfl #align list.sublists_len_succ_cons List.sublistsLen_succ_cons +theorem sublistsLen_one {α : Type*} (l : List α) : sublistsLen 1 l = l.reverse.map ([·]) := + l.rec (by rw [sublistsLen_succ_nil, reverse_nil, map_nil]) fun a s ih ↦ by + rw [sublistsLen_succ_cons, ih, reverse_cons, map_append, sublistsLen_zero]; rfl + @[simp] theorem length_sublistsLen {α : Type*} : ∀ (n) (l : List α), length (sublistsLen n l) = Nat.choose (length l) n diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index f06e54b530b97..c94f2c45bc695 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -2154,6 +2154,16 @@ lemma map_filter' {f : α → β} (hf : Injective f) (s : Multiset α) simp [(· ∘ ·), map_filter, hf.eq_iff] #align multiset.map_filter' Multiset.map_filter' +lemma card_filter_le_iff (s : Multiset α) (P : α → Prop) [DecidablePred P] (n : ℕ) : + card (s.filter P) ≤ n ↔ ∀ s' ≤ s, n < card s' → ∃ a ∈ s', ¬ P a := by + fconstructor + · intro H s' hs' s'_card + by_contra! rid + have card := card_le_of_le (monotone_filter_left P hs') |>.trans H + exact s'_card.not_le (filter_eq_self.mpr rid ▸ card) + · contrapose! + exact fun H ↦ ⟨s.filter P, filter_le _ _, H, fun a ha ↦ (mem_filter.mp ha).2⟩ + /-! ### Simultaneously filter and map elements of a multiset -/ diff --git a/Mathlib/Data/Multiset/Powerset.lean b/Mathlib/Data/Multiset/Powerset.lean index 5e3d047fc5f9b..821f85d657157 100644 --- a/Mathlib/Data/Multiset/Powerset.lean +++ b/Mathlib/Data/Multiset/Powerset.lean @@ -258,6 +258,10 @@ theorem powersetCard_cons (n : ℕ) (a : α) (s) : Quotient.inductionOn s fun l => by simp [powersetCard_coe'] #align multiset.powerset_len_cons Multiset.powersetCard_cons +theorem powersetCard_one (s : Multiset α) : powersetCard 1 s = s.map singleton := + Quotient.inductionOn s fun l ↦ by + simp [powersetCard_coe, sublistsLen_one, map_reverse, Function.comp] + @[simp] theorem mem_powersetCard {n : ℕ} {s t : Multiset α} : s ∈ powersetCard n t ↔ s ≤ t ∧ card s = n := Quotient.inductionOn t fun l => by simp [powersetCard_coe'] diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index 49d6f62404066..2223d72162f89 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -281,7 +281,7 @@ theorem testBit_two_pow_of_ne {n m : ℕ} (hm : n ≠ m) : testBit (2 ^ n) m = f cases' hm.lt_or_lt with hm hm · rw [Nat.div_eq_of_lt] · simp - · exact Nat.pow_lt_pow_of_lt_right one_lt_two hm + · exact pow_lt_pow_right one_lt_two hm · rw [pow_div hm.le zero_lt_two, ← tsub_add_cancel_of_le (succ_le_of_lt <| tsub_pos_of_lt hm)] -- Porting note: XXX why does this make it work? rw [(rfl : succ 0 = 1)] @@ -490,6 +490,6 @@ lemma shiftLeft_lt {x n m : ℕ} (h : x < 2 ^ n) : x <<< m < 2 ^ (n + m) := by lemma append_lt {x y n m} (hx : x < 2 ^ n) (hy : y < 2 ^ m) : y <<< n ||| x < 2 ^ (n + m) := by apply bitwise_lt · rw [add_comm]; apply shiftLeft_lt hy - · apply lt_of_lt_of_le hx <| pow_le_pow (le_succ _) (le_add_right _ _) + · apply lt_of_lt_of_le hx <| pow_le_pow_right (le_succ _) (le_add_right _ _) end Nat diff --git a/Mathlib/Data/Nat/Choose/Factorization.lean b/Mathlib/Data/Nat/Choose/Factorization.lean index 7f0dcf923109f..c81ab23f61e04 100644 --- a/Mathlib/Data/Nat/Choose/Factorization.lean +++ b/Mathlib/Data/Nat/Choose/Factorization.lean @@ -85,7 +85,7 @@ theorem factorization_choose_of_lt_three_mul (hp' : p ≠ 2) (hk : p ≤ k) (hk' n < 3 * p := hn _ ≤ p * p := mul_le_mul_right' this p _ = p ^ 2 := (sq p).symm - _ ≤ p ^ i := pow_le_pow hp.one_lt.le hi + _ ≤ p ^ i := pow_le_pow_right hp.one_lt.le hi rwa [mod_eq_of_lt (lt_of_le_of_lt hkn hn), mod_eq_of_lt (lt_of_le_of_lt tsub_le_self hn), add_tsub_cancel_of_le hkn] #align nat.factorization_choose_of_lt_three_mul Nat.factorization_choose_of_lt_three_mul diff --git a/Mathlib/Data/Nat/Factorial/Basic.lean b/Mathlib/Data/Nat/Factorial/Basic.lean index 0d7f963f8dda6..7a2b7b4d7626e 100644 --- a/Mathlib/Data/Nat/Factorial/Basic.lean +++ b/Mathlib/Data/Nat/Factorial/Basic.lean @@ -202,13 +202,9 @@ theorem add_factorial_le_factorial_add (i : ℕ) {n : ℕ} (n1 : 1 ≤ n) : i + #align nat.add_factorial_le_factorial_add Nat.add_factorial_le_factorial_add theorem factorial_mul_pow_sub_le_factorial {n m : ℕ} (hnm : n ≤ m) : n ! * n ^ (m - n) ≤ m ! := by - suffices n ! * (n + 1) ^ (m - n) ≤ m ! from by - apply LE.le.trans _ this - apply mul_le_mul_left - apply pow_le_pow_of_le_left (le_succ n) - have := @Nat.factorial_mul_pow_le_factorial n (m - n) - simp? [hnm] at this says simp only [ge_iff_le, hnm, add_tsub_cancel_of_le] at this - exact this + calc + _ ≤ n ! * (n + 1) ^ (m - n) := mul_le_mul_left' (Nat.pow_le_pow_left n.le_succ _) _ + _ ≤ _ := by simpa [hnm] using @Nat.factorial_mul_pow_le_factorial n (m - n) #align nat.factorial_mul_pow_sub_le_factorial Nat.factorial_mul_pow_sub_le_factorial end Factorial @@ -297,7 +293,7 @@ theorem ascFactorial_le_pow_add (n : ℕ) : ∀ k : ℕ, n.ascFactorial k ≤ (n rw [ascFactorial_succ, pow_succ, ← add_assoc, mul_comm _ (succ (n + k))] exact Nat.mul_le_mul_of_nonneg_left - ((ascFactorial_le_pow_add _ k).trans (Nat.pow_le_pow_of_le_left (le_succ _) _)) + ((ascFactorial_le_pow_add _ k).trans (Nat.pow_le_pow_left (le_succ _) _)) #align nat.asc_factorial_le_pow_add Nat.ascFactorial_le_pow_add theorem ascFactorial_lt_pow_add (n : ℕ) : ∀ {k : ℕ}, 2 ≤ k → n.ascFactorial k < (n + k) ^ k @@ -306,11 +302,8 @@ theorem ascFactorial_lt_pow_add (n : ℕ) : ∀ {k : ℕ}, 2 ≤ k → n.ascFact | k + 2 => fun _ => by rw [ascFactorial_succ, pow_succ] rw [add_assoc n (k + 1) 1, mul_comm <| (n + (k + 2)) ^ (k + 1)] - refine' - Nat.mul_lt_mul' le_rfl - ((ascFactorial_le_pow_add n _).trans_lt - (pow_lt_pow_of_lt_left (lt_add_one _) (succ_pos _))) - (succ_pos _) + exact mul_lt_mul_of_pos_left ((ascFactorial_le_pow_add n _).trans_lt $ + Nat.pow_lt_pow_left (lt_add_one _) k.succ_ne_zero) (succ_pos _) #align nat.asc_factorial_lt_pow_add Nat.ascFactorial_lt_pow_add theorem ascFactorial_pos (n k : ℕ) : 0 < n.ascFactorial k := @@ -409,7 +402,7 @@ theorem pow_sub_le_descFactorial (n : ℕ) : ∀ k : ℕ, (n + 1 - k) ^ k ≤ n. | k + 1 => by rw [descFactorial_succ, pow_succ, succ_sub_succ, mul_comm] apply Nat.mul_le_mul_of_nonneg_left - exact (le_trans (Nat.pow_le_pow_of_le_left (tsub_le_tsub_right (le_succ _) _) k) + exact (le_trans (Nat.pow_le_pow_left (tsub_le_tsub_right (le_succ _) _) k) (pow_sub_le_descFactorial n k)) #align nat.pow_sub_le_desc_factorial Nat.pow_sub_le_descFactorial @@ -423,7 +416,7 @@ theorem pow_sub_lt_descFactorial' {n : ℕ} : | k + 1 => fun h => by rw [descFactorial_succ, pow_succ, mul_comm] apply Nat.mul_lt_mul_of_pos_left - · refine' ((Nat.pow_le_pow_of_le_left (tsub_le_tsub_right (le_succ n) _) _).trans_lt _) + · refine' ((Nat.pow_le_pow_left (tsub_le_tsub_right (le_succ n) _) _).trans_lt _) rw [succ_sub_succ] exact pow_sub_lt_descFactorial' ((le_succ _).trans h) · apply tsub_pos_of_lt; apply h diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index eb01479ad42c0..9ffc3a06ff529 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -395,12 +395,8 @@ theorem ord_compl_mul (a b p : ℕ) : ord_compl[p] (a * b) = ord_compl[p] a * or /-- A crude upper bound on `n.factorization p` -/ theorem factorization_lt {n : ℕ} (p : ℕ) (hn : n ≠ 0) : n.factorization p < n := by by_cases pp : p.Prime - case neg => - simp only [factorization_eq_zero_of_non_prime n pp] - exact hn.bot_lt - rw [← pow_lt_iff_lt_right pp.two_le] - apply lt_of_le_of_lt (ord_proj_le p hn) - exact lt_of_lt_of_le (lt_two_pow n) (pow_le_pow_of_le_left pp.two_le n) + · exact (pow_lt_pow_iff_right pp.one_lt).1 $ (ord_proj_le p hn).trans_lt $ lt_pow_self pp.one_lt _ + · simpa only [factorization_eq_zero_of_non_prime n pp] using hn.bot_lt #align nat.factorization_lt Nat.factorization_lt /-- An upper bound on `n.factorization p` -/ @@ -408,7 +404,7 @@ theorem factorization_le_of_le_pow {n p b : ℕ} (hb : n ≤ p ^ b) : n.factoriz rcases eq_or_ne n 0 with (rfl | hn) · simp by_cases pp : p.Prime - · exact (pow_le_iff_le_right pp.two_le).1 (le_trans (ord_proj_le p hn) hb) + · exact (pow_le_pow_iff_right pp.one_lt).1 ((ord_proj_le p hn).trans hb) · simp [factorization_eq_zero_of_non_prime n pp] #align nat.factorization_le_of_le_pow Nat.factorization_le_of_le_pow @@ -702,8 +698,8 @@ theorem Ico_filter_pow_dvd_eq {n p b : ℕ} (pp : p.Prime) (hn : n ≠ 0) (hb : ext x simp only [Finset.mem_filter, mem_Ico, mem_Icc, and_congr_left_iff, and_congr_right_iff] rintro h1 - - simp [lt_of_pow_dvd_right hn pp.two_le h1, - (pow_le_iff_le_right pp.two_le).1 ((le_of_dvd hn.bot_lt h1).trans hb)] + exact iff_of_true (lt_of_pow_dvd_right hn pp.two_le h1) $ + (pow_le_pow_iff_right pp.one_lt).1 $ (le_of_dvd hn.bot_lt h1).trans hb #align nat.Ico_filter_pow_dvd_eq Nat.Ico_filter_pow_dvd_eq /-! ### Factorization and coprimes -/ @@ -763,10 +759,7 @@ def recOnPrimePow {P : ℕ → Sort*} (h0 : P 0) (h1 : P 1) · rw [Nat.mul_div_cancel' hpt] · rw [Nat.dvd_div_iff hpt, ← pow_succ] exact pow_succ_factorization_not_dvd (k + 1).succ_ne_zero hp - · rw [lt_mul_iff_one_lt_left Nat.succ_pos', one_lt_pow_iff htp.ne] - exact hp.one_lt - -- Porting note: was - -- simp [lt_mul_iff_one_lt_left Nat.succ_pos', one_lt_pow_iff htp.ne, hp.one_lt] + · simp [lt_mul_iff_one_lt_left Nat.succ_pos', one_lt_pow_iff htp.ne', hp.one_lt] #align nat.rec_on_prime_pow Nat.recOnPrimePow /-- Given `P 0`, `P 1`, and `P (p ^ n)` for positive prime powers, and a way to extend `P a` and diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index 003252fd98157..b247618d89f9c 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -140,7 +140,7 @@ theorem log_eq_iff {b m n : ℕ} (h : m ≠ 0 ∨ 1 < b ∧ n ≠ 0) : rw [not_and_or, not_lt, Ne.def, not_not] at hbn rcases hbn with (hb | rfl) · simpa only [log_of_left_le_one hb, hm.symm, false_iff_iff, not_and, not_lt] using - le_trans (pow_le_pow_of_le_one' hb m.le_succ) + le_trans (pow_le_pow_right_of_le_one' hb m.le_succ) · simpa only [log_zero_right, hm.symm, nonpos_iff_eq_zero, false_iff, not_and, not_lt, add_pos_iff, zero_lt_one, or_true, pow_eq_zero_iff] using pow_eq_zero #align nat.log_eq_iff Nat.log_eq_iff @@ -154,7 +154,7 @@ theorem log_eq_of_pow_le_of_lt_pow {b m n : ℕ} (h₁ : b ^ m ≤ n) (h₂ : n #align nat.log_eq_of_pow_le_of_lt_pow Nat.log_eq_of_pow_le_of_lt_pow theorem log_pow {b : ℕ} (hb : 1 < b) (x : ℕ) : log b (b ^ x) = x := - log_eq_of_pow_le_of_lt_pow le_rfl (pow_lt_pow hb x.lt_succ_self) + log_eq_of_pow_le_of_lt_pow le_rfl (pow_lt_pow_right hb x.lt_succ_self) #align nat.log_pow Nat.log_pow theorem log_eq_one_iff' {b n : ℕ} : log b n = 1 ↔ b ≤ n ∧ n < b * b := by @@ -195,7 +195,7 @@ theorem log_anti_left {b c n : ℕ} (hc : 1 < c) (hb : c ≤ b) : log b n ≤ lo rcases eq_or_ne n 0 with (rfl | hn); · rw [log_zero_right, log_zero_right] apply le_log_of_pow_le hc calc - c ^ log b n ≤ b ^ log b n := pow_le_pow_of_le_left' hb _ + c ^ log b n ≤ b ^ log b n := pow_le_pow_left' hb _ _ ≤ n := pow_log_le_self _ hn #align nat.log_anti_left Nat.log_anti_left @@ -340,7 +340,7 @@ theorem clog_anti_left {b c n : ℕ} (hc : 1 < c) (hb : c ≤ b) : clog b n ≤ rw [← le_pow_iff_clog_le (lt_of_lt_of_le hc hb)] calc n ≤ c ^ clog c n := le_pow_clog hc _ - _ ≤ b ^ clog c n := pow_le_pow_of_le_left hb _ + _ ≤ b ^ clog c n := Nat.pow_le_pow_left hb _ #align nat.clog_anti_left Nat.clog_anti_left theorem clog_monotone (b : ℕ) : Monotone (clog b) := fun _ _ => clog_mono_right _ diff --git a/Mathlib/Data/Nat/Pow.lean b/Mathlib/Data/Nat/Pow.lean index ba5c2d1bbc43b..8b23728b822ad 100644 --- a/Mathlib/Data/Nat/Pow.lean +++ b/Mathlib/Data/Nat/Pow.lean @@ -14,26 +14,29 @@ Results on the power operation on natural numbers. namespace Nat +variable {m n x y : ℕ} /-! ### `pow` -/ -- Porting note: the next two lemmas have moved into `Std`. +-- TODO: Rename `Nat.pow_le_pow_of_le_left` to `Nat.pow_le_pow_left`, protect it, remove the alias +-- TODO: Rename `Nat.pow_le_pow_of_le_right` to `Nat.pow_le_pow_right`, protect it, remove the alias --- The global `pow_le_pow_of_le_left` needs an extra hypothesis `0 ≤ x`. -#align nat.pow_le_pow_of_le_left Nat.pow_le_pow_of_le_left -#align nat.pow_le_pow_of_le_right Nat.pow_le_pow_of_le_right +-- The global `pow_le_pow_left` needs an extra hypothesis `0 ≤ x`. +protected alias pow_le_pow_left := pow_le_pow_of_le_left +protected alias pow_le_pow_right := pow_le_pow_of_le_right +#align nat.pow_le_pow_of_le_left Nat.pow_le_pow_left +#align nat.pow_le_pow_of_le_right Nat.pow_le_pow_right -theorem pow_lt_pow_of_lt_left {x y : ℕ} (H : x < y) {i} (h : 0 < i) : x ^ i < y ^ i := - _root_.pow_lt_pow_of_lt_left H (zero_le _) h -#align nat.pow_lt_pow_of_lt_left Nat.pow_lt_pow_of_lt_left +protected theorem pow_lt_pow_left (h : x < y) (hn : n ≠ 0) : x ^ n < y ^ n := + pow_lt_pow_left h (zero_le _) hn +#align nat.pow_lt_pow_of_lt_left Nat.pow_lt_pow_left -theorem pow_lt_pow_of_lt_right {x : ℕ} (H : 1 < x) {i j : ℕ} (h : i < j) : x ^ i < x ^ j := - pow_lt_pow H h -#align nat.pow_lt_pow_of_lt_right Nat.pow_lt_pow_of_lt_right +#align nat.pow_lt_pow_of_lt_right pow_lt_pow_right theorem pow_lt_pow_succ {p : ℕ} (h : 1 < p) (n : ℕ) : p ^ n < p ^ (n + 1) := - pow_lt_pow_of_lt_right h n.lt_succ_self + pow_lt_pow_right h n.lt_succ_self #align nat.pow_lt_pow_succ Nat.pow_lt_pow_succ theorem le_self_pow {n : ℕ} (hn : n ≠ 0) : ∀ m : ℕ, m ≤ m ^ n @@ -66,9 +69,9 @@ theorem one_le_two_pow (n : ℕ) : 1 ≤ 2 ^ n := one_le_pow n 2 (by decide) #align nat.one_le_two_pow Nat.one_le_two_pow -theorem one_lt_pow (n m : ℕ) (h₀ : 0 < n) (h₁ : 1 < m) : 1 < m ^ n := by +theorem one_lt_pow (n m : ℕ) (h₀ : n ≠ 0) (h₁ : 1 < m) : 1 < m ^ n := by rw [← one_pow n] - exact pow_lt_pow_of_lt_left h₁ h₀ + exact Nat.pow_lt_pow_left h₁ h₀ #align nat.one_lt_pow Nat.one_lt_pow theorem two_pow_pos (n : ℕ) : 0 < 2^n := Nat.pos_pow_of_pos _ (by decide) @@ -76,50 +79,32 @@ theorem two_pow_pos (n : ℕ) : 0 < 2^n := Nat.pos_pow_of_pos _ (by decide) theorem two_pow_succ (n : ℕ) : 2^(n + 1) = 2^n + 2^n := by simp [pow_succ, mul_two] theorem one_lt_pow' (n m : ℕ) : 1 < (m + 2) ^ (n + 1) := - one_lt_pow (n + 1) (m + 2) (succ_pos n) (Nat.lt_of_sub_eq_succ rfl) + one_lt_pow (n + 1) (m + 2) n.succ_ne_zero (Nat.lt_of_sub_eq_succ rfl) #align nat.one_lt_pow' Nat.one_lt_pow' @[simp] -theorem one_lt_pow_iff {k n : ℕ} (h : 0 ≠ k) : 1 < n ^ k ↔ 1 < n := by - rcases n with (rfl | n) - · cases k <;> simp [zero_pow_eq] - rcases n with (rfl | n) - · rw [← Nat.one_eq_succ_zero, one_pow] - refine' ⟨fun _ => one_lt_succ_succ n, fun _ => _⟩ - induction' k with k hk - · exact absurd rfl h - rcases k with (rfl | k) - · simp [← Nat.one_eq_succ_zero] - rw [pow_succ'] - exact one_lt_mul (one_lt_succ_succ _).le (hk (succ_ne_zero k).symm) +theorem one_lt_pow_iff {k n : ℕ} (h : k ≠ 0) : 1 < n ^ k ↔ 1 < n := + one_lt_pow_iff_of_nonneg (zero_le _) h #align nat.one_lt_pow_iff Nat.one_lt_pow_iff -theorem one_lt_two_pow (n : ℕ) (h₀ : 0 < n) : 1 < 2 ^ n := - one_lt_pow n 2 h₀ (by decide) +theorem one_lt_two_pow (n : ℕ) (h₀ : n ≠ 0) : 1 < 2 ^ n := one_lt_pow n 2 h₀ (by decide) #align nat.one_lt_two_pow Nat.one_lt_two_pow theorem one_lt_two_pow' (n : ℕ) : 1 < 2 ^ (n + 1) := - one_lt_pow (n + 1) 2 (succ_pos n) (by decide) + one_lt_pow (n + 1) 2 n.succ_ne_zero (by decide) #align nat.one_lt_two_pow' Nat.one_lt_two_pow' -theorem pow_right_strictMono {x : ℕ} (k : 2 ≤ x) : StrictMono fun n : ℕ => x ^ n := fun _ _ => - pow_lt_pow_of_lt_right k -#align nat.pow_right_strict_mono Nat.pow_right_strictMono +#align nat.pow_right_strict_mono pow_right_strictMono +#align nat.pow_le_iff_lt_right pow_le_pow_iff_right +#align nat.pow_lt_iff_lt_right pow_lt_pow_iff_right -theorem pow_le_iff_le_right {x m n : ℕ} (k : 2 ≤ x) : x ^ m ≤ x ^ n ↔ m ≤ n := - StrictMono.le_iff_le (pow_right_strictMono k) -#align nat.pow_le_iff_le_right Nat.pow_le_iff_le_right - -theorem pow_lt_iff_lt_right {x m n : ℕ} (k : 2 ≤ x) : x ^ m < x ^ n ↔ m < n := - StrictMono.lt_iff_lt (pow_right_strictMono k) -#align nat.pow_lt_iff_lt_right Nat.pow_lt_iff_lt_right - -theorem pow_right_injective {x : ℕ} (k : 2 ≤ x) : Function.Injective fun n : ℕ => x ^ n := - StrictMono.injective (pow_right_strictMono k) +protected lemma pow_right_injective (hx : 2 ≤ x) : Function.Injective (x ^ ·) := + StrictMono.injective (pow_right_strictMono hx) #align nat.pow_right_injective Nat.pow_right_injective -theorem pow_left_strictMono {m : ℕ} (k : 1 ≤ m) : StrictMono fun x : ℕ => x ^ m := fun _ _ h => - pow_lt_pow_of_lt_left h k +/-- See also `pow_left_strictMonoOn`. -/ +protected theorem pow_left_strictMono (hn : n ≠ 0) : StrictMono (. ^ n : ℕ → ℕ) := + fun _ _ h ↦ Nat.pow_lt_pow_left h hn #align nat.pow_left_strict_mono Nat.pow_left_strictMono theorem mul_lt_mul_pow_succ {n a q : ℕ} (a0 : 0 < a) (q1 : 1 < q) : n * q < a * q ^ (n + 1) := by @@ -127,25 +112,21 @@ theorem mul_lt_mul_pow_succ {n a q : ℕ} (a0 : 0 < a) (q1 : 1 < q) : n * q < a exact lt_mul_of_one_le_of_lt (Nat.succ_le_iff.mpr a0) (Nat.lt_pow_self q1 n) #align nat.mul_lt_mul_pow_succ Nat.mul_lt_mul_pow_succ -end Nat - -theorem StrictMono.nat_pow {n : ℕ} (hn : 1 ≤ n) {f : ℕ → ℕ} (hf : StrictMono f) : +theorem _root_.StrictMono.nat_pow {n : ℕ} (hn : n ≠ 0) {f : ℕ → ℕ} (hf : StrictMono f) : StrictMono fun m => f m ^ n := (Nat.pow_left_strictMono hn).comp hf #align strict_mono.nat_pow StrictMono.nat_pow -namespace Nat - -theorem pow_le_iff_le_left {m x y : ℕ} (k : 1 ≤ m) : x ^ m ≤ y ^ m ↔ x ≤ y := - StrictMono.le_iff_le (pow_left_strictMono k) -#align nat.pow_le_iff_le_left Nat.pow_le_iff_le_left +protected theorem pow_le_pow_iff_left (hm : m ≠ 0) : x ^ m ≤ y ^ m ↔ x ≤ y := + pow_le_pow_iff_left (zero_le _) (zero_le _) hm +#align nat.pow_le_iff_le_left Nat.pow_le_pow_iff_left -theorem pow_lt_iff_lt_left {m x y : ℕ} (k : 1 ≤ m) : x ^ m < y ^ m ↔ x < y := - StrictMono.lt_iff_lt (pow_left_strictMono k) -#align nat.pow_lt_iff_lt_left Nat.pow_lt_iff_lt_left +protected theorem pow_lt_pow_iff_left (hm : m ≠ 0) : x ^ m < y ^ m ↔ x < y := + pow_lt_pow_iff_left (zero_le _) (zero_le _) hm +#align nat.pow_lt_iff_lt_left Nat.pow_lt_pow_iff_left -theorem pow_left_injective {m : ℕ} (k : 1 ≤ m) : Function.Injective fun x : ℕ => x ^ m := - StrictMono.injective (pow_left_strictMono k) +theorem pow_left_injective (hm : m ≠ 0) : Function.Injective fun x : ℕ => x ^ m := + (Nat.pow_left_strictMono hm).injective #align nat.pow_left_injective Nat.pow_left_injective theorem sq_sub_sq (a b : ℕ) : a ^ 2 - b ^ 2 = (a + b) * (a - b) := by @@ -203,14 +184,14 @@ theorem pow_dvd_pow_iff_pow_le_pow {k l : ℕ} : ∀ {x : ℕ} (_ : 0 < x), x ^ · intro a cases' x with x · simp - · have le := (pow_le_iff_le_right (Nat.le_add_left _ _)).mp a + · have le := (pow_le_pow_iff_right $ by simp).mp a use (x + 2) ^ (l - k) rw [← pow_add, add_comm k, tsub_add_cancel_of_le le] #align nat.pow_dvd_pow_iff_pow_le_pow Nat.pow_dvd_pow_iff_pow_le_pow /-- If `1 < x`, then `x^k` divides `x^l` if and only if `k` is at most `l`. -/ theorem pow_dvd_pow_iff_le_right {x k l : ℕ} (w : 1 < x) : x ^ k ∣ x ^ l ↔ k ≤ l := by - rw [pow_dvd_pow_iff_pow_le_pow (lt_of_succ_lt w), pow_le_iff_le_right w] + rw [pow_dvd_pow_iff_pow_le_pow (lt_of_succ_lt w), pow_le_pow_iff_right w] #align nat.pow_dvd_pow_iff_le_right Nat.pow_dvd_pow_iff_le_right theorem pow_dvd_pow_iff_le_right' {b k l : ℕ} : (b + 2) ^ k ∣ (b + 2) ^ l ↔ k ≤ l := @@ -242,7 +223,7 @@ theorem pow_div {x m n : ℕ} (h : n ≤ m) (hx : 0 < x) : x ^ m / x ^ n = x ^ ( #align nat.pow_div Nat.pow_div theorem lt_of_pow_dvd_right {p i n : ℕ} (hn : n ≠ 0) (hp : 2 ≤ p) (h : p ^ i ∣ n) : i < n := by - rw [← pow_lt_iff_lt_right hp] + rw [← pow_lt_pow_iff_right (succ_le_iff.1 hp)] exact lt_of_le_of_lt (le_of_dvd hn.bot_lt h) (lt_pow_self (succ_le_iff.mp hp) n) #align nat.lt_of_pow_dvd_right Nat.lt_of_pow_dvd_right diff --git a/Mathlib/Data/Nat/Size.lean b/Mathlib/Data/Nat/Size.lean index 0feaee57900ea..606a5e1e4a053 100644 --- a/Mathlib/Data/Nat/Size.lean +++ b/Mathlib/Data/Nat/Size.lean @@ -95,7 +95,7 @@ theorem size_shiftLeft' {b m n} (h : shiftLeft' b m n ≠ 0) : simp only [zero_add, one_mul] at this obtain rfl : n = 0 := Nat.eq_zero_of_le_zero - (le_of_not_gt fun hn => ne_of_gt (pow_lt_pow_of_lt_right (by decide) hn) this) + (le_of_not_gt fun hn => ne_of_gt (pow_lt_pow_right (by decide) hn) this) rfl #align nat.size_shiftl' Nat.size_shiftLeft' @@ -148,7 +148,7 @@ theorem size_eq_zero {n : ℕ} : size n = 0 ↔ n = 0 := by #align nat.size_eq_zero Nat.size_eq_zero theorem size_pow {n : ℕ} : size (2 ^ n) = n + 1 := - le_antisymm (size_le.2 <| pow_lt_pow_of_lt_right (by decide) (lt_succ_self _)) + le_antisymm (size_le.2 <| pow_lt_pow_right (by decide) (lt_succ_self _)) (lt_size.2 <| le_rfl) #align nat.size_pow Nat.size_pow diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index 24f6b53a0e8a3..e324c6530e488 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -354,7 +354,7 @@ theorem sq_mul_squarefree_of_pos {n : ℕ} (hn : 0 < n) : refine' Nat.lt_le_asymm _ (Finset.le_max' S ((b * x) ^ 2) _) -- Porting note: these two goals were in the opposite order in Lean 3 · convert lt_mul_of_one_lt_right hlts - (one_lt_pow 2 x zero_lt_two (one_lt_iff_ne_zero_and_ne_one.mpr ⟨fun h => by simp_all, hx⟩)) + (one_lt_pow 2 x two_ne_zero (one_lt_iff_ne_zero_and_ne_one.mpr ⟨fun h => by simp_all, hx⟩)) using 1 rw [mul_pow] · simp_rw [hsa, Finset.mem_filter, Finset.mem_range] @@ -432,7 +432,7 @@ lemma prod_primeFactors_invOn_squarefree : theorem prod_factors_toFinset_of_squarefree {n : ℕ} (hn : Squarefree n) : ∏ p in n.factors.toFinset, p = n := by - rw [List.prod_toFinset _ hn.nodup_factors, List.map_id'', Nat.prod_factors hn.ne_zero] + rw [List.prod_toFinset _ hn.nodup_factors, List.map_id', Nat.prod_factors hn.ne_zero] end Nat diff --git a/Mathlib/Data/Nat/Totient.lean b/Mathlib/Data/Nat/Totient.lean index 76e663eda5af8..3cee8ceb16f2d 100644 --- a/Mathlib/Data/Nat/Totient.lean +++ b/Mathlib/Data/Nat/Totient.lean @@ -213,7 +213,7 @@ theorem totient_prime_pow_succ {p : ℕ} (hp : p.Prime) (n : ℕ) : φ (p ^ (n + rintro b ⟨h, rfl⟩ rw [pow_succ] exact (mul_lt_mul_right hp.pos).2 h - rw [card_sdiff h2, card_image_of_injOn (h1.injOn _), card_range, card_range, ← + rw [card_sdiff h2, Finset.card_image_of_injective _ h1, card_range, card_range, ← one_mul (p ^ n), pow_succ', ← tsub_mul, one_mul, mul_comm] #align nat.totient_prime_pow_succ Nat.totient_prime_pow_succ diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index 63ac9effd8277..b4a612239c28f 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -952,7 +952,7 @@ theorem castNum_shiftRight (m : Num) (n : Nat) : ↑(m >>> n) = (m : ℕ) >>> (n · rw [Nat.shiftRight_eq_div_pow] symm apply Nat.div_eq_of_lt - simp [@Nat.pow_lt_pow_of_lt_right 2 (by decide) 0 (n + 1) (Nat.succ_pos _)] + simp · trans apply IH change Nat.shiftRight m n = Nat.shiftRight (_root_.bit1 m) (n + 1) diff --git a/Mathlib/Data/Polynomial/Degree/CardPowDegree.lean b/Mathlib/Data/Polynomial/Degree/CardPowDegree.lean index be6609d0c9334..99474b5b3dd54 100644 --- a/Mathlib/Data/Polynomial/Degree/CardPowDegree.lean +++ b/Mathlib/Data/Polynomial/Degree/CardPowDegree.lean @@ -61,7 +61,7 @@ noncomputable def cardPowDegree : AbsoluteValue Fq[X] ℤ := · simp only [hpq, hp, hq, eq_self_iff_true, if_true, if_false] exact add_nonneg (pow_pos _).le (pow_pos _).le simp only [hpq, hp, hq, if_false] - refine' le_trans (pow_le_pow (by linarith) (Polynomial.natDegree_add_le _ _)) _ + refine' le_trans (pow_le_pow_right (by linarith) (Polynomial.natDegree_add_le _ _)) _ refine' le_trans (le_max_iff.mpr _) (max_le_add_of_nonneg (pow_nonneg (by linarith) _) (pow_nonneg (by linarith) _)) @@ -104,7 +104,7 @@ theorem cardPowDegree_isEuclidean : IsEuclidean (cardPowDegree : AbsoluteValue F · simp only [hp, hq, degree_zero, Ne.def, bot_lt_iff_ne_bot, degree_eq_bot, pow_pos, not_false_iff] · simp only [hp, hq, degree_zero, not_lt_bot, (pow_pos _).not_lt] - · rw [degree_eq_natDegree hp, degree_eq_natDegree hq, Nat.cast_lt, pow_lt_pow_iff] + · rw [degree_eq_natDegree hp, degree_eq_natDegree hq, Nat.cast_lt, pow_lt_pow_iff_right] exact mod_cast @Fintype.one_lt_card Fq _ _ } #align polynomial.card_pow_degree_is_euclidean Polynomial.cardPowDegree_isEuclidean diff --git a/Mathlib/Data/Polynomial/Degree/Lemmas.lean b/Mathlib/Data/Polynomial/Degree/Lemmas.lean index 6fb20d164aacd..90965bb8c1969 100644 --- a/Mathlib/Data/Polynomial/Degree/Lemmas.lean +++ b/Mathlib/Data/Polynomial/Degree/Lemmas.lean @@ -51,7 +51,7 @@ theorem natDegree_comp_le : natDegree (p.comp q) ≤ natDegree p * natDegree q : _ ≤ natDegree (C (coeff p n)) + n • degree q := (add_le_add degree_le_natDegree (degree_pow_le _ _)) _ ≤ natDegree (C (coeff p n)) + n • ↑(natDegree q) := - (add_le_add_left (nsmul_le_nsmul_of_le_right (@degree_le_natDegree _ _ q) n) _) + (add_le_add_left (nsmul_le_nsmul_right (@degree_le_natDegree _ _ q) n) _) _ = (n * natDegree q : ℕ) := by rw [natDegree_C, Nat.cast_zero, zero_add, nsmul_eq_mul]; simp diff --git a/Mathlib/Data/Polynomial/RingDivision.lean b/Mathlib/Data/Polynomial/RingDivision.lean index e7abfdaa98f36..f6aec117ed4ed 100644 --- a/Mathlib/Data/Polynomial/RingDivision.lean +++ b/Mathlib/Data/Polynomial/RingDivision.lean @@ -784,6 +784,10 @@ theorem roots_smul_nonzero (p : R[X]) (ha : a ≠ 0) : (a • p).roots = p.roots rw [smul_eq_C_mul, roots_C_mul _ ha] #align polynomial.roots_smul_nonzero Polynomial.roots_smul_nonzero +@[simp] +lemma roots_neg (p : R[X]) : (-p).roots = p.roots := by + rw [← neg_one_smul R p, roots_smul_nonzero p (neg_ne_zero.mpr one_ne_zero)] + theorem roots_list_prod (L : List R[X]) : (0 : R[X]) ∉ L → L.prod.roots = (L : Multiset R[X]).bind roots := List.recOn L (fun _ => roots_one) fun hd tl ih H => by @@ -1062,6 +1066,11 @@ theorem aroots_one [CommRing S] [IsDomain S] [Algebra T S] : (1 : T[X]).aroots S = 0 := aroots_C 1 +@[simp] +theorem aroots_neg [CommRing S] [IsDomain S] [Algebra T S] (p : T[X]) : + (-p).aroots S = p.aroots S := + by rw [aroots, Polynomial.map_neg, roots_neg] + @[simp] theorem aroots_C_mul [CommRing S] [IsDomain S] [Algebra T S] [NoZeroSMulDivisors T S] {a : T} (p : T[X]) (ha : a ≠ 0) : @@ -1123,6 +1132,15 @@ theorem rootSet_zero (S) [CommRing S] [IsDomain S] [Algebra T S] : (0 : T[X]).ro rw [← C_0, rootSet_C] #align polynomial.root_set_zero Polynomial.rootSet_zero +@[simp] +theorem rootSet_one (S) [CommRing S] [IsDomain S] [Algebra T S] : (1 : T[X]).rootSet S = ∅ := by + rw [← C_1, rootSet_C] + +@[simp] +theorem rootSet_neg (p : T[X]) (S) [CommRing S] [IsDomain S] [Algebra T S] : + (-p).rootSet S = p.rootSet S := by + rw [rootSet, aroots_neg, rootSet] + instance rootSetFintype (p : T[X]) (S : Type*) [CommRing S] [IsDomain S] [Algebra T S] : Fintype (p.rootSet S) := FinsetCoe.fintype _ @@ -1263,6 +1281,31 @@ theorem eq_of_monic_of_dvd_of_natDegree_le {R} [CommRing R] {p q : R[X]} (hp : p rw [hq.leadingCoeff, C_1, one_mul] #align polynomial.eq_of_monic_of_dvd_of_nat_degree_le Polynomial.eq_of_monic_of_dvd_of_natDegree_le +lemma eq_zero_of_natDegree_lt_card_of_eval_eq_zero {R} [CommRing R] [IsDomain R] + (p : R[X]) {ι} [Fintype ι] {f : ι → R} (hf : Function.Injective f) + (heval : ∀ i, p.eval (f i) = 0) (hcard : natDegree p < Fintype.card ι) : p = 0 := by + classical + by_contra hp + apply not_lt_of_le (le_refl (Finset.card p.roots.toFinset)) + calc + Finset.card p.roots.toFinset ≤ Multiset.card p.roots := Multiset.toFinset_card_le _ + _ ≤ natDegree p := Polynomial.card_roots' p + _ < Fintype.card ι := hcard + _ = Fintype.card (Set.range f) := (Set.card_range_of_injective hf).symm + _ = Finset.card (Finset.univ.image f) := by rw [← Set.toFinset_card, Set.toFinset_range] + _ ≤ Finset.card p.roots.toFinset := Finset.card_mono ?_ + intro _ + simp only [Finset.mem_image, Finset.mem_univ, true_and, Multiset.mem_toFinset, mem_roots', ne_eq, + IsRoot.def, forall_exists_index, hp, not_false_eq_true] + rintro x rfl + exact heval _ + +lemma eq_zero_of_natDegree_lt_card_of_eval_eq_zero' {R} [CommRing R] [IsDomain R] + (p : R[X]) (s : Finset R) (heval : ∀ i ∈ s, p.eval i = 0) (hcard : natDegree p < s.card) : + p = 0 := + eq_zero_of_natDegree_lt_card_of_eval_eq_zero p Subtype.val_injective + (fun i : s ↦ heval i i.prop) (hcard.trans_eq (Fintype.card_coe s).symm) + theorem isCoprime_X_sub_C_of_isUnit_sub {R} [CommRing R] {a b : R} (h : IsUnit (a - b)) : IsCoprime (X - C a) (X - C b) := ⟨-C h.unit⁻¹.val, C h.unit⁻¹.val, by diff --git a/Mathlib/Data/Real/ENNReal.lean b/Mathlib/Data/Real/ENNReal.lean index 16013ce9c274b..47f14748d308c 100644 --- a/Mathlib/Data/Real/ENNReal.lean +++ b/Mathlib/Data/Real/ENNReal.lean @@ -1030,7 +1030,7 @@ theorem pow_strictMono : ∀ {n : ℕ}, n ≠ 0 → StrictMono fun x : ℝ≥0 | (n + 1 + 1), _ => fun x y h => mul_lt_mul h (pow_strictMono n.succ_ne_zero h) #align ennreal.pow_strict_mono ENNReal.pow_strictMono -@[gcongr] protected theorem pow_lt_pow_of_lt_left (h : a < b) {n : ℕ} (hn : n ≠ 0) : +@[gcongr] protected theorem pow_lt_pow_left (h : a < b) {n : ℕ} (hn : n ≠ 0) : a ^ n < b ^ n := ENNReal.pow_strictMono hn h @@ -1977,13 +1977,13 @@ theorem Ioo_zero_top_eq_iUnion_Ico_zpow {y : ℝ≥0∞} (hy : 1 < y) (h'y : y theorem zpow_le_of_le {x : ℝ≥0∞} (hx : 1 ≤ x) {a b : ℤ} (h : a ≤ b) : x ^ a ≤ x ^ b := by induction' a with a a <;> induction' b with b b · simp only [Int.ofNat_eq_coe, zpow_ofNat] - exact pow_le_pow hx (Int.le_of_ofNat_le_ofNat h) + exact pow_le_pow_right hx (Int.le_of_ofNat_le_ofNat h) · apply absurd h (not_le_of_gt _) exact lt_of_lt_of_le (Int.negSucc_lt_zero _) (Int.ofNat_nonneg _) · simp only [zpow_negSucc, Int.ofNat_eq_coe, zpow_ofNat] refine' (ENNReal.inv_le_one.2 _).trans _ <;> exact one_le_pow_of_one_le' hx _ · simp only [zpow_negSucc, ENNReal.inv_le_inv] - apply pow_le_pow hx + apply pow_le_pow_right hx simpa only [← Int.ofNat_le, neg_le_neg_iff, Int.ofNat_add, Int.ofNat_one, Int.negSucc_eq] using h #align ennreal.zpow_le_of_le ENNReal.zpow_le_of_le diff --git a/Mathlib/Data/Real/Pi/Bounds.lean b/Mathlib/Data/Real/Pi/Bounds.lean index 4d05fd8800e2f..e4473ca0b6c14 100644 --- a/Mathlib/Data/Real/Pi/Bounds.lean +++ b/Mathlib/Data/Real/Pi/Bounds.lean @@ -45,11 +45,11 @@ theorem pi_lt_sqrtTwoAddSeries (n : ℕ) : · rw [div_le_iff'] · refine' le_trans pi_le_four _ simp only [show (4 : ℝ) = (2 : ℝ) ^ 2 by norm_num, mul_one] - apply pow_le_pow; norm_num; apply le_add_of_nonneg_left; apply Nat.zero_le + apply pow_le_pow_right; norm_num; apply le_add_of_nonneg_left; apply Nat.zero_le · apply pow_pos; norm_num apply add_le_add_left; rw [div_le_div_right] rw [le_div_iff, ← mul_pow] - refine' le_trans _ (le_of_eq (one_pow 3)); apply pow_le_pow_of_le_left + refine' le_trans _ (le_of_eq (one_pow 3)); apply pow_le_pow_left · apply le_of_lt; apply mul_pos; apply div_pos pi_pos; apply pow_pos; norm_num; apply pow_pos norm_num rw [← le_div_iff] diff --git a/Mathlib/Data/Real/Pi/Leibniz.lean b/Mathlib/Data/Real/Pi/Leibniz.lean index 644b510307692..a54c6ceb80ada 100644 --- a/Mathlib/Data/Real/Pi/Leibniz.lean +++ b/Mathlib/Data/Real/Pi/Leibniz.lean @@ -112,13 +112,13 @@ theorem tendsto_sum_pi_div_four : exact (le_add_of_nonneg_right (sq_nonneg x) : (1 : ℝ) ≤ _) have hbound1 : ∀ x ∈ Ico (U : ℝ) 1, |f' x| ≤ 1 := by rintro x ⟨hx_left, hx_right⟩ - have hincr := pow_le_pow_of_le_left (le_trans hU2 hx_left) (le_of_lt hx_right) (2 * k) + have hincr := pow_le_pow_left (le_trans hU2 hx_left) (le_of_lt hx_right) (2 * k) rw [one_pow (2 * k), ← abs_of_nonneg (le_trans hU2 hx_left)] at hincr rw [← abs_of_nonneg (le_trans hU2 hx_left)] at hx_right linarith [f'_bound x (mem_Icc.mpr (abs_le.mp (le_of_lt hx_right)))] have hbound2 : ∀ x ∈ Ico 0 (U : ℝ), |f' x| ≤ U ^ (2 * k) := by rintro x ⟨hx_left, hx_right⟩ - have hincr := pow_le_pow_of_le_left hx_left (le_of_lt hx_right) (2 * k) + have hincr := pow_le_pow_left hx_left (le_of_lt hx_right) (2 * k) rw [← abs_of_nonneg hx_left] at hincr hx_right rw [← abs_of_nonneg hU2] at hU1 hx_right exact (f'_bound x (mem_Icc.mpr (abs_le.mp (le_trans (le_of_lt hx_right) hU1)))).trans hincr diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index d61cf4bc106c3..9664197605460 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -476,7 +476,7 @@ theorem ncard_def (s : Set α) : s.ncard = ENat.toNat s.encard := rfl theorem Finite.cast_ncard_eq (hs : s.Finite) : s.ncard = s.encard := by rwa [ncard, ENat.coe_toNat_eq_self, ne_eq, encard_eq_top_iff, Set.Infinite, not_not] -@[simp] theorem Nat.card_coe_set_eq (s : Set α) : Nat.card s = s.ncard := by +theorem Nat.card_coe_set_eq (s : Set α) : Nat.card s = s.ncard := by obtain (h | h) := s.finite_or_infinite · have := h.fintype rw [ncard, h.encard_eq_coe_toFinset_card, Nat.card_eq_fintype_card, diff --git a/Mathlib/Data/Set/Finite.lean b/Mathlib/Data/Set/Finite.lean index d7ec3228f0c00..9484aa9de1cf5 100644 --- a/Mathlib/Data/Set/Finite.lean +++ b/Mathlib/Data/Set/Finite.lean @@ -286,11 +286,6 @@ protected theorem toFinset_compl [DecidableEq α] [Fintype α] (hs : s.Finite) ( simp #align set.finite.to_finset_compl Set.Finite.toFinset_compl --- porting note: was `@[simp]`, now `simp` can prove it -protected theorem toFinset_empty (h : (∅ : Set α).Finite) : h.toFinset = ∅ := - toFinite_toFinset _ -#align set.finite.to_finset_empty Set.Finite.toFinset_empty - protected theorem toFinset_univ [Fintype α] (h : (Set.univ : Set α).Finite) : h.toFinset = Finset.univ := by simp @@ -301,6 +296,10 @@ protected theorem toFinset_eq_empty {h : s.Finite} : h.toFinset = ∅ ↔ s = @toFinset_eq_empty _ _ h.fintype #align set.finite.to_finset_eq_empty Set.Finite.toFinset_eq_empty +protected theorem toFinset_empty (h : (∅ : Set α).Finite) : h.toFinset = ∅ := by + simp +#align set.finite.to_finset_empty Set.Finite.toFinset_empty + @[simp] protected theorem toFinset_eq_univ [Fintype α] {h : s.Finite} : h.toFinset = Finset.univ ↔ s = univ := @@ -786,6 +785,10 @@ theorem Finite.subset {s : Set α} (hs : s.Finite) {t : Set α} (ht : t ⊆ s) : apply toFinite #align set.finite.subset Set.Finite.subset +protected lemma Infinite.mono {s t : Set α} (h : s ⊆ t) : s.Infinite → t.Infinite := + mt fun ht ↦ ht.subset h +#align set.infinite.mono Set.Infinite.mono + theorem Finite.diff {s : Set α} (hs : s.Finite) (t : Set α) : (s \ t).Finite := by cases hs apply toFinite @@ -906,26 +909,34 @@ theorem Finite.of_finite_image {s : Set α} {f : α → β} (h : (f '' s).Finite Subtype.eq <| hi a.2 b.2 <| Subtype.ext_iff_val.1 eq⟩ #align set.finite.of_finite_image Set.Finite.of_finite_image -theorem finite_of_finite_preimage {f : α → β} {s : Set β} (h : (f ⁻¹' s).Finite) - (hs : s ⊆ range f) : s.Finite := by +section preimage +variable {f : α → β} {s : Set β} + +theorem finite_of_finite_preimage (h : (f ⁻¹' s).Finite) (hs : s ⊆ range f) : s.Finite := by rw [← image_preimage_eq_of_subset hs] exact Finite.image f h #align set.finite_of_finite_preimage Set.finite_of_finite_preimage -theorem Finite.of_preimage {f : α → β} {s : Set β} (h : (f ⁻¹' s).Finite) (hf : Surjective f) : - s.Finite := +theorem Finite.of_preimage (h : (f ⁻¹' s).Finite) (hf : Surjective f) : s.Finite := hf.image_preimage s ▸ h.image _ #align set.finite.of_preimage Set.Finite.of_preimage -theorem Finite.preimage {s : Set β} {f : α → β} (I : Set.InjOn f (f ⁻¹' s)) (h : s.Finite) : - (f ⁻¹' s).Finite := +theorem Finite.preimage (I : Set.InjOn f (f ⁻¹' s)) (h : s.Finite) : (f ⁻¹' s).Finite := (h.subset (image_preimage_subset f s)).of_finite_image I #align set.finite.preimage Set.Finite.preimage +protected lemma Infinite.preimage (hs : s.Infinite) (hf : s ⊆ range f) : (f ⁻¹' s).Infinite := + fun h ↦ hs $ finite_of_finite_preimage h hf + +lemma Infinite.preimage' (hs : (s ∩ range f).Infinite) : (f ⁻¹' s).Infinite := + (hs.preimage $ inter_subset_right _ _).mono $ preimage_mono $ inter_subset_left _ _ + theorem Finite.preimage_embedding {s : Set β} (f : α ↪ β) (h : s.Finite) : (f ⁻¹' s).Finite := h.preimage fun _ _ _ _ h' => f.injective h' #align set.finite.preimage_embedding Set.Finite.preimage_embedding +end preimage + theorem finite_lt_nat (n : ℕ) : Set.Finite { i | i < n } := toFinite _ #align set.finite_lt_nat Set.finite_lt_nat @@ -1237,9 +1248,8 @@ theorem empty_card : Fintype.card (∅ : Set α) = 0 := rfl #align set.empty_card Set.empty_card -@[simp] theorem empty_card' {h : Fintype.{u} (∅ : Set α)} : @Fintype.card (∅ : Set α) h = 0 := - Eq.trans (by congr; exact Subsingleton.elim _ _) empty_card + by simp #align set.empty_card' Set.empty_card' theorem card_fintypeInsertOfNotMem {a : α} (s : Set α) [Fintype s] (h : a ∉ s) : @@ -1362,10 +1372,6 @@ theorem Finite.infinite_compl [Infinite α] {s : Set α} (hs : s.Finite) : sᶜ. Set.infinite_univ (by simpa using hs.union h) #align set.finite.infinite_compl Set.Finite.infinite_compl -protected theorem Infinite.mono {s t : Set α} (h : s ⊆ t) : s.Infinite → t.Infinite := - mt fun ht => ht.subset h -#align set.infinite.mono Set.Infinite.mono - theorem Infinite.diff {s t : Set α} (hs : s.Infinite) (ht : t.Finite) : (s \ t).Infinite := fun h => hs <| h.of_diff ht #align set.infinite.diff Set.Infinite.diff @@ -1414,6 +1420,12 @@ theorem infinite_image2 (hfs : ∀ b ∈ t, InjOn (fun a => f a b) s) (hft : ∀ · exact ht.image2_right ha (hft _ ha) #align set.infinite_image2 Set.infinite_image2 +lemma finite_image2 (hfs : ∀ b ∈ t, InjOn (f · b) s) (hft : ∀ a ∈ s, InjOn (f a) t) : + (image2 f s t).Finite ↔ s.Finite ∧ t.Finite ∨ s = ∅ ∨ t = ∅ := by + rw [← not_infinite, infinite_image2 hfs hft] + simp [not_or, -not_and, not_and_or, not_nonempty_iff_eq_empty] + aesop + end Image2 theorem infinite_of_injOn_mapsTo {s : Set α} {t : Set β} {f : α → β} (hi : InjOn f s) diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index 7e4ea6308ce1a..344381d7f7b9b 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -428,6 +428,10 @@ theorem mapsTo_empty (f : α → β) (t : Set β) : MapsTo f ∅ t := empty_subset _ #align set.maps_to_empty Set.mapsTo_empty +/-- If `f` maps `s` to `t` and `s` is non-empty, `t` is non-empty. -/ +theorem MapsTo.nonempty (h : MapsTo f s t) (hs : s.Nonempty) : t.Nonempty := + (hs.image f).mono (mapsTo'.mp h) + theorem MapsTo.image_subset (h : MapsTo f s t) : f '' s ⊆ t := mapsTo'.1 h #align set.maps_to.image_subset Set.MapsTo.image_subset diff --git a/Mathlib/Data/Set/NAry.lean b/Mathlib/Data/Set/NAry.lean index c60dd9b7deaee..e162a2aa89e33 100644 --- a/Mathlib/Data/Set/NAry.lean +++ b/Mathlib/Data/Set/NAry.lean @@ -10,19 +10,19 @@ import Mathlib.Data.Set.Prod /-! # N-ary images of sets -This file defines `Set.image2`, the binary image of finsets. This is the finset version of -`Set.image2`. This is mostly useful to define pointwise operations. +This file defines `Set.image2`, the binary image of sets. +This is mostly useful to define pointwise operations and `Set.seq`. ## Notes -This file is very similar to the n-ary section of `Data.Set.Basic`, to `Order.Filter.NAry` and to +This file is very similar to `Data.Finset.NAry`, to `Order.Filter.NAry`, and to `Data.Option.NAry`. Please keep them in sync. -We do not define `Set.image3` as its only purpose would be to prove properties of `Set.image2` -and `Set.image2` already fulfills this task. +We also define `Set.image3`. +Its only purpose is to prove properties of `Set.image2`, +and it should be rarely used outside of this file. -/ - open Function namespace Set @@ -122,11 +122,7 @@ theorem image2_swap (s : Set α) (t : Set β) : image2 f s t = image2 (fun a b = variable {f} theorem image2_union_left : image2 f (s ∪ s') t = image2 f s t ∪ image2 f s' t := by - ext c - constructor - · rintro ⟨a, b, ha | ha, hb, rfl⟩ <;> [left; right] <;> exact ⟨_, _, ‹_›, ‹_›, rfl⟩ - · rintro (⟨_, _, _, _, rfl⟩ | ⟨_, _, _, _, rfl⟩) <;> refine' ⟨_, _, _, ‹_›, rfl⟩ <;> - simp [mem_union, *] + simp_rw [← image_prod, union_prod, image_union] #align set.image2_union_left Set.image2_union_left theorem image2_union_right : image2 f s (t ∪ t') = image2 f s t ∪ image2 f s t' := by @@ -181,14 +177,12 @@ theorem Subsingleton.image2 (hs : s.Subsingleton) (ht : t.Subsingleton) (f : α rw [← image_prod] apply (hs.prod ht).image -theorem image2_inter_subset_left : image2 f (s ∩ s') t ⊆ image2 f s t ∩ image2 f s' t := by - rintro _ ⟨a, b, ⟨h1a, h2a⟩, hb, rfl⟩ - constructor <;> exact ⟨_, _, ‹_›, ‹_›, rfl⟩ +theorem image2_inter_subset_left : image2 f (s ∩ s') t ⊆ image2 f s t ∩ image2 f s' t := + Monotone.map_inf_le (fun _ _ ↦ image2_subset_right) s s' #align set.image2_inter_subset_left Set.image2_inter_subset_left -theorem image2_inter_subset_right : image2 f s (t ∩ t') ⊆ image2 f s t ∩ image2 f s t' := by - rintro _ ⟨a, b, ha, ⟨h1b, h2b⟩, rfl⟩ - constructor <;> exact ⟨_, _, ‹_›, ‹_›, rfl⟩ +theorem image2_inter_subset_right : image2 f s (t ∩ t') ⊆ image2 f s t ∩ image2 f s t' := + Monotone.map_inf_le (fun _ _ ↦ image2_subset_left) t t' #align set.image2_inter_subset_right Set.image2_inter_subset_right @[simp] @@ -285,20 +279,12 @@ theorem image_image2 (f : α → β → γ) (g : γ → δ) : theorem image2_image_left (f : γ → β → δ) (g : α → γ) : image2 f (g '' s) t = image2 (fun a b => f (g a) b) s t := by - ext; constructor - · rintro ⟨_, b, ⟨a, ha, rfl⟩, hb, rfl⟩ - refine' ⟨a, b, ha, hb, rfl⟩ - · rintro ⟨a, b, ha, hb, rfl⟩ - refine' ⟨_, b, ⟨a, ha, rfl⟩, hb, rfl⟩ + ext; simp #align set.image2_image_left Set.image2_image_left theorem image2_image_right (f : α → γ → δ) (g : β → γ) : image2 f s (g '' t) = image2 (fun a b => f a (g b)) s t := by - ext; constructor - · rintro ⟨a, _, ha, ⟨b, hb, rfl⟩, rfl⟩ - refine' ⟨a, b, ha, hb, rfl⟩ - · rintro ⟨a, b, ha, hb, rfl⟩ - refine' ⟨a, _, ha, ⟨b, hb, rfl⟩, rfl⟩ + ext; simp #align set.image2_image_right Set.image2_image_right @[simp] diff --git a/Mathlib/Data/Set/Pointwise/Basic.lean b/Mathlib/Data/Set/Pointwise/Basic.lean index 7b88c72a59cad..207ccd9f79f16 100644 --- a/Mathlib/Data/Set/Pointwise/Basic.lean +++ b/Mathlib/Data/Set/Pointwise/Basic.lean @@ -344,7 +344,6 @@ theorem mul_mem_mul : a ∈ s → b ∈ t → a * b ∈ s * t := #align set.mul_mem_mul Set.mul_mem_mul #align set.add_mem_add Set.add_mem_add -/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[to_additive add_image_prod] theorem image_mul_prod : (fun x : α × α => x.fst * x.snd) '' s ×ˢ t = s * t := image_prod _ @@ -611,7 +610,6 @@ theorem div_mem_div : a ∈ s → b ∈ t → a / b ∈ s / t := #align set.div_mem_div Set.div_mem_div #align set.sub_mem_sub Set.sub_mem_sub -/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[to_additive sub_image_prod] theorem image_div_prod : (fun x : α × α => x.fst / x.snd) '' s ×ˢ t = s / t := image_prod _ @@ -1297,6 +1295,13 @@ theorem image_mul : m '' (s * t) = m '' s * m '' t := #align set.image_mul Set.image_mul #align set.image_add Set.image_add +@[to_additive] +lemma mul_subset_range {s t : Set β} (hs : s ⊆ range m) (ht : t ⊆ range m) : s * t ⊆ range m := by + rintro _ ⟨a, b, ha, hb, rfl⟩; + obtain ⟨a, rfl⟩ := hs ha + obtain ⟨b, rfl⟩ := ht hb + exact ⟨a * b, map_mul _ _ _⟩ + @[to_additive] theorem preimage_mul_preimage_subset {s t : Set β} : m ⁻¹' s * m ⁻¹' t ⊆ m ⁻¹' (s * t) := by rintro _ ⟨_, _, _, _, rfl⟩ @@ -1304,6 +1309,16 @@ theorem preimage_mul_preimage_subset {s t : Set β} : m ⁻¹' s * m ⁻¹' t #align set.preimage_mul_preimage_subset Set.preimage_mul_preimage_subset #align set.preimage_add_preimage_subset Set.preimage_add_preimage_subset +@[to_additive] +lemma preimage_mul (hm : Injective m) {s t : Set β} (hs : s ⊆ range m) (ht : t ⊆ range m) : + m ⁻¹' (s * t) = m ⁻¹' s * m ⁻¹' t := by + refine subset_antisymm ?_ (preimage_mul_preimage_subset m) + rintro a ⟨b, c, hb, hc, ha⟩ + obtain ⟨b, rfl⟩ := hs hb + obtain ⟨c, rfl⟩ := ht hc + simp only [← map_mul, hm.eq_iff] at ha + exact ⟨b, c, hb, hc, ha⟩ + end Mul section Group @@ -1316,6 +1331,13 @@ theorem image_div : m '' (s / t) = m '' s / m '' t := #align set.image_div Set.image_div #align set.image_sub Set.image_sub +@[to_additive] +lemma div_subset_range {s t : Set β} (hs : s ⊆ range m) (ht : t ⊆ range m) : s / t ⊆ range m := by + rintro _ ⟨a, b, ha, hb, rfl⟩; + obtain ⟨a, rfl⟩ := hs ha + obtain ⟨b, rfl⟩ := ht hb + exact ⟨a / b, map_div _ _ _⟩ + @[to_additive] theorem preimage_div_preimage_subset {s t : Set β} : m ⁻¹' s / m ⁻¹' t ⊆ m ⁻¹' (s / t) := by rintro _ ⟨_, _, _, _, rfl⟩ @@ -1323,6 +1345,16 @@ theorem preimage_div_preimage_subset {s t : Set β} : m ⁻¹' s / m ⁻¹' t #align set.preimage_div_preimage_subset Set.preimage_div_preimage_subset #align set.preimage_sub_preimage_subset Set.preimage_sub_preimage_subset +@[to_additive] +lemma preimage_div (hm : Injective m) {s t : Set β} (hs : s ⊆ range m) (ht : t ⊆ range m) : + m ⁻¹' (s / t) = m ⁻¹' s / m ⁻¹' t := by + refine subset_antisymm ?_ (preimage_div_preimage_subset m) + rintro a ⟨b, c, hb, hc, ha⟩ + obtain ⟨b, rfl⟩ := hs hb + obtain ⟨c, rfl⟩ := ht hc + simp only [← map_div, hm.eq_iff] at ha + exact ⟨b, c, hb, hc, ha⟩ + end Group @[to_additive] diff --git a/Mathlib/Data/Set/Pointwise/Finite.lean b/Mathlib/Data/Set/Pointwise/Finite.lean index b0d9811cfa2bd..f65d3b16657bb 100644 --- a/Mathlib/Data/Set/Pointwise/Finite.lean +++ b/Mathlib/Data/Set/Pointwise/Finite.lean @@ -135,6 +135,11 @@ theorem infinite_mul : (s * t).Infinite ↔ s.Infinite ∧ t.Nonempty ∨ t.Infi #align set.infinite_mul Set.infinite_mul #align set.infinite_add Set.infinite_add +@[to_additive] +lemma finite_mul : (s * t).Finite ↔ s.Finite ∧ t.Finite ∨ s = ∅ ∨ t = ∅ := + finite_image2 (fun _ _ ↦ (mul_left_injective _).injOn _) + fun _ _ ↦ (mul_right_injective _).injOn _ + end Cancel section Group diff --git a/Mathlib/Data/Sum/Basic.lean b/Mathlib/Data/Sum/Basic.lean index 5d3ce32aa1a4b..77b16d5e073b8 100644 --- a/Mathlib/Data/Sum/Basic.lean +++ b/Mathlib/Data/Sum/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Yury G. Kudryashov -/ import Mathlib.Logic.Function.Basic +import Mathlib.Tactic.MkIffOfInductiveProp #align_import data.sum.basic from "leanprover-community/mathlib"@"bd9851ca476957ea4549eb19b40e7b5ade9428cc" @@ -188,7 +189,9 @@ theorem swap_rightInverse : Function.RightInverse (@swap α β) swap := #align sum.get_left_swap Sum.getLeft?_swap #align sum.get_right_swap Sum.getRight?_swap -section LiftRel +mk_iff_of_inductive_prop Sum.LiftRel Sum.liftRel_iff + +namespace LiftRel #align sum.lift_rel Sum.LiftRel #align sum.lift_rel_inl_inl Sum.liftRel_inl_inl @@ -201,6 +204,36 @@ section LiftRel #align sum.lift_rel.swap Sum.LiftRel.swap #align sum.lift_rel_swap_iff Sum.liftRel_swap_iff +variable {r : α → γ → Prop} {s : β → δ → Prop} {x : Sum α β} {y : Sum γ δ} + {a : α} {b : β} {c : γ} {d : δ} + +theorem isLeft_congr (h : LiftRel r s x y) : x.isLeft ↔ y.isLeft := by cases h <;> rfl +theorem isRight_congr (h : LiftRel r s x y) : x.isRight ↔ y.isRight := by cases h <;> rfl + +theorem isLeft_left (h : LiftRel r s x (inl c)) : x.isLeft := by cases h; rfl +theorem isLeft_right (h : LiftRel r s (inl a) y) : y.isLeft := by cases h; rfl +theorem isRight_left (h : LiftRel r s x (inr d)) : x.isRight := by cases h; rfl +theorem isRight_right (h : LiftRel r s (inr b) y) : y.isRight := by cases h; rfl + +theorem exists_of_isLeft_left (h₁ : LiftRel r s x y) (h₂ : x.isLeft) : + ∃ a c, r a c ∧ x = inl a ∧ y = inl c := by + rcases isLeft_iff.mp h₂ with ⟨_, rfl⟩ + simp only [liftRel_iff, false_and, and_false, exists_false, or_false] at h₁ + exact h₁ + +theorem exists_of_isLeft_right (h₁ : LiftRel r s x y) (h₂ : y.isLeft) : + ∃ a c, r a c ∧ x = inl a ∧ y = inl c := exists_of_isLeft_left h₁ ((isLeft_congr h₁).mpr h₂) + +theorem exists_of_isRight_left (h₁ : LiftRel r s x y) (h₂ : x.isRight) : + ∃ b d, s b d ∧ x = inr b ∧ y = inr d := by + rcases isRight_iff.mp h₂ with ⟨_, rfl⟩ + simp only [liftRel_iff, false_and, and_false, exists_false, false_or] at h₁ + exact h₁ + +theorem exists_of_isRight_right (h₁ : LiftRel r s x y) (h₂ : y.isRight) : + ∃ b d, s b d ∧ x = inr b ∧ y = inr d := + exists_of_isRight_left h₁ ((isRight_congr h₁).mpr h₂) + end LiftRel section Lex diff --git a/Mathlib/FieldTheory/Finite/Basic.lean b/Mathlib/FieldTheory/Finite/Basic.lean index f94c126b11920..e7f3b3f42fee4 100644 --- a/Mathlib/FieldTheory/Finite/Basic.lean +++ b/Mathlib/FieldTheory/Finite/Basic.lean @@ -339,7 +339,7 @@ set_option linter.uppercaseLean3 false in theorem X_pow_card_pow_sub_X_natDegree_eq (hn : n ≠ 0) (hp : 1 < p) : (X ^ p ^ n - X : K'[X]).natDegree = p ^ n := - X_pow_card_sub_X_natDegree_eq K' <| Nat.one_lt_pow _ _ (Nat.pos_of_ne_zero hn) hp + X_pow_card_sub_X_natDegree_eq K' <| Nat.one_lt_pow _ _ hn hp set_option linter.uppercaseLean3 false in #align finite_field.X_pow_card_pow_sub_X_nat_degree_eq FiniteField.X_pow_card_pow_sub_X_natDegree_eq @@ -352,7 +352,7 @@ set_option linter.uppercaseLean3 false in #align finite_field.X_pow_card_sub_X_ne_zero FiniteField.X_pow_card_sub_X_ne_zero theorem X_pow_card_pow_sub_X_ne_zero (hn : n ≠ 0) (hp : 1 < p) : (X ^ p ^ n - X : K'[X]) ≠ 0 := - X_pow_card_sub_X_ne_zero K' <| Nat.one_lt_pow _ _ (Nat.pos_of_ne_zero hn) hp + X_pow_card_sub_X_ne_zero K' <| Nat.one_lt_pow _ _ hn hp set_option linter.uppercaseLean3 false in #align finite_field.X_pow_card_pow_sub_X_ne_zero FiniteField.X_pow_card_pow_sub_X_ne_zero diff --git a/Mathlib/FieldTheory/Finite/Polynomial.lean b/Mathlib/FieldTheory/Finite/Polynomial.lean index a10cbb36711a9..ce36f4958ce28 100644 --- a/Mathlib/FieldTheory/Finite/Polynomial.lean +++ b/Mathlib/FieldTheory/Finite/Polynomial.lean @@ -82,7 +82,7 @@ theorem degrees_indicator (c : σ → K) : refine' le_trans (degrees_prod _ _) (Finset.sum_le_sum fun s _ => _) refine' le_trans (degrees_sub _ _) _ rw [degrees_one, ← bot_eq_zero, bot_sup_eq] - refine' le_trans (degrees_pow _ _) (nsmul_le_nsmul_of_le_right _ _) + refine' le_trans (degrees_pow _ _) (nsmul_le_nsmul_right _ _) refine' le_trans (degrees_sub _ _) _ rw [degrees_C, ← bot_eq_zero, sup_bot_eq] exact degrees_X' _ diff --git a/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean b/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean new file mode 100644 index 0000000000000..acdd23f680982 --- /dev/null +++ b/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean @@ -0,0 +1,214 @@ +/- +Copyright (c) 2023 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.FieldTheory.Minpoly.IsIntegrallyClosed +import Mathlib.FieldTheory.PrimitiveElement +import Mathlib.FieldTheory.IsAlgClosed.Basic + +/-! +# Results about `minpoly R x / (X - C x)` + +## Main definition +- `minpolyDiv`: The polynomial `minpoly R x / (X - C x)`. + +We used the contents of this file to describe the dual basis of a powerbasis under the trace form. +See `traceForm_dualBasis_powerBasis_eq`. + +## Main results +- `span_coeff_minpolyDiv`: The coefficients of `minpolyDiv` spans `R`. +-/ + +open Polynomial BigOperators FiniteDimensional + +variable (R K) {L S} [CommRing R] [Field K] [Field L] [CommRing S] [Algebra R S] [Algebra K L] +variable (x : S) + +/-- `minpolyDiv R x : S[X]` for `x : S` is the polynomial `minpoly R x / (X - C x)`. -/ +noncomputable def minpolyDiv : S[X] := (minpoly R x).map (algebraMap R S) /ₘ (X - C x) + +lemma minpolyDiv_spec : + minpolyDiv R x * (X - C x) = (minpoly R x).map (algebraMap R S) := by + delta minpolyDiv + rw [mul_comm, mul_divByMonic_eq_iff_isRoot, IsRoot, eval_map, ← aeval_def, minpoly.aeval] + +lemma coeff_minpolyDiv (i) : coeff (minpolyDiv R x) i = + algebraMap R S (coeff (minpoly R x) (i + 1)) + coeff (minpolyDiv R x) (i + 1) * x := by + rw [← coeff_map, ← minpolyDiv_spec R x]; simp [mul_sub] + +variable (hx : IsIntegral R x) {R x} + +lemma minpolyDiv_ne_zero [Nontrivial S] : minpolyDiv R x ≠ 0 := by + intro e + have := minpolyDiv_spec R x + rw [e, zero_mul] at this + exact ((minpoly.monic hx).map (algebraMap R S)).ne_zero this.symm + +lemma minpolyDiv_eq_zero (hx : ¬IsIntegral R x) : minpolyDiv R x = 0 := by + delta minpolyDiv minpoly + rw [dif_neg hx, Polynomial.map_zero, zero_divByMonic] + +lemma minpolyDiv_monic : Monic (minpolyDiv R x) := by + nontriviality S + have := congr_arg leadingCoeff (minpolyDiv_spec R x) + rw [leadingCoeff_mul', ((minpoly.monic hx).map (algebraMap R S)).leadingCoeff] at this + · simpa using this + · simpa using minpolyDiv_ne_zero hx + +lemma eval_minpolyDiv_self : (minpolyDiv R x).eval x = aeval x (derivative <| minpoly R x) := by + rw [aeval_def, ← eval_map, ← derivative_map, ← minpolyDiv_spec R x]; simp + +lemma minpolyDiv_eval_eq_zero_of_ne_of_aeval_eq_zero [IsDomain S] + {y} (hxy : y ≠ x) (hy : aeval y (minpoly R x) = 0) : (minpolyDiv R x).eval y = 0 := by + rw [aeval_def, ← eval_map, ← minpolyDiv_spec R x] at hy + simp only [eval_mul, eval_sub, eval_X, eval_C, mul_eq_zero] at hy + exact hy.resolve_right (by rwa [sub_eq_zero]) + +lemma eval₂_minpolyDiv_of_eval₂_eq_zero {T} [CommRing T] + [IsDomain T] [DecidableEq T] {x y} + (σ : S →+* T) (hy : eval₂ (σ.comp (algebraMap R S)) y (minpoly R x) = 0) : + eval₂ σ y (minpolyDiv R x) = + if σ x = y then σ (aeval x (derivative <| minpoly R x)) else 0 := by + split_ifs with h + · rw [← h, eval₂_hom, eval_minpolyDiv_self] + · rw [← eval₂_map, ← minpolyDiv_spec] at hy + simpa [sub_eq_zero, Ne.symm h] using hy + +lemma eval₂_minpolyDiv_self {T} [CommRing T] [Algebra R T] [IsDomain T] [DecidableEq T] (x : S) + (σ₁ σ₂ : S →ₐ[R] T) : + eval₂ σ₁ (σ₂ x) (minpolyDiv R x) = + if σ₁ x = σ₂ x then σ₁ (aeval x (derivative <| minpoly R x)) else 0 := by + apply eval₂_minpolyDiv_of_eval₂_eq_zero + rw [AlgHom.comp_algebraMap, ← σ₂.comp_algebraMap, ← eval₂_map, ← RingHom.coe_coe, eval₂_hom, + eval_map, ← aeval_def, minpoly.aeval, map_zero] + +lemma eval_minpolyDiv_of_aeval_eq_zero [IsDomain S] [DecidableEq S] + {y} (hy : aeval y (minpoly R x) = 0) : + (minpolyDiv R x).eval y = if x = y then aeval x (derivative <| minpoly R x) else 0 := by + rw [eval, eval₂_minpolyDiv_of_eval₂_eq_zero]; rfl + exact hy + +lemma natDegree_minpolyDiv_succ [Nontrivial S] : + natDegree (minpolyDiv R x) + 1 = natDegree (minpoly R x) := by + rw [← (minpoly.monic hx).natDegree_map (algebraMap R S), ← minpolyDiv_spec, natDegree_mul'] + · simp + · simpa using minpolyDiv_ne_zero hx + +lemma natDegree_minpolyDiv : + natDegree (minpolyDiv R x) = natDegree (minpoly R x) - 1 := by + nontriviality S + by_cases hx : IsIntegral R x + · rw [← natDegree_minpolyDiv_succ hx]; rfl + · rw [minpolyDiv_eq_zero hx, minpoly.eq_zero hx]; rfl + +lemma natDegree_minpolyDiv_lt [Nontrivial S] : + natDegree (minpolyDiv R x) < natDegree (minpoly R x) := by + rw [← natDegree_minpolyDiv_succ hx] + exact Nat.lt.base _ + +lemma coeff_minpolyDiv_mem_adjoin (x : S) (i) : + coeff (minpolyDiv R x) i ∈ Algebra.adjoin R {x} := by + by_contra H + have : ∀ j, coeff (minpolyDiv R x) (i + j) ∉ Algebra.adjoin R {x} + · intro j; induction j with + | zero => exact H + | succ j IH => + intro H; apply IH + rw [coeff_minpolyDiv] + refine add_mem ?_ (mul_mem H (Algebra.self_mem_adjoin_singleton R x)) + exact Subalgebra.algebraMap_mem _ _ + apply this (natDegree (minpolyDiv R x) + 1) + rw [coeff_eq_zero_of_natDegree_lt] + · exact zero_mem _ + · refine (Nat.le_add_left _ i).trans_lt ?_ + rw [← add_assoc] + exact Nat.lt.base _ + +lemma minpolyDiv_eq_of_isIntegrallyClosed [IsDomain R] [IsIntegrallyClosed R] [IsDomain S] + [Algebra R K] [Algebra K S] [IsScalarTower R K S] [IsFractionRing R K] : + minpolyDiv R x = minpolyDiv K x := by + delta minpolyDiv + rw [IsScalarTower.algebraMap_eq R K S, ← map_map, + ← minpoly.isIntegrallyClosed_eq_field_fractions' _ hx] + +lemma coeff_minpolyDiv_sub_pow_mem_span {i} (hi : i ≤ natDegree (minpolyDiv R x)) : + coeff (minpolyDiv R x) (natDegree (minpolyDiv R x) - i) - x ^ i ∈ + Submodule.span R ((x ^ ·) '' Set.Iio i) := by + induction i with + | zero => simp [(minpolyDiv_monic hx).leadingCoeff] + | succ i IH => + rw [coeff_minpolyDiv, add_sub_assoc, pow_succ', ← sub_mul, Algebra.algebraMap_eq_smul_one] + refine add_mem ?_ ?_ + · apply Submodule.smul_mem + apply Submodule.subset_span + exact ⟨0, Nat.zero_lt_succ _, pow_zero _⟩ + · rw [Nat.succ_eq_add_one, ← tsub_tsub, tsub_add_cancel_of_le + (le_tsub_of_add_le_left (b := 1) hi)] + apply SetLike.le_def.mp ?_ + (Submodule.mul_mem_mul (IH ((Nat.le_succ _).trans hi)) + (Submodule.mem_span_singleton_self x)) + rw [Submodule.span_mul_span, Set.mul_singleton, Set.image_image] + apply Submodule.span_mono + rintro _ ⟨j, hj : j < i, rfl⟩ + exact ⟨j + 1, Nat.add_lt_of_lt_sub hj, pow_succ' x j⟩ + +lemma span_coeff_minpolyDiv : + Submodule.span R (Set.range (coeff (minpolyDiv R x))) = + Subalgebra.toSubmodule (Algebra.adjoin R {x}) := by + nontriviality S + classical + apply le_antisymm + · rw [Submodule.span_le] + rintro _ ⟨i, rfl⟩ + apply coeff_minpolyDiv_mem_adjoin + · rw [← Submodule.span_range_natDegree_eq_adjoin (minpoly.monic hx) (minpoly.aeval _ _), + Submodule.span_le] + simp only [Finset.coe_image, Finset.coe_range, Set.image_subset_iff] + intro i + apply Nat.strongInductionOn i + intro i hi hi' + have : coeff (minpolyDiv R x) (natDegree (minpolyDiv R x) - i) ∈ + Submodule.span R (Set.range (coeff (minpolyDiv R x))) := + Submodule.subset_span (Set.mem_range_self _) + rw [Set.mem_preimage, SetLike.mem_coe, ← Submodule.sub_mem_iff_right _ this] + refine SetLike.le_def.mp ?_ (coeff_minpolyDiv_sub_pow_mem_span hx ?_) + · rw [Submodule.span_le, Set.image_subset_iff] + intro j (hj : j < i) + exact hi j hj (lt_trans hj hi') + · rwa [← natDegree_minpolyDiv_succ hx, Set.mem_Iio, Nat.lt_succ_iff] at hi' + +section PowerBasis + +variable {K} + +lemma sum_smul_minpolyDiv_eq_X_pow (E) [Field E] [Algebra K E] [IsAlgClosed E] + [FiniteDimensional K L] [IsSeparable K L] + {x : L} (hxL : Algebra.adjoin K {x} = ⊤) {r : ℕ} (hr : r < finrank K L) : + ∑ σ : L →ₐ[K] E, ((x ^ r / aeval x (derivative <| minpoly K x)) • + minpolyDiv K x).map σ = (X ^ r : E[X]) := by + classical + rw [← sub_eq_zero] + have : Function.Injective (fun σ : L →ₐ[K] E ↦ σ x) := fun _ _ h => + AlgHom.ext_of_adjoin_eq_top hxL (fun _ hx ↦ hx ▸ h) + apply Polynomial.eq_zero_of_natDegree_lt_card_of_eval_eq_zero _ this + · intro σ + simp only [Polynomial.map_smul, map_div₀, map_pow, RingHom.coe_coe, eval_sub, eval_finset_sum, + eval_smul, eval_map, eval₂_minpolyDiv_self, this.eq_iff, smul_eq_mul, mul_ite, mul_zero, + Finset.sum_ite_eq', Finset.mem_univ, ite_true, eval_pow, eval_X] + rw [sub_eq_zero, div_mul_cancel] + rw [ne_eq, map_eq_zero_iff σ σ.toRingHom.injective] + exact (IsSeparable.separable _ _).aeval_derivative_ne_zero (minpoly.aeval _ _) + · refine (Polynomial.natDegree_sub_le _ _).trans_lt + (max_lt ((Polynomial.natDegree_sum_le _ _).trans_lt ?_) ?_) + · simp only [AlgEquiv.toAlgHom_eq_coe, Polynomial.map_smul, + map_div₀, map_pow, RingHom.coe_coe, AlgHom.coe_coe, Function.comp_apply, + Finset.mem_univ, forall_true_left, true_and, Finset.fold_max_lt, AlgHom.card] + refine ⟨finrank_pos, ?_⟩ + intro σ + exact ((Polynomial.natDegree_smul_le _ _).trans (natDegree_map_le _ _)).trans_lt + ((natDegree_minpolyDiv_lt (Algebra.IsIntegral.of_finite _ _ x)).trans_le + (minpoly.natDegree_le _)) + · rwa [natDegree_pow, natDegree_X, mul_one, AlgHom.card] + +end PowerBasis diff --git a/Mathlib/FieldTheory/Separable.lean b/Mathlib/FieldTheory/Separable.lean index b90d92843b8ba..d84753e080e0c 100644 --- a/Mathlib/FieldTheory/Separable.lean +++ b/Mathlib/FieldTheory/Separable.lean @@ -130,6 +130,19 @@ theorem Separable.map {p : R[X]} (h : p.Separable) {f : R →+* S} : (p.map f).S Polynomial.map_one]⟩ #align polynomial.separable.map Polynomial.Separable.map +theorem Separable.eval₂_derivative_ne_zero [Nontrivial S] (f : R →+* S) {p : R[X]} + (h : p.Separable) {x : S} (hx : p.eval₂ f x = 0) : + (derivative p).eval₂ f x ≠ 0 := by + intro hx' + obtain ⟨a, b, e⟩ := h + apply_fun Polynomial.eval₂ f x at e + simp only [eval₂_add, eval₂_mul, hx, mul_zero, hx', add_zero, eval₂_one, zero_ne_one] at e + +theorem Separable.aeval_derivative_ne_zero [Nontrivial S] [Algebra R S] {p : R[X]} + (h : p.Separable) {x : S} (hx : aeval x p = 0) : + aeval x (derivative p) ≠ 0 := + h.eval₂_derivative_ne_zero (algebraMap R S) hx + variable (p q : ℕ) theorem isUnit_of_self_mul_dvd_separable {p q : R[X]} (hp : p.Separable) (hq : q * q ∣ p) : @@ -492,7 +505,7 @@ variable (F K : Type*) [CommRing F] [Ring K] [Algebra F K] -- to allow non-algebraic extensions, then the definition of `IsGalois` must also be changed. /-- Typeclass for separable field extension: `K` is a separable field extension of `F` iff the minimal polynomial of every `x : K` is separable. This implies that `K/F` is an algebraic -extension, because the minimal polynomial of a non-integral element is 0, which is not +extension, because the minimal polynomial of a non-integral element is `0`, which is not separable. We define this for general (commutative) rings and only assume `F` and `K` are fields if this @@ -508,11 +521,17 @@ theorem IsSeparable.separable [IsSeparable F K] : ∀ x : K, (minpoly F x).Separ IsSeparable.separable' #align is_separable.separable IsSeparable.separable -theorem IsSeparable.isIntegral [IsSeparable F K] : ∀ x : K, IsIntegral F x := fun x ↦ by +variable {F} in +/-- If the minimal polynomial of `x : K` over `F` is separable, then `x` is integral over `F`, +because the minimal polynomial of a non-integral element is `0`, which is not separable. -/ +theorem Polynomial.Separable.isIntegral {x : K} (h : (minpoly F x).Separable) : IsIntegral F x := by cases subsingleton_or_nontrivial F · haveI := Module.subsingleton F K exact ⟨1, monic_one, Subsingleton.elim _ _⟩ - · exact of_not_not fun h ↦ not_separable_zero (minpoly.eq_zero h ▸ IsSeparable.separable F x) + · exact of_not_not fun h' ↦ not_separable_zero (minpoly.eq_zero h' ▸ h) + +theorem IsSeparable.isIntegral [IsSeparable F K] : + ∀ x : K, IsIntegral F x := fun x ↦ (IsSeparable.separable F x).isIntegral #align is_separable.is_integral IsSeparable.isIntegral variable {F} @@ -576,6 +595,19 @@ theorem IsSeparable.of_algHom (E' : Type*) [Field E'] [Algebra F E'] (f : E → exact isSeparable_tower_bot_of_isSeparable F E E' #align is_separable.of_alg_hom IsSeparable.of_algHom +lemma IsSeparable.of_equiv_equiv {A₁ B₁ A₂ B₂ : Type*} [Field A₁] [Field B₁] + [Field A₂] [Field B₂] [Algebra A₁ B₁] [Algebra A₂ B₂] (e₁ : A₁ ≃+* A₂) (e₂ : B₁ ≃+* B₂) + (he : RingHom.comp (algebraMap A₂ B₂) ↑e₁ = RingHom.comp ↑e₂ (algebraMap A₁ B₁)) + [IsSeparable A₁ B₁] : IsSeparable A₂ B₂ := by + letI := e₁.toRingHom.toAlgebra + letI := ((algebraMap A₁ B₁).comp e₁.symm.toRingHom).toAlgebra + haveI : IsScalarTower A₁ A₂ B₁ := IsScalarTower.of_algebraMap_eq + (fun x ↦ by simp [RingHom.algebraMap_toAlgebra]) + let e : B₁ ≃ₐ[A₂] B₂ := { e₂ with commutes' := fun r ↦ by simpa [RingHom.algebraMap_toAlgebra] + using FunLike.congr_fun he.symm (e₁.symm r) } + haveI := isSeparable_tower_top_of_isSeparable A₁ A₂ B₁ + exact IsSeparable.of_algHom _ _ e.symm.toAlgHom + end IsSeparableTower section CardAlgHom diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean new file mode 100644 index 0000000000000..e8f1d22de97be --- /dev/null +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -0,0 +1,127 @@ +/- +Copyright (c) 2023 Michael Rothgang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Rothgang +-/ + +import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners + +/-! +# Interior and boundary of a manifold +Define the interior and boundary of a manifold. + +## Main definitions +- **IsInteriorPoint x**: `p ∈ M` is an interior point if, for `φ` being the preferred chart at `x`, + `φ x` is an interior point of `φ.target`. +- **IsBoundaryPoint x**: `p ∈ M` is a boundary point if, `(extChartAt I x) x ∈ frontier (range I)`. +- **interior I M** is the **interior** of `M`, the set of its interior points. +- **boundary I M** is the **boundary** of `M`, the set of its boundary points. + +## Main results +- `univ_eq_interior_union_boundary`: `M` is the union of its interior and boundary +- `interior_boundary_disjoint`: interior and boundary of `M` are disjoint +- if `M` is boundaryless, every point is an interior point + +## Tags +manifold, interior, boundary + +## TODO +- `x` is an interior point iff *any* chart around `x` maps it to `interior (range I)`; +similarly for the boundary. +- the interior of `M` is open, hence the boundary is closed (and nowhere dense) + In finite dimensions, this requires e.g. the homology of spheres. +- the interior of `M` is a smooth manifold without boundary +- `boundary M` is a smooth submanifold (possibly with boundary and corners): +follows from the corresponding statement for the model with corners `I`; +this requires a definition of submanifolds +- if `M` is finite-dimensional, its boundary has measure zero + +-/ + +open Set + +-- Let `M` be a manifold with corners over the pair `(E, H)`. +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] + +namespace ModelWithCorners +/-- `p ∈ M` is an interior point of a manifold `M` iff its image in the extended chart +lies in the interior of the model space. -/ +def IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (range I) + +/-- `p ∈ M` is a boundary point of a manifold `M` iff its image in the extended chart +lies on the boundary of the model space. -/ +def IsBoundaryPoint (x : M) := extChartAt I x x ∈ frontier (range I) + +variable (M) in +/-- The **interior** of a manifold `M` is the set of its interior points. -/ +protected def interior : Set M := { x : M | I.IsInteriorPoint x } + +lemma isInteriorPoint_iff {x : M} : + I.IsInteriorPoint x ↔ extChartAt I x x ∈ interior (extChartAt I x).target := + ⟨fun h ↦ (chartAt H x).mem_interior_extend_target _ (mem_chart_target H x) h, + fun h ↦ PartialHomeomorph.interior_extend_target_subset_interior_range _ _ h⟩ + +variable (M) in +/-- The **boundary** of a manifold `M` is the set of its boundary points. -/ +protected def boundary : Set M := { x : M | I.IsBoundaryPoint x } + +lemma isBoundaryPoint_iff {x : M} : I.IsBoundaryPoint x ↔ extChartAt I x x ∈ frontier (range I) := + Iff.rfl + +/-- Every point is either an interior or a boundary point. -/ +lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by + by_cases h : extChartAt I x x ∈ interior (range I) + · exact Or.inl h + · right -- Otherwise, we have a boundary point. + rw [I.isBoundaryPoint_iff, ← closure_diff_interior, I.closed_range.closure_eq] + exact ⟨mem_range_self _, h⟩ + +/-- A manifold decomposes into interior and boundary. -/ +lemma interior_union_boundary_eq_univ : (I.interior M) ∪ (I.boundary M) = (univ : Set M) := + le_antisymm (fun _ _ ↦ trivial) (fun x _ ↦ I.isInteriorPoint_or_isBoundaryPoint x) + +/-- The interior and boundary of a manifold `M` are disjoint. -/ +lemma disjoint_interior_boundary : Disjoint (I.interior M) (I.boundary M) := by + by_contra h + -- Choose some x in the intersection of interior and boundary. + choose x hx using not_disjoint_iff.mp h + rcases hx with ⟨h1, h2⟩ + show (extChartAt I x) x ∈ (∅ : Set E) + rw [← disjoint_iff_inter_eq_empty.mp (disjoint_interior_frontier (s := range I))] + exact ⟨h1, h2⟩ + +/-- The boundary is the complement of the interior. -/ +lemma boundary_eq_complement_interior : I.boundary M = (I.interior M)ᶜ := by + apply (compl_unique ?_ I.interior_union_boundary_eq_univ).symm + exact disjoint_iff_inter_eq_empty.mp (I.disjoint_interior_boundary) + +lemma _root_.range_mem_nhds_isInteriorPoint {x : M} (h : I.IsInteriorPoint x) : + range I ∈ nhds (extChartAt I x x) := by + rw [mem_nhds_iff] + exact ⟨interior (range I), interior_subset, isOpen_interior, h⟩ + +section boundaryless +variable [I.Boundaryless] + +/-- If `I` is boundaryless, every point of `M` is an interior point. -/ +lemma isInteriorPoint {x : M} : I.IsInteriorPoint x := by + let r := ((chartAt H x).isOpen_extend_target I).interior_eq + have : extChartAt I x = (chartAt H x).extend I := rfl + rw [← this] at r + rw [ModelWithCorners.isInteriorPoint_iff, r] + exact PartialEquiv.map_source _ (mem_extChartAt_source _ _) + +/-- If `I` is boundaryless, `M` has full interior. -/ +lemma interior_eq_univ : I.interior M = univ := by + ext + refine ⟨fun _ ↦ trivial, fun _ ↦ I.isInteriorPoint⟩ + +/-- If `I` is boundaryless, `M` has empty boundary. -/ +lemma Boundaryless.boundary_eq_empty : I.boundary M = ∅ := by + rw [I.boundary_eq_complement_interior, I.interior_eq_univ, compl_empty_iff] + +end boundaryless +end ModelWithCorners diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 000615ba8de8d..198dab83bd691 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -976,6 +976,10 @@ theorem extend_target : (f.extend I).target = I.symm ⁻¹' f.target ∩ range I simp_rw [extend, PartialEquiv.trans_target, I.target_eq, I.toPartialEquiv_coe_symm, inter_comm] #align local_homeomorph.extend_target PartialHomeomorph.extend_target +lemma isOpen_extend_target [I.Boundaryless] : IsOpen (f.extend I).target := by + rw [extend_target, I.range_eq_univ, inter_univ] + exact I.continuous_symm.isOpen_preimage _ f.open_target + theorem mapsTo_extend (hs : s ⊆ f.source) : MapsTo (f.extend I) s ((f.extend I).symm ⁻¹' s ∩ range I) := by rw [mapsTo', extend_coe, extend_coe_symm, preimage_comp, ← I.image_eq, image_comp, @@ -1023,6 +1027,19 @@ theorem extend_target_mem_nhdsWithin {y : M} (hy : y ∈ f.source) : theorem extend_target_subset_range : (f.extend I).target ⊆ range I := by simp only [mfld_simps] #align local_homeomorph.extend_target_subset_range PartialHomeomorph.extend_target_subset_range +lemma interior_extend_target_subset_interior_range : + interior (f.extend I).target ⊆ interior (range I) := by + rw [f.extend_target, interior_inter, (f.open_target.preimage I.continuous_symm).interior_eq] + exact inter_subset_right _ _ + +/-- If `y ∈ f.target` and `I y ∈ interior (range I)`, + then `I y` is an interior point of `(I ∘ f).target`. -/ +lemma mem_interior_extend_target {y : H} (hy : y ∈ f.target) + (hy' : I y ∈ interior (range I)) : I y ∈ interior (f.extend I).target := by + rw [f.extend_target, interior_inter, (f.open_target.preimage I.continuous_symm).interior_eq, + mem_inter_iff, mem_preimage] + exact ⟨mem_of_eq_of_mem (I.left_inv (y)) hy, hy'⟩ + theorem nhdsWithin_extend_target_eq {y : M} (hy : y ∈ f.source) : 𝓝[(f.extend I).target] f.extend I y = 𝓝[range I] f.extend I y := (nhdsWithin_mono _ (extend_target_subset_range _ _)).antisymm <| diff --git a/Mathlib/GroupTheory/Archimedean.lean b/Mathlib/GroupTheory/Archimedean.lean index ac56be53d793d..9f6b570acc6c2 100644 --- a/Mathlib/GroupTheory/Archimedean.lean +++ b/Mathlib/GroupTheory/Archimedean.lean @@ -81,7 +81,7 @@ theorem AddSubgroup.exists_isLeast_pos {H : AddSubgroup G} (hbot : H ≠ ⊥) {a · exact hmin m hmn ⟨y, hyH, hm⟩ · refine disjoint_left.1 hd (sub_mem hxH hyH) ⟨sub_pos.2 hxy, sub_lt_iff_lt_add'.2 ?_⟩ calc x ≤ (n + 1) • a := hxn - _ ≤ (m + 1) • a := nsmul_le_nsmul h₀.le (add_le_add_right hnm _) + _ ≤ (m + 1) • a := nsmul_le_nsmul_left h₀.le (add_le_add_right hnm _) _ = m • a + a := succ_nsmul' _ _ _ < y + a := add_lt_add_right hm.1 _ diff --git a/Mathlib/GroupTheory/PGroup.lean b/Mathlib/GroupTheory/PGroup.lean index 4b61b066635fd..261210b0e8032 100644 --- a/Mathlib/GroupTheory/PGroup.lean +++ b/Mathlib/GroupTheory/PGroup.lean @@ -67,6 +67,8 @@ theorem iff_card [Fact p.Prime] [Fintype G] : IsPGroup p G ↔ ∃ n : ℕ, card exact (hq1.pow_eq_iff.mp (hg.symm.trans hk).symm).1.symm #align is_p_group.iff_card IsPGroup.iff_card +alias ⟨exists_card_eq, _⟩ := iff_card + section GIsPGroup variable (hG : IsPGroup p G) diff --git a/Mathlib/GroupTheory/Schreier.lean b/Mathlib/GroupTheory/Schreier.lean index e066b6ab84ca0..ec755ac9c39ce 100644 --- a/Mathlib/GroupTheory/Schreier.lean +++ b/Mathlib/GroupTheory/Schreier.lean @@ -196,7 +196,7 @@ theorem card_commutator_le_of_finite_commutatorSet [Finite (commutatorSet G)] : (Nat.pow_le_pow_of_le_right Finite.card_pos (rank_closureCommutatorRepresentatives_le G)) replace h2 := h2.trans (pow_dvd_pow _ (add_le_add_right (mul_le_mul_right' h1 _) 1)) rw [← pow_succ'] at h2 - refine' (Nat.le_of_dvd _ h2).trans (Nat.pow_le_pow_of_le_left h1 _) + refine' (Nat.le_of_dvd _ h2).trans (Nat.pow_le_pow_left h1 _) exact pow_pos (Nat.pos_of_ne_zero FiniteIndex.finiteIndex) _ #align subgroup.card_commutator_le_of_finite_commutator_set Subgroup.card_commutator_le_of_finite_commutatorSet diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 56dec72109afb..b07dd12d3c50c 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -177,22 +177,30 @@ theorem exists_inv_mem_iff_exists_mem {P : G → Prop} : @[to_additive] theorem mul_mem_cancel_right {x y : G} (h : x ∈ H) : y * x ∈ H ↔ y ∈ H := - -- Porting note: whut? why do we need this? - haveI : SubmonoidClass S G := SubgroupClass.toSubmonoidClass ⟨fun hba => by simpa using mul_mem hba (inv_mem h), fun hb => mul_mem hb h⟩ #align mul_mem_cancel_right mul_mem_cancel_right #align add_mem_cancel_right add_mem_cancel_right @[to_additive] theorem mul_mem_cancel_left {x y : G} (h : x ∈ H) : x * y ∈ H ↔ y ∈ H := - -- Porting note: whut? why do we need this? - have : SubmonoidClass S G := SubgroupClass.toSubmonoidClass ⟨fun hab => by simpa using mul_mem (inv_mem h) hab, mul_mem h⟩ #align mul_mem_cancel_left mul_mem_cancel_left #align add_mem_cancel_left add_mem_cancel_left namespace SubgroupClass +-- Here we assume H, K, and L are subgroups, but in fact any one of them +-- could be allowed to be a subsemigroup. +-- Counterexample where K and L are submonoids: H = ℤ, K = ℕ, L = -ℕ +-- Counterexample where H and K are submonoids: H = {n | n = 0 ∨ 3 ≤ n}, K = 3ℕ + 4ℕ, L = 5ℤ +@[to_additive] +theorem subset_union {H K L : S} : (H : Set G) ⊆ K ∪ L ↔ H ≤ K ∨ H ≤ L := by + refine ⟨fun h ↦ ?_, fun h x xH ↦ h.imp (· xH) (· xH)⟩ + rw [or_iff_not_imp_left, SetLike.not_le_iff_exists] + exact fun ⟨x, xH, xK⟩ y yH ↦ (h <| mul_mem xH yH).elim + ((h yH).resolve_left fun yK ↦ xK <| (mul_mem_cancel_right yK).mp ·) + (mul_mem_cancel_left <| (h xH).resolve_left xK).mp + /-- A subgroup of a group inherits an inverse. -/ @[to_additive "An additive subgroup of an `AddGroup` inherits an inverse."] instance inv {G : Type u_1} {S : Type u_2} [DivInvMonoid G] [SetLike S G] diff --git a/Mathlib/GroupTheory/Subsemigroup/Center.lean b/Mathlib/GroupTheory/Subsemigroup/Center.lean index 976df673386c3..dd4c15de90db2 100644 --- a/Mathlib/GroupTheory/Subsemigroup/Center.lean +++ b/Mathlib/GroupTheory/Subsemigroup/Center.lean @@ -56,6 +56,11 @@ structure IsMulCentral [Mul M] (z : M) : Prop where /-- associative property for right multiplication -/ right_assoc (a b : M) : (a * b) * z = a * (b * z) +-- TODO: these should have explicit arguments (mathlib4#9129) +attribute [mk_iff isMulCentral_iff] IsMulCentral +attribute [mk_iff isAddCentral_iff] IsAddCentral +attribute [to_additive existing] isMulCentral_iff + namespace IsMulCentral variable {a b c : M} [Mul M] diff --git a/Mathlib/GroupTheory/Sylow.lean b/Mathlib/GroupTheory/Sylow.lean index 830b5c9577724..3b95deaa7c007 100644 --- a/Mathlib/GroupTheory/Sylow.lean +++ b/Mathlib/GroupTheory/Sylow.lean @@ -653,6 +653,33 @@ theorem exists_subgroup_card_pow_prime [Fintype G] (p : ℕ) {n : ℕ} [Fact p.P ⟨K, hK.1⟩ #align sylow.exists_subgroup_card_pow_prime Sylow.exists_subgroup_card_pow_prime +lemma exists_subgroup_card_pow_prime_of_le_card {m p : ℕ} (hp : p.Prime) (h : IsPGroup p G) + (hm : p ^ m ≤ Nat.card G) : ∃ H : Subgroup G, Nat.card H = p ^ m := by + have : Fact p.Prime := ⟨hp⟩ + have : Finite G := Nat.finite_of_card_ne_zero $ by linarith [Nat.one_le_pow m p hp.pos] + cases nonempty_fintype G + obtain ⟨n, hn⟩ := h.exists_card_eq + simp_rw [Nat.card_eq_fintype_card] at hn hm ⊢ + refine exists_subgroup_card_pow_prime _ ?_ + rw [hn] at hm ⊢ + exact pow_dvd_pow _ $ (pow_le_pow_iff_right hp.one_lt).1 hm + +lemma exists_subgroup_le_card_pow_prime_of_le_card {m p : ℕ} (hp : p.Prime) (h : IsPGroup p G) + {H : Subgroup G} (hm : p ^ m ≤ Nat.card H) : ∃ H' ≤ H, Nat.card H' = p ^ m := by + obtain ⟨H', H'card⟩ := exists_subgroup_card_pow_prime_of_le_card hp (h.to_subgroup H) hm + refine ⟨H'.map H.subtype, map_subtype_le _, ?_⟩ + rw [← H'card] + let e : H' ≃* H'.map H.subtype := H'.equivMapOfInjective (Subgroup.subtype H) H.subtype_injective + exact Nat.card_congr e.symm.toEquiv + +lemma exists_subgroup_le_card_le {k p : ℕ} (hp : p.Prime) (h : IsPGroup p G) {H : Subgroup G} + (hk : k ≤ Nat.card H) (hk₀ : k ≠ 0) : ∃ H' ≤ H, Nat.card H' ≤ k ∧ k < p * Nat.card H' := by + obtain ⟨m, hmk, hkm⟩ : ∃ s, p ^ s ≤ k ∧ k < p ^ (s + 1) := + exists_nat_pow_near (Nat.one_le_iff_ne_zero.2 hk₀) hp.one_lt + obtain ⟨H', H'H, H'card⟩ := exists_subgroup_le_card_pow_prime_of_le_card hp h (hmk.trans hk) + refine ⟨H', H'H, ?_⟩ + simpa only [pow_succ, H'card] using And.intro hmk hkm + theorem pow_dvd_card_of_pow_dvd_card [Fintype G] {p n : ℕ} [hp : Fact p.Prime] (P : Sylow p G) (hdvd : p ^ n ∣ card G) : p ^ n ∣ card P := (hp.1.coprime_pow_of_not_dvd diff --git a/Mathlib/Init/Data/Int/CompLemmas.lean b/Mathlib/Init/Data/Int/CompLemmas.lean index ad16299e0b411..7fb5a1885379c 100644 --- a/Mathlib/Init/Data/Int/CompLemmas.lean +++ b/Mathlib/Init/Data/Int/CompLemmas.lean @@ -95,7 +95,7 @@ theorem zero_le_ofNat (n : ℕ) : 0 ≤ ofNat n := theorem ne_of_natAbs_ne_natAbs_of_nonneg {a b : ℤ} (ha : 0 ≤ a) (hb : 0 ≤ b) (h : natAbs a ≠ natAbs b) : a ≠ b := fun h => by have : (natAbs a : ℤ) = natAbs b := by - rwa [ofNat_natAbs_eq_of_nonneg _ ha, ofNat_natAbs_eq_of_nonneg _ hb] + rwa [natAbs_of_nonneg ha, natAbs_of_nonneg hb] injection this contradiction #align int.ne_of_nat_abs_ne_nat_abs_of_nonneg Int.ne_of_natAbs_ne_natAbs_of_nonneg diff --git a/Mathlib/Lean/Expr/Basic.lean b/Mathlib/Lean/Expr/Basic.lean index 20350fd2aac08..3a14668d7fd41 100644 --- a/Mathlib/Lean/Expr/Basic.lean +++ b/Mathlib/Lean/Expr/Basic.lean @@ -200,44 +200,6 @@ def getAppApps (e : Expr) : Array Expr := let nargs := e.getAppNumArgs getAppAppsAux e (mkArray nargs dummy) (nargs-1) -/-- Turn an expression that is a natural number literal into a natural number. -/ -def natLit! : Expr → Nat - | lit (Literal.natVal v) => v - | _ => panic! "nat literal expected" - -/-- Turn an expression that is a constructor of `Int` applied to a natural number literal -into an integer. -/ -def intLit! (e : Expr) : Int := - if e.isAppOfArity ``Int.ofNat 1 then - e.appArg!.natLit! - else if e.isAppOfArity ``Int.negOfNat 1 then - .negOfNat e.appArg!.natLit! - else - panic! "not a raw integer literal" - -/-- -Check if an expression is a "natural number in normal form", -i.e. of the form `OfNat n`, where `n` matches `.lit (.natVal lit)` for some `lit`. -and if so returns `lit`. --/ --- Note that an `Expr.lit (.natVal n)` is not considered in normal form! -def nat? (e : Expr) : Option Nat := do - guard <| e.isAppOfArity ``OfNat.ofNat 3 - let lit (.natVal n) := e.appFn!.appArg! | none - n - - -/-- -Check if an expression is an "integer in normal form", -i.e. either a natural number in normal form, or the negation of one, -and if so returns the integer. --/ -def int? (e : Expr) : Option Int := - if e.isAppOfArity ``Neg.neg 3 then - (- ·) <$> e.appArg!.nat? - else - e.nat? - /-- Check if an expression is a "rational in normal form", i.e. either an integer number in normal form, diff --git a/Mathlib/Lean/Meta.lean b/Mathlib/Lean/Meta.lean index b08a6e9862aa5..606f13de63ee4 100644 --- a/Mathlib/Lean/Meta.lean +++ b/Mathlib/Lean/Meta.lean @@ -32,10 +32,6 @@ def «let» (g : MVarId) (h : Name) (v : Expr) (t : Option Expr := .none) : MetaM (FVarId × MVarId) := do (← g.define h (← t.getDM (inferType v)) v).intro1P -/-- Short-hand for applying a constant to the goal. -/ -def applyConst (mvar : MVarId) (c : Name) (cfg : ApplyConfig := {}) : MetaM (List MVarId) := do - mvar.apply (← mkConstWithFreshMVarLevels c) cfg - /-- Has the effect of `refine ⟨e₁,e₂,⋯, ?_⟩`. -/ def existsi (mvar : MVarId) (es : List Expr) : MetaM MVarId := do @@ -130,13 +126,6 @@ end Lean.MVarId namespace Lean.Meta -/-- Return local hypotheses which are not "implementation detail", as `Expr`s. -/ -def getLocalHyps [Monad m] [MonadLCtx m] : m (Array Expr) := do - let mut hs := #[] - for d in ← getLCtx do - if !d.isImplementationDetail then hs := hs.push d.toExpr - return hs - /-- Count how many local hypotheses appear in an expression. -/ def countLocalHypsUsed [Monad m] [MonadLCtx m] [MonadMCtx m] (e : Expr) : m Nat := do let e' ← instantiateMVars e diff --git a/Mathlib/LinearAlgebra/AdicCompletion.lean b/Mathlib/LinearAlgebra/AdicCompletion.lean index 03809a972c51d..0de18c42ca6a3 100644 --- a/Mathlib/LinearAlgebra/AdicCompletion.lean +++ b/Mathlib/LinearAlgebra/AdicCompletion.lean @@ -90,7 +90,7 @@ In fact, this is only complete if the ideal is finitely generated. -/ def adicCompletion : Submodule R (∀ n : ℕ, M ⧸ (I ^ n • ⊤ : Submodule R M)) where carrier := { f | ∀ {m n} (h : m ≤ n), liftQ _ (mkQ _) (by rw [ker_mkQ] - exact smul_mono (Ideal.pow_le_pow h) le_rfl) + exact smul_mono (Ideal.pow_le_pow_right h) le_rfl) (f n) = f m } zero_mem' hmn := by rw [Pi.zero_apply, Pi.zero_apply, LinearMap.map_zero] add_mem' hf hg m n hmn := by diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean index 4dd0d618b4196..79fd91e5fd90b 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean @@ -367,44 +367,6 @@ theorem equivOfIsometry_refl : end Map -variable (Q) - -/-- If the quadratic form of a vector is invertible, then so is that vector. -/ -def invertibleιOfInvertible (m : M) [Invertible (Q m)] : Invertible (ι Q m) where - invOf := ι Q (⅟ (Q m) • m) - invOf_mul_self := by - rw [map_smul, smul_mul_assoc, ι_sq_scalar, Algebra.smul_def, ← map_mul, invOf_mul_self, map_one] - mul_invOf_self := by - rw [map_smul, mul_smul_comm, ι_sq_scalar, Algebra.smul_def, ← map_mul, invOf_mul_self, map_one] -#align clifford_algebra.invertible_ι_of_invertible CliffordAlgebra.invertibleιOfInvertible - -/-- For a vector with invertible quadratic form, $v^{-1} = \frac{v}{Q(v)}$ -/ -theorem invOf_ι (m : M) [Invertible (Q m)] [Invertible (ι Q m)] : - ⅟ (ι Q m) = ι Q (⅟ (Q m) • m) := by - letI := invertibleιOfInvertible Q m - convert (rfl : ⅟ (ι Q m) = _) -#align clifford_algebra.inv_of_ι CliffordAlgebra.invOf_ι - -theorem isUnit_ι_of_isUnit {m : M} (h : IsUnit (Q m)) : IsUnit (ι Q m) := by - cases h.nonempty_invertible - letI := invertibleιOfInvertible Q m - exact isUnit_of_invertible (ι Q m) -#align clifford_algebra.is_unit_ι_of_is_unit CliffordAlgebra.isUnit_ι_of_isUnit - -/-- $aba^{-1}$ is a vector. -/ -theorem ι_mul_ι_mul_invOf_ι (a b : M) [Invertible (ι Q a)] [Invertible (Q a)] : - ι Q a * ι Q b * ⅟ (ι Q a) = ι Q ((⅟ (Q a) * QuadraticForm.polar Q a b) • a - b) := by - rw [invOf_ι, map_smul, mul_smul_comm, ι_mul_ι_mul_ι, ← map_smul, smul_sub, smul_smul, smul_smul, - invOf_mul_self, one_smul] -#align clifford_algebra.ι_mul_ι_mul_inv_of_ι CliffordAlgebra.ι_mul_ι_mul_invOf_ι - -/-- $a^{-1}ba$ is a vector. -/ -theorem invOf_ι_mul_ι_mul_ι (a b : M) [Invertible (ι Q a)] [Invertible (Q a)] : - ⅟ (ι Q a) * ι Q b * ι Q a = ι Q ((⅟ (Q a) * QuadraticForm.polar Q a b) • a - b) := by - rw [invOf_ι, map_smul, smul_mul_assoc, smul_mul_assoc, ι_mul_ι_mul_ι, ← map_smul, smul_sub, - smul_smul, smul_smul, invOf_mul_self, one_smul] -#align clifford_algebra.inv_of_ι_mul_ι_mul_ι CliffordAlgebra.invOf_ι_mul_ι_mul_ι - end CliffordAlgebra namespace TensorAlgebra diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Inversion.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Inversion.lean new file mode 100644 index 0000000000000..923d4985f9062 --- /dev/null +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Inversion.lean @@ -0,0 +1,76 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.LinearAlgebra.CliffordAlgebra.Contraction + +/-! # Results about inverses in Clifford algebras + +This contains some basic results about the inversion of vectors, related to the fact that +$ι(m)^{-1} = \frac{ι(m)}{Q(m)}$. +-/ + +variable {R M : Type*} +variable [CommRing R] [AddCommGroup M] [Module R M] {Q : QuadraticForm R M} + +namespace CliffordAlgebra + +variable (Q) + +/-- If the quadratic form of a vector is invertible, then so is that vector. -/ +def invertibleιOfInvertible (m : M) [Invertible (Q m)] : Invertible (ι Q m) where + invOf := ι Q (⅟ (Q m) • m) + invOf_mul_self := by + rw [map_smul, smul_mul_assoc, ι_sq_scalar, Algebra.smul_def, ← map_mul, invOf_mul_self, map_one] + mul_invOf_self := by + rw [map_smul, mul_smul_comm, ι_sq_scalar, Algebra.smul_def, ← map_mul, invOf_mul_self, map_one] +#align clifford_algebra.invertible_ι_of_invertible CliffordAlgebra.invertibleιOfInvertible + +/-- For a vector with invertible quadratic form, $v^{-1} = \frac{v}{Q(v)}$ -/ +theorem invOf_ι (m : M) [Invertible (Q m)] [Invertible (ι Q m)] : + ⅟ (ι Q m) = ι Q (⅟ (Q m) • m) := by + letI := invertibleιOfInvertible Q m + convert (rfl : ⅟ (ι Q m) = _) +#align clifford_algebra.inv_of_ι CliffordAlgebra.invOf_ι + +theorem isUnit_ι_of_isUnit {m : M} (h : IsUnit (Q m)) : IsUnit (ι Q m) := by + cases h.nonempty_invertible + letI := invertibleιOfInvertible Q m + exact isUnit_of_invertible (ι Q m) +#align clifford_algebra.is_unit_ι_of_is_unit CliffordAlgebra.isUnit_ι_of_isUnit + +/-- $aba^{-1}$ is a vector. -/ +theorem ι_mul_ι_mul_invOf_ι (a b : M) [Invertible (ι Q a)] [Invertible (Q a)] : + ι Q a * ι Q b * ⅟ (ι Q a) = ι Q ((⅟ (Q a) * QuadraticForm.polar Q a b) • a - b) := by + rw [invOf_ι, map_smul, mul_smul_comm, ι_mul_ι_mul_ι, ← map_smul, smul_sub, smul_smul, smul_smul, + invOf_mul_self, one_smul] +#align clifford_algebra.ι_mul_ι_mul_inv_of_ι CliffordAlgebra.ι_mul_ι_mul_invOf_ι + +/-- $a^{-1}ba$ is a vector. -/ +theorem invOf_ι_mul_ι_mul_ι (a b : M) [Invertible (ι Q a)] [Invertible (Q a)] : + ⅟ (ι Q a) * ι Q b * ι Q a = ι Q ((⅟ (Q a) * QuadraticForm.polar Q a b) • a - b) := by + rw [invOf_ι, map_smul, smul_mul_assoc, smul_mul_assoc, ι_mul_ι_mul_ι, ← map_smul, smul_sub, + smul_smul, smul_smul, invOf_mul_self, one_smul] +#align clifford_algebra.inv_of_ι_mul_ι_mul_ι CliffordAlgebra.invOf_ι_mul_ι_mul_ι + +section +variable [Invertible (2 : R)] + +/-- Over a ring where `2` is invertible, `Q m` is invertible whenever `ι Q m`. -/ +def invertibleOfInvertibleι (m : M) [Invertible (ι Q m)] : Invertible (Q m) := + ExteriorAlgebra.invertibleAlgebraMapEquiv M (Q m) <| + .algebraMapOfInvertibleAlgebraMap (equivExterior Q).toLinearMap (by simp) <| + .copy (.mul ‹Invertible (ι Q m)› ‹Invertible (ι Q m)›) _ (ι_sq_scalar _ _).symm + +theorem isUnit_of_isUnit_ι {m : M} (h : IsUnit (ι Q m)) : IsUnit (Q m) := by + cases h.nonempty_invertible + letI := invertibleOfInvertibleι Q m + exact isUnit_of_invertible (Q m) + +@[simp] theorem isUnit_ι_iff {m : M} : IsUnit (ι Q m) ↔ IsUnit (Q m) := + ⟨isUnit_of_isUnit_ι Q, isUnit_ι_of_isUnit Q⟩ + +end + +end CliffordAlgebra diff --git a/Mathlib/LinearAlgebra/Dimension.lean b/Mathlib/LinearAlgebra/Dimension.lean index facbf8108c688..498fb191118c6 100644 --- a/Mathlib/LinearAlgebra/Dimension.lean +++ b/Mathlib/LinearAlgebra/Dimension.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro, Johannes Hölzl, Sander Dahmen, Scott Morrison -/ import Mathlib.Algebra.Module.BigOperators import Mathlib.LinearAlgebra.Basis.VectorSpace +import Mathlib.Algebra.Module.Torsion import Mathlib.LinearAlgebra.DFinsupp import Mathlib.LinearAlgebra.FreeModule.Basic import Mathlib.LinearAlgebra.InvariantBasisNumber @@ -78,7 +79,7 @@ noncomputable section universe u v v' v'' u₁' w w' -variable {K : Type u} {V V₁ V₂ V₃ : Type v} {V' V'₁ : Type v'} {V'' : Type v''} +variable {K R : Type u} {V V₁ V₂ V₃ : Type v} {V' V'₁ : Type v'} {V'' : Type v''} variable {ι : Type w} {ι' : Type w'} {η : Type u₁'} {φ : η → Type*} @@ -102,8 +103,6 @@ this is the same as the dimension of the space (i.e. the cardinality of any basi In particular this agrees with the usual notion of the dimension of a vector space. -The definition is marked as protected to avoid conflicts with `_root_.rank`, -the rank of a linear map. -/ protected irreducible_def Module.rank : Cardinal := ⨆ ι : { s : Set V // LinearIndependent K ((↑) : s → V) }, (#ι.1) @@ -113,7 +112,7 @@ end section -variable {R : Type u} [Ring R] +variable [Ring R] variable {M : Type v} [AddCommGroup M] [Module R M] @@ -129,7 +128,7 @@ theorem LinearMap.lift_rank_le_of_injective (f : M →ₗ[R] M') (i : Injective apply ciSup_mono' (Cardinal.bddAbove_range.{v', v} _) rintro ⟨s, li⟩ refine' ⟨⟨f '' s, _⟩, Cardinal.lift_mk_le'.mpr ⟨(Equiv.Set.image f s i).toEmbedding⟩⟩ - exact (li.map' _ <| LinearMap.ker_eq_bot.mpr i).image + exact (li.map' _ <| LinearMap.ker_eq_bot_of_injective i).image #align linear_map.lift_rank_le_of_injective LinearMap.lift_rank_le_of_injective theorem LinearMap.rank_le_of_injective (f : M →ₗ[R] M₁) (i : Injective f) : @@ -348,8 +347,7 @@ lemma basis_finite_of_finite_spans (w : Set M) (hw : w.Finite) (s : span R w = rw [k] exact mem_top -- giving the desire contradiction. - refine' b.linearIndependent.not_mem_span_image _ k' - exact nm + exact b.linearIndependent.not_mem_span_image nm k' #align basis_fintype_of_finite_spans basis_finite_of_finite_spansₓ -- From [Les familles libres maximales d'un module ont-elles le meme cardinal?][lazarus1973] @@ -383,18 +381,10 @@ theorem union_support_maximal_linearIndependent_eq_range_basis {ι : Type w} (b rfl rw [← e] at p exact r' p - have inj' : Injective v' := by - rintro (_ | k) (_ | k) z - · rfl - · exfalso - exact r' ⟨k, z.symm⟩ - · exfalso - exact r' ⟨k, z⟩ - · congr - exact i.injective z -- The key step in the proof is checking that this strictly larger family is linearly independent. have i' : LinearIndependent R ((↑) : range v' → M) := by - rw [linearIndependent_subtype_range inj', linearIndependent_iff] + apply LinearIndependent.to_subtype_range + rw [linearIndependent_iff] intro l z rw [Finsupp.total_option] at z simp only [Option.elim'] at z @@ -419,7 +409,7 @@ theorem union_support_maximal_linearIndependent_eq_range_basis {ι : Type w} (b · simp only [l₀, Finsupp.coe_zero, Pi.zero_apply] · erw [FunLike.congr_fun l₁ a] simp only [Finsupp.coe_zero, Pi.zero_apply] - dsimp [LinearIndependent.Maximal] at m + rw [LinearIndependent.Maximal] at m specialize m (range v') i' r exact r'' m #align union_support_maximal_linear_independent_eq_range_basis union_support_maximal_linearIndependent_eq_range_basis @@ -471,7 +461,7 @@ variable {R : Type u} {M : Type v} variable [Ring R] [AddCommGroup M] [Module R M] -@[simp] +@[nontriviality, simp] theorem rank_subsingleton [Subsingleton R] : Module.rank R M = 1 := by haveI := Module.subsingleton R M have : Nonempty { s : Set M // LinearIndependent R ((↑) : s → M) } := @@ -513,7 +503,35 @@ theorem rank_zero_iff_forall_zero : Module.rank R M = 0 ↔ ∀ x : M, x = 0 := rw [← rank_top, this, rank_bot] #align rank_zero_iff_forall_zero rank_zero_iff_forall_zero -/-- See `rank_subsingleton` for the reason that `Nontrivial R` is needed. -/ +/-- See `rank_zero_iff` for a stronger version with `NoZeroSMulDivisor R M`. -/ +lemma rank_eq_zero_iff {R M} [Ring R] [AddCommGroup M] [Module R M] : + Module.rank R M = 0 ↔ ∀ x : M, ∃ a : R, a ≠ 0 ∧ a • x = 0 := by + nontriviality R + constructor + · contrapose! + rintro ⟨x, hx⟩ + rw [← Cardinal.one_le_iff_ne_zero] + have : LinearIndependent R (fun _ : Unit ↦ x) + · exact linearIndependent_iff.mpr (fun l hl ↦ Finsupp.unique_ext <| not_not.mp fun H ↦ + hx _ H ((Finsupp.total_unique _ _ _).symm.trans hl)) + simpa using cardinal_lift_le_rank_of_linearIndependent this + · intro h + rw [← le_zero_iff, Module.rank_def] + apply ciSup_le' + intro ⟨s, hs⟩ + rw [nonpos_iff_eq_zero, Cardinal.mk_eq_zero_iff, ← not_nonempty_iff] + rintro ⟨i : s⟩ + obtain ⟨a, ha, ha'⟩ := h i + apply ha + simpa using FunLike.congr_fun (linearIndependent_iff.mp hs (Finsupp.single i a) (by simpa)) i + +lemma rank_eq_zero_iff_isTorsion {R M} [CommRing R] [IsDomain R] [AddCommGroup M] [Module R M] : + Module.rank R M = 0 ↔ Module.IsTorsion R M := by + rw [Module.IsTorsion, rank_eq_zero_iff] + simp [mem_nonZeroDivisors_iff_ne_zero] + +/-- See `rank_subsingleton` for the reason that `Nontrivial R` is needed. +Also see `rank_eq_zero_iff` for the version without `NoZeroSMulDivisor R M`. -/ theorem rank_zero_iff : Module.rank R M = 0 ↔ Subsingleton M := rank_zero_iff_forall_zero.trans (subsingleton_iff_forall_eq 0).symm #align rank_zero_iff rank_zero_iff @@ -682,14 +700,15 @@ then any linearly independent family `v : ι → M` contained in the span of some finite `w : Set M`, is itself finite. -/ -def linearIndependentFintypeOfLeSpanFintype {ι : Type*} (v : ι → M) (i : LinearIndependent R v) - (w : Set M) [Fintype w] (s : range v ≤ span R w) : Fintype ι := - fintypeOfFinsetCardLe (Fintype.card w) fun t => by +lemma LinearIndependent.finite_of_le_span_finite {ι : Type*} (v : ι → M) (i : LinearIndependent R v) + (w : Set M) [Finite w] (s : range v ≤ span R w) : Finite ι := + letI := Fintype.ofFinite w + Fintype.finite <| fintypeOfFinsetCardLe (Fintype.card w) fun t => by let v' := fun x : (t : Set ι) => v x have i' : LinearIndependent R v' := i.comp _ Subtype.val_injective have s' : range v' ≤ span R w := (range_comp_subset_range _ _).trans s simpa using linearIndependent_le_span_aux' v' i' w s' -#align linear_independent_fintype_of_le_span_fintype linearIndependentFintypeOfLeSpanFintype +#align linear_independent_fintype_of_le_span_fintype LinearIndependent.finite_of_le_span_finite /-- If `R` satisfies the strong rank condition, then for any linearly independent family `v : ι → M` @@ -698,7 +717,8 @@ the cardinality of `ι` is bounded by the cardinality of `w`. -/ theorem linearIndependent_le_span' {ι : Type*} (v : ι → M) (i : LinearIndependent R v) (w : Set M) [Fintype w] (s : range v ≤ span R w) : #ι ≤ Fintype.card w := by - haveI : Fintype ι := linearIndependentFintypeOfLeSpanFintype v i w s + haveI : Finite ι := i.finite_of_le_span_finite v w s + letI := Fintype.ofFinite ι rw [Cardinal.mk_fintype] simp only [Cardinal.natCast_le] exact linearIndependent_le_span_aux' v i w s @@ -734,8 +754,8 @@ theorem linearIndependent_le_infinite_basis {ι : Type w} (b : Basis ι R M) [In obtain ⟨s, w : Infinite ↑(Φ ⁻¹' {s})⟩ := Cardinal.exists_infinite_fiber Φ h (by infer_instance) let v' := fun k : Φ ⁻¹' {s} => v k have i' : LinearIndependent R v' := i.comp _ Subtype.val_injective - have w' : Fintype (Φ ⁻¹' {s}) := by - apply linearIndependentFintypeOfLeSpanFintype v' i' (s.image b) + have w' : Finite (Φ ⁻¹' {s}) := by + apply i'.finite_of_le_span_finite v' (s.image b) rintro m ⟨⟨p, ⟨rfl⟩⟩, rfl⟩ simp only [SetLike.mem_coe, Subtype.coe_mk, Finset.coe_image] apply Basis.mem_span_repr_support @@ -880,10 +900,11 @@ theorem rank_span_set {s : Set M} (hs : LinearIndependent R (fun x => x : s → finite free module `M`. A property is true for all submodules of `M` if it satisfies the following "inductive step": the property is true for a submodule `N` if it's true for all submodules `N'` of `N` with the property that there exists `0 ≠ x ∈ N` such that the sum `N' + Rx` is direct. -/ -def Submodule.inductionOnRank [IsDomain R] [Fintype ι] (b : Basis ι R M) +def Submodule.inductionOnRank [IsDomain R] [Finite ι] (b : Basis ι R M) (P : Submodule R M → Sort*) (ih : ∀ N : Submodule R M, (∀ N' ≤ N, ∀ x ∈ N, (∀ (c : R), ∀ y ∈ N', c • x + y = (0 : M) → c = 0) → P N') → P N) (N : Submodule R M) : P N := + letI := Fintype.ofFinite ι Submodule.inductionOnRankAux b P ih (Fintype.card ι) N fun hs hli => by simpa using b.card_le_card_of_linearIndependent hli #align submodule.induction_on_rank Submodule.inductionOnRank @@ -1071,9 +1092,9 @@ end Free section DivisionRing -variable [DivisionRing K] +variable [DivisionRing K] [Ring R] [StrongRankCondition R] -variable [AddCommGroup V] [Module K V] +variable [AddCommGroup V] [Module K V] [Module R V] variable [AddCommGroup V'] [Module K V'] @@ -1086,19 +1107,25 @@ theorem Basis.finite_ofVectorSpaceIndex_of_rank_lt_aleph0 (h : Module.rank K V < #align basis.finite_of_vector_space_index_of_rank_lt_aleph_0 Basis.finite_ofVectorSpaceIndex_of_rank_lt_aleph0 -- TODO how far can we generalise this? --- When `s` is finite, we could prove this for any ring satisfying the strong rank condition --- using `linearIndependent_le_span'` theorem rank_span_le (s : Set V) : Module.rank K (span K s) ≤ #s := by obtain ⟨b, hb, hsab, hlib⟩ := exists_linearIndependent K s convert Cardinal.mk_le_mk_of_subset hb rw [← hsab, rank_span_set hlib] #align rank_span_le rank_span_le -theorem rank_span_of_finset (s : Finset V) : Module.rank K (span K (↑s : Set V)) < ℵ₀ := - calc - Module.rank K (span K (↑s : Set V)) ≤ #(↑s : Set V) := rank_span_le (s : Set V) - _ = s.card := by rw [Finset.coe_sort_coe, Cardinal.mk_coe_finset] - _ < ℵ₀ := Cardinal.nat_lt_aleph0 _ +theorem rank_span_le_of_finite {s : Set V} (hs : s.Finite) : Module.rank R (span R s) ≤ #s := by + rw [Module.rank_def] + apply ciSup_le' + rintro ⟨t, ht⟩ + letI := hs.fintype + simpa [Cardinal.mk_image_eq Subtype.val_injective] using linearIndependent_le_span' _ + (ht.map (f := Submodule.subtype _) (by simp)).image s (fun x ↦ by aesop) + +theorem rank_span_finset_le (s : Finset V) : Module.rank R (span R (s : Set V)) ≤ s.card := by + simpa using rank_span_le_of_finite s.finite_toSet + +theorem rank_span_of_finset (s : Finset V) : Module.rank R (span R (s : Set V)) < ℵ₀ := + (rank_span_finset_le s).trans_lt (Cardinal.nat_lt_aleph0 _) #align rank_span_of_finset rank_span_of_finset theorem rank_quotient_add_rank (p : Submodule K V) : diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 3d7b2b16f3151..9574160cc460d 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Fabian Glöckle, Kyle Miller -/ import Mathlib.LinearAlgebra.FiniteDimensional +import Mathlib.LinearAlgebra.FreeModule.Finite.Basic +import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition import Mathlib.LinearAlgebra.Projection import Mathlib.LinearAlgebra.SesquilinearForm -import Mathlib.RingTheory.Finiteness -import Mathlib.LinearAlgebra.FreeModule.Finite.Basic import Mathlib.RingTheory.TensorProduct #align_import linear_algebra.dual from "leanprover-community/mathlib"@"b1c017582e9f18d8494e5c18602a8cb4a6f843ac" @@ -229,17 +229,11 @@ theorem LinearMap.dualMap_injective_of_surjective {f : M₁ →ₗ[R] M₂} (hf #align linear_map.dual_map_injective_of_surjective LinearMap.dualMap_injective_of_surjective /-- The `Linear_equiv` version of `LinearMap.dualMap`. -/ -def LinearEquiv.dualMap (f : M₁ ≃ₗ[R] M₂) : Dual R M₂ ≃ₗ[R] Dual R M₁ := - { f.toLinearMap.dualMap with - invFun := f.symm.toLinearMap.dualMap - left_inv := by - intro φ; ext x - simp only [LinearMap.dualMap_apply, LinearEquiv.coe_toLinearMap, LinearMap.toFun_eq_coe, - LinearEquiv.apply_symm_apply] - right_inv := by - intro φ; ext x - simp only [LinearMap.dualMap_apply, LinearEquiv.coe_toLinearMap, LinearMap.toFun_eq_coe, - LinearEquiv.symm_apply_apply] } +def LinearEquiv.dualMap (f : M₁ ≃ₗ[R] M₂) : Dual R M₂ ≃ₗ[R] Dual R M₁ where + __ := f.toLinearMap.dualMap + invFun := f.symm.toLinearMap.dualMap + left_inv φ := LinearMap.ext fun x ↦ congr_arg φ (f.right_inv x) + right_inv φ := LinearMap.ext fun x ↦ congr_arg φ (f.left_inv x) #align linear_equiv.dual_map LinearEquiv.dualMap @[simp] @@ -347,12 +341,11 @@ theorem toDual_eq_equivFun [Fintype ι] (m : M) (i : ι) : b.toDual m (b i) = b. rw [b.equivFun_apply, toDual_eq_repr] #align basis.to_dual_eq_equiv_fun Basis.toDual_eq_equivFun -theorem toDual_inj (m : M) (a : b.toDual m = 0) : m = 0 := by - rw [← mem_bot R, ← b.repr.ker, mem_ker, LinearEquiv.coe_coe] - apply Finsupp.ext - intro b - rw [← toDual_eq_repr, a] - rfl +theorem toDual_injective : Injective b.toDual := fun x y h ↦ b.ext_elem_iff.mpr fun i ↦ by + simp_rw [← toDual_eq_repr]; exact FunLike.congr_fun h _ + +theorem toDual_inj (m : M) (a : b.toDual m = 0) : m = 0 := + b.toDual_injective (by rwa [_root_.map_zero]) #align basis.to_dual_inj Basis.toDual_inj -- Porting note: broken dot notation lean4#1910 LinearMap.ker @@ -362,10 +355,8 @@ theorem toDual_ker : LinearMap.ker b.toDual = ⊥ := -- Porting note: broken dot notation lean4#1910 LinearMap.range theorem toDual_range [Finite ι] : LinearMap.range b.toDual = ⊤ := by - cases nonempty_fintype ι refine' eq_top_iff'.2 fun f => _ - rw [LinearMap.mem_range] - let lin_comb : ι →₀ R := Finsupp.equivFunOnFinite.symm fun i => f.toFun (b i) + let lin_comb : ι →₀ R := Finsupp.equivFunOnFinite.symm fun i => f (b i) refine' ⟨Finsupp.total ι M R b lin_comb, b.ext fun i => _⟩ rw [b.toDual_eq_repr _ i, repr_total b] rfl @@ -499,15 +490,9 @@ theorem total_coord [CommRing R] [AddCommGroup M] [Module R M] [Finite ι] (b : rw [← coe_dualBasis, total_dualBasis] #align basis.total_coord Basis.total_coord --- Porting note: universes very dodgy in Cardinals... theorem dual_rank_eq [CommRing K] [AddCommGroup V] [Module K V] [Finite ι] (b : Basis ι K V) : Cardinal.lift.{uK,uV} (Module.rank K V) = Module.rank K (Dual K V) := by - classical - cases nonempty_fintype ι - have := LinearEquiv.lift_rank_eq b.toDualEquiv - rw [Cardinal.lift_umax.{uV,uK}] at this - rw [this, ← Cardinal.lift_umax] - apply Cardinal.lift_id + classical rw [← lift_umax.{uV,uK}, b.toDualEquiv.lift_rank_eq, lift_id'.{uV,uK}] #align basis.dual_rank_eq Basis.dual_rank_eq end Basis @@ -533,17 +518,15 @@ theorem eval_ker : LinearMap.ker (eval K V) = ⊥ := by theorem map_eval_injective : (Submodule.map (eval K V)).Injective := by apply Submodule.map_injective_of_injective rw [← LinearMap.ker_eq_bot] - apply eval_ker K V + exact eval_ker K V #align module.map_eval_injective Module.map_eval_injective --- elaborates faster than `exact` theorem comap_eval_surjective : (Submodule.comap (eval K V)).Surjective := by apply Submodule.comap_surjective_of_injective rw [← LinearMap.ker_eq_bot] - apply eval_ker K V + exact eval_ker K V #align module.comap_eval_surjective Module.comap_eval_surjective --- elaborates faster than `exact` end section @@ -585,6 +568,17 @@ theorem nontrivial_dual_iff : instance instNontrivialDual [Nontrivial V] : Nontrivial (Dual K V) := (nontrivial_dual_iff K).mpr inferInstance +theorem finite_dual_iff : Finite K (Dual K V) ↔ Finite K V := by + constructor <;> intro h + · obtain ⟨⟨ι, b⟩⟩ := Module.Free.exists_basis (R := K) (M := V) + nontriviality K + obtain ⟨⟨s, span_s⟩⟩ := h + classical + haveI := (b.linearIndependent.map' _ b.toDual_ker).finite_of_le_span_finite _ s ?_ + · exact Finite.of_basis b + · rw [span_s]; apply le_top + · infer_instance + end theorem dual_rank_eq [Module.Finite K V] : @@ -609,9 +603,9 @@ Any finitely-generated free module (and thus any finite-dimensional vector space See `Module.IsReflexive.of_finite_of_free`. -/ class IsReflexive : Prop where /-- A reflexive module is one for which the natural map to its double dual is a bijection. -/ - bijective_dual_eval' : Bijective $ Dual.eval R M + bijective_dual_eval' : Bijective (Dual.eval R M) -lemma bijective_dual_eval [IsReflexive R M] : Bijective $ Dual.eval R M := +lemma bijective_dual_eval [IsReflexive R M] : Bijective (Dual.eval R M) := IsReflexive.bijective_dual_eval' instance IsReflexive.of_finite_of_free [Finite R M] [Free R M] : IsReflexive R M where @@ -622,7 +616,7 @@ variable [IsReflexive R M] /-- The bijection between a reflexive module and its double dual, bundled as a `LinearEquiv`. -/ def evalEquiv : M ≃ₗ[R] Dual R (Dual R M) := - LinearEquiv.ofBijective _ $ bijective_dual_eval R M + LinearEquiv.ofBijective _ (bijective_dual_eval R M) #align module.eval_equiv Module.evalEquiv @[simp] lemma evalEquiv_toLinearMap : evalEquiv R M = Dual.eval R M := rfl @@ -630,7 +624,7 @@ def evalEquiv : M ≃ₗ[R] Dual R (Dual R M) := @[simp] lemma evalEquiv_apply (m : M) : evalEquiv R M m = Dual.eval R M m := rfl -@[simp] lemma apply_evalEquiv_symm_apply (f : Dual R M) (g : Dual R $ Dual R M) : +@[simp] lemma apply_evalEquiv_symm_apply (f : Dual R M) (g : Dual R (Dual R M)) : f ((evalEquiv R M).symm g) = g f := by set m := (evalEquiv R M).symm g rw [← (evalEquiv R M).apply_symm_apply g, evalEquiv_apply, Dual.eval_apply] @@ -1034,28 +1028,12 @@ theorem dualCoannihilator_top (W : Subspace K V) : @[simp] theorem dualAnnihilator_dualCoannihilator_eq {W : Subspace K V} : W.dualAnnihilator.dualCoannihilator = W := by - refine' le_antisymm _ (le_dualAnnihilator_dualCoannihilator _) - intro v + refine le_antisymm (fun v ↦ Function.mtr ?_) (le_dualAnnihilator_dualCoannihilator _) simp only [mem_dualAnnihilator, mem_dualCoannihilator] - contrapose! - intro hv - obtain ⟨W', hW⟩ := Submodule.exists_isCompl W - obtain ⟨⟨w, w'⟩, rfl, -⟩ := existsUnique_add_of_isCompl_prod hW v - have hw'n : (w' : V) ∉ W := by - contrapose! hv - exact Submodule.add_mem W w.2 hv - have hw'nz : w' ≠ 0 := by - rintro rfl - exact hw'n (Submodule.zero_mem W) - rw [Ne.def, ← Module.forall_dual_apply_eq_zero_iff K w'] at hw'nz - push_neg at hw'nz - obtain ⟨φ, hφ⟩ := hw'nz - exists ((LinearMap.ofIsComplProd hW).comp (LinearMap.inr _ _ _)) φ - simp only [coe_comp, coe_inr, Function.comp_apply, ofIsComplProd_apply, map_add, - ofIsCompl_left_apply, zero_apply, ofIsCompl_right_apply, zero_add, Ne.def] - refine' ⟨_, hφ⟩ - intro v hv - apply LinearMap.ofIsCompl_left_apply hW ⟨v, hv⟩ + rw [← Quotient.mk_eq_zero W, ← Module.forall_dual_apply_eq_zero_iff K] + push_neg + refine fun ⟨φ, hφ⟩ ↦ ⟨φ.comp W.mkQ, fun w hw ↦ ?_, hφ⟩ + rw [comp_apply, mkQ_apply, (Quotient.mk_eq_zero W).mpr hw, φ.map_zero] #align subspace.dual_annihilator_dual_coannihilator_eq Subspace.dualAnnihilator_dualCoannihilator_eq -- exact elaborates slowly @@ -1081,31 +1059,27 @@ theorem dualAnnihilator_le_dualAnnihilator_iff {W W' : Subspace K V} : #align subspace.dual_annihilator_le_dual_annihilator_iff Subspace.dualAnnihilator_le_dualAnnihilator_iff theorem dualAnnihilator_inj {W W' : Subspace K V} : - W.dualAnnihilator = W'.dualAnnihilator ↔ W = W' := by - constructor - · apply (dualAnnihilatorGci K V).l_injective - · rintro rfl - rfl + W.dualAnnihilator = W'.dualAnnihilator ↔ W = W' := + ⟨fun h ↦ (dualAnnihilatorGci K V).l_injective h, congr_arg _⟩ #align subspace.dual_annihilator_inj Subspace.dualAnnihilator_inj /-- Given a subspace `W` of `V` and an element of its dual `φ`, `dualLift W φ` is an arbitrary extension of `φ` to an element of the dual of `V`. That is, `dualLift W φ` sends `w ∈ W` to `φ x` and `x` in a chosen complement of `W` to `0`. -/ noncomputable def dualLift (W : Subspace K V) : Module.Dual K W →ₗ[K] Module.Dual K V := - let h := Classical.indefiniteDescription _ W.exists_isCompl - (LinearMap.ofIsComplProd h.2).comp (LinearMap.inl _ _ _) + (Classical.choose <| W.subtype.exists_leftInverse_of_injective W.ker_subtype).dualMap #align subspace.dual_lift Subspace.dualLift variable {W : Subspace K V} @[simp] -theorem dualLift_of_subtype {φ : Module.Dual K W} (w : W) : W.dualLift φ (w : V) = φ w := by - erw [ofIsCompl_left_apply _ w] - rfl +theorem dualLift_of_subtype {φ : Module.Dual K W} (w : W) : W.dualLift φ (w : V) = φ w := + congr_arg φ <| FunLike.congr_fun + (Classical.choose_spec <| W.subtype.exists_leftInverse_of_injective W.ker_subtype) w #align subspace.dual_lift_of_subtype Subspace.dualLift_of_subtype theorem dualLift_of_mem {φ : Module.Dual K W} {w : V} (hw : w ∈ W) : W.dualLift φ w = φ ⟨w, hw⟩ := - by convert dualLift_of_subtype ⟨w, hw⟩ + dualLift_of_subtype ⟨w, hw⟩ #align subspace.dual_lift_of_mem Subspace.dualLift_of_mem @[simp] @@ -1173,12 +1147,21 @@ open FiniteDimensional variable {V₁ : Type*} [AddCommGroup V₁] [Module K V₁] -instance instModuleDualFiniteDimensional [H : FiniteDimensional K V] : +instance instModuleDualFiniteDimensional [FiniteDimensional K V] : FiniteDimensional K (Module.Dual K V) := by infer_instance #align subspace.module.dual.finite_dimensional Subspace.instModuleDualFiniteDimensional -variable [FiniteDimensional K V] [FiniteDimensional K V₁] +@[simp] +theorem dual_finrank_eq : finrank K (Module.Dual K V) = finrank K V := by + by_cases h : Module.Finite K V + · classical exact LinearEquiv.finrank_eq (Basis.ofVectorSpace K V).toDualEquiv.symm + rw [finrank_eq_zero_of_basis_imp_false, finrank_eq_zero_of_basis_imp_false] + · exact fun _ b ↦ h (Module.Finite.of_basis b) + · exact fun _ b ↦ h ((Module.finite_dual_iff K).mp <| Module.Finite.of_basis b) +#align subspace.dual_finrank_eq Subspace.dual_finrank_eq + +variable [FiniteDimensional K V] theorem dualAnnihilator_dualAnnihilator_eq (W : Subspace K V) : W.dualAnnihilator.dualAnnihilator = Module.mapEvalEquiv K V W := by @@ -1187,13 +1170,6 @@ theorem dualAnnihilator_dualAnnihilator_eq (W : Subspace K V) : rwa [← OrderIso.symm_apply_eq] #align subspace.dual_annihilator_dual_annihilator_eq Subspace.dualAnnihilator_dualAnnihilator_eq --- TODO(kmill): https://github.com/leanprover-community/mathlib/pull/17521#discussion_r1083241963 -@[simp] -theorem dual_finrank_eq : finrank K (Module.Dual K V) = finrank K V := by - classical - exact LinearEquiv.finrank_eq (Basis.ofVectorSpace K V).toDualEquiv.symm -#align subspace.dual_finrank_eq Subspace.dual_finrank_eq - /-- The quotient by the dual is isomorphic to its dual annihilator. -/ -- Porting note: broken dot notation lean4#1910 LinearMap.range noncomputable def quotDualEquivAnnihilator (W : Subspace K V) : @@ -1250,15 +1226,10 @@ variable (f : M₁ →ₗ[R] M₂) -- Porting note: broken dot notation lean4#1910 LinearMap.ker theorem ker_dualMap_eq_dualAnnihilator_range : LinearMap.ker f.dualMap = f.range.dualAnnihilator := by - ext φ; constructor <;> intro hφ - · rw [mem_ker] at hφ - rw [Submodule.mem_dualAnnihilator] - rintro y ⟨x, rfl⟩ - rw [← dualMap_apply, hφ, zero_apply] - · ext x - rw [dualMap_apply] - rw [Submodule.mem_dualAnnihilator] at hφ - exact hφ (f x) ⟨x, rfl⟩ + ext + simp_rw [mem_ker, ext_iff, Submodule.mem_dualAnnihilator, + ← SetLike.mem_coe, range_coe, Set.forall_range_iff] + rfl #align linear_map.ker_dual_map_eq_dual_annihilator_range LinearMap.ker_dualMap_eq_dualAnnihilator_range -- Porting note: broken dot notation lean4#1910 LinearMap.range @@ -1332,7 +1303,7 @@ theorem range_dualMap_mkQ_eq (W : Submodule R M) : constructor · rintro ⟨ψ, rfl⟩ have := LinearMap.mem_range_self W.mkQ.dualMap ψ - simpa only [ker_mkQ] using LinearMap.range_dualMap_le_dualAnnihilator_ker W.mkQ this + simpa only [ker_mkQ] using W.mkQ.range_dualMap_le_dualAnnihilator_ker this · intro hφ exists W.dualCopairing ⟨φ, hφ⟩ #align submodule.range_dual_map_mkq_eq Submodule.range_dualMap_mkQ_eq @@ -1348,13 +1319,7 @@ def dualQuotEquivDualAnnihilator (W : Submodule R M) : (W.mkQ.dualMap.codRestrict W.dualAnnihilator fun φ => -- Porting note: broken dot notation lean4#1910 LinearMap.mem_range_self W.range_dualMap_mkQ_eq ▸ LinearMap.mem_range_self W.mkQ.dualMap φ) - W.dualCopairing - (by - ext - rfl) - (by - ext - rfl) + W.dualCopairing (by ext; rfl) (by ext; rfl) #align submodule.dual_quot_equiv_dual_annihilator Submodule.dualQuotEquivDualAnnihilator @[simp] @@ -1382,18 +1347,8 @@ open Submodule -- Porting note: broken dot notation lean4#1910 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 := by - rw [← f.ker.range_dualMap_mkQ_eq] - let f' := LinearMap.quotKerEquivOfSurjective f hf - trans LinearMap.range (f.dualMap.comp f'.symm.dualMap.toLinearMap) - · rw [LinearMap.range_comp_of_range_eq_top] - apply LinearEquiv.range - · apply congr_arg - ext φ x - simp only [LinearMap.coe_comp, LinearEquiv.coe_toLinearMap, LinearMap.dualMap_apply, - LinearEquiv.dualMap_apply, mkQ_apply, LinearMap.quotKerEquivOfSurjective, - LinearEquiv.trans_symm, LinearEquiv.trans_apply, LinearEquiv.ofTop_symm_apply, - LinearMap.quotKerEquivRange_symm_apply_image, mkQ_apply, Function.comp] + (hf : Function.Surjective f) : LinearMap.range f.dualMap = f.ker.dualAnnihilator := + ((f.quotKerEquivOfSurjective hf).dualMap.range_comp _).trans f.ker.range_dualMap_mkQ_eq #align linear_map.range_dual_map_eq_dual_annihilator_ker_of_surjective LinearMap.range_dualMap_eq_dualAnnihilator_ker_of_surjective -- Note, this can be specialized to the case where `R` is an injective `R`-module, or when @@ -1402,16 +1357,15 @@ theorem range_dualMap_eq_dualAnnihilator_ker_of_subtype_range_surjective (f : M (hf : Function.Surjective f.range.subtype.dualMap) : LinearMap.range f.dualMap = f.ker.dualAnnihilator := by have rr_surj : Function.Surjective f.rangeRestrict := by - rw [← LinearMap.range_eq_top, LinearMap.range_rangeRestrict] + rw [← range_eq_top, range_rangeRestrict] have := range_dualMap_eq_dualAnnihilator_ker_of_surjective f.rangeRestrict rr_surj convert this using 1 -- Porting note: broken dot notation lean4#1910 - · change LinearMap.range - ((Submodule.subtype <| LinearMap.range f).comp f.rangeRestrict).dualMap = _ - rw [← LinearMap.dualMap_comp_dualMap, LinearMap.range_comp_of_range_eq_top] - rwa [LinearMap.range_eq_top] + · change range ((range f).subtype.comp f.rangeRestrict).dualMap = _ + rw [← dualMap_comp_dualMap, range_comp_of_range_eq_top] + rwa [range_eq_top] · apply congr_arg - exact (LinearMap.ker_rangeRestrict f).symm + exact (ker_rangeRestrict f).symm #align linear_map.range_dual_map_eq_dual_annihilator_ker_of_subtype_range_surjective LinearMap.range_dualMap_eq_dualAnnihilator_ker_of_subtype_range_surjective end LinearMap @@ -1433,15 +1387,9 @@ theorem dualPairing_nondegenerate : (dualPairing K V₁).Nondegenerate := #align linear_map.dual_pairing_nondegenerate LinearMap.dualPairing_nondegenerate theorem dualMap_surjective_of_injective {f : V₁ →ₗ[K] V₂} (hf : Function.Injective f) : - Function.Surjective f.dualMap := by - intro φ - let f' := LinearEquiv.ofInjective f hf - use Subspace.dualLift (range f) (f'.symm.dualMap φ) - ext x - rw [LinearMap.dualMap_apply, Subspace.dualLift_of_mem (mem_range_self f x), - LinearEquiv.dualMap_apply] - congr 1 - exact LinearEquiv.symm_apply_apply f' x + Function.Surjective f.dualMap := fun φ ↦ + have ⟨f', hf'⟩ := f.exists_leftInverse_of_injective (ker_eq_bot.mpr hf) + ⟨φ.comp f', ext fun x ↦ congr(φ <| $hf' x)⟩ #align linear_map.dual_map_surjective_of_injective LinearMap.dualMap_surjective_of_injective -- Porting note: broken dot notation lean4#1910 LinearMap.range @@ -1455,8 +1403,8 @@ theorem range_dualMap_eq_dualAnnihilator_ker (f : V₁ →ₗ[K] V₂) : @[simp] theorem dualMap_surjective_iff {f : V₁ →ₗ[K] V₂} : Function.Surjective f.dualMap ↔ Function.Injective f := by - rw [← LinearMap.range_eq_top, range_dualMap_eq_dualAnnihilator_ker, ← - Submodule.dualAnnihilator_bot, Subspace.dualAnnihilator_inj, LinearMap.ker_eq_bot] + rw [← LinearMap.range_eq_top, range_dualMap_eq_dualAnnihilator_ker, + ← Submodule.dualAnnihilator_bot, Subspace.dualAnnihilator_inj, LinearMap.ker_eq_bot] #align linear_map.dual_map_surjective_iff LinearMap.dualMap_surjective_iff end LinearMap @@ -1523,7 +1471,7 @@ theorem dualAnnihilator_inf_eq (W W' : Subspace K V₁) : theorem dualAnnihilator_iInf_eq {ι : Type*} [Finite ι] (W : ι → Subspace K V₁) : (⨅ i : ι, W i).dualAnnihilator = ⨆ i : ι, (W i).dualAnnihilator := by revert ι - refine' @Finite.induction_empty_option _ _ _ _ + apply Finite.induction_empty_option · intro α β h hyp W rw [← h.iInf_comp, hyp _, ← h.iSup_comp] · intro W @@ -1555,41 +1503,27 @@ section FiniteDimensional open FiniteDimensional LinearMap -variable [FiniteDimensional K V₂] - namespace LinearMap --- TODO(kmill) remove finite_dimensional if possible --- see https://github.com/leanprover-community/mathlib/pull/17521#discussion_r1083242551 @[simp] theorem finrank_range_dualMap_eq_finrank_range (f : V₁ →ₗ[K] V₂) : -- Porting note: broken dot notation lean4#1910 finrank K (LinearMap.range f.dualMap) = finrank K (LinearMap.range f) := by - have that := Submodule.finrank_quotient_add_finrank (LinearMap.range f) - -- Porting note: Again LinearEquiv.finrank_eq needs help - let equiv := (Subspace.quotEquivAnnihilator <| LinearMap.range f) - have eq := LinearEquiv.finrank_eq (R := K) (M := (V₂ ⧸ range f)) - (M₂ := { x // x ∈ Submodule.dualAnnihilator (range f) }) equiv - rw [eq, ← ker_dualMap_eq_dualAnnihilator_range] at that - -- Porting note: cannot convert at `this`? - conv_rhs at that => rw [← Subspace.dual_finrank_eq] - refine' add_left_injective (finrank K <| LinearMap.ker f.dualMap) _ - change _ + _ = _ + _ - rw [finrank_range_add_finrank_ker f.dualMap, add_comm, that] + rw [congr_arg dualMap (show f = (range f).subtype.comp f.rangeRestrict by rfl), + ← dualMap_comp_dualMap, range_comp, + range_eq_top.mpr (dualMap_surjective_of_injective (range f).injective_subtype), + Submodule.map_top, finrank_range_of_inj, Subspace.dual_finrank_eq] + exact dualMap_injective_of_surjective (range_eq_top.mp f.range_rangeRestrict) #align linear_map.finrank_range_dual_map_eq_finrank_range LinearMap.finrank_range_dualMap_eq_finrank_range /-- `f.dualMap` is injective if and only if `f` is surjective -/ @[simp] theorem dualMap_injective_iff {f : V₁ →ₗ[K] V₂} : Function.Injective f.dualMap ↔ Function.Surjective f := by - refine' ⟨_, fun h => dualMap_injective_of_surjective h⟩ - rw [← range_eq_top, ← ker_eq_bot] - intro h - apply Submodule.eq_top_of_finrank_eq - rw [← finrank_eq_zero] at h - rw [← add_zero (FiniteDimensional.finrank K <| LinearMap.range f), ← h, ← - LinearMap.finrank_range_dualMap_eq_finrank_range, LinearMap.finrank_range_add_finrank_ker, - Subspace.dual_finrank_eq] + refine ⟨Function.mtr fun not_surj inj ↦ ?_, dualMap_injective_of_surjective⟩ + rw [← range_eq_top, ← Ne, ← lt_top_iff_ne_top] at not_surj + obtain ⟨φ, φ0, range_le_ker⟩ := (range f).exists_le_ker_of_lt_top not_surj + exact φ0 (inj <| ext fun x ↦ range_le_ker ⟨x, rfl⟩) #align linear_map.dual_map_injective_iff LinearMap.dualMap_injective_iff /-- `f.dualMap` is bijective if and only if `f` is -/ @@ -1599,6 +1533,28 @@ theorem dualMap_bijective_iff {f : V₁ →ₗ[K] V₂} : simp_rw [Function.Bijective, dualMap_surjective_iff, dualMap_injective_iff, and_comm] #align linear_map.dual_map_bijective_iff LinearMap.dualMap_bijective_iff +variable {B : V₁ →ₗ[K] V₂ →ₗ[K] K} + +open Function + +theorem flip_injective_iff₁ [FiniteDimensional K V₁] : Injective B.flip ↔ Surjective B := by + rw [← dualMap_surjective_iff, ← (evalEquiv K V₁).toEquiv.surjective_comp]; rfl + +theorem flip_injective_iff₂ [FiniteDimensional K V₂] : Injective B.flip ↔ Surjective B := by + rw [← dualMap_injective_iff]; exact (evalEquiv K V₂).toEquiv.injective_comp B.dualMap + +theorem flip_surjective_iff₁ [FiniteDimensional K V₁] : Surjective B.flip ↔ Injective B := + flip_injective_iff₂.symm + +theorem flip_surjective_iff₂ [FiniteDimensional K V₂] : Surjective B.flip ↔ Injective B := + flip_injective_iff₁.symm + +theorem flip_bijective_iff₁ [FiniteDimensional K V₁] : Bijective B.flip ↔ Bijective B := by + simp_rw [Bijective, flip_injective_iff₁, flip_surjective_iff₁, and_comm] + +theorem flip_bijective_iff₂ [FiniteDimensional K V₂] : Bijective B.flip ↔ Bijective B := + flip_bijective_iff₁.symm + end LinearMap end FiniteDimensional diff --git a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean index 8fcfe34337fcd..ae3280cadf923 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean @@ -68,7 +68,7 @@ theorem iSup_generalizedEigenspace_eq_top [IsAlgClosed K] [FiniteDimensional K V cases' n with n -- If the vector space is 0-dimensional, the result is trivial. · rw [← top_le_iff] - simp only [finrank_eq_zero.1 (Eq.trans (finrank_top _ _) h_dim), bot_le] + simp only [Submodule.finrank_eq_zero.1 (Eq.trans (finrank_top _ _) h_dim), bot_le] -- Otherwise the vector space is nontrivial. · haveI : Nontrivial V := finrank_pos_iff.1 (by rw [h_dim]; apply Nat.zero_lt_succ) -- Hence, `f` has an eigenvalue `μ₀`. diff --git a/Mathlib/LinearAlgebra/FiniteDimensional.lean b/Mathlib/LinearAlgebra/FiniteDimensional.lean index 5547226bed321..098d690e09217 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional.lean @@ -67,13 +67,8 @@ in `Mathlib.LinearAlgebra.Dimension`. Not all results have been ported yet. You should not assume that there has been any effort to state lemmas as generally as possible. -One of the characterizations of finite-dimensionality is in terms of finite generation. This -property is currently defined only for submodules, so we express it through the fact that the -maximal submodule (which, as a set, coincides with the whole space) is finitely generated. This is -not very convenient to use, although there are some helper functions. However, this becomes very -convenient when speaking of submodules which are finite-dimensional, as this notion coincides with -the fact that the submodule is finitely generated (as a submodule of the whole space). This -equivalence is proved in `Submodule.fg_iff_finiteDimensional`. +Plenty of the results hold for general fg modules or notherian modules, and they can be found in +`Mathlib.LinearAlgebra.FreeModule.Finite.Rank` and `Mathlib.RingTheory.Noetherian`. -/ @@ -115,13 +110,12 @@ theorem of_surjective (f : V →ₗ[K] V₂) (w : Function.Surjective f) [Finite variable (K V) instance finiteDimensional_pi {ι : Type*} [Finite ι] : FiniteDimensional K (ι → K) := - iff_fg.1 isNoetherian_pi + Finite.pi #align finite_dimensional.finite_dimensional_pi FiniteDimensional.finiteDimensional_pi instance finiteDimensional_pi' {ι : Type*} [Finite ι] (M : ι → Type*) [∀ i, AddCommGroup (M i)] - [∀ i, Module K (M i)] [I : ∀ i, FiniteDimensional K (M i)] : FiniteDimensional K (∀ i, M i) := - haveI : ∀ i : ι, IsNoetherian K (M i) := fun i => iff_fg.2 (I i) - iff_fg.1 isNoetherian_pi + [∀ i, Module K (M i)] [∀ i, FiniteDimensional K (M i)] : FiniteDimensional K (∀ i, M i) := + Finite.pi #align finite_dimensional.finite_dimensional_pi' FiniteDimensional.finiteDimensional_pi' /-- A finite dimensional vector space over a finite field is finite -/ @@ -138,19 +132,14 @@ theorem finite_of_finite [Finite K] [FiniteDimensional K V] : Finite V := by variable {K V} /-- If a vector space has a finite basis, then it is finite-dimensional. -/ -theorem of_fintype_basis {ι : Type w} [Finite ι] (h : Basis ι K V) : FiniteDimensional K V := by - classical - cases nonempty_fintype ι - exact ⟨⟨Finset.univ.image h, by - convert h.span_eq - simp⟩⟩ +theorem of_fintype_basis {ι : Type w} [Finite ι] (h : Basis ι K V) : FiniteDimensional K V := + Module.Finite.of_basis h #align finite_dimensional.of_fintype_basis FiniteDimensional.of_fintype_basis /-- If a vector space is `FiniteDimensional`, all bases are indexed by a finite type -/ noncomputable def fintypeBasisIndex {ι : Type*} [FiniteDimensional K V] (b : Basis ι K V) : Fintype ι := - letI : IsNoetherian K V := IsNoetherian.iff_fg.2 inferInstance - IsNoetherian.fintypeBasisIndex b + @Fintype.ofFinite _ (Module.Free.finite_basis b) #align finite_dimensional.fintype_basis_index FiniteDimensional.fintypeBasisIndex /-- If a vector space is `FiniteDimensional`, `Basis.ofVectorSpace` is indexed by @@ -181,7 +170,7 @@ instance finiteDimensional_submodule [FiniteDimensional K V] (S : Submodule K V) /-- A quotient of a finite-dimensional space is also finite-dimensional. -/ instance finiteDimensional_quotient [FiniteDimensional K V] (S : Submodule K V) : FiniteDimensional K (V ⧸ S) := - Module.Finite.of_surjective (Submodule.mkQ S) <| surjective_quot_mk _ + Module.Finite.quotient K S #align finite_dimensional.finite_dimensional_quotient FiniteDimensional.finiteDimensional_quotient variable (K V) @@ -195,152 +184,43 @@ theorem finrank_eq_rank' [FiniteDimensional K V] : (finrank K V : Cardinal.{v}) variable {K V} theorem finrank_of_infinite_dimensional (h : ¬FiniteDimensional K V) : finrank K V = 0 := - dif_neg <| mt IsNoetherian.iff_rank_lt_aleph0.2 <| (not_iff_not.2 iff_fg).2 h + FiniteDimensional.finrank_of_not_finite h #align finite_dimensional.finrank_of_infinite_dimensional FiniteDimensional.finrank_of_infinite_dimensional -theorem finiteDimensional_of_finrank (h : 0 < finrank K V) : FiniteDimensional K V := by - contrapose h - simp [finrank_of_infinite_dimensional h] +theorem finiteDimensional_of_finrank (h : 0 < finrank K V) : FiniteDimensional K V := + Module.finite_of_finrank_pos h #align finite_dimensional.finite_dimensional_of_finrank FiniteDimensional.finiteDimensional_of_finrank theorem finiteDimensional_of_finrank_eq_succ {n : ℕ} (hn : finrank K V = n.succ) : FiniteDimensional K V := - finiteDimensional_of_finrank <| by rw [hn]; exact n.succ_pos + Module.finite_of_finrank_eq_succ hn #align finite_dimensional.finite_dimensional_of_finrank_eq_succ FiniteDimensional.finiteDimensional_of_finrank_eq_succ /-- We can infer `FiniteDimensional K V` in the presence of `[Fact (finrank K V = n + 1)]`. Declare this as a local instance where needed. -/ -theorem fact_finiteDimensional_of_finrank_eq_succ (n : ℕ) [Fact (finrank K V = n + 1)] : +theorem fact_finiteDimensional_of_finrank_eq_succ (n : ℕ) [hn : Fact (finrank K V = n + 1)] : FiniteDimensional K V := - finiteDimensional_of_finrank <| by convert Nat.succ_pos n; apply Fact.out + finiteDimensional_of_finrank_eq_succ hn.out #align finite_dimensional.fact_finite_dimensional_of_finrank_eq_succ FiniteDimensional.fact_finiteDimensional_of_finrank_eq_succ theorem finiteDimensional_iff_of_rank_eq_nsmul {W} [AddCommGroup W] [Module K W] {n : ℕ} (hn : n ≠ 0) (hVW : Module.rank K V = n • Module.rank K W) : - FiniteDimensional K V ↔ FiniteDimensional K W := by - simp only [FiniteDimensional, ← IsNoetherian.iff_fg, IsNoetherian.iff_rank_lt_aleph0, hVW, - Cardinal.nsmul_lt_aleph0_iff_of_ne_zero hn] + FiniteDimensional K V ↔ FiniteDimensional K W := + Module.finite_iff_of_rank_eq_nsmul hn hVW #align finite_dimensional.finite_dimensional_iff_of_rank_eq_nsmul FiniteDimensional.finiteDimensional_iff_of_rank_eq_nsmul /-- If a vector space is finite-dimensional, then the cardinality of any basis is equal to its `finrank`. -/ theorem finrank_eq_card_basis' [FiniteDimensional K V] {ι : Type w} (h : Basis ι K V) : - (finrank K V : Cardinal.{w}) = #ι := by - haveI : IsNoetherian K V := iff_fg.2 inferInstance - haveI : Fintype ι := fintypeBasisIndex h - rw [Cardinal.mk_fintype, finrank_eq_card_basis h] + (finrank K V : Cardinal.{w}) = #ι := + Module.mk_finrank_eq_card_basis h #align finite_dimensional.finrank_eq_card_basis' FiniteDimensional.finrank_eq_card_basis' -/-- Given a basis of a division ring over itself indexed by a type `ι`, then `ι` is `Unique`. -/ -noncomputable def _root_.Basis.unique {ι : Type*} (b : Basis ι K K) : Unique ι := by - have A : Cardinal.mk ι = ↑(FiniteDimensional.finrank K K) := - (FiniteDimensional.finrank_eq_card_basis' b).symm - -- porting note: replace `algebraMap.coe_one` with `Nat.cast_one` - simp only [Cardinal.eq_one_iff_unique, FiniteDimensional.finrank_self, Nat.cast_one] at A - exact Nonempty.some ((unique_iff_subsingleton_and_nonempty _).2 A) -#align basis.unique Basis.unique - -variable (K V) - -/-- A finite dimensional vector space has a basis indexed by `Fin (finrank K V)`. -/ -noncomputable def finBasis [FiniteDimensional K V] : Basis (Fin (finrank K V)) K V := - have h : Fintype.card (@finsetBasisIndex K V _ _ _ (iff_fg.2 inferInstance)) = finrank K V := - (finrank_eq_card_basis (@finsetBasis K V _ _ _ (iff_fg.2 inferInstance))).symm - (@finsetBasis K V _ _ _ (iff_fg.2 inferInstance)).reindex (Fintype.equivFinOfCardEq h) -#align finite_dimensional.fin_basis FiniteDimensional.finBasis - -/-- An `n`-dimensional vector space has a basis indexed by `Fin n`. -/ -noncomputable def finBasisOfFinrankEq [FiniteDimensional K V] {n : ℕ} (hn : finrank K V = n) : - Basis (Fin n) K V := - (finBasis K V).reindex (Fin.castIso hn).toEquiv -#align finite_dimensional.fin_basis_of_finrank_eq FiniteDimensional.finBasisOfFinrankEq - -variable {K V} - -/-- A module with dimension 1 has a basis with one element. -/ -noncomputable def basisUnique (ι : Type*) [Unique ι] (h : finrank K V = 1) : Basis ι K V := - haveI : FiniteDimensional _ _ := - finiteDimensional_of_finrank (_root_.zero_lt_one.trans_le h.symm.le) - (finBasisOfFinrankEq K V h).reindex (Equiv.equivOfUnique _ _) -#align finite_dimensional.basis_unique FiniteDimensional.basisUnique - -@[simp] -theorem basisUnique.repr_eq_zero_iff {ι : Type*} [Unique ι] {h : finrank K V = 1} {v : V} {i : ι} : - (basisUnique ι h).repr v i = 0 ↔ v = 0 := - ⟨fun hv => - (basisUnique ι h).repr.map_eq_zero_iff.mp (Finsupp.ext fun j => Subsingleton.elim i j ▸ hv), - fun hv => by rw [hv, LinearEquiv.map_zero, Finsupp.zero_apply]⟩ -#align finite_dimensional.basis_unique.repr_eq_zero_iff FiniteDimensional.basisUnique.repr_eq_zero_iff - -theorem cardinal_mk_le_finrank_of_linearIndependent [FiniteDimensional K V] {ι : Type w} {b : ι → V} - (h : LinearIndependent K b) : #ι ≤ finrank K V := by - rw [← lift_le.{max v w}] - simpa [← finrank_eq_rank', -finrank_eq_rank] using - cardinal_lift_le_rank_of_linearIndependent h -#align finite_dimensional.cardinal_mk_le_finrank_of_linear_independent FiniteDimensional.cardinal_mk_le_finrank_of_linearIndependent - -theorem fintype_card_le_finrank_of_linearIndependent [FiniteDimensional K V] {ι : Type*} - [Fintype ι] {b : ι → V} (h : LinearIndependent K b) : Fintype.card ι ≤ finrank K V := by - simpa using cardinal_mk_le_finrank_of_linearIndependent h -#align finite_dimensional.fintype_card_le_finrank_of_linear_independent FiniteDimensional.fintype_card_le_finrank_of_linearIndependent - -theorem finset_card_le_finrank_of_linearIndependent [FiniteDimensional K V] {b : Finset V} - (h : LinearIndependent K (fun x => x : b → V)) : b.card ≤ finrank K V := by - rw [← Fintype.card_coe] - exact fintype_card_le_finrank_of_linearIndependent h -#align finite_dimensional.finset_card_le_finrank_of_linear_independent FiniteDimensional.finset_card_le_finrank_of_linearIndependent - theorem lt_aleph0_of_linearIndependent {ι : Type w} [FiniteDimensional K V] {v : ι → V} - (h : LinearIndependent K v) : #ι < ℵ₀ := by - apply Cardinal.lift_lt.1 - apply lt_of_le_of_lt - apply cardinal_lift_le_rank_of_linearIndependent h - rw [← finrank_eq_rank, Cardinal.lift_aleph0, Cardinal.lift_natCast] - apply Cardinal.nat_lt_aleph0 + (h : LinearIndependent K v) : #ι < ℵ₀ := + Module.Finite.lt_aleph0_of_linearIndependent h #align finite_dimensional.lt_aleph_0_of_linear_independent FiniteDimensional.lt_aleph0_of_linearIndependent -theorem _root_.LinearIndependent.finite [FiniteDimensional K V] {b : Set V} - (h : LinearIndependent K fun x : b => (x : V)) : b.Finite := - Cardinal.lt_aleph0_iff_set_finite.mp (FiniteDimensional.lt_aleph0_of_linearIndependent h) -#align linear_independent.finite LinearIndependent.finite - -theorem not_linearIndependent_of_infinite {ι : Type w} [inf : Infinite ι] [FiniteDimensional K V] - (v : ι → V) : ¬LinearIndependent K v := by - intro h_lin_indep - have : ¬ℵ₀ ≤ #ι := not_le.mpr (lt_aleph0_of_linearIndependent h_lin_indep) - have : ℵ₀ ≤ #ι := infinite_iff.mp inf - contradiction -#align finite_dimensional.not_linear_independent_of_infinite FiniteDimensional.not_linearIndependent_of_infinite - -/-- A finite dimensional space has positive `finrank` iff it has a nonzero element. -/ -theorem finrank_pos_iff_exists_ne_zero [FiniteDimensional K V] : 0 < finrank K V ↔ ∃ x : V, x ≠ 0 := - Iff.trans - (by - rw [← finrank_eq_rank] - norm_cast) - (@rank_pos_iff_exists_ne_zero K V _ _ _ _ _) -#align finite_dimensional.finrank_pos_iff_exists_ne_zero FiniteDimensional.finrank_pos_iff_exists_ne_zero - -/-- A finite dimensional space has positive `finrank` iff it is nontrivial. -/ -theorem finrank_pos_iff [FiniteDimensional K V] : 0 < finrank K V ↔ Nontrivial V := - Iff.trans - (by rw [← finrank_eq_rank]; norm_cast) - (rank_pos_iff_nontrivial (R := K)) -#align finite_dimensional.finrank_pos_iff FiniteDimensional.finrank_pos_iff - -/-- A nontrivial finite dimensional space has positive `finrank`. -/ -theorem finrank_pos [FiniteDimensional K V] [h : Nontrivial V] : 0 < finrank K V := - finrank_pos_iff.mpr h -#align finite_dimensional.finrank_pos FiniteDimensional.finrank_pos - -/-- A finite dimensional space has zero `finrank` iff it is a subsingleton. -This is the `finrank` version of `rank_zero_iff`. -/ -theorem finrank_zero_iff [FiniteDimensional K V] : finrank K V = 0 ↔ Subsingleton V := - Iff.trans - (by rw [← finrank_eq_rank]; norm_cast) - (rank_zero_iff (R := K)) -#align finite_dimensional.finrank_zero_iff FiniteDimensional.finrank_zero_iff - /-- If a submodule has maximal dimension in a finite dimensional space, then it is equal to the whole space. -/ theorem _root_.Submodule.eq_top_of_finrank_eq [FiniteDimensional K V] {S : Submodule K V} @@ -378,17 +258,17 @@ instance finiteDimensional_self : FiniteDimensional K K := by infer_instance /-- The submodule generated by a finite set is finite-dimensional. -/ theorem span_of_finite {A : Set V} (hA : Set.Finite A) : FiniteDimensional K (Submodule.span K A) := - iff_fg.1 <| isNoetherian_span_of_finite K hA + Module.Finite.span_of_finite K hA #align finite_dimensional.span_of_finite FiniteDimensional.span_of_finite /-- The submodule generated by a single element is finite-dimensional. -/ instance span_singleton (x : V) : FiniteDimensional K (K ∙ x) := - span_of_finite K <| Set.finite_singleton _ + Module.Finite.span_singleton K x #align finite_dimensional.span_singleton FiniteDimensional.span_singleton /-- The submodule generated by a finset is finite-dimensional. -/ instance span_finset (s : Finset V) : FiniteDimensional K (span K (s : Set V)) := - span_of_finite K <| s.finite_toSet + Module.Finite.span_finset K s #align finite_dimensional.span_finset FiniteDimensional.span_finset /-- Pushforwards of finite-dimensional submodules are finite-dimensional. -/ @@ -398,172 +278,12 @@ instance (f : V →ₗ[K] V₂) (p : Submodule K V) [FiniteDimensional K p] : variable {K} -theorem _root_.CompleteLattice.Independent.subtype_ne_bot_le_finrank_aux [FiniteDimensional K V] - {ι : Type w} {p : ι → Submodule K V} (hp : CompleteLattice.Independent p) : - #{ i // p i ≠ ⊥ } ≤ (finrank K V : Cardinal.{w}) := by - suffices Cardinal.lift.{v} #{ i // p i ≠ ⊥ } ≤ Cardinal.lift.{v} (finrank K V : Cardinal.{w}) by - rwa [Cardinal.lift_le] at this - calc - Cardinal.lift.{v} #{ i // p i ≠ ⊥ } ≤ Cardinal.lift.{w} (Module.rank K V) := - hp.subtype_ne_bot_le_rank - _ = Cardinal.lift.{w} (finrank K V : Cardinal.{v}) := by rw [finrank_eq_rank] - _ = Cardinal.lift.{v} (finrank K V : Cardinal.{w}) := by simp -#align complete_lattice.independent.subtype_ne_bot_le_finrank_aux CompleteLattice.Independent.subtype_ne_bot_le_finrank_aux - -/-- If `p` is an independent family of subspaces of a finite-dimensional space `V`, then the -number of nontrivial subspaces in the family `p` is finite. -/ -noncomputable def _root_.CompleteLattice.Independent.fintypeNeBotOfFiniteDimensional - [FiniteDimensional K V] {ι : Type w} {p : ι → Submodule K V} - (hp : CompleteLattice.Independent p) : Fintype { i : ι // p i ≠ ⊥ } := by - suffices #{ i // p i ≠ ⊥ } < (ℵ₀ : Cardinal.{w}) by - rw [Cardinal.lt_aleph0_iff_fintype] at this - exact this.some - refine' lt_of_le_of_lt hp.subtype_ne_bot_le_finrank_aux _ - simp [Cardinal.nat_lt_aleph0] -#align complete_lattice.independent.fintype_ne_bot_of_finite_dimensional CompleteLattice.Independent.fintypeNeBotOfFiniteDimensional - -/-- If `p` is an independent family of subspaces of a finite-dimensional space `V`, then the -number of nontrivial subspaces in the family `p` is bounded above by the dimension of `V`. - -Note that the `Fintype` hypothesis required here can be provided by -`CompleteLattice.Independent.fintypeNeBotOfFiniteDimensional`. -/ -theorem _root_.CompleteLattice.Independent.subtype_ne_bot_le_finrank [FiniteDimensional K V] - {ι : Type w} {p : ι → Submodule K V} (hp : CompleteLattice.Independent p) - [Fintype { i // p i ≠ ⊥ }] : - Fintype.card { i // p i ≠ ⊥ } ≤ finrank K V := by simpa using hp.subtype_ne_bot_le_finrank_aux -#align complete_lattice.independent.subtype_ne_bot_le_finrank CompleteLattice.Independent.subtype_ne_bot_le_finrank - section open BigOperators open Finset -/-- If a finset has cardinality larger than the dimension of the space, -then there is a nontrivial linear relation amongst its elements. --/ -theorem exists_nontrivial_relation_of_rank_lt_card [FiniteDimensional K V] {t : Finset V} - (h : finrank K V < t.card) : ∃ f : V → K, ∑ e in t, f e • e = 0 ∧ ∃ x ∈ t, f x ≠ 0 := by - classical - have := mt finset_card_le_finrank_of_linearIndependent (by simpa using h) - rw [not_linearIndependent_iff] at this - obtain ⟨s, g, sum, z, zm, nonzero⟩ := this - -- Now we have to extend `g` to all of `t`, then to all of `V`. - let f : V → K := fun x => if h : x ∈ t then if (⟨x, h⟩ : t) ∈ s then g ⟨x, h⟩ else 0 else 0 - -- and finally clean up the mess caused by the extension. - refine' ⟨f, _, _⟩ - · dsimp - rw [← (sum)] -- porting note: need parens to disambiguate - fapply sum_bij_ne_zero fun v hvt _ => (⟨v, hvt⟩ : { v // v ∈ t }) - · intro v hvt H - dsimp - rw [dif_pos hvt] at H - contrapose! H - rw [if_neg H, zero_smul] - · intro _ _ _ _ _ _ - exact Subtype.mk.inj - · intro b hbs hb - use b - simpa only [hbs, exists_prop, dif_pos, Finset.mk_coe, and_true_iff, if_true, Finset.coe_mem, - eq_self_iff_true, exists_prop_of_true, Ne.def] using hb - · intro a h₁ - dsimp - rw [dif_pos h₁] - intro h₂ - rw [if_pos] - contrapose! h₂ - rw [if_neg h₂, zero_smul] - · refine' ⟨z, z.2, _⟩ - dsimp only - erw [dif_pos z.2, if_pos] <;> rwa [Subtype.coe_eta] -#align finite_dimensional.exists_nontrivial_relation_of_rank_lt_card FiniteDimensional.exists_nontrivial_relation_of_rank_lt_card - -/-- If a finset has cardinality larger than `finrank + 1`, -then there is a nontrivial linear relation amongst its elements, -such that the coefficients of the relation sum to zero. --/ -theorem exists_nontrivial_relation_sum_zero_of_rank_succ_lt_card [FiniteDimensional K V] - {t : Finset V} (h : finrank K V + 1 < t.card) : - ∃ f : V → K, ∑ e in t, f e • e = 0 ∧ ∑ e in t, f e = 0 ∧ ∃ x ∈ t, f x ≠ 0 := by - classical - -- Pick an element x₀ ∈ t, - have card_pos : 0 < t.card := lt_trans (Nat.succ_pos _) h - obtain ⟨x₀, m⟩ := (Finset.card_pos.1 card_pos).bex - -- and apply the previous lemma to the {xᵢ - x₀} - let shift : V ↪ V := ⟨fun x => x - x₀, sub_left_injective⟩ - let t' := (t.erase x₀).map shift - have h' : finrank K V < t'.card := by - simp only [card_map, Finset.card_erase_of_mem m] - exact Nat.lt_pred_iff.mpr h - -- to obtain a function `g`. - obtain ⟨g, gsum, x₁, x₁_mem, nz⟩ := exists_nontrivial_relation_of_rank_lt_card h' - -- Then obtain `f` by translating back by `x₀`, - -- and setting the value of `f` at `x₀` to ensure `∑ e in t, f e = 0`. - let f : V → K := fun z => if z = x₀ then -∑ z in t.erase x₀, g (z - x₀) else g (z - x₀) - refine' ⟨f, _, _, _⟩ - -- After this, it's a matter of verifying the properties, - -- based on the corresponding properties for `g`. - · show (∑ e : V in t, f e • e) = 0 - -- We prove this by splitting off the `x₀` term of the sum, - -- which is itself a sum over `t.erase x₀`, - -- combining the two sums, and - -- observing that after reindexing we have exactly - -- ∑ (x : V) in t', g x • x = 0. - simp only - conv_lhs => - apply_congr - rfl - rw [ite_smul] - rw [Finset.sum_ite] - conv => - congr - congr - apply_congr - -- Porting note: the next two steps used to work by `simp [filter_eq', m]` - erw [filter_eq'] - simp [m] - conv => - congr - congr - rfl - apply_congr - simp [filter_ne'] - rw [sum_singleton, neg_smul, add_comm, ← sub_eq_add_neg, sum_smul, ← sum_sub_distrib] - simp only [← smul_sub] - -- At the end we have to reindex the sum, so we use `change` to - -- express the summand using `shift`. - change (∑ x : V in t.erase x₀, (fun e => g e • e) (shift x)) = 0 - -- porting note: last argument can't be inferred - rw [← sum_map _ shift (fun e => g e • e)] - exact gsum - · show (∑ e : V in t, f e) = 0 - -- Again we split off the `x₀` term, - -- observing that it exactly cancels the other terms. - rw [← insert_erase m, sum_insert (not_mem_erase x₀ t)] - dsimp - rw [if_pos rfl] - conv_lhs => - congr - rfl - apply_congr - rfl - rw [if_neg (show _ ≠ x₀ from (mem_erase.mp ‹_›).1)] - exact neg_add_self _ - · show ∃ (x : V), x ∈ t ∧ f x ≠ 0 - -- We can use x₁ + x₀. - refine' ⟨x₁ + x₀, _, _⟩ - · rw [Finset.mem_map] at x₁_mem - rcases x₁_mem with ⟨x₁, x₁_mem, rfl⟩ - rw [mem_erase] at x₁_mem - simp only [x₁_mem, sub_add_cancel, Function.Embedding.coeFn_mk] - · dsimp only - rwa [if_neg, add_sub_cancel] - rw [add_left_eq_self] - rintro rfl - simp only [sub_eq_zero, exists_prop, Finset.mem_map, Embedding.coeFn_mk, eq_self_iff_true, - mem_erase, not_true, exists_eq_right, Ne.def, false_and_iff] at x₁_mem -#align finite_dimensional.exists_nontrivial_relation_sum_zero_of_rank_succ_lt_card FiniteDimensional.exists_nontrivial_relation_sum_zero_of_rank_succ_lt_card - section variable {L : Type*} [LinearOrderedField L] @@ -574,12 +294,14 @@ variable {W : Type v} [AddCommGroup W] [Module L W] available when working over an ordered field: we can ensure a positive coefficient, not just a nonzero coefficient. -/ -theorem exists_relation_sum_zero_pos_coefficient_of_rank_succ_lt_card [FiniteDimensional L W] +theorem exists_relation_sum_zero_pos_coefficient_of_finrank_succ_lt_card [FiniteDimensional L W] {t : Finset W} (h : finrank L W + 1 < t.card) : ∃ f : W → L, ∑ e in t, f e • e = 0 ∧ ∑ e in t, f e = 0 ∧ ∃ x ∈ t, 0 < f x := by - obtain ⟨f, sum, total, nonzero⟩ := exists_nontrivial_relation_sum_zero_of_rank_succ_lt_card h + obtain ⟨f, sum, total, nonzero⟩ := + Module.exists_nontrivial_relation_sum_zero_of_finrank_succ_lt_card h exact ⟨f, sum, total, exists_pos_of_sum_zero_of_exists_nonzero f total nonzero⟩ -#align finite_dimensional.exists_relation_sum_zero_pos_coefficient_of_rank_succ_lt_card FiniteDimensional.exists_relation_sum_zero_pos_coefficient_of_rank_succ_lt_card +#align finite_dimensional.exists_relation_sum_zero_pos_coefficient_of_rank_succ_lt_card FiniteDimensional.exists_relation_sum_zero_pos_coefficient_of_finrank_succ_lt_card + end @@ -589,8 +311,8 @@ end @[simps repr_apply] noncomputable def basisSingleton (ι : Type*) [Unique ι] (h : finrank K V = 1) (v : V) (hv : v ≠ 0) : Basis ι K V := - let b := basisUnique ι h - let h : b.repr v default ≠ 0 := mt basisUnique.repr_eq_zero_iff.mp hv + let b := FiniteDimensional.basisUnique ι h + let h : b.repr v default ≠ 0 := mt FiniteDimensional.basisUnique_repr_eq_zero_iff.mp hv Basis.ofRepr { toFun := fun w => Finsupp.single default (b.repr w default / b.repr v default) invFun := fun f => f default • v @@ -632,26 +354,19 @@ variable [DivisionRing K] [AddCommGroup V] [Module K V] open FiniteDimensional theorem finiteDimensional_of_rank_eq_nat {n : ℕ} (h : Module.rank K V = n) : - FiniteDimensional K V := by - rw [FiniteDimensional, ← IsNoetherian.iff_fg, IsNoetherian.iff_rank_lt_aleph0, h] - exact nat_lt_aleph0 n + FiniteDimensional K V := + Module.finite_of_rank_eq_nat h #align finite_dimensional_of_rank_eq_nat finiteDimensional_of_rank_eq_nat -- TODO: generalize to free modules over general rings. theorem finiteDimensional_of_rank_eq_zero (h : Module.rank K V = 0) : FiniteDimensional K V := - finiteDimensional_of_rank_eq_nat <| h.trans Nat.cast_zero.symm + Module.finite_of_rank_eq_zero h #align finite_dimensional_of_rank_eq_zero finiteDimensional_of_rank_eq_zero theorem finiteDimensional_of_rank_eq_one (h : Module.rank K V = 1) : FiniteDimensional K V := - finiteDimensional_of_rank_eq_nat <| h.trans Nat.cast_one.symm + Module.finite_of_rank_eq_one h #align finite_dimensional_of_rank_eq_one finiteDimensional_of_rank_eq_one -theorem finrank_eq_zero_of_rank_eq_zero [FiniteDimensional K V] (h : Module.rank K V = 0) : - finrank K V = 0 := by - convert finrank_eq_rank K V - rw [h]; norm_cast -#align finrank_eq_zero_of_rank_eq_zero finrank_eq_zero_of_rank_eq_zero - variable (K V) instance finiteDimensional_bot : FiniteDimensional K (⊥ : Submodule K V) := @@ -660,27 +375,6 @@ instance finiteDimensional_bot : FiniteDimensional K (⊥ : Submodule K V) := variable {K V} -theorem bot_eq_top_of_rank_eq_zero (h : Module.rank K V = 0) : (⊥ : Submodule K V) = ⊤ := by - haveI : FiniteDimensional _ _ := finiteDimensional_of_rank_eq_zero h - apply eq_top_of_finrank_eq - rw [finrank_bot, finrank_eq_zero_of_rank_eq_zero h] -#align bot_eq_top_of_rank_eq_zero bot_eq_top_of_rank_eq_zero - -@[simp] -theorem rank_eq_zero {S : Submodule K V} : Module.rank K S = 0 ↔ S = ⊥ := - ⟨fun h => - (Submodule.eq_bot_iff _).2 fun x hx => - congr_arg Subtype.val <| - ((Submodule.eq_bot_iff _).1 <| Eq.symm <| bot_eq_top_of_rank_eq_zero h) ⟨x, hx⟩ - Submodule.mem_top, - fun h => by rw [h, rank_bot]⟩ -#align rank_eq_zero rank_eq_zero - -@[simp] -theorem finrank_eq_zero {S : Submodule K V} [FiniteDimensional K S] : finrank K S = 0 ↔ S = ⊥ := by - rw [← rank_eq_zero, ← finrank_eq_rank, ← @Nat.cast_zero Cardinal, Cardinal.natCast_inj] -#align finrank_eq_zero finrank_eq_zero - end ZeroRank namespace Submodule @@ -826,7 +520,7 @@ variable [DivisionRing K] [AddCommGroup V] [Module K V] instance finiteDimensional_finsupp {ι : Type*} [Finite ι] [FiniteDimensional K V] : FiniteDimensional K (ι →₀ V) := - (Finsupp.linearEquivFunOnFinite K V ι).symm.finiteDimensional + Module.Finite.finsupp #align finite_dimensional_finsupp finiteDimensional_finsupp end @@ -1080,7 +774,7 @@ section variable [DivisionRing K] [AddCommGroup V] [Module K V] theorem finrank_zero_iff_forall_zero [FiniteDimensional K V] : finrank K V = 0 ↔ ∀ x : V, x = 0 := - finrank_zero_iff.trans (subsingleton_iff_forall_eq 0) + FiniteDimensional.finrank_zero_iff.trans (subsingleton_iff_forall_eq 0) #align finrank_zero_iff_forall_zero finrank_zero_iff_forall_zero /-- If `ι` is an empty type and `V` is zero-dimensional, there is a unique `ι`-indexed basis. -/ @@ -1105,7 +799,7 @@ theorem injective_iff_surjective_of_finrank_eq_finrank [FiniteDimensional K V] · rw [h, finrank_bot, add_zero, H] at this exact eq_top_of_finrank_eq this · rw [h, finrank_top, H] at this - exact finrank_eq_zero.1 (add_right_injective _ this) + exact Submodule.finrank_eq_zero.1 (add_right_injective _ this) #align linear_map.injective_iff_surjective_of_finrank_eq_finrank LinearMap.injective_iff_surjective_of_finrank_eq_finrank theorem ker_eq_bot_iff_range_eq_top_of_finrank_eq_finrank [FiniteDimensional K V] @@ -1334,7 +1028,7 @@ theorem finrank_eq_one_iff (ι : Type*) [Unique ι] : finrank K V = 1 ↔ Nonemp constructor · intro h haveI := finiteDimensional_of_finrank (_root_.zero_lt_one.trans_le h.symm.le) - exact ⟨basisUnique ι h⟩ + exact ⟨FiniteDimensional.basisUnique ι h⟩ · rintro ⟨b⟩ simpa using finrank_eq_card_basis b #align finrank_eq_one_iff finrank_eq_one_iff diff --git a/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean b/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean index 32fb7ab5b68c2..7426961e297b8 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean @@ -31,6 +31,13 @@ section Ring variable [Ring R] [AddCommGroup M] [Module R M] [Module.Free R M] +/-- If a free module is finite, then any arbitrary basis is finite. -/ +lemma finite_basis {R M} [Ring R] [Nontrivial R] [AddCommGroup M] [Module R M] + {ι} [Module.Finite R M] (b : Basis ι R M) : + _root_.Finite ι := + let ⟨s, hs⟩ := ‹Module.Finite R M› + basis_finite_of_finite_spans (↑s) s.finite_toSet hs b + /-- If a free module is finite, then the arbitrary basis is finite. -/ noncomputable instance ChooseBasisIndex.fintype [Module.Finite R M] : Fintype (Module.Free.ChooseBasisIndex R M) := by @@ -39,8 +46,7 @@ noncomputable instance ChooseBasisIndex.fintype [Module.Finite R M] : · have := Module.subsingleton R M rw [ChooseBasisIndex] infer_instance - · obtain ⟨s, hs⟩ := id ‹Module.Finite R M› - exact basis_finite_of_finite_spans (↑s) s.finite_toSet hs (chooseBasis _ _) + · exact finite_basis (chooseBasis _ _) #align module.free.choose_basis_index.fintype Module.Free.ChooseBasisIndex.fintype end Ring @@ -54,7 +60,7 @@ variable [AddCommGroup N] [Module R N] [Module.Free R N] variable {R} /-- A free module with a basis indexed by a `Fintype` is finite. -/ -theorem _root_.Module.Finite.of_basis {R M ι : Type*} [CommRing R] [AddCommGroup M] [Module R M] +theorem _root_.Module.Finite.of_basis {R M ι : Type*} [Semiring R] [AddCommMonoid M] [Module R M] [_root_.Finite ι] (b : Basis ι R M) : Module.Finite R M := by cases nonempty_fintype ι classical diff --git a/Mathlib/LinearAlgebra/FreeModule/Finite/Rank.lean b/Mathlib/LinearAlgebra/FreeModule/Finite/Rank.lean index 7d962c45a5e76..046c7e91812e0 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Finite/Rank.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Finite/Rank.lean @@ -19,7 +19,8 @@ This is a basic API for the rank of finite free modules. --TODO: many results from `LinearAlgebra/FiniteDimensional` should be moved here. -universe u v w +--TODO: This file contains many misplaced lemmas and should be reorganized. +universe u v w v' variable (R : Type u) (M : Type v) (N : Type w) @@ -91,11 +92,15 @@ theorem finrank_eq_card_chooseBasisIndex : simp [finrank, rank_eq_card_chooseBasisIndex] #align finite_dimensional.finrank_eq_card_choose_basis_index FiniteDimensional.finrank_eq_card_chooseBasisIndex +@[simp] +theorem finrank_finsupp {ι : Type v} [Fintype ι] : finrank R (ι →₀ M) = card ι * finrank R M := by + rw [finrank, finrank, rank_finsupp, ← mk_toNat_eq_card, toNat_mul, toNat_lift, toNat_lift] + /-- The finrank of `(ι →₀ R)` is `Fintype.card ι`. -/ @[simp] -theorem finrank_finsupp {ι : Type v} [Fintype ι] : finrank R (ι →₀ R) = card ι := by +theorem finrank_finsupp_self {ι : Type v} [Fintype ι] : finrank R (ι →₀ R) = card ι := by rw [finrank, rank_finsupp_self, ← mk_toNat_eq_card, toNat_lift] -#align finite_dimensional.finrank_finsupp FiniteDimensional.finrank_finsupp +#align finite_dimensional.finrank_finsupp FiniteDimensional.finrank_finsupp_self /-- The finrank of `(ι → R)` is `Fintype.card ι`. -/ theorem finrank_pi {ι : Type v} [Fintype ι] : finrank R (ι → R) = card ι := by @@ -227,3 +232,492 @@ theorem Submodule.finrank_le_finrank_of_le {s t : Submodule R M} [Module.Finite #align submodule.finrank_le_finrank_of_le Submodule.finrank_le_finrank_of_le end + + +open Cardinal Submodule Module Function FiniteDimensional + +attribute [local instance] nontrivial_of_invariantBasisNumber + +variable {R} {V : Type v} + +open IsNoetherian + +section DivisionRing + +variable [AddCommGroup V] {V₂ : Type v'} [AddCommGroup V₂] +variable [Ring R] [StrongRankCondition R] [Module R V] [Module.Free R V] + +/-- See `FiniteDimensional.rank_lt_aleph0` for the inverse direction without `Module.Free R V`. -/ +lemma Module.rank_lt_alpeh0_iff : + Module.rank R V < ℵ₀ ↔ Module.Finite R V := by + rw [Free.rank_eq_card_chooseBasisIndex, mk_lt_aleph0_iff] + exact ⟨fun h ↦ Finite.of_basis (Free.chooseBasis R V), + fun I ↦ Finite.of_fintype (Free.ChooseBasisIndex R V)⟩ + +theorem FiniteDimensional.finrank_of_not_finite + (h : ¬Module.Finite R V) : + finrank R V = 0 := + dif_neg (rank_lt_alpeh0_iff.not.mpr h) + +theorem Module.finite_of_finrank_pos (h : 0 < finrank R V) : + Module.Finite R V := by + contrapose h + simp [finrank_of_not_finite h] + +theorem Module.finite_of_finrank_eq_succ {n : ℕ} + (hn : finrank R V = n.succ) : Module.Finite R V := + Module.finite_of_finrank_pos <| by rw [hn]; exact n.succ_pos + +theorem Module.finite_iff_of_rank_eq_nsmul {W} [AddCommGroup W] + [Module R W] [Module.Free R W] {n : ℕ} (hn : n ≠ 0) + (hVW : Module.rank R V = n • Module.rank R W) : + Module.Finite R V ↔ Module.Finite R W := by + simp only [← rank_lt_alpeh0_iff, hVW, nsmul_lt_aleph0_iff_of_ne_zero hn] + +/-- If a free module is of finite rank, then the cardinality of any basis is equal to its +`finrank`. -/ +theorem Module.mk_finrank_eq_card_basis [Module.Finite R V] + {ι : Type w} (h : Basis ι R V) : (finrank R V : Cardinal.{w}) = #ι := by + cases @nonempty_fintype _ (Free.finite_basis h) + rw [Cardinal.mk_fintype, finrank_eq_card_basis h] + +/-- Given a basis of a ring over itself indexed by a type `ι`, then `ι` is `Unique`. -/ +noncomputable def Basis.unique {ι : Type*} (b : Basis ι R R) : Unique ι := by + have A : Cardinal.mk ι = ↑(FiniteDimensional.finrank R R) := + (Module.mk_finrank_eq_card_basis b).symm + -- porting note: replace `algebraMap.coe_one` with `Nat.cast_one` + simp only [Cardinal.eq_one_iff_unique, FiniteDimensional.finrank_self, Nat.cast_one] at A + exact Nonempty.some ((unique_iff_subsingleton_and_nonempty _).2 A) +#align basis.unique Basis.unique + +namespace FiniteDimensional +variable (R V) + +/-- A finite rank free module has a basis indexed by `Fin (finrank R V)`. -/ +noncomputable def finBasis [Module.Finite R V] : + Basis (Fin (finrank R V)) R V := + (Free.chooseBasis R V).reindex (Fintype.equivFinOfCardEq + (finrank_eq_card_chooseBasisIndex R V).symm) +#align finite_dimensional.fin_basis FiniteDimensional.finBasis + +/-- A rank `n` free module has a basis indexed by `Fin n`. -/ +noncomputable def finBasisOfFinrankEq [Module.Finite R V] + {n : ℕ} (hn : finrank R V = n) : + Basis (Fin n) R V := + (finBasis R V).reindex (Fin.castIso hn).toEquiv +#align finite_dimensional.fin_basis_of_finrank_eq FiniteDimensional.finBasisOfFinrankEq + +variable {R V} + +/-- A free module with rank 1 has a basis with one element. -/ +noncomputable def basisUnique (ι : Type*) [Unique ι] + (h : finrank R V = 1) : + Basis ι R V := + haveI : Module.Finite R V := + Module.finite_of_finrank_pos (_root_.zero_lt_one.trans_le h.symm.le) + (finBasisOfFinrankEq R V h).reindex (Equiv.equivOfUnique _ _) +#align finite_dimensional.basis_unique FiniteDimensional.basisUnique + +@[simp] +theorem basisUnique_repr_eq_zero_iff {ι : Type*} [Unique ι] + {h : finrank R V = 1} {v : V} {i : ι} : + (basisUnique ι h).repr v i = 0 ↔ v = 0 := + ⟨fun hv => + (basisUnique ι h).repr.map_eq_zero_iff.mp (Finsupp.ext fun j => Subsingleton.elim i j ▸ hv), + fun hv => by rw [hv, LinearEquiv.map_zero, Finsupp.zero_apply]⟩ +#align finite_dimensional.basis_unique.repr_eq_zero_iff FiniteDimensional.basisUnique_repr_eq_zero_iff + +theorem cardinal_mk_le_finrank_of_linearIndependent [Module.Finite R V] + {ι : Type w} {b : ι → V} (h : LinearIndependent R b) : #ι ≤ finrank R V := by + rw [← lift_le.{max v w}] + simpa only [← finrank_eq_rank, lift_natCast, lift_le_nat_iff] using + cardinal_lift_le_rank_of_linearIndependent h +#align finite_dimensional.cardinal_mk_le_finrank_of_linear_independent FiniteDimensional.cardinal_mk_le_finrank_of_linearIndependent + +theorem fintype_card_le_finrank_of_linearIndependent [Module.Finite R V] + {ι : Type*} [Fintype ι] {b : ι → V} (h : LinearIndependent R b) : + Fintype.card ι ≤ finrank R V := by + simpa using cardinal_mk_le_finrank_of_linearIndependent h +#align finite_dimensional.fintype_card_le_finrank_of_linear_independent FiniteDimensional.fintype_card_le_finrank_of_linearIndependent + +theorem finset_card_le_finrank_of_linearIndependent [Module.Finite R V] + {b : Finset V} (h : LinearIndependent R (fun x => x : b → V)) : + b.card ≤ finrank R V := by + rw [← Fintype.card_coe] + exact fintype_card_le_finrank_of_linearIndependent h +#align finite_dimensional.finset_card_le_finrank_of_linear_independent FiniteDimensional.finset_card_le_finrank_of_linearIndependent + +end FiniteDimensional + +theorem Module.Finite.lt_aleph0_of_linearIndependent {ι : Type w} + [Module.Finite R V] {v : ι → V} (h : LinearIndependent R v) : #ι < ℵ₀ := by + apply Cardinal.lift_lt.1 + apply lt_of_le_of_lt + apply cardinal_lift_le_rank_of_linearIndependent h + rw [← finrank_eq_rank, Cardinal.lift_aleph0, Cardinal.lift_natCast] + apply Cardinal.nat_lt_aleph0 + +theorem LinearIndependent.finite [Module.Finite R V] {ι : Type*} {f : ι → V} + (h : LinearIndependent R f) : Finite ι := + Cardinal.lt_aleph0_iff_finite.1 <| Module.Finite.lt_aleph0_of_linearIndependent h + +theorem LinearIndependent.setFinite [Module.Finite R V] {b : Set V} + (h : LinearIndependent R fun x : b => (x : V)) : b.Finite := + Cardinal.lt_aleph0_iff_set_finite.mp (Module.Finite.lt_aleph0_of_linearIndependent h) +#align linear_independent.finite LinearIndependent.setFinite + +theorem Module.Finite.not_linearIndependent_of_infinite {ι : Type*} [Infinite ι] [Module.Finite R V] + (v : ι → V) : ¬LinearIndependent R v := mt LinearIndependent.finite <| @not_finite _ _ +#align finite_dimensional.not_linear_independent_of_infinite Module.Finite.not_linearIndependent_of_infinite + +/-- A finite rank torsion-free module has positive `finrank` iff it has a nonzero element. -/ +theorem FiniteDimensional.finrank_pos_iff_exists_ne_zero + [Module.Finite R V] [NoZeroSMulDivisors R V] : + 0 < finrank R V ↔ ∃ x : V, x ≠ 0 := + Iff.trans + (by + rw [← finrank_eq_rank] + norm_cast) + (@rank_pos_iff_exists_ne_zero R V _ _ _ _ _) +#align finite_dimensional.finrank_pos_iff_exists_ne_zero FiniteDimensional.finrank_pos_iff_exists_ne_zero + +/-- An `R`-finite torsion-free module has positive `finrank` iff it is nontrivial. -/ +theorem FiniteDimensional.finrank_pos_iff [Module.Finite R V] [NoZeroSMulDivisors R V] : + 0 < finrank R V ↔ Nontrivial V := + Iff.trans + (by rw [← finrank_eq_rank]; norm_cast) + (rank_pos_iff_nontrivial (R := R)) +#align finite_dimensional.finrank_pos_iff FiniteDimensional.finrank_pos_iff + +/-- A nontrivial finite dimensional space has positive `finrank`. -/ +theorem FiniteDimensional.finrank_pos + [Module.Finite R V] [NoZeroSMulDivisors R V] [h : Nontrivial V] : + 0 < finrank R V := + finrank_pos_iff.mpr h +#align finite_dimensional.finrank_pos FiniteDimensional.finrank_pos + +/-- See `FiniteDimensional.finrank_zero_iff` + for the stronger version with `NoZeroSMulDivisors R V`. -/ +theorem FiniteDimensional.finrank_eq_zero_iff [Module.Finite R V] : + finrank R V = 0 ↔ ∀ x : V, ∃ a : R, a ≠ 0 ∧ a • x = 0 := + Iff.trans + (by rw [← finrank_eq_rank]; norm_cast) + (rank_eq_zero_iff (R := R)) + +/-- The `StrongRankCondition` is automatic. See `commRing_strongRankCondition`. -/ +theorem FiniteDimensional.finrank_eq_zero_iff_isTorsion {R} [CommRing R] [StrongRankCondition R] + [IsDomain R] [Module R V] [Module.Finite R V] : + finrank R V = 0 ↔ IsTorsion R V := + Iff.trans + (by rw [← finrank_eq_rank]; norm_cast) + (rank_eq_zero_iff_isTorsion (R := R)) + +/-- A finite dimensional space has zero `finrank` iff it is a subsingleton. +This is the `finrank` version of `rank_zero_iff`. -/ +theorem FiniteDimensional.finrank_zero_iff [Module.Finite R V] [NoZeroSMulDivisors R V] : + finrank R V = 0 ↔ Subsingleton V := + Iff.trans + (by rw [← finrank_eq_rank]; norm_cast) + (rank_zero_iff (R := R)) +#align finite_dimensional.finrank_zero_iff FiniteDimensional.finrank_zero_iff + +variable (R K) + +/-- The submodule generated by a finite set is `R`-finite. -/ +theorem Module.Finite.span_of_finite {A : Set V} (hA : Set.Finite A) : + Module.Finite R (Submodule.span R A) := + ⟨(Submodule.fg_top _).mpr ⟨hA.toFinset, hA.coe_toFinset.symm ▸ rfl⟩⟩ + +/-- The submodule generated by a single element is `R`-finite. -/ +instance Module.Finite.span_singleton (x : V) : Module.Finite R (R ∙ x) := + Module.Finite.span_of_finite R <| Set.finite_singleton _ + +/-- The submodule generated by a finset is `R`-finite. -/ +instance Module.Finite.span_finset (s : Finset V) : Module.Finite R (span R (s : Set V)) := + ⟨(Submodule.fg_top _).mpr ⟨s, rfl⟩⟩ + +variable {K R} + +theorem CompleteLattice.Independent.subtype_ne_bot_le_finrank_aux [Module.Finite R V] + [NoZeroSMulDivisors R V] + {ι : Type w} {p : ι → Submodule R V} (hp : CompleteLattice.Independent p) : + #{ i // p i ≠ ⊥ } ≤ (finrank R V : Cardinal.{w}) := by + suffices Cardinal.lift.{v} #{ i // p i ≠ ⊥ } ≤ Cardinal.lift.{v} (finrank R V : Cardinal.{w}) by + rwa [Cardinal.lift_le] at this + calc + Cardinal.lift.{v} #{ i // p i ≠ ⊥ } ≤ Cardinal.lift.{w} (Module.rank R V) := + hp.subtype_ne_bot_le_rank + _ = Cardinal.lift.{w} (finrank R V : Cardinal.{v}) := by rw [finrank_eq_rank] + _ = Cardinal.lift.{v} (finrank R V : Cardinal.{w}) := by simp +#align complete_lattice.independent.subtype_ne_bot_le_finrank_aux CompleteLattice.Independent.subtype_ne_bot_le_finrank_aux + +/-- If `p` is an independent family of submodules of a `R`-finite module `V`, then the +number of nontrivial subspaces in the family `p` is finite. -/ +noncomputable def CompleteLattice.Independent.fintypeNeBotOfFiniteDimensional + [Module.Finite R V] [NoZeroSMulDivisors R V] {ι : Type w} {p : ι → Submodule R V} + (hp : CompleteLattice.Independent p) : Fintype { i : ι // p i ≠ ⊥ } := by + suffices #{ i // p i ≠ ⊥ } < (ℵ₀ : Cardinal.{w}) by + rw [Cardinal.lt_aleph0_iff_fintype] at this + exact this.some + refine' lt_of_le_of_lt hp.subtype_ne_bot_le_finrank_aux _ + simp [Cardinal.nat_lt_aleph0] +#align complete_lattice.independent.fintype_ne_bot_of_finite_dimensional CompleteLattice.Independent.fintypeNeBotOfFiniteDimensional + +/-- If `p` is an independent family of submodules of a `R`-finite module `V`, then the +number of nontrivial subspaces in the family `p` is bounded above by the dimension of `V`. + +Note that the `Fintype` hypothesis required here can be provided by +`CompleteLattice.Independent.fintypeNeBotOfFiniteDimensional`. -/ +theorem CompleteLattice.Independent.subtype_ne_bot_le_finrank + [Module.Finite R V] [NoZeroSMulDivisors R V] + {ι : Type w} {p : ι → Submodule R V} (hp : CompleteLattice.Independent p) + [Fintype { i // p i ≠ ⊥ }] : + Fintype.card { i // p i ≠ ⊥ } ≤ finrank R V := by simpa using hp.subtype_ne_bot_le_finrank_aux +#align complete_lattice.independent.subtype_ne_bot_le_finrank CompleteLattice.Independent.subtype_ne_bot_le_finrank + +section + +open BigOperators + +open Finset + +/-- If a finset has cardinality larger than the rank of a module, +then there is a nontrivial linear relation amongst its elements. +-/ +-- TODO: golf this +theorem Module.exists_nontrivial_relation_of_finrank_lt_card + [Module.Finite R V] {t : Finset V} + (h : finrank R V < t.card) : ∃ f : V → R, ∑ e in t, f e • e = 0 ∧ ∃ x ∈ t, f x ≠ 0 := by + classical + have := mt FiniteDimensional.finset_card_le_finrank_of_linearIndependent (by simpa using h) + rw [_root_.not_linearIndependent_iff] at this + obtain ⟨s, g, sum, z, zm, nonzero⟩ := this + -- Now we have to extend `g` to all of `t`, then to all of `V`. + let f : V → R := fun x => if h : x ∈ t then if (⟨x, h⟩ : t) ∈ s then g ⟨x, h⟩ else 0 else 0 + -- and finally clean up the mess caused by the extension. + refine' ⟨f, _, _⟩ + · dsimp + rw [← (sum)] -- porting note: need parens to disambiguate + fapply sum_bij_ne_zero fun v hvt _ => (⟨v, hvt⟩ : { v // v ∈ t }) + · intro v hvt H + dsimp + rw [dif_pos hvt] at H + contrapose! H + rw [if_neg H, zero_smul] + · intro _ _ _ _ _ _ + exact Subtype.mk.inj + · intro b hbs hb + use b + simpa only [hbs, exists_prop, dif_pos, Finset.mk_coe, and_true_iff, if_true, Finset.coe_mem, + eq_self_iff_true, exists_prop_of_true, Ne.def] using hb + · intro a h₁ + dsimp + rw [dif_pos h₁] + intro h₂ + rw [if_pos] + contrapose! h₂ + rw [if_neg h₂, zero_smul] + · refine' ⟨z, z.2, _⟩ + dsimp only + erw [dif_pos z.2, if_pos] <;> rwa [Subtype.coe_eta] +#align finite_dimensional.exists_nontrivial_relation_of_rank_lt_card Module.exists_nontrivial_relation_of_finrank_lt_card + +/-- If a finset has cardinality larger than `finrank + 1`, +then there is a nontrivial linear relation amongst its elements, +such that the coefficients of the relation sum to zero. +-/ +-- TODO: golf this +theorem Module.exists_nontrivial_relation_sum_zero_of_finrank_succ_lt_card [Module.Finite R V] + {t : Finset V} (h : finrank R V + 1 < t.card) : + ∃ f : V → R, ∑ e in t, f e • e = 0 ∧ ∑ e in t, f e = 0 ∧ ∃ x ∈ t, f x ≠ 0 := by + classical + -- Pick an element x₀ ∈ t, + have card_pos : 0 < t.card := lt_trans (Nat.succ_pos _) h + obtain ⟨x₀, m⟩ := (Finset.card_pos.1 card_pos).bex + -- and apply the previous lemma to the {xᵢ - x₀} + let shift : V ↪ V := ⟨fun x => x - x₀, sub_left_injective⟩ + let t' := (t.erase x₀).map shift + have h' : finrank R V < t'.card := by + simp only [card_map, Finset.card_erase_of_mem m] + exact Nat.lt_pred_iff.mpr h + -- to obtain a function `g`. + obtain ⟨g, gsum, x₁, x₁_mem, nz⟩ := exists_nontrivial_relation_of_finrank_lt_card h' + -- Then obtain `f` by translating back by `x₀`, + -- and setting the value of `f` at `x₀` to ensure `∑ e in t, f e = 0`. + let f : V → R := fun z => if z = x₀ then -∑ z in t.erase x₀, g (z - x₀) else g (z - x₀) + refine' ⟨f, _, _, _⟩ + -- After this, it's a matter of verifying the properties, + -- based on the corresponding properties for `g`. + · show (∑ e : V in t, f e • e) = 0 + -- We prove this by splitting off the `x₀` term of the sum, + -- which is itself a sum over `t.erase x₀`, + -- combining the two sums, and + -- observing that after reindexing we have exactly + -- ∑ (x : V) in t', g x • x = 0. + simp only + conv_lhs => + apply_congr + rfl + rw [ite_smul] + rw [Finset.sum_ite] + conv => + congr + congr + apply_congr + -- Porting note: the next two steps used to work by `simp [filter_eq', m]` + erw [filter_eq'] + simp [m] + conv => + congr + congr + rfl + apply_congr + simp [filter_ne'] + rw [sum_singleton, neg_smul, add_comm, ← sub_eq_add_neg, sum_smul, ← sum_sub_distrib] + simp only [← smul_sub] + -- At the end we have to reindex the sum, so we use `change` to + -- express the summand using `shift`. + change (∑ x : V in t.erase x₀, (fun e => g e • e) (shift x)) = 0 + -- porting note: last argument can't be inferred + rw [← sum_map _ shift (fun e => g e • e)] + exact gsum + · show (∑ e : V in t, f e) = 0 + -- Again we split off the `x₀` term, + -- observing that it exactly cancels the other terms. + rw [← insert_erase m, sum_insert (not_mem_erase x₀ t)] + dsimp + rw [if_pos rfl] + conv_lhs => + congr + rfl + apply_congr + rfl + rw [if_neg (show _ ≠ x₀ from (mem_erase.mp ‹_›).1)] + exact neg_add_self _ + · show ∃ (x : V), x ∈ t ∧ f x ≠ 0 + -- We can use x₁ + x₀. + refine' ⟨x₁ + x₀, _, _⟩ + · rw [Finset.mem_map] at x₁_mem + rcases x₁_mem with ⟨x₁, x₁_mem, rfl⟩ + rw [mem_erase] at x₁_mem + simp only [x₁_mem, sub_add_cancel, Function.Embedding.coeFn_mk] + · dsimp only + rwa [if_neg, add_sub_cancel] + rw [add_left_eq_self] + rintro rfl + simp only [sub_eq_zero, exists_prop, Finset.mem_map, Embedding.coeFn_mk, eq_self_iff_true, + mem_erase, not_true, exists_eq_right, Ne.def, false_and_iff] at x₁_mem +#align finite_dimensional.exists_nontrivial_relation_sum_zero_of_rank_succ_lt_card Module.exists_nontrivial_relation_sum_zero_of_finrank_succ_lt_card + +end + +end DivisionRing + +section ZeroRank + +variable [Ring R] [StrongRankCondition R] [AddCommGroup V] [Module R V] + +attribute [local instance] nontrivial_of_invariantBasisNumber + +open FiniteDimensional + +theorem Module.finite_of_rank_eq_nat [Module.Free R V] {n : ℕ} (h : Module.rank R V = n) : + Module.Finite R V := by + have := Cardinal.mk_lt_aleph0_iff.mp + (((Free.rank_eq_card_chooseBasisIndex R V).symm.trans h).trans_lt (nat_lt_aleph0 n)) + exact Module.Finite.of_basis (Free.chooseBasis R V) + +theorem Module.finite_of_rank_eq_zero [NoZeroSMulDivisors R V] + (h : Module.rank R V = 0) : + Module.Finite R V := by + rw [rank_zero_iff] at h + infer_instance + +theorem Module.finite_of_rank_eq_one [Module.Free R V] (h : Module.rank R V = 1) : + Module.Finite R V := + Module.finite_of_rank_eq_nat <| h.trans Nat.cast_one.symm + +theorem FiniteDimensional.finrank_eq_zero_of_rank_eq_zero (h : Module.rank R V = 0) : + finrank R V = 0 := by + delta finrank + rw [h, zero_toNat] +#align finrank_eq_zero_of_rank_eq_zero FiniteDimensional.finrank_eq_zero_of_rank_eq_zero + +variable (R V) + +instance Module.finite_bot : Module.Finite R (⊥ : Submodule R V) := inferInstance + +variable {R V} + +theorem Submodule.bot_eq_top_of_rank_eq_zero [NoZeroSMulDivisors R V] (h : Module.rank R V = 0) : + (⊥ : Submodule R V) = ⊤ := by + rw [rank_zero_iff] at h + exact Subsingleton.elim _ _ +#align bot_eq_top_of_rank_eq_zero Submodule.bot_eq_top_of_rank_eq_zero + +@[simp] +theorem Submodule.rank_eq_zero [NoZeroSMulDivisors R V] {S : Submodule R V} : + Module.rank R S = 0 ↔ S = ⊥ := + ⟨fun h => + (Submodule.eq_bot_iff _).2 fun x hx => + congr_arg Subtype.val <| + ((Submodule.eq_bot_iff _).1 <| Eq.symm <| Submodule.bot_eq_top_of_rank_eq_zero h) ⟨x, hx⟩ + Submodule.mem_top, + fun h => by rw [h, rank_bot]⟩ +#align rank_eq_zero Submodule.rank_eq_zero + +@[simp] +theorem Submodule.finrank_eq_zero [NoZeroSMulDivisors R V] {S : Submodule R V} [Module.Finite R S] : + finrank R S = 0 ↔ S = ⊥ := by + rw [← Submodule.rank_eq_zero, ← finrank_eq_rank, ← @Nat.cast_zero Cardinal, Cardinal.natCast_inj] +#align finrank_eq_zero Submodule.finrank_eq_zero + +end ZeroRank + +namespace Submodule + +open IsNoetherian FiniteDimensional + +variable [AddCommGroup V] [Ring R] [StrongRankCondition R] [Module R V] + +theorem fg_iff_finite (s : Submodule R V) : s.FG ↔ Module.Finite R s := + (finite_def.trans (fg_top s)).symm + +/-- The sup of two fg submodules is finite. Also see `Submodule.FG.sup`. -/ +instance finite_sup (S₁ S₂ : Submodule R V) [h₁ : Module.Finite R S₁] + [h₂ : Module.Finite R S₂] : Module.Finite R (S₁ ⊔ S₂ : Submodule R V) := by + rw [finite_def] at * + exact (fg_top _).2 (((fg_top S₁).1 h₁).sup ((fg_top S₂).1 h₂)) + +/-- The submodule generated by a finite supremum of finite dimensional submodules is +finite-dimensional. + +Note that strictly this only needs `∀ i ∈ s, FiniteDimensional K (S i)`, but that doesn't +work well with typeclass search. -/ +instance finite_finset_sup {ι : Type*} (s : Finset ι) (S : ι → Submodule R V) + [∀ i, Module.Finite R (S i)] : Module.Finite R (s.sup S : Submodule R V) := by + refine' + @Finset.sup_induction _ _ _ _ s S (fun i => Module.Finite R ↑i) (Module.finite_bot R V) + _ fun i _ => by infer_instance + · intro S₁ hS₁ S₂ hS₂ + exact Submodule.finite_sup S₁ S₂ + +/-- The submodule generated by a supremum of finite dimensional submodules, indexed by a finite +sort is finite-dimensional. -/ +instance finite_iSup {ι : Sort*} [Finite ι] (S : ι → Submodule R V) + [∀ i, Module.Finite R (S i)] : Module.Finite R ↑(⨆ i, S i) := by + cases nonempty_fintype (PLift ι) + rw [← iSup_plift_down, ← Finset.sup_univ_eq_iSup] + exact Submodule.finite_finset_sup _ _ + +end Submodule + +section + +variable [Ring R] [StrongRankCondition R] [AddCommGroup V] [Module R V] + +instance Module.Finite.finsupp {ι : Type*} [_root_.Finite ι] [Module.Finite R V] : + Module.Finite R (ι →₀ V) := + Module.Finite.equiv (Finsupp.linearEquivFunOnFinite R V ι).symm + +end diff --git a/Mathlib/LinearAlgebra/Matrix/Hermitian.lean b/Mathlib/LinearAlgebra/Matrix/Hermitian.lean index 25b5cb610bd7e..0c81e187da734 100644 --- a/Mathlib/LinearAlgebra/Matrix/Hermitian.lean +++ b/Mathlib/LinearAlgebra/Matrix/Hermitian.lean @@ -131,13 +131,18 @@ section AddMonoid variable [AddMonoid α] [StarAddMonoid α] [AddMonoid β] [StarAddMonoid β] -/-- A diagonal matrix is hermitian if the entries are self-adjoint -/ +/-- A diagonal matrix is hermitian if the entries are self-adjoint (as a vector) -/ theorem isHermitian_diagonal_of_self_adjoint [DecidableEq n] (v : n → α) (h : IsSelfAdjoint v) : (diagonal v).IsHermitian := (-- TODO: add a `pi.has_trivial_star` instance and remove the `funext` diagonal_conjTranspose v).trans <| congr_arg _ h #align matrix.is_hermitian_diagonal_of_self_adjoint Matrix.isHermitian_diagonal_of_self_adjoint +/-- A diagonal matrix is hermitian if each diagonal entry is self-adjoint -/ +lemma isHermitian_diagonal_iff [DecidableEq n]{d : n → α} : + IsHermitian (diagonal d) ↔ (∀ i : n, IsSelfAdjoint (d i)) := by + simp [isSelfAdjoint_iff, IsHermitian, conjTranspose, diagonal_transpose, diagonal_map] + /-- A diagonal matrix is hermitian if the entries have the trivial `star` operation (such as on the reals). -/ @[simp] diff --git a/Mathlib/LinearAlgebra/Matrix/PosDef.lean b/Mathlib/LinearAlgebra/Matrix/PosDef.lean index 9f0a7a1b85de7..a71114b77a639 100644 --- a/Mathlib/LinearAlgebra/Matrix/PosDef.lean +++ b/Mathlib/LinearAlgebra/Matrix/PosDef.lean @@ -11,20 +11,25 @@ import Mathlib.LinearAlgebra.QuadraticForm.Basic /-! # Positive Definite Matrices This file defines positive (semi)definite matrices and connects the notion to positive definiteness -of quadratic forms. +of quadratic forms. Most results require `𝕜 = ℝ` or `ℂ`. -## Main definition +## Main definitions * `Matrix.PosDef` : a matrix `M : Matrix n n 𝕜` is positive definite if it is hermitian and `xᴴMx` is greater than zero for all nonzero `x`. * `Matrix.PosSemidef` : a matrix `M : Matrix n n 𝕜` is positive semidefinite if it is hermitian and `xᴴMx` is nonnegative for all `x`. +## Main results + +* `Matrix.posSemidef_iff_eq_transpose_mul_self` : a matrix `M : Matrix n n 𝕜` is positive + semidefinite iff it has the form `Bᴴ * B` for some `B`. +* `Matrix.PosSemidef.sqrt` : the unique positive semidefinite square root of a positive semidefinite + matrix. (See `Matrix.PosSemidef.eq_sqrt_of_sq_eq` for the proof of uniqueness.) -/ open scoped ComplexOrder - namespace Matrix variable {m n R 𝕜 : Type*} @@ -37,32 +42,201 @@ open scoped Matrix ## Positive semidefinite matrices -/ -/-- A matrix `M : Matrix n n R` is positive semidefinite if it is hermitian - and `xᴴMx` is nonnegative for all `x`. -/ +/-- A matrix `M : Matrix n n R` is positive semidefinite if it is Hermitian and `xᴴ * M * x` is +nonnegative for all `x`. -/ def PosSemidef (M : Matrix n n R) := M.IsHermitian ∧ ∀ x : n → R, 0 ≤ dotProduct (star x) (M.mulVec x) #align matrix.pos_semidef Matrix.PosSemidef -theorem PosSemidef.re_dotProduct_nonneg {M : Matrix n n 𝕜} (hM : M.PosSemidef) (x : n → 𝕜) : +/-- A diagonal matrix is positive semidefinite iff its diagonal entries are nonnegative. -/ +lemma posSemidef_diagonal_iff [DecidableEq n] {d : n → R} : + PosSemidef (diagonal d) ↔ (∀ i : n, 0 ≤ d i) := by + refine ⟨fun ⟨_, hP⟩ i ↦ by simpa using hP (Pi.single i 1), ?_⟩ + refine fun hd ↦ ⟨isHermitian_diagonal_iff.2 <| fun i ↦ IsSelfAdjoint.of_nonneg (hd i), ?_⟩ + refine fun x ↦ Finset.sum_nonneg fun i _ ↦ ?_ + simpa only [mulVec_diagonal, mul_assoc] using conjugate_nonneg (hd i) _ + +namespace PosSemidef + +theorem isHermitian {M : Matrix n n R} (hM : M.PosSemidef) : M.IsHermitian := + hM.1 + +theorem re_dotProduct_nonneg {M : Matrix n n 𝕜} (hM : M.PosSemidef) (x : n → 𝕜) : 0 ≤ IsROrC.re (dotProduct (star x) (M.mulVec x)) := IsROrC.nonneg_iff.mp (hM.2 _) |>.1 -theorem PosSemidef.submatrix {M : Matrix n n R} (hM : M.PosSemidef) (e : m ≃ n) : +lemma conjTranspose_mul_mul_same {A : Matrix n n R} (hA : PosSemidef A) + {m : Type*} [Fintype m] (B : Matrix n m R) : + PosSemidef (Bᴴ * A * B) := by + constructor + · exact isHermitian_conjTranspose_mul_mul B hA.1 + · intro x + simpa only [star_mulVec, dotProduct_mulVec, vecMul_vecMul] using hA.2 (mulVec B x) + +lemma mul_mul_conjTranspose_same {A : Matrix n n R} (hA : PosSemidef A) + {m : Type*} [Fintype m] (B : Matrix m n R): + PosSemidef (B * A * Bᴴ) := by + simpa only [conjTranspose_conjTranspose] using hA.conjTranspose_mul_mul_same Bᴴ + +theorem submatrix {M : Matrix n n R} (hM : M.PosSemidef) (e : m → n) : (M.submatrix e e).PosSemidef := by - refine' ⟨hM.1.submatrix e, fun x => _⟩ - have : (M.submatrix (⇑e) e).mulVec x = (M.mulVec fun i : n => x (e.symm i)) ∘ e := by - ext i - dsimp only [(· ∘ ·), mulVec, dotProduct] - rw [Finset.sum_bij' (fun i _ => e i) _ _ fun i _ => e.symm i] <;> - simp only [eq_self_iff_true, imp_true_iff, Equiv.symm_apply_apply, Finset.mem_univ, - submatrix_apply, Equiv.apply_symm_apply] - rw [this] - convert hM.2 fun i => x (e.symm i) using 3 - unfold dotProduct - rw [Finset.sum_bij' (fun i _ => e i) _ _ fun i _ => e.symm i] <;> - simp + classical + rw [(by simp : M = 1 * M * 1), submatrix_mul (he₂ := Function.bijective_id), + submatrix_mul (he₂ := Function.bijective_id), submatrix_id_id] + simpa only [conjTranspose_submatrix, conjTranspose_one] using + conjTranspose_mul_mul_same hM (Matrix.submatrix 1 id e) #align matrix.pos_semidef.submatrix Matrix.PosSemidef.submatrix +theorem transpose {M : Matrix n n R} (hM : M.PosSemidef) : Mᵀ.PosSemidef := by + refine ⟨IsHermitian.transpose hM.1, fun x => ?_⟩ + convert hM.2 (star x) using 1 + rw [mulVec_transpose, Matrix.dotProduct_mulVec, star_star, dotProduct_comm] + +theorem conjTranspose {M : Matrix n n R} (hM : M.PosSemidef) : Mᴴ.PosSemidef := hM.1.symm ▸ hM + +protected lemma zero : PosSemidef (0 : Matrix n n R) := + ⟨isHermitian_zero, by simp⟩ + +protected lemma one [DecidableEq n] : PosSemidef (1 : Matrix n n R) := + ⟨isHermitian_one, fun x => by + rw [one_mulVec]; exact Fintype.sum_nonneg fun i => star_mul_self_nonneg _⟩ + +protected lemma pow [DecidableEq n] {M : Matrix n n R} (hM : M.PosSemidef) (k : ℕ) : + PosSemidef (M ^ k) := + match k with + | 0 => .one + | 1 => by simpa using hM + | (k + 2) => by + rw [pow_succ', pow_succ] + simpa only [hM.isHermitian.eq] using (hM.pow k).mul_mul_conjTranspose_same M + +protected lemma inv [DecidableEq n] {M : Matrix n n R} (hM : M.PosSemidef) : M⁻¹.PosSemidef := by + by_cases h : IsUnit M.det + · have := (conjTranspose_mul_mul_same hM M⁻¹).conjTranspose + rwa [mul_nonsing_inv_cancel_right _ _ h, conjTranspose_conjTranspose] at this + · rw [nonsing_inv_apply_not_isUnit _ h] + exact .zero + +protected lemma zpow [DecidableEq n] {M : Matrix n n R} (hM : M.PosSemidef) (z : ℤ) : + (M ^ z).PosSemidef := by + obtain ⟨n, rfl | rfl⟩ := z.eq_nat_or_neg + · simpa using hM.pow n + · simpa using (hM.pow n).inv + +/-- The eigenvalues of a positive semi-definite matrix are non-negative -/ +lemma eigenvalues_nonneg [DecidableEq n] {A : Matrix n n 𝕜} + (hA : Matrix.PosSemidef A) (i : n) : 0 ≤ hA.1.eigenvalues i := + (hA.re_dotProduct_nonneg _).trans_eq (hA.1.eigenvalues_eq _).symm + +section sqrt + +variable [DecidableEq n] {A : Matrix n n 𝕜} (hA : PosSemidef A) + +/-- The positive semidefinite square root of a positive semidefinite matrix -/ +@[pp_dot] +noncomputable def sqrt : Matrix n n 𝕜 := + hA.1.eigenvectorMatrix * diagonal ((↑) ∘ Real.sqrt ∘ hA.1.eigenvalues) * hA.1.eigenvectorMatrixᴴ + +open Lean PrettyPrinter.Delaborator SubExpr in +/-- Custom elaborator to produce output like `(_ : PosSemidef A).sqrt` in the goal view. -/ +@[delab app.Matrix.PosSemidef.sqrt] +def delabSqrt : Delab := + whenPPOption getPPNotation <| + whenNotPPOption getPPAnalysisSkip <| + withOptionAtCurrPos `pp.analysis.skip true do + let e ← getExpr + guard <| e.isAppOfArity ``Matrix.PosSemidef.sqrt 7 + let optionsPerPos ← withNaryArg 6 do + return (← read).optionsPerPos.setBool (← getPos) `pp.proofs.withType true + withTheReader Context ({· with optionsPerPos}) delab + +-- test for custom elaborator +/-- +info: (_ : PosSemidef A).sqrt : Matrix n n 𝕜 +-/ +#guard_msgs in +#check (id hA).sqrt + +lemma posSemidef_sqrt : PosSemidef hA.sqrt := by + apply PosSemidef.mul_mul_conjTranspose_same + refine posSemidef_diagonal_iff.mpr fun i ↦ ?_ + rw [Function.comp_apply, IsROrC.nonneg_iff] + constructor + · simp only [IsROrC.ofReal_re] + exact Real.sqrt_nonneg _ + · simp only [IsROrC.ofReal_im] + +@[simp] +lemma sq_sqrt : hA.sqrt ^ 2 = A := by + let C := hA.1.eigenvectorMatrix + let E := diagonal ((↑) ∘ Real.sqrt ∘ hA.1.eigenvalues : n → 𝕜) + suffices : C * (E * (Cᴴ * C) * E) * Cᴴ = A + · rw [Matrix.PosSemidef.sqrt, pow_two] + change (C * E * Cᴴ) * (C * E * Cᴴ) = A + simpa only [← mul_assoc] using this + have : Cᴴ * C = 1 + · rw [Matrix.IsHermitian.conjTranspose_eigenvectorMatrix, mul_eq_one_comm] + exact hA.1.eigenvectorMatrix_mul_inv + rw [this, mul_one] + have : E * E = diagonal ((↑) ∘ hA.1.eigenvalues) + · rw [diagonal_mul_diagonal] + refine congr_arg _ (funext fun v ↦ ?_) -- why doesn't "congr with v" work? + simp [← pow_two, ← IsROrC.ofReal_pow, Real.sq_sqrt (hA.eigenvalues_nonneg v)] + rw [this] + convert hA.1.spectral_theorem'.symm + apply Matrix.IsHermitian.conjTranspose_eigenvectorMatrix + +@[simp] +lemma sqrt_mul_self : hA.sqrt * hA.sqrt = A := by rw [← pow_two, sq_sqrt] + +lemma eq_of_sq_eq_sq {B : Matrix n n 𝕜} (hB : PosSemidef B) (hAB : A ^ 2 = B ^ 2) : A = B := by + /- This is deceptively hard, much more difficult than the positive *definite* case. We follow a + clever proof due to Koeber and Schäfer. The idea is that if `A ≠ B`, then `A - B` has a nonzero + real eigenvalue, with eigenvector `v`. Then a manipulation using the identity + `A ^ 2 - B ^ 2 = A * (A - B) + (A - B) * B` leads to the conclusion that + `⟨v, A v⟩ + ⟨v, B v⟩ = 0`. Since `A, B` are positive semidefinite, both terms must be zero. Thus + `⟨v, (A - B) v⟩ = 0`, but this is a nonzero scalar multiple of `⟨v, v⟩`, contradiction. -/ + by_contra h_ne + let ⟨v, t, ht, hv, hv'⟩ := (hA.1.sub hB.1).exists_eigenvector_of_ne_zero (sub_ne_zero.mpr h_ne) + have h_sum : 0 = t * (star v ⬝ᵥ mulVec A v + star v ⬝ᵥ mulVec B v) + · calc + 0 = star v ⬝ᵥ mulVec (A ^ 2 - B ^ 2) v := by rw [hAB, sub_self, zero_mulVec, dotProduct_zero] + _ = star v ⬝ᵥ mulVec A (mulVec (A - B) v) + star v ⬝ᵥ mulVec (A - B) (mulVec B v) := by + rw [mulVec_mulVec, mulVec_mulVec, ← dotProduct_add, ← add_mulVec, mul_sub, sub_mul, + add_sub, sub_add_cancel, pow_two, pow_two] + _ = t * (star v ⬝ᵥ mulVec A v) + vecMul (star v) (A - B)ᴴ ⬝ᵥ mulVec B v := by + rw [hv', mulVec_smul, dotProduct_smul, IsROrC.real_smul_eq_coe_mul, + dotProduct_mulVec _ (A - B), hA.1.sub hB.1] + _ = t * (star v ⬝ᵥ mulVec A v + star v ⬝ᵥ mulVec B v) := by + simp_rw [← star_mulVec, hv', mul_add, ← IsROrC.real_smul_eq_coe_mul, ← smul_dotProduct] + congr 2 with i + simp only [Pi.star_apply, Pi.smul_apply, IsROrC.real_smul_eq_coe_mul, star_mul', + IsROrC.star_def, IsROrC.conj_ofReal] + replace h_sum : star v ⬝ᵥ mulVec A v + star v ⬝ᵥ mulVec B v = 0 + · rw [eq_comm, ← mul_zero (t : 𝕜)] at h_sum + exact mul_left_cancel₀ (IsROrC.ofReal_ne_zero.mpr ht) h_sum + have h_van : star v ⬝ᵥ mulVec A v = 0 ∧ star v ⬝ᵥ mulVec B v = 0 + · refine ⟨le_antisymm ?_ (hA.2 v), le_antisymm ?_ (hB.2 v)⟩ + · rw [add_comm, add_eq_zero_iff_eq_neg] at h_sum + simpa only [h_sum, neg_nonneg] using hB.2 v + · simpa only [add_eq_zero_iff_eq_neg.mp h_sum, neg_nonneg] using hA.2 v + have aux : star v ⬝ᵥ mulVec (A - B) v = 0 + · rw [sub_mulVec, dotProduct_sub, h_van.1, h_van.2, sub_zero] + rw [hv', dotProduct_smul, IsROrC.real_smul_eq_coe_mul, ← mul_zero ↑t] at aux + exact hv <| Matrix.dotProduct_star_self_eq_zero.mp <| mul_left_cancel₀ + (IsROrC.ofReal_ne_zero.mpr ht) aux + +lemma sqrt_sq : (hA.pow 2 : PosSemidef (A ^ 2)).sqrt = A := + (hA.pow 2).posSemidef_sqrt.eq_of_sq_eq_sq hA (hA.pow 2).sq_sqrt + +lemma eq_sqrt_of_sq_eq {B : Matrix n n 𝕜} (hB : PosSemidef B) (hAB : A ^ 2 = B) : A = hB.sqrt := by + subst B + rw [hA.sqrt_sq] + +end sqrt + +end PosSemidef + @[simp] theorem posSemidef_submatrix_equiv {M : Matrix n n R} (e : m ≃ n) : (M.submatrix e e).PosSemidef ↔ M.PosSemidef := @@ -70,27 +244,55 @@ theorem posSemidef_submatrix_equiv {M : Matrix n n R} (e : m ≃ n) : #align matrix.pos_semidef_submatrix_equiv Matrix.posSemidef_submatrix_equiv /-- The conjugate transpose of a matrix mulitplied by the matrix is positive semidefinite -/ -theorem posSemidef_conjTranspose_mul_self (A : Matrix m n R) : Matrix.PosSemidef (Aᴴ * A) := by +theorem posSemidef_conjTranspose_mul_self (A : Matrix m n R) : PosSemidef (Aᴴ * A) := by refine ⟨isHermitian_transpose_mul_self _, fun x => ?_⟩ rw [← mulVec_mulVec, dotProduct_mulVec, vecMul_conjTranspose, star_star] exact Finset.sum_nonneg fun i _ => star_mul_self_nonneg _ /-- A matrix multiplied by its conjugate transpose is positive semidefinite -/ -theorem posSemidef_self_mul_conjTranspose (A : Matrix m n R) : Matrix.PosSemidef (A * Aᴴ) := +theorem posSemidef_self_mul_conjTranspose (A : Matrix m n R) : PosSemidef (A * Aᴴ) := by simpa only [conjTranspose_conjTranspose] using posSemidef_conjTranspose_mul_self Aᴴ -/-- The eigenvalues of a positive semi-definite matrix are non-negative -/ -lemma PosSemidef.eigenvalues_nonneg [DecidableEq n] {A : Matrix n n 𝕜} - (hA : Matrix.PosSemidef A) (i : n) : 0 ≤ hA.1.eigenvalues i := - (hA.re_dotProduct_nonneg _).trans_eq (hA.1.eigenvalues_eq _).symm - lemma eigenvalues_conjTranspose_mul_self_nonneg (A : Matrix m n 𝕜) [DecidableEq n] (i : n) : 0 ≤ (isHermitian_transpose_mul_self A).eigenvalues i := - (Matrix.posSemidef_conjTranspose_mul_self _).eigenvalues_nonneg _ + (posSemidef_conjTranspose_mul_self _).eigenvalues_nonneg _ lemma eigenvalues_self_mul_conjTranspose_nonneg (A : Matrix m n 𝕜) [DecidableEq m] (i : m) : 0 ≤ (isHermitian_mul_conjTranspose_self A).eigenvalues i := - (Matrix.posSemidef_self_mul_conjTranspose _).eigenvalues_nonneg _ + (posSemidef_self_mul_conjTranspose _).eigenvalues_nonneg _ + +/-- A matrix is positive semidefinite if and only if it has the form `Bᴴ * B` for some `B`. -/ +lemma posSemidef_iff_eq_transpose_mul_self [DecidableEq n] {A : Matrix n n 𝕜} : + PosSemidef A ↔ ∃ (B : Matrix n n 𝕜), A = Bᴴ * B := by + refine ⟨fun hA ↦ ⟨hA.sqrt, ?_⟩, fun ⟨B, hB⟩ ↦ (hB ▸ posSemidef_conjTranspose_mul_self B)⟩ + simp_rw [← PosSemidef.sq_sqrt hA, pow_two] + rw [hA.posSemidef_sqrt.1] + +lemma IsHermitian.posSemidef_of_eigenvalues_nonneg [DecidableEq n] {A : Matrix n n 𝕜} + (hA : IsHermitian A) (h : ∀ i : n, 0 ≤ hA.eigenvalues i) : PosSemidef A := by + simp_rw [hA.conjTranspose_eigenvectorMatrix.symm ▸ hA.spectral_theorem'] + refine (posSemidef_diagonal_iff.mpr fun i ↦ ?_).mul_mul_conjTranspose_same _ + rw [IsROrC.le_iff_re_im] + simpa using h i + +/-- For `A` positive semidefinite, we have `x⋆ A x = 0` iff `A x = 0`. -/ +theorem PosSemidef.dotProduct_mulVec_zero_iff [DecidableEq n] + {A : Matrix n n 𝕜} (hA : PosSemidef A) (x : n → 𝕜) : + star x ⬝ᵥ mulVec A x = 0 ↔ mulVec A x = 0 := by + constructor + · obtain ⟨B, rfl⟩ := posSemidef_iff_eq_transpose_mul_self.mp hA + rw [← Matrix.mulVec_mulVec, dotProduct_mulVec, + vecMul_conjTranspose, star_star, dotProduct_star_self_eq_zero] + intro h0 + rw [h0, mulVec_zero] + · intro h0 + rw [h0, dotProduct_zero] + +/-- For `A` positive semidefinite, we have `x⋆ A x = 0` iff `A x = 0` (linear maps version). -/ +theorem PosSemidef.toLinearMap₂'_zero_iff [DecidableEq n] + {A : Matrix n n 𝕜} (hA : PosSemidef A) (x : n → 𝕜) : + Matrix.toLinearMap₂' A (star x) x = 0 ↔ Matrix.toLin' A x = 0 := by + simpa only [toLinearMap₂'_apply', toLin'_apply] using hA.dotProduct_mulVec_zero_iff x /-! ## Positive definite matrices diff --git a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean index 08bde350392b3..8f03821568c01 100644 --- a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean +++ b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean @@ -161,6 +161,28 @@ lemma rank_eq_rank_diagonal : A.rank = (Matrix.diagonal hA.eigenvalues).rank := lemma rank_eq_card_non_zero_eigs : A.rank = Fintype.card {i // hA.eigenvalues i ≠ 0} := by rw [rank_eq_rank_diagonal hA, Matrix.rank_diagonal] +/-- The entries of `eigenvectorBasis` are eigenvectors. -/ +lemma mulVec_eigenvectorBasis (i : n) : + mulVec A (hA.eigenvectorBasis i) = hA.eigenvalues i • hA.eigenvectorBasis i := by + have := congr_arg (· * hA.eigenvectorMatrix) hA.spectral_theorem' + simp only [mul_assoc, mul_eq_one_comm.mp hA.eigenvectorMatrix_mul_inv, mul_one] at this + ext1 j + have := congr_fun (congr_fun this j) i + simp only [mul_diagonal, Function.comp_apply] at this + convert this using 1 + rw [mul_comm, Pi.smul_apply, IsROrC.real_smul_eq_coe_mul, hA.eigenvectorMatrix_apply] + +/-- A nonzero Hermitian matrix has an eigenvector with nonzero eigenvalue. -/ +lemma exists_eigenvector_of_ne_zero (h_ne : A ≠ 0) : + ∃ (v : n → 𝕜) (t : ℝ), t ≠ 0 ∧ v ≠ 0 ∧ mulVec A v = t • v := by + have : hA.eigenvalues ≠ 0 + · contrapose! h_ne + have := hA.spectral_theorem' + rwa [h_ne, Pi.comp_zero, IsROrC.ofReal_zero, (by rfl : Function.const n (0 : 𝕜) = fun _ ↦ 0), + diagonal_zero, mul_zero, zero_mul] at this + obtain ⟨i, hi⟩ := Function.ne_iff.mp this + exact ⟨_, _, hi, hA.eigenvectorBasis.orthonormal.ne_zero i, hA.mulVec_eigenvectorBasis i⟩ + end IsHermitian end Matrix diff --git a/Mathlib/LinearAlgebra/Orientation.lean b/Mathlib/LinearAlgebra/Orientation.lean index 3a2c45063f181..7caee38d5b01a 100644 --- a/Mathlib/LinearAlgebra/Orientation.lean +++ b/Mathlib/LinearAlgebra/Orientation.lean @@ -410,9 +410,7 @@ theorem map_eq_iff_det_pos (x : Orientation R M ι) (f : M ≃ₗ[R] M) (h : Fintype.card ι = finrank R M) : Orientation.map ι f x = x ↔ 0 < LinearMap.det (f : M →ₗ[R] M) := by cases isEmpty_or_nonempty ι - · have H : finrank R M = 0 := by - refine' h.symm.trans _ - convert @Fintype.card_of_isEmpty ι _ + · have H : finrank R M = 0 := h.symm.trans Fintype.card_eq_zero simp [LinearMap.det_eq_one_of_finrank_eq_zero H] rw [map_eq_det_inv_smul _ _ h, units_inv_smul, units_smul_eq_self_iff, LinearEquiv.coe_det] #align orientation.map_eq_iff_det_pos Orientation.map_eq_iff_det_pos @@ -424,9 +422,7 @@ theorem map_eq_neg_iff_det_neg (x : Orientation R M ι) (f : M ≃ₗ[R] M) (h : Fintype.card ι = finrank R M) : Orientation.map ι f x = -x ↔ LinearMap.det (f : M →ₗ[R] M) < 0 := by cases isEmpty_or_nonempty ι - · have H : finrank R M = 0 := by - refine' h.symm.trans _ - convert @Fintype.card_of_isEmpty ι _ + · have H : finrank R M = 0 := h.symm.trans Fintype.card_eq_zero simp [LinearMap.det_eq_one_of_finrank_eq_zero H, Module.Ray.ne_neg_self x] have H : 0 < finrank R M := by rw [← h] diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean new file mode 100644 index 0000000000000..0d2eefaed5243 --- /dev/null +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean @@ -0,0 +1,263 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.Data.Int.Order.Units +import Mathlib.Data.ZMod.IntUnitsPower +import Mathlib.RingTheory.TensorProduct +import Mathlib.LinearAlgebra.DirectSum.TensorProduct +import Mathlib.Algebra.DirectSum.Algebra + +/-! +# Graded tensor products over graded algebras + +The graded tensor product $A \hat\otimes_R B$ is imbued with a multiplication defined on homogeneous +tensors by: + +$$(a \otimes b) \cdot (a' \otimes b') = (-1)^{\deg a' \deg b} (a \cdot a') \otimes (b \cdot b')$$ + +where $A$ and $B$ are algebras graded by `ℕ`, `ℤ`, or `ZMod 2` (or more generally, any index +that satisfies `Module ι (Additive ℤˣ)`). + +The results for internally-graded algebras (via `GradedAlgebra`) are elsewhere, as is the type +`GradedTensorProduct`. + +## Main results + +* `TensorProduct.gradedComm`: the symmetric braiding operator on the tensor product of + externally-graded rings. +* `TensorProduct.gradedMul`: the previously-described multiplication on externally-graded rings, as + a bilinear map. + +## Implementation notes + +Rather than implementing the multiplication directly as above, we first implement the canonical +non-trivial braiding sending $a \otimes b$ to $(-1)^{\deg a' \deg b} (b \otimes a)$, as the +multiplication follows trivially from this after some point-free nonsense. + +## References + +* https://math.stackexchange.com/q/202718/1896 +* [*Algebra I*, Bourbaki : Chapter III, §4.7, example (2)][bourbaki1989] + +-/ + +suppress_compilation + +open scoped TensorProduct DirectSum + +variable {R ι A B : Type*} + +namespace TensorProduct + +variable [CommSemiring ι] [Module ι (Additive ℤˣ)] [DecidableEq ι] +variable (𝒜 : ι → Type*) (ℬ : ι → Type*) +variable [CommRing R] +variable [∀ i, AddCommGroup (𝒜 i)] [∀ i, AddCommGroup (ℬ i)] +variable [∀ i, Module R (𝒜 i)] [∀ i, Module R (ℬ i)] +variable [DirectSum.GRing 𝒜] [DirectSum.GRing ℬ] +variable [DirectSum.GAlgebra R 𝒜] [DirectSum.GAlgebra R ℬ] + +-- this helps with performance +instance (i : ι × ι) : Module R (𝒜 (Prod.fst i) ⊗[R] ℬ (Prod.snd i)) := + TensorProduct.leftModule + +open DirectSum (lof) + +variable (R) + +section gradedComm + +local notation "𝒜ℬ" => (fun i : ι × ι => 𝒜 (Prod.fst i) ⊗[R] ℬ (Prod.snd i)) +local notation "ℬ𝒜" => (fun i : ι × ι => ℬ (Prod.fst i) ⊗[R] 𝒜 (Prod.snd i)) + +/-- Auxliary construction used to build `TensorProduct.gradedComm`. + +This operates on direct sums of tensors instead of tensors of direct sums. -/ +def gradedCommAux : DirectSum _ 𝒜ℬ →ₗ[R] DirectSum _ ℬ𝒜 := by + refine DirectSum.toModule R _ _ fun i => ?_ + have o := DirectSum.lof R _ ℬ𝒜 i.swap + have s : ℤˣ := ((-1 : ℤˣ)^(i.1* i.2 : ι) : ℤˣ) + exact (s • o) ∘ₗ (TensorProduct.comm R _ _).toLinearMap + +@[simp] +theorem gradedCommAux_lof_tmul (i j : ι) (a : 𝒜 i) (b : ℬ j) : + gradedCommAux R 𝒜 ℬ (lof R _ 𝒜ℬ (i, j) (a ⊗ₜ b)) = + (-1 : ℤˣ)^(j * i) • lof R _ ℬ𝒜 (j, i) (b ⊗ₜ a) := by + rw [gradedCommAux] + dsimp + simp [mul_comm i j] + +@[simp] +theorem gradedCommAux_comp_gradedCommAux : + gradedCommAux R 𝒜 ℬ ∘ₗ gradedCommAux R ℬ 𝒜 = LinearMap.id := by + ext i a b + dsimp + rw [gradedCommAux_lof_tmul, LinearMap.map_smul_of_tower, gradedCommAux_lof_tmul, smul_smul, + mul_comm i.2 i.1, Int.units_mul_self, one_smul] + +/-- The braiding operation for tensor products of externally `ι`-graded algebras. + +This sends $a ⊗ b$ to $(-1)^{\deg a' \deg b} (b ⊗ a)$. -/ +def gradedComm : + (⨁ i, 𝒜 i) ⊗[R] (⨁ i, ℬ i) ≃ₗ[R] (⨁ i, ℬ i) ⊗[R] (⨁ i, 𝒜 i) := by + refine TensorProduct.directSum R 𝒜 ℬ ≪≫ₗ ?_ ≪≫ₗ (TensorProduct.directSum R ℬ 𝒜).symm + exact LinearEquiv.ofLinear (gradedCommAux _ _ _) (gradedCommAux _ _ _) + (gradedCommAux_comp_gradedCommAux _ _ _) (gradedCommAux_comp_gradedCommAux _ _ _) + +/-- The braiding is symmetric. -/ +@[simp] +theorem gradedComm_symm : (gradedComm R 𝒜 ℬ).symm = gradedComm R ℬ 𝒜 := by + rw [gradedComm, gradedComm, LinearEquiv.trans_symm, LinearEquiv.symm_symm] + ext + rfl + +theorem gradedComm_of_tmul_of (i j : ι) (a : 𝒜 i) (b : ℬ j) : + gradedComm R 𝒜 ℬ (lof R _ 𝒜 i a ⊗ₜ lof R _ ℬ j b) = + (-1 : ℤˣ)^(j * i) • (lof R _ ℬ _ b ⊗ₜ lof R _ 𝒜 _ a) := by + rw [gradedComm] + dsimp only [LinearEquiv.trans_apply, LinearEquiv.ofLinear_apply] + rw [TensorProduct.directSum_lof_tmul_lof, gradedCommAux_lof_tmul, Units.smul_def, + zsmul_eq_smul_cast R, map_smul, TensorProduct.directSum_symm_lof_tmul, + ← zsmul_eq_smul_cast, ← Units.smul_def] + +theorem gradedComm_tmul_of_zero (a : ⨁ i, 𝒜 i) (b : ℬ 0) : + gradedComm R 𝒜 ℬ (a ⊗ₜ lof R _ ℬ 0 b) = lof R _ ℬ _ b ⊗ₜ a := by + suffices + (gradedComm R 𝒜 ℬ).toLinearMap ∘ₗ + (TensorProduct.mk R (⨁ i, 𝒜 i) (⨁ i, ℬ i)).flip (lof R _ ℬ 0 b) = + TensorProduct.mk R _ _ (lof R _ ℬ 0 b) from + FunLike.congr_fun this a + ext i a + dsimp + rw [gradedComm_of_tmul_of, zero_mul, uzpow_zero, one_smul] + +theorem gradedComm_of_zero_tmul (a : 𝒜 0) (b : ⨁ i, ℬ i) : + gradedComm R 𝒜 ℬ (lof R _ 𝒜 0 a ⊗ₜ b) = b ⊗ₜ lof R _ 𝒜 _ a := by + suffices + (gradedComm R 𝒜 ℬ).toLinearMap ∘ₗ (TensorProduct.mk R (⨁ i, 𝒜 i) (⨁ i, ℬ i)) (lof R _ 𝒜 0 a) = + (TensorProduct.mk R _ _).flip (lof R _ 𝒜 0 a) from + FunLike.congr_fun this b + ext i b + dsimp + rw [gradedComm_of_tmul_of, mul_zero, uzpow_zero, one_smul] + +theorem gradedComm_tmul_one (a : ⨁ i, 𝒜 i) : gradedComm R 𝒜 ℬ (a ⊗ₜ 1) = 1 ⊗ₜ a := + gradedComm_tmul_of_zero _ _ _ _ _ + +theorem gradedComm_one_tmul (b : ⨁ i, ℬ i) : gradedComm R 𝒜 ℬ (1 ⊗ₜ b) = b ⊗ₜ 1 := + gradedComm_of_zero_tmul _ _ _ _ _ + +@[simp, nolint simpNF] -- linter times out +theorem gradedComm_one : gradedComm R 𝒜 ℬ 1 = 1 := + gradedComm_one_tmul _ _ _ _ + +theorem gradedComm_tmul_algebraMap (a : ⨁ i, 𝒜 i) (r : R) : + gradedComm R 𝒜 ℬ (a ⊗ₜ algebraMap R _ r) = algebraMap R _ r ⊗ₜ a := + gradedComm_tmul_of_zero _ _ _ _ _ + +theorem gradedComm_algebraMap_tmul (r : R) (b : ⨁ i, ℬ i) : + gradedComm R 𝒜 ℬ (algebraMap R _ r ⊗ₜ b) = b ⊗ₜ algebraMap R _ r := + gradedComm_of_zero_tmul _ _ _ _ _ + +theorem gradedComm_algebraMap (r : R) : + gradedComm R 𝒜 ℬ (algebraMap R _ r) = algebraMap R _ r := + (gradedComm_algebraMap_tmul R 𝒜 ℬ r 1).trans (Algebra.TensorProduct.algebraMap_apply' r).symm + +end gradedComm + +open TensorProduct (assoc map) in +/-- The multiplication operation for tensor products of externally `ι`-graded algebras. -/ +noncomputable irreducible_def gradedMul : + letI AB := DirectSum _ 𝒜 ⊗[R] DirectSum _ ℬ + letI : Module R AB := TensorProduct.leftModule + AB →ₗ[R] AB →ₗ[R] AB := by + refine TensorProduct.curry ?_ + refine map (LinearMap.mul' R (⨁ i, 𝒜 i)) (LinearMap.mul' R (⨁ i, ℬ i)) ∘ₗ ?_ + refine (assoc R _ _ _).symm.toLinearMap ∘ₗ .lTensor _ ?_ ∘ₗ (assoc R _ _ _).toLinearMap + refine (assoc R _ _ _).toLinearMap ∘ₗ .rTensor _ ?_ ∘ₗ (assoc R _ _ _).symm.toLinearMap + exact (gradedComm _ _ _).toLinearMap + +theorem tmul_of_gradedMul_of_tmul (j₁ i₂ : ι) + (a₁ : ⨁ i, 𝒜 i) (b₁ : ℬ j₁) (a₂ : 𝒜 i₂) (b₂ : ⨁ i, ℬ i) : + gradedMul R 𝒜 ℬ (a₁ ⊗ₜ lof R _ ℬ j₁ b₁) (lof R _ 𝒜 i₂ a₂ ⊗ₜ b₂) = + (-1 : ℤˣ)^(j₁ * i₂) • ((a₁ * lof R _ 𝒜 _ a₂) ⊗ₜ (lof R _ ℬ _ b₁ * b₂)) := by + rw [gradedMul] + dsimp only [curry_apply, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, assoc_tmul, + map_tmul, LinearMap.id_coe, id_eq, assoc_symm_tmul, LinearMap.rTensor_tmul, + LinearMap.lTensor_tmul] + rw [mul_comm j₁ i₂, gradedComm_of_tmul_of] + -- the tower smul lemmas elaborate too slowly + rw [Units.smul_def, Units.smul_def, zsmul_eq_smul_cast R, zsmul_eq_smul_cast R] + rw [← smul_tmul', map_smul, tmul_smul, map_smul, map_smul] + dsimp + +variable {R} + +theorem algebraMap_gradedMul (r : R) (x : (⨁ i, 𝒜 i) ⊗[R] (⨁ i, ℬ i)) : + gradedMul R 𝒜 ℬ (algebraMap R _ r ⊗ₜ 1) x = r • x := by + suffices gradedMul R 𝒜 ℬ (algebraMap R _ r ⊗ₜ 1) = DistribMulAction.toLinearMap R _ r by + exact FunLike.congr_fun this x + ext ia a ib b + dsimp + erw [tmul_of_gradedMul_of_tmul] + rw [zero_mul, uzpow_zero, one_smul, smul_tmul'] + erw [one_mul, _root_.Algebra.smul_def] + +theorem one_gradedMul (x : (⨁ i, 𝒜 i) ⊗[R] (⨁ i, ℬ i)) : + gradedMul R 𝒜 ℬ 1 x = x := by + simpa only [_root_.map_one, one_smul] using algebraMap_gradedMul 𝒜 ℬ 1 x + +theorem gradedMul_algebraMap (x : (⨁ i, 𝒜 i) ⊗[R] (⨁ i, ℬ i)) (r : R) : + gradedMul R 𝒜 ℬ x (algebraMap R _ r ⊗ₜ 1) = r • x := by + suffices (gradedMul R 𝒜 ℬ).flip (algebraMap R _ r ⊗ₜ 1) = DistribMulAction.toLinearMap R _ r by + exact FunLike.congr_fun this x + ext + dsimp + erw [tmul_of_gradedMul_of_tmul] + rw [mul_zero, uzpow_zero, one_smul, smul_tmul'] + erw [mul_one, _root_.Algebra.smul_def, Algebra.commutes] + rfl + +theorem gradedMul_one (x : (⨁ i, 𝒜 i) ⊗[R] (⨁ i, ℬ i)) : + gradedMul R 𝒜 ℬ x 1 = x := by + simpa only [_root_.map_one, one_smul] using gradedMul_algebraMap 𝒜 ℬ x 1 + +theorem gradedMul_assoc (x y z : DirectSum _ 𝒜 ⊗[R] DirectSum _ ℬ) : + gradedMul R 𝒜 ℬ (gradedMul R 𝒜 ℬ x y) z = gradedMul R 𝒜 ℬ x (gradedMul R 𝒜 ℬ y z) := by + let mA := gradedMul R 𝒜 ℬ + -- restate as an equality of morphisms so that we can use `ext` + suffices LinearMap.llcomp R _ _ _ mA ∘ₗ mA = + (LinearMap.llcomp R _ _ _ LinearMap.lflip <| LinearMap.llcomp R _ _ _ mA.flip ∘ₗ mA).flip by + exact FunLike.congr_fun (FunLike.congr_fun (FunLike.congr_fun this x) y) z + ext ixa xa ixb xb iya ya iyb yb iza za izb zb + dsimp + simp_rw [tmul_of_gradedMul_of_tmul, Units.smul_def, zsmul_eq_smul_cast R, + LinearMap.map_smul₂, LinearMap.map_smul, DirectSum.lof_eq_of, DirectSum.of_mul_of, + ← DirectSum.lof_eq_of R, tmul_of_gradedMul_of_tmul, DirectSum.lof_eq_of, ← DirectSum.of_mul_of, + ← DirectSum.lof_eq_of R, mul_assoc] + simp_rw [← zsmul_eq_smul_cast R, ← Units.smul_def, smul_smul, ← uzpow_add, add_mul, mul_add] + congr 2 + abel + +theorem gradedComm_gradedMul (x y : DirectSum _ 𝒜 ⊗[R] DirectSum _ ℬ) : + gradedComm R 𝒜 ℬ (gradedMul R 𝒜 ℬ x y) + = gradedMul R ℬ 𝒜 (gradedComm R 𝒜 ℬ x) (gradedComm R 𝒜 ℬ y) := by + suffices (gradedMul R 𝒜 ℬ).compr₂ (gradedComm R 𝒜 ℬ).toLinearMap + = (gradedMul R ℬ 𝒜 ∘ₗ (gradedComm R 𝒜 ℬ).toLinearMap).compl₂ + (gradedComm R 𝒜 ℬ).toLinearMap from + LinearMap.congr_fun₂ this x y + ext i₁ a₁ j₁ b₁ i₂ a₂ j₂ b₂ + dsimp + rw [gradedComm_of_tmul_of, gradedComm_of_tmul_of, tmul_of_gradedMul_of_tmul] + simp_rw [Units.smul_def, zsmul_eq_smul_cast R, map_smul, LinearMap.smul_apply] + simp_rw [← zsmul_eq_smul_cast R, ← Units.smul_def, DirectSum.lof_eq_of, DirectSum.of_mul_of, + ← DirectSum.lof_eq_of R, gradedComm_of_tmul_of, tmul_of_gradedMul_of_tmul, smul_smul, + DirectSum.lof_eq_of, ← DirectSum.of_mul_of, ← DirectSum.lof_eq_of R] + simp_rw [← uzpow_add, mul_add, add_mul, mul_comm i₁ j₂] + congr 1 + abel_nf + rw [two_nsmul, uzpow_add, uzpow_add, Int.units_mul_self, one_mul] + +end TensorProduct diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean new file mode 100644 index 0000000000000..8988e8863e9f3 --- /dev/null +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean @@ -0,0 +1,380 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.LinearAlgebra.TensorProduct.Graded.External +import Mathlib.RingTheory.GradedAlgebra.Basic + +/-! +# Graded tensor products over graded algebras + +The graded tensor product $A \hat\otimes_R B$ is imbued with a multiplication defined on homogeneous +tensors by: + +$$(a \otimes b) \cdot (a' \otimes b') = (-1)^{\deg a' \deg b} (a \cdot a') \otimes (b \cdot b')$$ + +where $A$ and $B$ are algebras graded by `ℕ`, `ℤ`, or `ι` (or more generally, any index +that satisfies `Module ι (Additive ℤˣ)`). + +## Main results + +* `GradedTensorProduct R 𝒜 ℬ`: for families of submodules of `A` and `B` that form a graded algebra, + this is a type alias for `A ⊗[R] B` with the appropriate multiplication. +* `GradedTensorProduct.instAlgebra`: the ring structure induced by this multiplication. +* `GradedTensorProduct.liftEquiv`: a universal property for graded tensor products + +## Notation + +* `𝒜 ᵍ⊗[R] ℬ` is notation for `GradedTensorProduct R 𝒜 ℬ`. +* `a ᵍ⊗ₜ b` is notation for `GradedTensorProduct.tmul _ a b`. + +## References + +* https://math.stackexchange.com/q/202718/1896 +* [*Algebra I*, Bourbaki : Chapter III, §4.7, example (2)][bourbaki1989] + +## Implementation notes + +We cannot put the multiplication on `A ⊗[R] B` directly as it would conflict with the existing +multiplication defined without the $(-1)^{\deg a' \deg b}$ term. Furthermore, the ring `A` may not +have a unique graduation, and so we need the chosen graduation `𝒜` to appear explicitly in the +type. + +## TODO + +* Show that the tensor product of graded algebras is itself a graded algebra. +* Determine if replacing the synonym with a single-field structure improves performance. +-/ + +suppress_compilation + +open scoped TensorProduct + +variable {R ι A B : Type*} + +variable [CommSemiring ι] [Module ι (Additive ℤˣ)] [DecidableEq ι] +variable [CommRing R] [Ring A] [Ring B] [Algebra R A] [Algebra R B] +variable (𝒜 : ι → Submodule R A) (ℬ : ι → Submodule R B) +variable [GradedAlgebra 𝒜] [GradedAlgebra ℬ] + +open DirectSum + + +variable (R) in +/-- A Type synonym for `A ⊗[R] B`, but with multiplication as `TensorProduct.gradedMul`. + +This has notation `𝒜 ᵍ⊗[R] ℬ`. -/ +@[nolint unusedArguments] +def GradedTensorProduct + (𝒜 : ι → Submodule R A) (ℬ : ι → Submodule R B) + [GradedAlgebra 𝒜] [GradedAlgebra ℬ] : + Type _ := + A ⊗[R] B + +namespace GradedTensorProduct + +open TensorProduct + +@[inherit_doc GradedTensorProduct] +scoped[TensorProduct] notation:100 𝒜 " ᵍ⊗[" R "] " ℬ:100 => GradedTensorProduct R 𝒜 ℬ + +instance instAddCommGroupWithOne : AddCommGroupWithOne (𝒜 ᵍ⊗[R] ℬ) := + Algebra.TensorProduct.instAddCommGroupWithOne +instance : Module R (𝒜 ᵍ⊗[R] ℬ) := TensorProduct.leftModule + +variable (R) in +/-- The casting equivalence to move between regular and graded tensor products. -/ +def of : A ⊗[R] B ≃ₗ[R] 𝒜 ᵍ⊗[R] ℬ := LinearEquiv.refl _ _ + +@[simp] +theorem of_one : of R 𝒜 ℬ 1 = 1 := rfl + +@[simp] +theorem of_symm_one : (of R 𝒜 ℬ).symm 1 = 1 := rfl + +-- for dsimp +@[simp, nolint simpNF] +theorem of_symm_of (x : A ⊗[R] B) : (of R 𝒜 ℬ).symm (of R 𝒜 ℬ x) = x := rfl + +-- for dsimp +@[simp, nolint simpNF] +theorem symm_of_of (x : 𝒜 ᵍ⊗[R] ℬ) : of R 𝒜 ℬ ((of R 𝒜 ℬ).symm x) = x := rfl + +/-- Two linear maps from the graded tensor product agree if they agree on the underlying tensor +product. -/ +@[ext] +theorem hom_ext {M} [AddCommMonoid M] [Module R M] ⦃f g : 𝒜 ᵍ⊗[R] ℬ →ₗ[R] M⦄ + (h : f ∘ₗ of R 𝒜 ℬ = (g ∘ₗ of R 𝒜 ℬ : A ⊗[R] B →ₗ[R] M)) : + f = g := + h + +variable (R) {𝒜 ℬ} in +/-- The graded tensor product of two elements of graded rings. -/ +abbrev tmul (a : A) (b : B) : 𝒜 ᵍ⊗[R] ℬ := of R 𝒜 ℬ (a ⊗ₜ b) + +@[inherit_doc] +notation:100 x " ᵍ⊗ₜ" y:100 => tmul _ x y + +@[inherit_doc] +notation:100 x " ᵍ⊗ₜ[" R "] " y:100 => tmul R x y + +variable (R) in +/-- An auxiliary construction to move between the graded tensor product of internally-graded objects +and the tensor product of direct sums. -/ +noncomputable def auxEquiv : (𝒜 ᵍ⊗[R] ℬ) ≃ₗ[R] (⨁ i, 𝒜 i) ⊗[R] (⨁ i, ℬ i) := + let fA := (decomposeAlgEquiv 𝒜).toLinearEquiv + let fB := (decomposeAlgEquiv ℬ).toLinearEquiv + (of R 𝒜 ℬ).symm.trans (TensorProduct.congr fA fB) + +@[simp] theorem auxEquiv_tmul (a : A) (b : B) : + auxEquiv R 𝒜 ℬ (a ᵍ⊗ₜ b) = decompose 𝒜 a ⊗ₜ decompose ℬ b := rfl + +@[simp] theorem auxEquiv_one : auxEquiv R 𝒜 ℬ 1 = 1 := by + rw [← of_one, Algebra.TensorProduct.one_def, auxEquiv_tmul 𝒜 ℬ, DirectSum.decompose_one, + DirectSum.decompose_one, Algebra.TensorProduct.one_def] + +@[simp, nolint simpNF] -- simpNF linter crashes +theorem auxEquiv_symm_one : (auxEquiv R 𝒜 ℬ).symm 1 = 1 := + (LinearEquiv.symm_apply_eq _).mpr (auxEquiv_one _ _).symm + +/-- Auxiliary construction used to build the `Mul` instance and get distributivity of `+` and +`\smul`. -/ +noncomputable def mulHom : (𝒜 ᵍ⊗[R] ℬ) →ₗ[R] (𝒜 ᵍ⊗[R] ℬ) →ₗ[R] (𝒜 ᵍ⊗[R] ℬ) := by + letI fAB1 := auxEquiv R 𝒜 ℬ + have := ((gradedMul R (𝒜 ·) (ℬ ·)).compl₁₂ fAB1.toLinearMap fAB1.toLinearMap).compr₂ + fAB1.symm.toLinearMap + exact this + +theorem mulHom_apply (x y : 𝒜 ᵍ⊗[R] ℬ) : + mulHom 𝒜 ℬ x y + = (auxEquiv R 𝒜 ℬ).symm (gradedMul R (𝒜 ·) (ℬ ·) (auxEquiv R 𝒜 ℬ x) (auxEquiv R 𝒜 ℬ y)) := + rfl + +/-- The multipication on the graded tensor product. + +See `GradedTensorProduct.coe_mul_coe` for a characterization on pure tensors. -/ +instance : Mul (𝒜 ᵍ⊗[R] ℬ) where mul x y := mulHom 𝒜 ℬ x y + +theorem mul_def (x y : 𝒜 ᵍ⊗[R] ℬ) : x * y = mulHom 𝒜 ℬ x y := rfl + +@[simp] +theorem auxEquiv_mul (x y : 𝒜 ᵍ⊗[R] ℬ) : + auxEquiv R 𝒜 ℬ (x * y) = gradedMul R (𝒜 ·) (ℬ ·) (auxEquiv R 𝒜 ℬ x) (auxEquiv R 𝒜 ℬ y) := + LinearEquiv.eq_symm_apply _ |>.mp rfl + +instance instMonoid : Monoid (𝒜 ᵍ⊗[R] ℬ) where + mul_one x := by + rw [mul_def, mulHom_apply, auxEquiv_one, gradedMul_one, LinearEquiv.symm_apply_apply] + one_mul x := by + rw [mul_def, mulHom_apply, auxEquiv_one, one_gradedMul, LinearEquiv.symm_apply_apply] + mul_assoc x y z := by + simp_rw [mul_def, mulHom_apply, LinearEquiv.apply_symm_apply] + rw [gradedMul_assoc] + +instance instRing : Ring (𝒜 ᵍ⊗[R] ℬ) where + __ := instAddCommGroupWithOne 𝒜 ℬ + __ := instMonoid 𝒜 ℬ + right_distrib x y z := by simp_rw [mul_def, LinearMap.map_add₂] + left_distrib x y z := by simp_rw [mul_def, map_add] + mul_zero x := by simp_rw [mul_def, map_zero] + zero_mul x := by simp_rw [mul_def, LinearMap.map_zero₂] + +/-- The characterization of this multiplication on partially homogenous elements. -/ +theorem tmul_coe_mul_coe_tmul {j₁ i₂ : ι} (a₁ : A) (b₁ : ℬ j₁) (a₂ : 𝒜 i₂) (b₂ : B) : + (a₁ ᵍ⊗ₜ[R] (b₁ : B) * (a₂ : A) ᵍ⊗ₜ[R] b₂ : 𝒜 ᵍ⊗[R] ℬ) = + (-1 : ℤˣ)^(j₁ * i₂) • ((a₁ * a₂ : A) ᵍ⊗ₜ (b₁ * b₂ : B)) := by + dsimp only [mul_def, mulHom_apply, of_symm_of] + dsimp [auxEquiv, tmul] + erw [decompose_coe, decompose_coe] + simp_rw [← lof_eq_of R] + rw [tmul_of_gradedMul_of_tmul] + simp_rw [lof_eq_of R] + rw [LinearEquiv.symm_symm] + rw [@Units.smul_def _ _ (_) (_), zsmul_eq_smul_cast R, map_smul, map_smul, + ← zsmul_eq_smul_cast R, ← @Units.smul_def _ _ (_) (_)] + rw [congr_symm_tmul] + dsimp + simp_rw [decompose_symm_mul, decompose_symm_of, Equiv.symm_apply_apply] + +/-- A special case for when `b₁` has grade 0. -/ +theorem tmul_zero_coe_mul_coe_tmul {i₂ : ι} (a₁ : A) (b₁ : ℬ 0) (a₂ : 𝒜 i₂) (b₂ : B) : + (a₁ ᵍ⊗ₜ[R] (b₁ : B) * (a₂ : A) ᵍ⊗ₜ[R] b₂ : 𝒜 ᵍ⊗[R] ℬ) = + ((a₁ * a₂ : A) ᵍ⊗ₜ (b₁ * b₂ : B)) := by + rw [tmul_coe_mul_coe_tmul, zero_mul, uzpow_zero, one_smul] + +/-- A special case for when `a₂` has grade 0. -/ +theorem tmul_coe_mul_zero_coe_tmul {j₁ : ι} (a₁ : A) (b₁ : ℬ j₁) (a₂ : 𝒜 0) (b₂ : B) : + (a₁ ᵍ⊗ₜ[R] (b₁ : B) * (a₂ : A) ᵍ⊗ₜ[R] b₂ : 𝒜 ᵍ⊗[R] ℬ) = + ((a₁ * a₂ : A) ᵍ⊗ₜ (b₁ * b₂ : B)) := by + rw [tmul_coe_mul_coe_tmul, mul_zero, uzpow_zero, one_smul] + +theorem tmul_one_mul_coe_tmul {i₂ : ι} (a₁ : A) (a₂ : 𝒜 i₂) (b₂ : B) : + (a₁ ᵍ⊗ₜ[R] (1 : B) * (a₂ : A) ᵍ⊗ₜ[R] b₂ : 𝒜 ᵍ⊗[R] ℬ) = (a₁ * a₂ : A) ᵍ⊗ₜ (b₂ : B) := by + convert tmul_zero_coe_mul_coe_tmul 𝒜 ℬ a₁ (@GradedMonoid.GOne.one _ (ℬ ·) _ _) a₂ b₂ + rw [SetLike.coe_gOne, one_mul] + +theorem tmul_coe_mul_one_tmul {j₁ : ι} (a₁ : A) (b₁ : ℬ j₁) (b₂ : B) : + (a₁ ᵍ⊗ₜ[R] (b₁ : B) * (1 : A) ᵍ⊗ₜ[R] b₂ : 𝒜 ᵍ⊗[R] ℬ) = (a₁ : A) ᵍ⊗ₜ (b₁ * b₂ : B) := by + convert tmul_coe_mul_zero_coe_tmul 𝒜 ℬ a₁ b₁ (@GradedMonoid.GOne.one _ (𝒜 ·) _ _) b₂ + rw [SetLike.coe_gOne, mul_one] + +theorem tmul_one_mul_one_tmul (a₁ : A) (b₂ : B) : + (a₁ ᵍ⊗ₜ[R] (1 : B) * (1 : A) ᵍ⊗ₜ[R] b₂ : 𝒜 ᵍ⊗[R] ℬ) = (a₁ : A) ᵍ⊗ₜ (b₂ : B) := by + convert tmul_coe_mul_zero_coe_tmul 𝒜 ℬ + a₁ (@GradedMonoid.GOne.one _ (ℬ ·) _ _) (@GradedMonoid.GOne.one _ (𝒜 ·) _ _) b₂ + · rw [SetLike.coe_gOne, mul_one] + · rw [SetLike.coe_gOne, one_mul] + +/-- The ring morphism `A →+* A ⊗[R] B` sending `a` to `a ⊗ₜ 1`. -/ +@[simps] +def includeLeftRingHom : A →+* 𝒜 ᵍ⊗[R] ℬ where + toFun a := a ᵍ⊗ₜ 1 + map_zero' := by simp + map_add' := by simp [tmul, TensorProduct.add_tmul] + map_one' := rfl + map_mul' a₁ a₂ := by + dsimp + classical + rw [← DirectSum.sum_support_decompose 𝒜 a₂, Finset.mul_sum] + simp_rw [tmul, sum_tmul, map_sum, Finset.mul_sum] + congr + ext i + rw [← SetLike.coe_gOne ℬ, tmul_coe_mul_coe_tmul, zero_mul, uzpow_zero, one_smul, + SetLike.coe_gOne, one_mul] + +instance instAlgebra : Algebra R (𝒜 ᵍ⊗[R] ℬ) where + toRingHom := (includeLeftRingHom 𝒜 ℬ).comp (algebraMap R A) + commutes' r x := by + dsimp [mul_def, mulHom_apply, auxEquiv_tmul] + simp_rw [DirectSum.decompose_algebraMap, DirectSum.decompose_one, algebraMap_gradedMul, + gradedMul_algebraMap] + smul_def' r x := by + dsimp [mul_def, mulHom_apply, auxEquiv_tmul] + simp_rw [DirectSum.decompose_algebraMap, DirectSum.decompose_one, algebraMap_gradedMul, + map_smul, LinearEquiv.symm_apply_apply] + +lemma algebraMap_def (r : R) : algebraMap R (𝒜 ᵍ⊗[R] ℬ) r = algebraMap R A r ᵍ⊗ₜ[R] 1 := rfl + +theorem tmul_algebraMap_mul_coe_tmul {i₂ : ι} (a₁ : A) (r : R) (a₂ : 𝒜 i₂) (b₂ : B) : + (a₁ ᵍ⊗ₜ[R] algebraMap R B r * (a₂ : A) ᵍ⊗ₜ[R] b₂ : 𝒜 ᵍ⊗[R] ℬ) + = (a₁ * a₂ : A) ᵍ⊗ₜ (algebraMap R B r * b₂ : B) := + tmul_zero_coe_mul_coe_tmul 𝒜 ℬ a₁ (GAlgebra.toFun (A := (ℬ ·)) r) a₂ b₂ + +theorem tmul_coe_mul_algebraMap_tmul {j₁ : ι} (a₁ : A) (b₁ : ℬ j₁) (r : R) (b₂ : B) : + (a₁ ᵍ⊗ₜ[R] (b₁ : B) * algebraMap R A r ᵍ⊗ₜ[R] b₂ : 𝒜 ᵍ⊗[R] ℬ) + = (a₁ * algebraMap R A r : A) ᵍ⊗ₜ (b₁ * b₂ : B) := + tmul_coe_mul_zero_coe_tmul 𝒜 ℬ a₁ b₁ (GAlgebra.toFun (A := (𝒜 ·)) r) b₂ + +/-- The algebra morphism `A →ₐ[R] A ⊗[R] B` sending `a` to `a ⊗ₜ 1`. -/ +@[simps!] +def includeLeft : A →ₐ[R] 𝒜 ᵍ⊗[R] ℬ where + toRingHom := includeLeftRingHom 𝒜 ℬ + commutes' _ := rfl + +/-- The algebra morphism `B →ₐ[R] A ⊗[R] B` sending `b` to `1 ⊗ₜ b`. -/ +@[simps!] +def includeRight : B →ₐ[R] (𝒜 ᵍ⊗[R] ℬ) := + AlgHom.ofLinearMap (R := R) (A := B) (B := 𝒜 ᵍ⊗[R] ℬ) + (f := { + toFun := fun b => 1 ᵍ⊗ₜ b + map_add' := by simp [tmul, TensorProduct.tmul_add] + map_smul' := by simp [tmul, TensorProduct.tmul_smul] }) + (map_one := rfl) + (map_mul := by + rw [LinearMap.map_mul_iff] + refine DirectSum.decompose_lhom_ext ℬ fun i₁ => ?_ + ext b₁ b₂ : 2 + dsimp + rw [tmul_coe_mul_one_tmul]) + +lemma algebraMap_def' (r : R) : algebraMap R (𝒜 ᵍ⊗[R] ℬ) r = 1 ᵍ⊗ₜ[R] algebraMap R B r := + (includeRight 𝒜 ℬ).commutes r |>.symm + +variable {C} [Ring C] [Algebra R C] + +/-- The forwards direction of the universal property; an algebra morphism out of the graded tensor +product can be assembed from maps on each component that (anti)commute on pure elements of the +corresponding graded algebras. -/ +def lift (f : A →ₐ[R] C) (g : B →ₐ[R] C) + (h_anti_commutes : ∀ ⦃i j⦄ (a : 𝒜 i) (b : ℬ j), f a * g b = (-1 : ℤˣ)^(j * i) • (g b * f a)) : + (𝒜 ᵍ⊗[R] ℬ) →ₐ[R] C := + AlgHom.ofLinearMap + (LinearMap.mul' R C + ∘ₗ (TensorProduct.map f.toLinearMap g.toLinearMap) + ∘ₗ ((of R 𝒜 ℬ).symm : 𝒜 ᵍ⊗[R] ℬ →ₗ[R] A ⊗[R] B)) + (by dsimp [Algebra.TensorProduct.one_def]; simp only [_root_.map_one, mul_one]) + (by + rw [LinearMap.map_mul_iff] + ext a₁ : 3 + refine DirectSum.decompose_lhom_ext ℬ fun j₁ => ?_ + ext b₁ : 3 + refine DirectSum.decompose_lhom_ext 𝒜 fun i₂ => ?_ + ext a₂ b₂ : 2 + dsimp + rw [tmul_coe_mul_coe_tmul] + rw [@Units.smul_def _ _ (_) (_), zsmul_eq_smul_cast R, map_smul, map_smul, map_smul] + rw [← zsmul_eq_smul_cast R, ← @Units.smul_def _ _ (_) (_)] + rw [of_symm_of, map_tmul, LinearMap.mul'_apply] + simp_rw [AlgHom.toLinearMap_apply, _root_.map_mul] + simp_rw [mul_assoc (f a₁), ← mul_assoc _ _ (g b₂), h_anti_commutes, mul_smul_comm, + smul_mul_assoc, smul_smul, Int.units_mul_self, one_smul]) + +@[simp] +theorem lift_tmul (f : A →ₐ[R] C) (g : B →ₐ[R] C) + (h_anti_commutes : ∀ ⦃i j⦄ (a : 𝒜 i) (b : ℬ j), f a * g b = (-1 : ℤˣ)^(j * i) • (g b * f a)) + (a : A) (b : B) : + lift 𝒜 ℬ f g h_anti_commutes (a ᵍ⊗ₜ b) = f a * g b := + rfl + +/-- The universal property of the graded tensor product; every algebra morphism uniquely factors +as a pair of algebra morphisms that anticommute with respect to the grading. -/ +def liftEquiv : + { fg : (A →ₐ[R] C) × (B →ₐ[R] C) // + ∀ ⦃i j⦄ (a : 𝒜 i) (b : ℬ j), fg.1 a * fg.2 b = (-1 : ℤˣ)^(j * i) • (fg.2 b * fg.1 a)} ≃ + ((𝒜 ᵍ⊗[R] ℬ) →ₐ[R] C) where + toFun fg := lift 𝒜 ℬ _ _ fg.prop + invFun F := ⟨(F.comp (includeLeft 𝒜 ℬ), F.comp (includeRight 𝒜 ℬ)), fun i j a b => by + dsimp + rw [← F.map_mul, ← F.map_mul, tmul_coe_mul_coe_tmul, one_mul, mul_one, AlgHom.map_smul_of_tower, + tmul_one_mul_one_tmul, smul_smul, Int.units_mul_self, one_smul]⟩ + left_inv fg := by ext <;> (dsimp; simp only [_root_.map_one, mul_one, one_mul]) + right_inv F := by + apply AlgHom.toLinearMap_injective + ext + dsimp + rw [← F.map_mul, tmul_one_mul_one_tmul] + +/-- Two algebra morphism from the graded tensor product agree if their compositions with the left +and right inclusions agree. -/ +@[ext] +lemma algHom_ext ⦃f g : (𝒜 ᵍ⊗[R] ℬ) →ₐ[R] C⦄ + (ha : f.comp (includeLeft 𝒜 ℬ) = g.comp (includeLeft 𝒜 ℬ)) + (hb : f.comp (includeRight 𝒜 ℬ) = g.comp (includeRight 𝒜 ℬ)) : f = g := + (liftEquiv 𝒜 ℬ).symm.injective <| Subtype.ext <| Prod.ext ha hb + +/-- The non-trivial symmetric braiding, sending $a \otimes b$ to +$(-1)^{\deg a' \deg b} (b \otimes a)$. -/ +def comm : (𝒜 ᵍ⊗[R] ℬ) ≃ₐ[R] (ℬ ᵍ⊗[R] 𝒜) := + AlgEquiv.ofLinearEquiv + (auxEquiv R 𝒜 ℬ ≪≫ₗ gradedComm R _ _ ≪≫ₗ (auxEquiv R ℬ 𝒜).symm) + (by + dsimp + simp_rw [auxEquiv_one, gradedComm_one, auxEquiv_symm_one]) + (fun x y => by + dsimp + simp_rw [auxEquiv_mul, gradedComm_gradedMul, LinearEquiv.symm_apply_eq, + ← gradedComm_gradedMul, auxEquiv_mul, LinearEquiv.apply_symm_apply, gradedComm_gradedMul]) + +@[simp] lemma auxEquiv_comm (x : 𝒜 ᵍ⊗[R] ℬ) : + auxEquiv R ℬ 𝒜 (comm 𝒜 ℬ x) = gradedComm R (𝒜 ·) (ℬ ·) (auxEquiv R 𝒜 ℬ x) := + LinearEquiv.eq_symm_apply _ |>.mp rfl + +@[simp] lemma comm_coe_tmul_coe {i j : ι} (a : 𝒜 i) (b : ℬ j) : + comm 𝒜 ℬ (a ᵍ⊗ₜ b) = (-1 : ℤˣ)^(j * i) • (b ᵍ⊗ₜ a : ℬ ᵍ⊗[R] 𝒜) := + (auxEquiv R ℬ 𝒜).injective <| by + simp_rw [auxEquiv_comm, auxEquiv_tmul, decompose_coe, ← lof_eq_of R, gradedComm_of_tmul_of, + @Units.smul_def _ _ (_) (_), zsmul_eq_smul_cast R, map_smul, auxEquiv_tmul, decompose_coe, + lof_eq_of] + +end GradedTensorProduct diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index 07b37c5bd9770..3a0489aaaa9b7 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -1107,6 +1107,19 @@ instance Prod.borelSpace [SecondCountableTopologyEither α β] : ⟨le_antisymm prod_le_borel_prod OpensMeasurableSpace.borel_le⟩ #align prod.borel_space Prod.borelSpace +/-- Given a measurable embedding to a Borel space which is also a topological embedding, then the +source space is also a Borel space. -/ +lemma MeasurableEmbedding.borelSpace {α β : Type*} [MeasurableSpace α] [TopologicalSpace α] + [MeasurableSpace β] [TopologicalSpace β] [hβ : BorelSpace β] {e : α → β} + (h'e : MeasurableEmbedding e) (h''e : Inducing e) : + BorelSpace α := by + constructor + have : MeasurableSpace.comap e (borel β) = ‹_› := by simpa [hβ.measurable_eq] using h'e.comap_eq + rw [← this, ← borel_comap, h''e.induced] + +instance _root_.ULift.instBorelSpace [BorelSpace α] : BorelSpace (ULift α) := + MeasurableEquiv.ulift.measurableEmbedding.borelSpace Homeomorph.ulift.inducing + instance DiscreteMeasurableSpace.toBorelSpace {α : Type*} [TopologicalSpace α] [DiscreteTopology α] [MeasurableSpace α] [DiscreteMeasurableSpace α] : BorelSpace α := by constructor; ext; simp [MeasurableSpace.measurableSet_generateFrom, measurableSet_discrete] @@ -1334,7 +1347,7 @@ theorem measurableSet_of_mem_nhdsWithin_Ioi {s : Set α} (h : ∀ x ∈ s, s ∈ lemma measurableSet_bddAbove_range {ι} [Countable ι] {f : ι → δ → α} (hf : ∀ i, Measurable (f i)) : MeasurableSet {b | BddAbove (range (fun i ↦ f i b))} := by rcases isEmpty_or_nonempty α with hα|hα - · have : ∀ b, range (fun i ↦ f i b) = ∅ := fun b ↦ Iff.mp toFinset_eq_empty rfl + · have : ∀ b, range (fun i ↦ f i b) = ∅ := fun b ↦ eq_empty_of_isEmpty _ simp [this] have A : ∀ (i : ι) (c : α), MeasurableSet {x | f i x ≤ c} := by intro i c diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index 291aa15ea12b8..6e6367e19080c 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -412,8 +412,9 @@ instance {α : ι → Type*} [∀ i, MeasureSpace (α i)] [∀ i, SigmaFinite (v SigmaFinite (volume : Measure (∀ i, α i)) := pi.sigmaFinite _ -theorem pi_of_empty {α : Type*} [IsEmpty α] {β : α → Type*} {m : ∀ a, MeasurableSpace (β a)} - (μ : ∀ a : α, Measure (β a)) (x : ∀ a, β a := isEmptyElim) : Measure.pi μ = dirac x := by +theorem pi_of_empty {α : Type*} [Fintype α] [IsEmpty α] {β : α → Type*} + {m : ∀ a, MeasurableSpace (β a)} (μ : ∀ a : α, Measure (β a)) (x : ∀ a, β a := isEmptyElim) : + Measure.pi μ = dirac x := by haveI : ∀ a, SigmaFinite (μ a) := isEmptyElim refine' pi_eq fun s _ => _ rw [Fintype.prod_empty, dirac_apply_of_mem] @@ -421,17 +422,16 @@ theorem pi_of_empty {α : Type*} [IsEmpty α] {β : α → Type*} {m : ∀ a, Me #align measure_theory.measure.pi_of_empty MeasureTheory.Measure.pi_of_empty @[simp] -theorem pi_empty_univ {α : Type*} {β : α → Type*} [IsEmpty α] {m : ∀ α, MeasurableSpace (β α)} - (μ : ∀ a : α, Measure (β a)) : Measure.pi μ (Set.univ) = 1 := by - rw [pi_of_empty, measure_univ] +theorem pi_empty_univ {α : Type*} [Fintype α] [IsEmpty α] {β : α → Type*} + {m : ∀ α, MeasurableSpace (β α)} (μ : ∀ a : α, Measure (β a)) : + Measure.pi μ (Set.univ) = 1 := by + rw [pi_of_empty, measure_univ] theorem pi_eval_preimage_null {i : ι} {s : Set (α i)} (hs : μ i s = 0) : Measure.pi μ (eval i ⁻¹' s) = 0 := by -- WLOG, `s` is measurable rcases exists_measurable_superset_of_null hs with ⟨t, hst, _, hμt⟩ - suffices : Measure.pi μ (eval i ⁻¹' t) = 0 - exact measure_mono_null (preimage_mono hst) this - clear! s + suffices Measure.pi μ (eval i ⁻¹' t) = 0 from measure_mono_null (preimage_mono hst) this -- Now rewrite it as `Set.pi`, and apply `pi_pi` rw [← univ_pi_update_univ, pi_pi] apply Finset.prod_eq_zero (Finset.mem_univ i) @@ -874,7 +874,7 @@ theorem volume_preserving_finTwoArrow (α : Type u) [MeasureSpace α] measurePreserving_finTwoArrow volume #align measure_theory.volume_preserving_fin_two_arrow MeasureTheory.volume_preserving_finTwoArrow -theorem measurePreserving_pi_empty {ι : Type u} {α : ι → Type v} [IsEmpty ι] +theorem measurePreserving_pi_empty {ι : Type u} {α : ι → Type v} [Fintype ι] [IsEmpty ι] {m : ∀ i, MeasurableSpace (α i)} (μ : ∀ i, Measure (α i)) : MeasurePreserving (MeasurableEquiv.ofUniqueOfUnique (∀ i, α i) Unit) (Measure.pi μ) (Measure.dirac ()) := by @@ -883,7 +883,7 @@ theorem measurePreserving_pi_empty {ι : Type u} {α : ι → Type v} [IsEmpty rw [Measure.pi_of_empty, Measure.map_dirac e.measurable] #align measure_theory.measure_preserving_pi_empty MeasureTheory.measurePreserving_pi_empty -theorem volume_preserving_pi_empty {ι : Type u} (α : ι → Type v) [IsEmpty ι] +theorem volume_preserving_pi_empty {ι : Type u} (α : ι → Type v) [Fintype ι] [IsEmpty ι] [∀ i, MeasureSpace (α i)] : MeasurePreserving (MeasurableEquiv.ofUniqueOfUnique (∀ i, α i) Unit) volume volume := measurePreserving_pi_empty fun _ => volume diff --git a/Mathlib/MeasureTheory/Covering/Differentiation.lean b/Mathlib/MeasureTheory/Covering/Differentiation.lean index 7358f734bdfbb..392596f7e5807 100644 --- a/Mathlib/MeasureTheory/Covering/Differentiation.lean +++ b/Mathlib/MeasureTheory/Covering/Differentiation.lean @@ -587,7 +587,6 @@ theorem withDensity_le_mul {s : Set α} (hs : MeasurableSet s) {t : ℝ≥0} (ht zpow_add₀ t_ne_zero'] conv_rhs => rw [← mul_one (t ^ n)] gcongr - · apply NNReal.zpow_pos t_ne_zero' -- TODO: positivity extension for `zpow` rw [zpow_neg_one] exact inv_lt_one ht calc diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index f9e1855e9d3ad..2da64473aadd6 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -132,8 +132,7 @@ instance haveLebesgueDecomposition_smul_right (μ ν : Measure α) [HaveLebesgue refine ⟨⟨μ.singularPart ν, r⁻¹ • μ.rnDeriv ν⟩, ?_, ?_, ?_⟩ · change Measurable (r⁻¹ • μ.rnDeriv ν) exact hmeas.const_smul _ - · refine MutuallySingular.mono_ac hsing AbsolutelyContinuous.rfl ?_ - exact absolutelyContinuous_of_le_smul le_rfl + · exact hsing.mono_ac AbsolutelyContinuous.rfl smul_absolutelyContinuous · have : r⁻¹ • rnDeriv μ ν = ((r⁻¹ : ℝ≥0) : ℝ≥0∞) • rnDeriv μ ν := by simp [ENNReal.smul_def] rw [this, withDensity_smul _ hmeas, ENNReal.smul_def r, withDensity_smul_measure, ← smul_assoc, smul_eq_mul, ENNReal.coe_inv hr, ENNReal.inv_mul_cancel, one_smul] @@ -388,8 +387,8 @@ theorem singularPart_smul_right (μ ν : Measure α) (r : ℝ≥0) (hr : r ≠ 0 μ.singularPart (r • ν) = μ.singularPart ν := by by_cases hl : HaveLebesgueDecomposition μ ν · refine (eq_singularPart ((measurable_rnDeriv μ ν).const_smul r⁻¹) ?_ ?_).symm - · refine (mutuallySingular_singularPart μ ν).mono_ac AbsolutelyContinuous.rfl ?_ - exact absolutelyContinuous_of_le_smul le_rfl + · exact (mutuallySingular_singularPart μ ν).mono_ac AbsolutelyContinuous.rfl + smul_absolutelyContinuous · rw [ENNReal.smul_def r, withDensity_smul_measure, ← withDensity_smul] swap; · exact (measurable_rnDeriv _ _).const_smul _ convert haveLebesgueDecomposition_add μ ν @@ -546,11 +545,8 @@ See also `rnDeriv_smul_right'`, which requires sigma-finite `ν` and `μ`. -/ theorem rnDeriv_smul_right (ν μ : Measure α) [IsFiniteMeasure ν] [ν.HaveLebesgueDecomposition μ] {r : ℝ≥0} (hr : r ≠ 0) : ν.rnDeriv (r • μ) =ᵐ[μ] r⁻¹ • ν.rnDeriv μ := by - suffices ν.rnDeriv (r • μ) =ᵐ[r • μ] r⁻¹ • ν.rnDeriv μ by - suffices hμ : μ ≪ r • μ by exact hμ.ae_le this - refine absolutelyContinuous_of_le_smul (c := r⁻¹) ?_ - rw [← ENNReal.coe_inv hr, ← ENNReal.smul_def, ← smul_assoc, smul_eq_mul, - inv_mul_cancel hr, one_smul] + refine (absolutelyContinuous_smul $ ENNReal.coe_ne_zero.2 hr).ae_le + (?_ : ν.rnDeriv (r • μ) =ᵐ[r • μ] r⁻¹ • ν.rnDeriv μ) rw [← withDensity_eq_iff] rotate_left · exact (measurable_rnDeriv _ _).aemeasurable @@ -1019,11 +1015,8 @@ See also `rnDeriv_smul_right`, which has no hypothesis on `μ` but requires fini theorem rnDeriv_smul_right' (ν μ : Measure α) [SigmaFinite ν] [SigmaFinite μ] {r : ℝ≥0} (hr : r ≠ 0) : ν.rnDeriv (r • μ) =ᵐ[μ] r⁻¹ • ν.rnDeriv μ := by - suffices ν.rnDeriv (r • μ) =ᵐ[r • μ] r⁻¹ • ν.rnDeriv μ by - suffices hμ : μ ≪ r • μ by exact hμ.ae_le this - refine absolutelyContinuous_of_le_smul (c := r⁻¹) ?_ - rw [← ENNReal.coe_inv hr, ← ENNReal.smul_def, ← smul_assoc, smul_eq_mul, - inv_mul_cancel hr, one_smul] + refine (absolutelyContinuous_smul $ ENNReal.coe_ne_zero.2 hr).ae_le + (?_ : ν.rnDeriv (r • μ) =ᵐ[r • μ] r⁻¹ • ν.rnDeriv μ) rw [← withDensity_eq_iff_of_sigmaFinite] · simp_rw [ENNReal.smul_def] rw [withDensity_smul _ (measurable_rnDeriv _ _)] diff --git a/Mathlib/MeasureTheory/Group/AddCircle.lean b/Mathlib/MeasureTheory/Group/AddCircle.lean index 861b5ac182994..6f2dd41d25695 100644 --- a/Mathlib/MeasureTheory/Group/AddCircle.lean +++ b/Mathlib/MeasureTheory/Group/AddCircle.lean @@ -74,7 +74,7 @@ theorem isAddFundamentalDomain_of_ae_ball (I : Set <| AddCircle T) (u x : AddCir rw [dist_eq_norm, add_sub_cancel, div_mul_eq_div_div, ← add_div, ← add_div, add_self_div_two, div_le_iff' (by positivity : 0 < (n : ℝ)), ← nsmul_eq_mul] refine' (le_add_order_smul_norm_of_isOfFinAddOrder (hu.of_mem_zmultiples hg) hg').trans - (nsmul_le_nsmul (norm_nonneg g) _) + (nsmul_le_nsmul_left (norm_nonneg g) _) exact Nat.le_of_dvd (addOrderOf_pos_iff.mpr hu) (addOrderOf_dvd_of_mem_zmultiples hg) · -- `∀ (g : G), QuasiMeasurePreserving (VAdd.vadd g) volume volume` exact fun g => quasiMeasurePreserving_add_left (G := AddCircle T) volume g diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index 864de31622f76..91450aa3457ae 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -252,12 +252,12 @@ theorem tendsto_set_integral_pow_smul_of_unique_maximum_of_isCompact_of_measure_ exact lt_of_le_of_lt (measure_mono (inter_subset_right _ _)) hs.measure_lt_top · exact (I n).mono (inter_subset_right _ _) le_rfl · intro x hx - exact pow_le_pow_of_le_left t'_pos.le (le_of_lt (hv hx)) _ + exact pow_le_pow_left t'_pos.le (le_of_lt (hv hx)) _ _ ≤ ∫ y in s, c y ^ n ∂μ := set_integral_mono_set (I n) (J n) (eventually_of_forall (inter_subset_right _ _)) simp_rw [← div_eq_inv_mul, div_pow, div_div] apply div_le_div (pow_nonneg t_pos n) _ _ B - · exact pow_le_pow_of_le_left (hnc _ hx.1) (ht x hx) _ + · exact pow_le_pow_left (hnc _ hx.1) (ht x hx) _ · apply mul_pos (pow_pos (t_pos.trans_lt tt') _) (ENNReal.toReal_pos (hμ v v_open x₀_v).ne' _) have : μ (v ∩ s) ≤ μ s := measure_mono (inter_subset_right _ _) exact ne_of_lt (lt_of_le_of_lt this hs.measure_lt_top) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 1a4cbb71b804b..b26484b265459 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -449,6 +449,22 @@ theorem measurable_unit [MeasurableSpace α] (f : Unit → α) : Measurable f := measurable_from_top #align measurable_unit measurable_unit +section ULift +variable [MeasurableSpace α] + +instance _root_.ULift.instMeasurableSpace : MeasurableSpace (ULift α) := + ‹MeasurableSpace α›.map ULift.up + +lemma measurable_down : Measurable (ULift.down : ULift α → α) := fun _ ↦ id +lemma measurable_up : Measurable (ULift.up : α → ULift α) := fun _ ↦ id + +@[simp] lemma measurableSet_preimage_down {s : Set α} : + MeasurableSet (ULift.down ⁻¹' s) ↔ MeasurableSet s := Iff.rfl +@[simp] lemma measurableSet_preimage_up {s : Set (ULift α)} : + MeasurableSet (ULift.up ⁻¹' s) ↔ MeasurableSet s := Iff.rfl + +end ULift + section Nat variable [MeasurableSpace α] @@ -1541,6 +1557,10 @@ protected def cast {α β} [i₁ : MeasurableSpace α] [i₂ : MeasurableSpace exact measurable_id #align measurable_equiv.cast MeasurableEquiv.cast +/-- Measurable equivalence between `ULift α` and `α`. -/ +def ulift.{u, v} {α : Type u} [MeasurableSpace α] : ULift.{v, u} α ≃ᵐ α := + ⟨Equiv.ulift, measurable_down, measurable_up⟩ + protected theorem measurable_comp_iff {f : β → γ} (e : α ≃ᵐ β) : Measurable (f ∘ e) ↔ Measurable f := Iff.intro diff --git a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean index 8058435dc2ced..0cebd49896375 100644 --- a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean @@ -51,11 +51,6 @@ lemma aemeasurable_of_map_neZero {mβ : MeasurableSpace β} {μ : Measure α} namespace AEMeasurable -protected theorem nullMeasurable (h : AEMeasurable f μ) : NullMeasurable f μ := - let ⟨_g, hgm, hg⟩ := h - hgm.nullMeasurable.congr hg.symm -#align ae_measurable.null_measurable AEMeasurable.nullMeasurable - lemma mono_ac (hf : AEMeasurable f ν) (hμν : μ ≪ ν) : AEMeasurable f μ := ⟨hf.mk f, hf.measurable_mk, hμν.ae_le hf.ae_eq_mk⟩ diff --git a/Mathlib/MeasureTheory/Measure/Dirac.lean b/Mathlib/MeasureTheory/Measure/Dirac.lean index cecc9e24da629..ffead1aa256a9 100644 --- a/Mathlib/MeasureTheory/Measure/Dirac.lean +++ b/Mathlib/MeasureTheory/Measure/Dirac.lean @@ -13,14 +13,12 @@ In this file we define the Dirac measure `MeasureTheory.Measure.dirac a` and prove some basic facts about it. -/ -set_option autoImplicit true - open Function Set open scoped ENNReal Classical noncomputable section -variable [MeasurableSpace α] [MeasurableSpace β] {s : Set α} +variable {α β δ : Type*} [MeasurableSpace α] [MeasurableSpace β] {s : Set α} {a : α} namespace MeasureTheory @@ -140,11 +138,15 @@ theorem ae_eq_dirac [MeasurableSingletonClass α] {a : α} (f : α → δ) : f =ᵐ[dirac a] const α (f a) := by simp [Filter.EventuallyEq] #align measure_theory.ae_eq_dirac MeasureTheory.ae_eq_dirac -instance Measure.dirac.isProbabilityMeasure [MeasurableSpace α] {x : α} : - IsProbabilityMeasure (dirac x) := +instance Measure.dirac.isProbabilityMeasure {x : α} : IsProbabilityMeasure (dirac x) := ⟨dirac_apply_of_mem <| mem_univ x⟩ #align measure_theory.measure.dirac.is_probability_measure MeasureTheory.Measure.dirac.isProbabilityMeasure +/-! Extra instances to short-circuit type class resolution -/ + +instance Measure.dirac.instIsFiniteMeasure {a : α} : IsFiniteMeasure (dirac a) := inferInstance +instance Measure.dirac.instSigmaFinite {a : α} : SigmaFinite (dirac a) := inferInstance + theorem restrict_dirac' (hs : MeasurableSet s) [Decidable (a ∈ s)] : (Measure.dirac a).restrict s = if a ∈ s then Measure.dirac a else 0 := by split_ifs with has diff --git a/Mathlib/MeasureTheory/Measure/Hausdorff.lean b/Mathlib/MeasureTheory/Measure/Hausdorff.lean index 9dcee3830bdd6..a01a11bffec8f 100644 --- a/Mathlib/MeasureTheory/Measure/Hausdorff.lean +++ b/Mathlib/MeasureTheory/Measure/Hausdorff.lean @@ -1019,7 +1019,7 @@ theorem hausdorffMeasure_pi_real {ι : Type*} [Fintype ι] : apply Finset.sum_le_sum fun i _ => _ simp only [ENNReal.rpow_nat_cast] intros i _ - exact pow_le_pow_of_le_left' (hn i) _ + exact pow_le_pow_left' (hn i) _ · isBoundedDefault _ = liminf (fun n : ℕ => ∏ i : ι, (⌈((b i : ℝ) - a i) * n⌉₊ : ℝ≥0∞) / n) atTop := by simp only [Finset.card_univ, Nat.cast_prod, one_mul, Fintype.card_fin, Finset.sum_const, diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index 7b4f47521df6d..03dc54465341c 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -193,7 +193,7 @@ theorem addHaar_submodule {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] convert s.sub_mem hym hyn using 1 simp only [sub_smul, neg_sub_neg, add_sub_add_right_eq_sub] have H : c ^ n - c ^ m ≠ 0 := by - simpa only [sub_eq_zero, Ne.def] using (strictAnti_pow cpos cone).injective.ne hmn.symm + simpa only [sub_eq_zero, Ne.def] using (pow_right_strictAnti cpos cone).injective.ne hmn.symm have : x ∈ s := by convert s.smul_mem (c ^ n - c ^ m)⁻¹ A rw [smul_smul, inv_mul_cancel H, one_smul] @@ -448,7 +448,7 @@ theorem addHaar_ball_of_pos (x : E) {r : ℝ} (hr : 0 < r) : theorem addHaar_ball_mul [Nontrivial E] (x : E) {r : ℝ} (hr : 0 ≤ r) (s : ℝ) : μ (ball x (r * s)) = ENNReal.ofReal (r ^ finrank ℝ E) * μ (ball 0 s) := by rcases hr.eq_or_lt with (rfl | h) - · simp only [zero_pow (finrank_pos (K := ℝ) (V := E)), measure_empty, zero_mul, + · simp only [zero_pow (finrank_pos (R := ℝ) (V := E)), measure_empty, zero_mul, ENNReal.ofReal_zero, ball_zero] · exact addHaar_ball_mul_of_pos μ x h s #align measure_theory.measure.add_haar_ball_mul MeasureTheory.Measure.addHaar_ball_mul diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index c7f1103d074b8..44c3e38c25fa5 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -211,6 +211,10 @@ theorem tsum_measure_preimage_singleton {s : Set β} (hs : s.Countable) {f : α rw [← Set.biUnion_preimage_singleton, measure_biUnion hs (pairwiseDisjoint_fiber f s) hf] #align measure_theory.tsum_measure_preimage_singleton MeasureTheory.tsum_measure_preimage_singleton +lemma measure_preimage_eq_zero_iff_of_countable {s : Set β} {f : α → β} (hs : s.Countable) : + μ (f ⁻¹' s) = 0 ↔ ∀ x ∈ s, μ (f ⁻¹' {x}) = 0 := by + rw [← biUnion_preimage_singleton, measure_biUnion_null_iff hs] + /-- If `s` is a `Finset`, then the measure of its preimage can be found as the sum of measures of the fibers `f ⁻¹' {y}`. -/ theorem sum_measure_preimage_singleton (s : Finset β) {f : α → β} @@ -287,11 +291,20 @@ theorem measure_eq_measure_larger_of_between_null_diff {s₁ s₂ s₃ : Set α} (measure_eq_measure_of_between_null_diff h12 h23 h_nulldiff).2 #align measure_theory.measure_eq_measure_larger_of_between_null_diff MeasureTheory.measure_eq_measure_larger_of_between_null_diff -theorem measure_compl (h₁ : MeasurableSet s) (h_fin : μ s ≠ ∞) : μ sᶜ = μ univ - μ s := by - rw [compl_eq_univ_diff] - exact measure_diff (subset_univ s) h₁ h_fin +lemma measure_compl₀ (h : NullMeasurableSet s μ) (hs : μ s ≠ ∞) : + μ sᶜ = μ Set.univ - μ s := by + rw [← measure_add_measure_compl₀ h, ENNReal.add_sub_cancel_left hs] + +theorem measure_compl (h₁ : MeasurableSet s) (h_fin : μ s ≠ ∞) : μ sᶜ = μ univ - μ s := + measure_compl₀ h₁.nullMeasurableSet h_fin #align measure_theory.measure_compl MeasureTheory.measure_compl +lemma measure_inter_conull' (ht : μ (s \ t) = 0) : μ (s ∩ t) = μ s := by + rw [← diff_compl, measure_diff_null']; rwa [← diff_eq] + +lemma measure_inter_conull (ht : μ tᶜ = 0) : μ (s ∩ t) = μ s := by + rw [← diff_compl, measure_diff_null ht] + @[simp] theorem union_ae_eq_left_iff_ae_subset : (s ∪ t : Set α) =ᵐ[μ] s ↔ t ≤ᵐ[μ] s := by rw [ae_le_set] @@ -852,6 +865,10 @@ instance instIsCentralScalar [SMul Rᵐᵒᵖ ℝ≥0∞] [IsCentralScalar R ℝ end SMul +instance instNoZeroSMulDivisors [Zero R] [SMulWithZero R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] + [NoZeroSMulDivisors R ℝ≥0∞] : NoZeroSMulDivisors R (Measure α) where + eq_zero_or_eq_zero_of_smul_eq_zero h := by simpa [Ne.def, @ext_iff', forall_or_left] using h + instance instMulAction [Monoid R] [MulAction R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] [MeasurableSpace α] : MulAction R (Measure α) := Injective.mulAction _ toOuterMeasure_injective smul_toOuterMeasure @@ -1219,6 +1236,8 @@ protected theorem map_smul_nnreal (c : ℝ≥0) (μ : Measure α) (f : α → β μ.map_smul (c : ℝ≥0∞) f #align measure_theory.measure.map_smul_nnreal MeasureTheory.Measure.map_smul_nnreal +variable {f : α → β} + lemma map_apply₀ {f : α → β} (hf : AEMeasurable f μ) {s : Set β} (hs : NullMeasurableSet s (map f μ)) : μ.map f s = μ (f ⁻¹' s) := by rw [map, dif_pos hf, mapₗ, dif_pos hf.measurable_mk] at hs ⊢ @@ -1228,24 +1247,33 @@ lemma map_apply₀ {f : α → β} (hf : AEMeasurable f μ) {s : Set β} /-- We can evaluate the pushforward on measurable sets. For non-measurable sets, see `MeasureTheory.Measure.le_map_apply` and `MeasurableEquiv.map_apply`. -/ @[simp] -theorem map_apply_of_aemeasurable {f : α → β} (hf : AEMeasurable f μ) {s : Set β} - (hs : MeasurableSet s) : μ.map f s = μ (f ⁻¹' s) := - map_apply₀ hf hs.nullMeasurableSet +theorem map_apply_of_aemeasurable (hf : AEMeasurable f μ) {s : Set β} (hs : MeasurableSet s) : + μ.map f s = μ (f ⁻¹' s) := map_apply₀ hf hs.nullMeasurableSet #align measure_theory.measure.map_apply_of_ae_measurable MeasureTheory.Measure.map_apply_of_aemeasurable @[simp] -theorem map_apply {f : α → β} (hf : Measurable f) {s : Set β} (hs : MeasurableSet s) : +theorem map_apply (hf : Measurable f) {s : Set β} (hs : MeasurableSet s) : μ.map f s = μ (f ⁻¹' s) := map_apply_of_aemeasurable hf.aemeasurable hs #align measure_theory.measure.map_apply MeasureTheory.Measure.map_apply -theorem map_toOuterMeasure {f : α → β} (hf : AEMeasurable f μ) : +theorem map_toOuterMeasure (hf : AEMeasurable f μ) : (μ.map f).toOuterMeasure = (OuterMeasure.map f μ.toOuterMeasure).trim := by rw [← trimmed, OuterMeasure.trim_eq_trim_iff] intro s hs rw [map_apply_of_aemeasurable hf hs, OuterMeasure.map_apply] #align measure_theory.measure.map_to_outer_measure MeasureTheory.Measure.map_toOuterMeasure +@[simp] lemma map_eq_zero_iff (hf : AEMeasurable f μ) : μ.map f = 0 ↔ μ = 0 := by + simp_rw [← measure_univ_eq_zero, map_apply_of_aemeasurable hf .univ, preimage_univ] + +@[simp] lemma mapₗ_eq_zero_iff (hf : Measurable f) : Measure.mapₗ f μ = 0 ↔ μ = 0 := by + rw [mapₗ_apply_of_measurable hf, map_eq_zero_iff hf.aemeasurable] + +lemma map_ne_zero_iff (hf : AEMeasurable f μ) : μ.map f ≠ 0 ↔ μ ≠ 0 := (map_eq_zero_iff hf).not +lemma mapₗ_ne_zero_iff (hf : Measurable f) : Measure.mapₗ f μ ≠ 0 ↔ μ ≠ 0 := + (mapₗ_eq_zero_iff hf).not + @[simp] theorem map_id : map id μ = μ := ext fun _ => map_apply measurable_id @@ -1615,11 +1643,19 @@ lemma add_right (h1 : μ ≪ ν) (ν' : Measure α) : μ ≪ ν + ν' := by end AbsolutelyContinuous +alias absolutelyContinuous_refl := AbsolutelyContinuous.refl +alias absolutelyContinuous_rfl := AbsolutelyContinuous.rfl + theorem absolutelyContinuous_of_le_smul {μ' : Measure α} {c : ℝ≥0∞} (hμ'_le : μ' ≤ c • μ) : μ' ≪ μ := (Measure.absolutelyContinuous_of_le hμ'_le).trans (Measure.AbsolutelyContinuous.rfl.smul c) #align measure_theory.measure.absolutely_continuous_of_le_smul MeasureTheory.Measure.absolutelyContinuous_of_le_smul +lemma smul_absolutelyContinuous {c : ℝ≥0∞} : c • μ ≪ μ := absolutelyContinuous_of_le_smul le_rfl + +lemma absolutelyContinuous_smul {c : ℝ≥0∞} (hc : c ≠ 0) : μ ≪ c • μ := by + simp [AbsolutelyContinuous, hc] + theorem ae_le_iff_absolutelyContinuous : μ.ae ≤ ν.ae ↔ μ ≪ ν := ⟨fun h s => by rw [measure_zero_iff_ae_nmem, measure_zero_iff_ae_nmem] @@ -1870,6 +1906,15 @@ open Measure open MeasureTheory +protected theorem _root_.AEMeasurable.nullMeasurable {f : α → β} (h : AEMeasurable f μ) : + NullMeasurable f μ := + let ⟨_g, hgm, hg⟩ := h; hgm.nullMeasurable.congr hg.symm +#align ae_measurable.null_measurable AEMeasurable.nullMeasurable + +lemma _root_.AEMeasurable.nullMeasurableSet_preimage {f : α → β} {s : Set β} + (hf : AEMeasurable f μ) (hs : MeasurableSet s) : NullMeasurableSet (f ⁻¹' s) μ := + hf.nullMeasurable hs + /-- The preimage of a null measurable set under a (quasi) measure preserving map is a null measurable set. -/ theorem NullMeasurableSet.preimage {ν : Measure β} {f : α → β} {t : Set β} diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean index 35e842e781051..a68d762f3c371 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean @@ -295,6 +295,9 @@ theorem measure_sUnion_null_iff {S : Set (Set α)} (hS : S.Countable) : μ.toOuterMeasure.sUnion_null_iff hS #align measure_theory.measure_sUnion_null_iff MeasureTheory.measure_sUnion_null_iff +lemma measure_null_iff_singleton {s : Set α} (hs : s.Countable) : μ s = 0 ↔ ∀ x ∈ s, μ {x} = 0 := by + rw [← measure_biUnion_null_iff hs, biUnion_of_singleton] + theorem measure_union_le (s₁ s₂ : Set α) : μ (s₁ ∪ s₂) ≤ μ s₁ + μ s₂ := μ.toOuterMeasure.union _ _ #align measure_theory.measure_union_le MeasureTheory.measure_union_le @@ -428,7 +431,11 @@ theorem all_ae_of {ι : Sort _} {p : α → ι → Prop} (hp : ∀ᵐ a ∂μ, ∀ᵐ a ∂μ, p a i := by filter_upwards [hp] with a ha using ha i -theorem ae_ball_iff {S : Set ι} (hS : S.Countable) {p : ∀ (_x : α), ∀ i ∈ S, Prop} : +lemma ae_iff_of_countable [Countable α] {p : α → Prop} : (∀ᵐ x ∂μ, p x) ↔ ∀ x, μ {x} ≠ 0 → p x := by + rw [ae_iff, measure_null_iff_singleton] + exacts [forall_congr' fun _ ↦ not_imp_comm, Set.to_countable _] + +theorem ae_ball_iff {S : Set ι} (hS : S.Countable) {p : α → ∀ i ∈ S, Prop} : (∀ᵐ x ∂μ, ∀ i (hi : i ∈ S), p x i hi) ↔ ∀ i (hi : i ∈ S), ∀ᵐ x ∂μ, p x i hi := eventually_countable_ball hS #align measure_theory.ae_ball_iff MeasureTheory.ae_ball_iff diff --git a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean index d2d7e1cc86237..a6450db04283a 100644 --- a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean @@ -135,6 +135,8 @@ instance (μ : ProbabilityMeasure Ω) : IsProbabilityMeasure (μ : Measure Ω) : -- porting note: syntactic tautology because of the way coercions work in Lean 4 #noalign measure_theory.probability_measure.coe_fn_eq_to_nnreal_coe_fn_to_measure +@[simp, norm_cast] lemma coe_mk (μ : Measure Ω) (hμ) : toMeasure ⟨μ, hμ⟩ = μ := rfl + @[simp] theorem val_eq_to_measure (ν : ProbabilityMeasure Ω) : ν.val = (ν : Measure Ω) := rfl diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index f957f76d25591..bc3dacabea3b3 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -21,7 +21,7 @@ We introduce the following typeclasses for measures: -/ open scoped ENNReal NNReal Topology -open Set MeasureTheory Measure Filter MeasurableSpace ENNReal +open Set MeasureTheory Measure Filter Function MeasurableSpace ENNReal variable {α β δ ι : Type*} @@ -218,6 +218,9 @@ export MeasureTheory.IsProbabilityMeasure (measure_univ) attribute [simp] IsProbabilityMeasure.measure_univ +lemma isProbabilityMeasure_iff : IsProbabilityMeasure μ ↔ μ univ = 1 := + ⟨fun _ ↦ measure_univ, IsProbabilityMeasure.mk⟩ + instance (priority := 100) IsProbabilityMeasure.toIsFiniteMeasure (μ : Measure α) [IsProbabilityMeasure μ] : IsFiniteMeasure μ := ⟨by simp only [measure_univ, ENNReal.one_lt_top]⟩ @@ -276,6 +279,26 @@ theorem prob_compl_eq_one_iff [IsProbabilityMeasure μ] (hs : MeasurableSet s) : μ sᶜ = 1 ↔ μ s = 0 := by rw [← prob_compl_eq_zero_iff hs.compl, compl_compl] #align measure_theory.prob_compl_eq_one_iff MeasureTheory.prob_compl_eq_one_iff +variable [IsProbabilityMeasure μ] {p : α → Prop} {f : β → α} + +lemma mem_ae_iff_prob_eq_one (hs : MeasurableSet s) : s ∈ μ.ae ↔ μ s = 1 := + mem_ae_iff.trans $ prob_compl_eq_zero_iff hs + +lemma ae_iff_prob_eq_one (hp : Measurable p) : (∀ᵐ a ∂μ, p a) ↔ μ {a | p a} = 1 := + mem_ae_iff_prob_eq_one hp.setOf + +lemma isProbabilityMeasure_comap (hf : Injective f) (hf' : ∀ᵐ a ∂μ, a ∈ range f) + (hf'' : ∀ s, MeasurableSet s → MeasurableSet (f '' s)) : + IsProbabilityMeasure (μ.comap f) where + measure_univ := by + rw [comap_apply _ hf hf'' _ MeasurableSet.univ, + ← mem_ae_iff_prob_eq_one (hf'' _ MeasurableSet.univ)] + simpa + +protected lemma _root_.MeasurableEmbedding.isProbabilityMeasure_comap (hf : MeasurableEmbedding f) + (hf' : ∀ᵐ a ∂μ, a ∈ range f) : IsProbabilityMeasure (μ.comap f) := + isProbabilityMeasure_comap hf.injective hf' hf.measurableSet_image' + end IsProbabilityMeasure section NoAtoms diff --git a/Mathlib/ModelTheory/Basic.lean b/Mathlib/ModelTheory/Basic.lean index bd65a19d64dd6..24b4de9f95466 100644 --- a/Mathlib/ModelTheory/Basic.lean +++ b/Mathlib/ModelTheory/Basic.lean @@ -96,7 +96,7 @@ theorem lift_mk {i : ℕ} : Cardinal.lift.{v,u} #(Sequence₂ a₀ a₁ a₂ i) = #(Sequence₂ (ULift.{v,u} a₀) (ULift.{v,u} a₁) (ULift.{v,u} a₂) i) := by rcases i with (_ | _ | _ | i) <;> - simp only [Sequence₂, mk_uLift, mk_fintype, Fintype.card_of_isEmpty, Nat.cast_zero, lift_zero] + simp only [Sequence₂, mk_uLift, mk_fintype, Nat.cast_zero, lift_zero, Fintype.card_pempty] #align first_order.sequence₂.lift_mk FirstOrder.Sequence₂.lift_mk @[simp] diff --git a/Mathlib/NumberTheory/Bertrand.lean b/Mathlib/NumberTheory/Bertrand.lean index c7293c44fb0e9..6423b972525a4 100644 --- a/Mathlib/NumberTheory/Bertrand.lean +++ b/Mathlib/NumberTheory/Bertrand.lean @@ -178,13 +178,13 @@ theorem centralBinom_le_of_no_bertrand_prime (n : ℕ) (n_big : 2 < n) · exact pow_factorization_choose_le (mul_pos two_pos n_pos) have : (Finset.Icc 1 (sqrt (2 * n))).card = sqrt (2 * n) := by rw [card_Icc, Nat.add_sub_cancel] rw [Finset.prod_const] - refine' pow_le_pow n2_pos ((Finset.card_le_of_subset fun x hx => _).trans this.le) + refine' pow_le_pow_right n2_pos ((Finset.card_le_of_subset fun x hx => _).trans this.le) obtain ⟨h1, h2⟩ := Finset.mem_filter.1 hx exact Finset.mem_Icc.mpr ⟨(Finset.mem_filter.1 h1).2.one_lt.le, h2⟩ · refine' le_trans _ (primorial_le_4_pow (2 * n / 3)) refine' (Finset.prod_le_prod' fun p hp => (_ : f p ≤ p)).trans _ · obtain ⟨h1, h2⟩ := Finset.mem_filter.1 hp - refine' (pow_le_pow (Finset.mem_filter.1 h1).2.one_lt.le _).trans (pow_one p).le + refine' (pow_le_pow_right (Finset.mem_filter.1 h1).2.one_lt.le _).trans (pow_one p).le exact Nat.factorization_choose_le_one (sqrt_lt'.mp <| not_le.1 h2) refine' Finset.prod_le_prod_of_subset_of_one_le' (Finset.filter_subset _ _) _ exact fun p hp _ => (Finset.mem_filter.1 hp).2.one_lt.le diff --git a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean index 814f1a5c0c94a..d8d7ad664d974 100644 --- a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean +++ b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean @@ -165,7 +165,7 @@ theorem cardPowDegree_anti_archimedean {x y z : Fq[X]} {a : ℤ} (hxy : cardPowD cardPowDegree_nonzero _ hyz'] have : (1 : ℤ) ≤ Fintype.card Fq := mod_cast (@Fintype.one_lt_card Fq _ _).le simp only [Int.cast_pow, Int.cast_ofNat, le_max_iff] - refine' Or.imp (pow_le_pow this) (pow_le_pow this) _ + refine' Or.imp (pow_le_pow_right this) (pow_le_pow_right this) _ rw [natDegree_le_iff_degree_le, natDegree_le_iff_degree_le, ← le_max_iff, ← degree_eq_natDegree hxy', ← degree_eq_natDegree hyz'] convert degree_add_le (x - y) (y - z) using 2 diff --git a/Mathlib/NumberTheory/ClassNumber/Finite.lean b/Mathlib/NumberTheory/ClassNumber/Finite.lean index 05b7d9e9e45a5..662c91b16e1e0 100644 --- a/Mathlib/NumberTheory/ClassNumber/Finite.lean +++ b/Mathlib/NumberTheory/ClassNumber/Finite.lean @@ -121,7 +121,7 @@ theorem norm_lt {T : Type*} [LinearOrderedRing T] (a : S) {y : T} apply (Int.cast_le.mpr (norm_le abv bS a hy')).trans_lt simp only [Int.cast_mul, Int.cast_pow] apply mul_lt_mul' le_rfl - · exact pow_lt_pow_of_lt_left this (Int.cast_nonneg.mpr y'_nonneg) (Fintype.card_pos_iff.mpr ⟨i⟩) + · exact pow_lt_pow_left this (Int.cast_nonneg.mpr y'_nonneg) (@Fintype.card_ne_zero _ _ ⟨i⟩) · exact pow_nonneg (Int.cast_nonneg.mpr y'_nonneg) _ · exact Int.cast_pos.mpr (normBound_pos abv bS) #align class_group.norm_lt ClassGroup.norm_lt diff --git a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean index 9ef15c8d2cbd4..e8762737fa5d2 100644 --- a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean +++ b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean @@ -447,7 +447,7 @@ theorem sub_one_norm_two {k : ℕ} (hζ : IsPrimitiveRoot ζ (2 ^ k)) (hk : 2 have : 2 < (2 : ℕ+) ^ k := by simp only [← coe_lt_coe, one_coe, pow_coe] nth_rw 1 [← pow_one 2] - exact pow_lt_pow one_lt_two (lt_of_lt_of_le one_lt_two hk) + exact pow_lt_pow_right one_lt_two (lt_of_lt_of_le one_lt_two hk) -- Porting note: `simpa using hirr` was `simp [hirr]`_ replace hirr : Irreducible (cyclotomic ((2 : ℕ+) ^ k : ℕ+) K) := by simpa using hirr -- Porting note: `simpa using hζ` was `simp [hζ]`_ diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 9f85d351a8bdd..d628a349a5c8e 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -333,7 +333,7 @@ theorem Prime.properDivisors {p : ℕ} (pp : p.Prime) : properDivisors p = {1} : #align nat.prime.proper_divisors Nat.Prime.properDivisors theorem divisors_prime_pow {p : ℕ} (pp : p.Prime) (k : ℕ) : - divisors (p ^ k) = (Finset.range (k + 1)).map ⟨(p ^ ·), pow_right_injective pp.two_le⟩ := by + divisors (p ^ k) = (Finset.range (k + 1)).map ⟨(p ^ ·), Nat.pow_right_injective pp.two_le⟩ := by ext a rw [mem_divisors_prime_pow pp] simp [Nat.lt_succ, eq_comm] @@ -438,15 +438,15 @@ theorem mem_properDivisors_prime_pow {p : ℕ} (pp : p.Prime) (k : ℕ) {x : ℕ intro a constructor <;> intro h · rcases h with ⟨_h_left, rfl, h_right⟩ - rw [pow_lt_pow_iff pp.one_lt] at h_right + rw [pow_lt_pow_iff_right pp.one_lt] at h_right exact ⟨h_right, by rfl⟩ · rcases h with ⟨h_left, rfl⟩ - rw [pow_lt_pow_iff pp.one_lt] + rw [pow_lt_pow_iff_right pp.one_lt] simp [h_left, le_of_lt] #align nat.mem_proper_divisors_prime_pow Nat.mem_properDivisors_prime_pow theorem properDivisors_prime_pow {p : ℕ} (pp : p.Prime) (k : ℕ) : - properDivisors (p ^ k) = (Finset.range k).map ⟨HPow.hPow p, pow_right_injective pp.two_le⟩ := by + properDivisors (p ^ k) = (Finset.range k).map ⟨(p ^ ·), Nat.pow_right_injective pp.two_le⟩ := by ext a simp only [mem_properDivisors, Nat.isUnit_iff, mem_map, mem_range, Function.Embedding.coeFn_mk, pow_eq] diff --git a/Mathlib/NumberTheory/EulerProduct/Basic.lean b/Mathlib/NumberTheory/EulerProduct/Basic.lean index ea3a17a21c0e8..40aeee76afd65 100644 --- a/Mathlib/NumberTheory/EulerProduct/Basic.lean +++ b/Mathlib/NumberTheory/EulerProduct/Basic.lean @@ -187,7 +187,7 @@ theorem eulerProduct_completely_multiplicative {f : ℕ →*₀ F} (hsum : Summa simp_rw [map_pow] refine (tsum_geometric_of_norm_lt_1 <| summable_geometric_iff_norm_lt_1.mp ?_).symm refine Summable.of_norm ?_ - convert hsum.comp_injective <| pow_right_injective (prime_of_mem_primesBelow hN).one_lt + convert hsum.comp_injective <| Nat.pow_right_injective (prime_of_mem_primesBelow hN).one_lt simp only [norm_pow, Function.comp_apply, map_pow] end CompletelyMultiplicative diff --git a/Mathlib/NumberTheory/FLT/Four.lean b/Mathlib/NumberTheory/FLT/Four.lean index dcc656362a681..f14ef036a5507 100644 --- a/Mathlib/NumberTheory/FLT/Four.lean +++ b/Mathlib/NumberTheory/FLT/Four.lean @@ -101,7 +101,7 @@ theorem coprime_of_minimal {a b c : ℤ} (h : Minimal a b c) : IsCoprime a b := (Fermat42.mul (Int.coe_nat_ne_zero.mpr (Nat.Prime.ne_zero hp))).mpr h.1 apply Nat.le_lt_asymm (h.2 _ _ _ hf) rw [Int.natAbs_mul, lt_mul_iff_one_lt_left, Int.natAbs_pow, Int.natAbs_ofNat] - · exact Nat.one_lt_pow _ _ zero_lt_two (Nat.Prime.one_lt hp) + · exact Nat.one_lt_pow _ _ two_ne_zero (Nat.Prime.one_lt hp) · exact Nat.pos_of_ne_zero (Int.natAbs_ne_zero.2 (ne_zero hf)) #align fermat_42.coprime_of_minimal Fermat42.coprime_of_minimal diff --git a/Mathlib/NumberTheory/FermatPsp.lean b/Mathlib/NumberTheory/FermatPsp.lean index b85301cd8da69..7a2db88a60b5d 100644 --- a/Mathlib/NumberTheory/FermatPsp.lean +++ b/Mathlib/NumberTheory/FermatPsp.lean @@ -133,14 +133,11 @@ theorem fermatPsp_base_one {n : ℕ} (h₁ : 1 < n) (h₂ : ¬n.Prime) : FermatP -- pseudoprimes section HelperLemmas -private theorem pow_gt_exponent {a : ℕ} (b : ℕ) (h : 2 ≤ a) : b < a ^ b := - lt_of_lt_of_le (Nat.lt_two_pow b) <| Nat.pow_le_pow_of_le_left h _ - private theorem a_id_helper {a b : ℕ} (ha : 2 ≤ a) (hb : 2 ≤ b) : 2 ≤ (a ^ b - 1) / (a - 1) := by change 1 < _ have h₁ : a - 1 ∣ a ^ b - 1 := by simpa only [one_pow] using nat_sub_dvd_pow_sub_pow a 1 b rw [Nat.lt_div_iff_mul_lt h₁, mul_one, tsub_lt_tsub_iff_right (Nat.le_of_succ_le ha)] - exact self_lt_pow (Nat.lt_of_succ_le ha) hb + exact lt_self_pow (Nat.lt_of_succ_le ha) hb private theorem b_id_helper {a b : ℕ} (ha : 2 ≤ a) (hb : 2 < b) : 2 ≤ (a ^ b + 1) / (a + 1) := by rw [Nat.le_div_iff_mul_le (Nat.zero_lt_succ _)] @@ -148,7 +145,7 @@ private theorem b_id_helper {a b : ℕ} (ha : 2 ≤ a) (hb : 2 < b) : 2 ≤ (a ^ calc 2 * a + 1 ≤ a ^ 2 * a := by nlinarith _ = a ^ 3 := by rw [pow_succ a 2] - _ ≤ a ^ b := pow_le_pow (Nat.le_of_succ_le ha) hb + _ ≤ a ^ b := pow_le_pow_right (Nat.le_of_succ_le ha) hb private theorem AB_id_helper (b p : ℕ) (_ : 2 ≤ b) (hp : Odd p) : (b ^ p - 1) / (b - 1) * ((b ^ p + 1) / (b + 1)) = (b ^ (2 * p) - 1) / (b ^ 2 - 1) := by @@ -331,7 +328,7 @@ private theorem psp_from_prime_gt_p {b : ℕ} (b_ge_two : 2 ≤ b) {p : ℕ} (p_ have : 2 ≤ 2 * p - 2 := le_tsub_of_add_le_left (show 4 ≤ 2 * p by linarith) have : 2 + p ≤ 2 * p := by linarith have : p ≤ 2 * p - 2 := le_tsub_of_add_le_left this - exact Nat.lt_of_le_of_lt this (pow_gt_exponent _ b_ge_two) + exact this.trans_lt (lt_pow_self b_ge_two _) /-- For all positive bases, there exist infinite **Fermat pseudoprimes** to that base. Given in this form: for all numbers `b ≥ 1` and `m`, there exists a pseudoprime `n` to base `b` such @@ -350,7 +347,7 @@ theorem exists_infinite_pseudoprimes {b : ℕ} (h : 1 ≤ b) (m : ℕ) : cases' h with p hp cases' hp with hp₁ hp₂ have h₁ : 0 < b := pos_of_gt (Nat.succ_le_iff.mp b_ge_two) - have h₂ : 4 ≤ b ^ 2 := pow_le_pow_of_le_left' b_ge_two 2 + have h₂ : 4 ≤ b ^ 2 := pow_le_pow_left' b_ge_two 2 have h₃ : 0 < b ^ 2 - 1 := tsub_pos_of_lt (gt_of_ge_of_gt h₂ (by norm_num)) have h₄ : 0 < b * (b ^ 2 - 1) := mul_pos h₁ h₃ have h₅ : b * (b ^ 2 - 1) < p := by linarith diff --git a/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean b/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean index 5669975897223..90e07877c06d7 100644 --- a/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean +++ b/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean @@ -155,7 +155,7 @@ theorem aux_calc (n : ℕ) {m : ℝ} (hm : 2 ≤ m) : any_goals exact pow_pos (zero_lt_two.trans_le hm) _ -- `2 ≤ m ^ n!` is a consequence of monotonicity of exponentiation at `2 ≤ m`. exact _root_.trans (_root_.trans hm (pow_one _).symm.le) - (pow_mono (one_le_two.trans hm) n.factorial_pos) + (pow_right_mono (one_le_two.trans hm) n.factorial_pos) _ = 1 / (m ^ n !) ^ n := congr_arg (1 / ·) (pow_mul m n ! n) #align liouville_number.aux_calc LiouvilleNumber.aux_calc diff --git a/Mathlib/NumberTheory/LucasLehmer.lean b/Mathlib/NumberTheory/LucasLehmer.lean index 0d162b815dea4..23d293c6403cf 100644 --- a/Mathlib/NumberTheory/LucasLehmer.lean +++ b/Mathlib/NumberTheory/LucasLehmer.lean @@ -53,7 +53,7 @@ theorem mersenne_pos {p : ℕ} (h : 0 < p) : 0 < mersenne p := by theorem one_lt_mersenne {p : ℕ} (hp : 1 < p) : 1 < mersenne p := lt_tsub_iff_right.2 <| calc 1 + 1 = 2 ^ 1 := by rw [one_add_one_eq_two, pow_one] - _ < 2 ^ p := Nat.pow_lt_pow_of_lt_right one_lt_two hp + _ < 2 ^ p := pow_lt_pow_right one_lt_two hp @[simp] theorem succ_mersenne (k : ℕ) : mersenne k + 1 = 2 ^ k := by @@ -94,27 +94,27 @@ def sMod (p : ℕ) : ℕ → ℤ | i + 1 => (sMod p i ^ 2 - 2) % (2 ^ p - 1) #align lucas_lehmer.s_mod LucasLehmer.sMod -theorem mersenne_int_pos {p : ℕ} (hp : 0 < p) : (0 : ℤ) < 2 ^ p - 1 := +theorem mersenne_int_pos {p : ℕ} (hp : p ≠ 0) : (0 : ℤ) < 2 ^ p - 1 := sub_pos.2 <| mod_cast Nat.one_lt_two_pow p hp -theorem mersenne_int_ne_zero (p : ℕ) (w : 0 < p) : (2 ^ p - 1 : ℤ) ≠ 0 := - (mersenne_int_pos w).ne' +theorem mersenne_int_ne_zero (p : ℕ) (hp : p ≠ 0) : (2 ^ p - 1 : ℤ) ≠ 0 := + (mersenne_int_pos hp).ne' #align lucas_lehmer.mersenne_int_ne_zero LucasLehmer.mersenne_int_ne_zero -theorem sMod_nonneg (p : ℕ) (w : 0 < p) (i : ℕ) : 0 ≤ sMod p i := by +theorem sMod_nonneg (p : ℕ) (hp : p ≠ 0) (i : ℕ) : 0 ≤ sMod p i := by cases i <;> dsimp [sMod] · exact sup_eq_right.mp rfl · apply Int.emod_nonneg - exact mersenne_int_ne_zero p w + exact mersenne_int_ne_zero p hp #align lucas_lehmer.s_mod_nonneg LucasLehmer.sMod_nonneg theorem sMod_mod (p i : ℕ) : sMod p i % (2 ^ p - 1) = sMod p i := by cases i <;> simp [sMod] #align lucas_lehmer.s_mod_mod LucasLehmer.sMod_mod -theorem sMod_lt (p : ℕ) (w : 0 < p) (i : ℕ) : sMod p i < 2 ^ p - 1 := by +theorem sMod_lt (p : ℕ) (hp : p ≠ 0) (i : ℕ) : sMod p i < 2 ^ p - 1 := by rw [← sMod_mod] - refine (Int.emod_lt _ (mersenne_int_ne_zero p w)).trans_eq ?_ - exact abs_of_nonneg (mersenne_int_pos w).le + refine (Int.emod_lt _ (mersenne_int_ne_zero p hp)).trans_eq ?_ + exact abs_of_nonneg (mersenne_int_pos hp).le #align lucas_lehmer.s_mod_lt LucasLehmer.sMod_lt theorem sZMod_eq_s (p' : ℕ) (i : ℕ) : sZMod (p' + 2) i = (s i : ZMod (2 ^ (p' + 2) - 1)) := by @@ -155,8 +155,8 @@ theorem residue_eq_zero_iff_sMod_eq_zero (p : ℕ) (w : 1 < p) : simp only [ge_iff_le, ZMod.int_cast_zmod_eq_zero_iff_dvd, gt_iff_lt, zero_lt_two, pow_pos, cast_pred, cast_pow, cast_ofNat] at h apply Int.eq_zero_of_dvd_of_nonneg_of_lt _ _ h <;> clear h - · apply sMod_nonneg _ (Nat.lt_of_succ_lt w) - · exact sMod_lt _ (Nat.lt_of_succ_lt w) (p - 2) + · exact sMod_nonneg _ (by positivity) _ + · exact sMod_lt _ (by positivity) _ · intro h rw [h] simp diff --git a/Mathlib/NumberTheory/MaricaSchoenheim.lean b/Mathlib/NumberTheory/MaricaSchoenheim.lean index 805b31d94ae71..a9d9220e50076 100644 --- a/Mathlib/NumberTheory/MaricaSchoenheim.lean +++ b/Mathlib/NumberTheory/MaricaSchoenheim.lean @@ -45,7 +45,7 @@ lemma grahamConjecture_of_squarefree {n : ℕ} (f : ℕ → ℕ) (hf' : ∀ k < _ ≤ (Ioo 0 n).card := card_le_card_of_inj_on (fun s ↦ ∏ p in s, p) ?_ ?_ _ = n - 1 := by rw [card_Ioo, tsub_zero] _ < n := tsub_lt_self hn.bot_lt zero_lt_one - · rw [card_image_of_injOn, card_Iio] + · rw [Finset.card_image_of_injOn, card_Iio] simpa using prod_primeFactors_invOn_squarefree.2.injOn.comp hf.injOn hf' · simp only [forall_mem_diffs, forall_image, mem_Ioo, mem_Iio] rintro i hi j hj diff --git a/Mathlib/NumberTheory/Multiplicity.lean b/Mathlib/NumberTheory/Multiplicity.lean index 022dc93070d8c..97ce4daa73e39 100644 --- a/Mathlib/NumberTheory/Multiplicity.lean +++ b/Mathlib/NumberTheory/Multiplicity.lean @@ -228,13 +228,13 @@ theorem Nat.pow_sub_pow {x y : ℕ} (hxy : p ∣ x - y) (hx : ¬p ∣ x) (n : multiplicity p (x ^ n - y ^ n) = multiplicity p (x - y) + multiplicity p n := by obtain hyx | hyx := le_total y x · iterate 2 rw [← Int.coe_nat_multiplicity] - rw [Int.ofNat_sub (Nat.pow_le_pow_of_le_left hyx n)] + rw [Int.ofNat_sub (Nat.pow_le_pow_left hyx n)] rw [← Int.coe_nat_dvd] at hxy hx rw [Int.coe_nat_sub hyx] at * push_cast at * exact Int.pow_sub_pow hp hp1 hxy hx n · simp only [Nat.sub_eq_zero_iff_le.mpr hyx, - Nat.sub_eq_zero_iff_le.mpr (Nat.pow_le_pow_of_le_left hyx n), multiplicity.zero, + Nat.sub_eq_zero_iff_le.mpr (Nat.pow_le_pow_left hyx n), multiplicity.zero, PartENat.top_add] #align multiplicity.nat.pow_sub_pow multiplicity.Nat.pow_sub_pow @@ -362,7 +362,7 @@ theorem Nat.two_pow_sub_pow {x y : ℕ} (hxy : 2 ∣ x - y) (hx : ¬2 ∣ x) {n multiplicity 2 (x + y) + multiplicity 2 (x - y) + multiplicity 2 n := by obtain hyx | hyx := le_total y x · iterate 3 rw [← multiplicity.Int.coe_nat_multiplicity] - simp only [Int.ofNat_sub hyx, Int.ofNat_sub (pow_le_pow_of_le_left' hyx _), Int.ofNat_add, + simp only [Int.ofNat_sub hyx, Int.ofNat_sub (pow_le_pow_left' hyx _), Int.ofNat_add, Int.coe_nat_pow] rw [← Int.coe_nat_dvd] at hx rw [← Int.coe_nat_dvd, Int.ofNat_sub hyx] at hxy @@ -370,7 +370,7 @@ theorem Nat.two_pow_sub_pow {x y : ℕ} (hxy : 2 ∣ x - y) (hx : ¬2 ∣ x) {n rw [← multiplicity.Int.coe_nat_multiplicity] rfl · simp only [Nat.sub_eq_zero_iff_le.mpr hyx, - Nat.sub_eq_zero_iff_le.mpr (pow_le_pow_of_le_left' hyx n), multiplicity.zero, + Nat.sub_eq_zero_iff_le.mpr (pow_le_pow_left' hyx n), multiplicity.zero, PartENat.top_add, PartENat.add_top] #align nat.two_pow_sub_pow Nat.two_pow_sub_pow @@ -378,29 +378,29 @@ namespace padicValNat variable {x y : ℕ} -theorem pow_two_sub_pow (hyx : y < x) (hxy : 2 ∣ x - y) (hx : ¬2 ∣ x) {n : ℕ} (hn : 0 < n) +theorem pow_two_sub_pow (hyx : y < x) (hxy : 2 ∣ x - y) (hx : ¬2 ∣ x) {n : ℕ} (hn : n ≠ 0) (hneven : Even n) : padicValNat 2 (x ^ n - y ^ n) + 1 = padicValNat 2 (x + y) + padicValNat 2 (x - y) + padicValNat 2 n := by simp only [← PartENat.natCast_inj, Nat.cast_add] iterate 4 rw [padicValNat_def, PartENat.natCast_get] · convert Nat.two_pow_sub_pow hxy hx hneven using 2 - · exact hn + · exact hn.bot_lt · exact Nat.sub_pos_of_lt hyx · linarith - · simp only [tsub_pos_iff_lt, pow_lt_pow_of_lt_left hyx (@zero_le' _ y _) hn] + · simp only [tsub_pos_iff_lt, pow_lt_pow_left hyx zero_le' hn] #align padic_val_nat.pow_two_sub_pow padicValNat.pow_two_sub_pow variable {p : ℕ} [hp : Fact p.Prime] (hp1 : Odd p) -theorem pow_sub_pow (hyx : y < x) (hxy : p ∣ x - y) (hx : ¬p ∣ x) {n : ℕ} (hn : 0 < n) : +theorem pow_sub_pow (hyx : y < x) (hxy : p ∣ x - y) (hx : ¬p ∣ x) {n : ℕ} (hn : n ≠ 0) : padicValNat p (x ^ n - y ^ n) = padicValNat p (x - y) + padicValNat p n := by rw [← PartENat.natCast_inj, Nat.cast_add] iterate 3 rw [padicValNat_def, PartENat.natCast_get] · exact multiplicity.Nat.pow_sub_pow hp.out hp1 hxy hx n - · exact hn + · exact hn.bot_lt · exact Nat.sub_pos_of_lt hyx - · exact Nat.sub_pos_of_lt (Nat.pow_lt_pow_of_lt_left hyx hn) + · exact Nat.sub_pos_of_lt (Nat.pow_lt_pow_left hyx hn) #align padic_val_nat.pow_sub_pow padicValNat.pow_sub_pow theorem pow_add_pow (hxy : p ∣ x + y) (hx : ¬p ∣ x) {n : ℕ} (hn : Odd n) : diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index 0c9c5896b60fc..fc9a2847b6041 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -146,7 +146,7 @@ instance : Free ℤ (𝓞 K) := IsIntegralClosure.module_free ℤ ℚ K (𝓞 K) instance : IsLocalization (Algebra.algebraMapSubmonoid (𝓞 K) ℤ⁰) K := - IsIntegralClosure.isLocalization ℤ ℚ K (𝓞 K) + IsIntegralClosure.isLocalization_of_isSeparable ℤ ℚ K (𝓞 K) /-- A ℤ-basis of the ring of integers of `K`. -/ noncomputable def basis : Basis (Free.ChooseBasisIndex ℤ (𝓞 K)) ℤ (𝓞 K) := diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index cd8f2f3b5d52b..f7ab06d0e7103 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -167,7 +167,7 @@ theorem abs_discr_ge (h : 1 < finrank ℚ K) : suffices ∀ n, 2 ≤ n → (4 / 9 : ℝ) * (3 * π / 4) ^ n ≤ a n by refine le_trans (this (finrank ℚ K) h) ?_ refine div_le_div_of_le_left (by positivity) (by positivity) ?_ - refine mul_le_mul_of_nonneg_right (pow_le_pow ?_ ?_) (by positivity) + refine mul_le_mul_of_nonneg_right (pow_le_pow_right ?_ ?_) (by positivity) · rw [_root_.le_div_iff Real.pi_pos, one_mul] exact Real.pi_le_four · rw [← card_add_two_mul_card_eq_rank, mul_comm] @@ -197,7 +197,7 @@ theorem discr_gt_one (h : 1 < finrank ℚ K) : 2 < |discr K| := by exact Real.pi_gt_three refine Int.cast_lt.mp <| lt_of_lt_of_le ?_ (abs_discr_ge h) rw [← _root_.div_lt_iff' (by positivity), Int.int_cast_ofNat] - refine lt_of_lt_of_le ?_ (pow_le_pow (n := 2) h₁ h) + refine lt_of_lt_of_le ?_ (pow_le_pow_right (n := 2) h₁ h) rw [div_pow, _root_.lt_div_iff (by norm_num), mul_pow] norm_num rw [ ← _root_.div_lt_iff' (by positivity), show (72:ℝ) / 9 = 8 by norm_num] diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index afa71267c9089..e60329ec1eeb2 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -334,7 +334,7 @@ theorem seq_next {x : 𝓞 K} (hx : x ≠ 0) : _ ≤ ∏ w : InfinitePlace K, (g w : ℝ) ^ mult w := by refine prod_le_prod ?_ ?_ · exact fun _ _ => pow_nonneg (by positivity) _ - · exact fun w _ => pow_le_pow_of_le_left (by positivity) (le_of_lt (h_yle w)) (mult w) + · exact fun w _ => pow_le_pow_left (by positivity) (le_of_lt (h_yle w)) (mult w) _ ≤ (B : ℝ) := by simp_rw [← NNReal.coe_pow, ← NNReal.coe_prod] exact le_of_eq (congr_arg toReal h_gprod) diff --git a/Mathlib/NumberTheory/Padics/Hensel.lean b/Mathlib/NumberTheory/Padics/Hensel.lean index bb4b424e11a41..7cb36d60c312d 100644 --- a/Mathlib/NumberTheory/Padics/Hensel.lean +++ b/Mathlib/NumberTheory/Padics/Hensel.lean @@ -279,7 +279,7 @@ private theorem T_pos : T > 0 := by private theorem newton_seq_succ_dist_weak (n : ℕ) : ‖newton_seq (n + 2) - newton_seq (n + 1)‖ < ‖F.eval a‖ / ‖F.derivative.eval a‖ := have : 2 ≤ 2 ^ (n + 1) := by - have := pow_le_pow (by norm_num : 1 ≤ 2) (Nat.le_add_left _ _ : 1 ≤ n + 1) + have := pow_le_pow_right (by norm_num : 1 ≤ 2) (Nat.le_add_left _ _ : 1 ≤ n + 1) simpa using this calc ‖newton_seq (n + 2) - newton_seq (n + 1)‖ ≤ ‖F.derivative.eval a‖ * T ^ 2 ^ (n + 1) := @@ -288,7 +288,7 @@ private theorem newton_seq_succ_dist_weak (n : ℕ) : (mul_le_mul_of_nonneg_left (pow_le_pow_of_le_one (norm_nonneg _) (le_of_lt (T_lt_one hnorm)) this) (norm_nonneg _)) _ < ‖F.derivative.eval a‖ * T ^ 1 := - (mul_lt_mul_of_pos_left (pow_lt_pow_of_lt_one (T_pos hnorm hnsol) + (mul_lt_mul_of_pos_left (pow_lt_pow_right_of_lt_one (T_pos hnorm hnsol) (T_lt_one hnorm) (by norm_num)) (deriv_norm_pos hnorm)) _ = ‖F.eval a‖ / ‖F.derivative.eval a‖ := by rw [T_gen, sq, pow_one, norm_div, ← mul_div_assoc, PadicInt.padic_norm_e_of_padicInt, @@ -301,7 +301,7 @@ private theorem newton_seq_dist_aux (n : ℕ) : | 0 => by simp [T_pow_nonneg, mul_nonneg] | k + 1 => have : 2 ^ n ≤ 2 ^ (n + k) := by - apply pow_le_pow + apply pow_le_pow_right norm_num apply Nat.le_add_right calc diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index 93deff94a45b4..55d1d44380bdb 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -520,7 +520,7 @@ theorem norm_le_pow_iff_le_valuation (x : ℤ_[p]) (hx : x ≠ 0) (n : ℕ) : refine pow_pos ?_ m exact mod_cast hp.1.pos rw [inv_le_inv (aux _) (aux _)] - have : p ^ n ≤ p ^ k ↔ n ≤ k := (pow_strictMono_right hp.1.one_lt).le_iff_le + have : p ^ n ≤ p ^ k ↔ n ≤ k := (pow_right_strictMono hp.1.one_lt).le_iff_le rw [← this] norm_cast #align padic_int.norm_le_pow_iff_le_valuation PadicInt.norm_le_pow_iff_le_valuation diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index f953bec7900da..30ff28630323b 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -640,7 +640,7 @@ theorem range_pow_padicValNat_subset_divisors' {n : ℕ} [hp : Fact p.Prime] : obtain ⟨k, hk, rfl⟩ := ht rw [Finset.mem_erase, Nat.mem_divisors] refine' ⟨_, (pow_dvd_pow p <| succ_le_iff.2 hk).trans pow_padicValNat_dvd, hn⟩ - exact (Nat.one_lt_pow _ _ k.succ_pos hp.out.one_lt).ne' + exact (Nat.one_lt_pow _ _ k.succ_ne_zero hp.out.one_lt).ne' #align range_pow_padic_val_nat_subset_divisors' range_pow_padicValNat_subset_divisors' /-- The `p`-adic valuation of `(p * n)!` is `n` more than that of `n!`. -/ diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index 3991537649c74..2caefc0436765 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -310,7 +310,7 @@ theorem appr_lt (x : ℤ_[p]) (n : ℕ) : x.appr n < p ^ n := by induction' n with n ih generalizing x · simp only [appr, zero_eq, _root_.pow_zero, zero_lt_one] simp only [appr, map_natCast, ZMod.nat_cast_self, RingHom.map_pow, Int.natAbs, RingHom.map_mul] - have hp : p ^ n < p ^ (n + 1) := by apply pow_lt_pow hp_prime.1.one_lt (lt_add_one n) + have hp : p ^ n < p ^ (n + 1) := by apply pow_lt_pow_right hp_prime.1.one_lt (lt_add_one n) split_ifs with h · apply lt_trans (ih _) hp · calc diff --git a/Mathlib/NumberTheory/PellMatiyasevic.lean b/Mathlib/NumberTheory/PellMatiyasevic.lean index 1215ab257c134..61f8b6204c7d1 100644 --- a/Mathlib/NumberTheory/PellMatiyasevic.lean +++ b/Mathlib/NumberTheory/PellMatiyasevic.lean @@ -965,7 +965,7 @@ theorem eq_pow_of_pell {m n k} : refine' eq_pow_of_pell_lem hn.ne' hk.ne' _ calc n ^ k ≤ n ^ w := Nat.pow_le_pow_of_le_right hn kw - _ < (w + 1) ^ w := (Nat.pow_lt_pow_of_lt_left (Nat.lt_succ_of_le nw) wpos) + _ < (w + 1) ^ w := (Nat.pow_lt_pow_left (Nat.lt_succ_of_le nw) wpos.ne') _ ≤ a := xn_ge_a_pow w1 w lift (2 * a * n - n * n - 1 : ℤ) to ℕ using (Nat.cast_nonneg _).trans nt.le with t te have tm : x ≡ y * (a - n) + n ^ k [MOD t] := by @@ -995,7 +995,7 @@ theorem eq_pow_of_pell {m n k} : have hnka : n ^ k < xn hw1 j calc n ^ k ≤ n ^ j := Nat.pow_le_pow_of_le_right hn0 (le_trans kw wj) - _ < (w + 1) ^ j := (Nat.pow_lt_pow_of_lt_left (Nat.lt_succ_of_le nw) hj0) + _ < (w + 1) ^ j := (Nat.pow_lt_pow_left (Nat.lt_succ_of_le nw) hj0.ne') _ ≤ xn hw1 j := xn_ge_a_pow hw1 j have nt : (↑(n ^ k) : ℤ) < 2 * xn hw1 j * n - n * n - 1 := eq_pow_of_pell_lem hn0.ne' hk0.ne' hnka diff --git a/Mathlib/NumberTheory/RamificationInertia.lean b/Mathlib/NumberTheory/RamificationInertia.lean index 07c91a0fa56c1..3f5f11f09d59b 100644 --- a/Mathlib/NumberTheory/RamificationInertia.lean +++ b/Mathlib/NumberTheory/RamificationInertia.lean @@ -88,7 +88,7 @@ theorem ramificationIdx_spec {n : ℕ} (hle : map f p ≤ P ^ n) (hgt : ¬map f have : Q n := by intro k hk refine le_of_not_lt fun hnk => ?_ - exact hgt (hk.trans (Ideal.pow_le_pow hnk)) + exact hgt (hk.trans (Ideal.pow_le_pow_right hnk)) rw [ramificationIdx_eq_find ⟨n, this⟩] refine le_antisymm (Nat.find_min' _ this) (le_of_not_gt fun h : Nat.find _ < n => ?_) obtain this' := Nat.find_spec ⟨n, this⟩ @@ -101,7 +101,7 @@ theorem ramificationIdx_lt {n : ℕ} (hgt : ¬map f p ≤ P ^ n) : ramificationI · rw [Nat.lt_succ_iff] have : ∀ k, map f p ≤ P ^ k → k ≤ n := by refine fun k hk => le_of_not_lt fun hnk => ?_ - exact hgt (hk.trans (Ideal.pow_le_pow hnk)) + exact hgt (hk.trans (Ideal.pow_le_pow_right hnk)) rw [ramificationIdx_eq_find ⟨n, this⟩] exact Nat.find_min' ⟨n, this⟩ this #align ideal.ramification_idx_lt Ideal.ramificationIdx_lt @@ -486,7 +486,7 @@ theorem Quotient.algebraMap_quotient_of_ramificationIdx_neZero (x : R) : def powQuotSuccInclusion (i : ℕ) : Ideal.map (Ideal.Quotient.mk (P ^ e)) (P ^ (i + 1)) →ₗ[R ⧸ p] Ideal.map (Ideal.Quotient.mk (P ^ e)) (P ^ i) where - toFun x := ⟨x, Ideal.map_mono (Ideal.pow_le_pow i.le_succ) x.2⟩ + toFun x := ⟨x, Ideal.map_mono (Ideal.pow_le_pow_right i.le_succ) x.2⟩ map_add' _ _ := rfl map_smul' _ _ := rfl #align ideal.pow_quot_succ_inclusion Ideal.powQuotSuccInclusion @@ -553,7 +553,7 @@ theorem quotientToQuotientRangePowQuotSucc_injective [IsDomain S] [IsDedekindDom Function.Injective (quotientToQuotientRangePowQuotSucc f p P a_mem) := fun x => Quotient.inductionOn' x fun x y => Quotient.inductionOn' y fun y h => by - have Pe_le_Pi1 : P ^ e ≤ P ^ (i + 1) := Ideal.pow_le_pow hi + have Pe_le_Pi1 : P ^ e ≤ P ^ (i + 1) := Ideal.pow_le_pow_right hi simp only [Submodule.Quotient.mk''_eq_mk, quotientToQuotientRangePowQuotSucc_mk, Submodule.Quotient.eq, LinearMap.mem_range, Subtype.ext_iff, Subtype.coe_mk, Submodule.coe_sub] at h ⊢ @@ -573,7 +573,7 @@ theorem quotientToQuotientRangePowQuotSucc_surjective [IsDomain S] [IsDedekindDo (a_not_mem : a ∉ P ^ (i + 1)) : Function.Surjective (quotientToQuotientRangePowQuotSucc f p P a_mem) := by rintro ⟨⟨⟨x⟩, hx⟩⟩ - have Pe_le_Pi : P ^ e ≤ P ^ i := Ideal.pow_le_pow hi.le + have Pe_le_Pi : P ^ e ≤ P ^ i := Ideal.pow_le_pow_right hi.le rw [Submodule.Quotient.quot_mk_eq_mk, Ideal.Quotient.mk_eq_mk, Ideal.mem_quotient_iff_mem_sup, sup_eq_left.mpr Pe_le_Pi] at hx suffices hx' : x ∈ Ideal.span {a} ⊔ P ^ (i + 1) @@ -606,7 +606,7 @@ noncomputable def quotientRangePowQuotSuccInclusionEquiv [IsDomain S] [IsDedekin ≃ₗ[R ⧸ p] S ⧸ P := by choose a a_mem a_not_mem using SetLike.exists_of_lt - (Ideal.strictAnti_pow P hP (Ideal.IsPrime.ne_top inferInstance) (le_refl i.succ)) + (Ideal.pow_right_strictAnti P hP (Ideal.IsPrime.ne_top inferInstance) (le_refl i.succ)) refine' (LinearEquiv.ofBijective _ ⟨_, _⟩).symm · exact quotientToQuotientRangePowQuotSucc f p P a_mem · exact quotientToQuotientRangePowQuotSucc_injective f p P hi a_mem a_not_mem diff --git a/Mathlib/NumberTheory/SumFourSquares.lean b/Mathlib/NumberTheory/SumFourSquares.lean index 442a3ed8f7aad..1dc0944b7b3c9 100644 --- a/Mathlib/NumberTheory/SumFourSquares.lean +++ b/Mathlib/NumberTheory/SumFourSquares.lean @@ -71,7 +71,7 @@ theorem lt_of_sum_four_squares_eq_mul {a b c d k m : ℕ} 2 ^ 2 * (k * ↑m) = ∑ i : Fin 4, (2 * ![a, b, c, d] i) ^ 2 := by simp [← h, Fin.sum_univ_succ, mul_add, mul_pow, add_assoc] _ < ∑ _i : Fin 4, m ^ 2 := Finset.sum_lt_sum_of_nonempty Finset.univ_nonempty fun i _ ↦ by - refine pow_lt_pow_of_lt_left ?_ (zero_le _) two_pos + refine pow_lt_pow_left ?_ (zero_le _) two_ne_zero fin_cases i <;> assumption _ = 2 ^ 2 * (m * m) := by simp; ring diff --git a/Mathlib/Order/Filter/Archimedean.lean b/Mathlib/Order/Filter/Archimedean.lean index 7039000f2e27a..2a77bbd6d33c2 100644 --- a/Mathlib/Order/Filter/Archimedean.lean +++ b/Mathlib/Order/Filter/Archimedean.lean @@ -226,7 +226,7 @@ theorem Tendsto.atTop_nsmul_const {f : α → ℕ} (hr : 0 < r) (hf : Tendsto f Tendsto (fun x => f x • r) l atTop := by refine' tendsto_atTop.mpr fun s => _ obtain ⟨n : ℕ, hn : s ≤ n • r⟩ := Archimedean.arch s hr - exact (tendsto_atTop.mp hf n).mono fun a ha => hn.trans (nsmul_le_nsmul hr.le ha) + exact (tendsto_atTop.mp hf n).mono fun a ha => hn.trans (nsmul_le_nsmul_left hr.le ha) #align filter.tendsto.at_top_nsmul_const Filter.Tendsto.atTop_nsmul_const end LinearOrderedCancelAddCommMonoid diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index 85cad5e9e9889..ed455c726dda7 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -721,7 +721,7 @@ theorem Tendsto.nsmul_atTop (hf : Tendsto f l atTop) {n : ℕ} (hn : 0 < n) : calc y ≤ f x := hy _ = 1 • f x := (one_nsmul _).symm - _ ≤ n • f x := nsmul_le_nsmul h₀ hn + _ ≤ n • f x := nsmul_le_nsmul_left h₀ hn #align filter.tendsto.nsmul_at_top Filter.Tendsto.nsmul_atTop theorem Tendsto.nsmul_atBot (hf : Tendsto f l atBot) {n : ℕ} (hn : 0 < n) : diff --git a/Mathlib/Order/SuccPred/CompleteLinearOrder.lean b/Mathlib/Order/SuccPred/CompleteLinearOrder.lean new file mode 100644 index 0000000000000..3496364a6b58f --- /dev/null +++ b/Mathlib/Order/SuccPred/CompleteLinearOrder.lean @@ -0,0 +1,123 @@ +/- +Copyright (c) 2023 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.Order.SuccPred.Limit +import Mathlib.Order.ConditionallyCompleteLattice.Basic + +/-! + +# Relation between `IsSuccLimit` and `iSup` in (conditionally) complete linear orders. + +-/ + +open Order + +variable {ι α : Type*} + +section ConditionallyCompleteLinearOrder +variable [ConditionallyCompleteLinearOrder α] [Nonempty ι] {f : ι → α} {s : Set α} {x : α} + +lemma csSup_mem_of_not_isSuccLimit + (hne : s.Nonempty) (hbdd : BddAbove s) (hlim : ¬ IsSuccLimit (sSup s)) : + sSup s ∈ s := by + obtain ⟨y, hy⟩ := not_forall_not.mp hlim + obtain ⟨i, his, hi⟩ := exists_lt_of_lt_csSup hne hy.lt + exact eq_of_le_of_not_lt (le_csSup hbdd his) (hy.2 hi) ▸ his + +lemma csInf_mem_of_not_isPredLimit + (hne : s.Nonempty) (hbdd : BddBelow s) (hlim : ¬ IsPredLimit (sInf s)) : + sInf s ∈ s := by + obtain ⟨y, hy⟩ := not_forall_not.mp hlim + obtain ⟨i, his, hi⟩ := exists_lt_of_csInf_lt hne hy.lt + exact eq_of_le_of_not_lt (csInf_le hbdd his) (hy.2 · hi) ▸ his + +lemma exists_eq_ciSup_of_not_isSuccLimit + (hf : BddAbove (Set.range f)) (hf' : ¬ IsSuccLimit (⨆ i, f i)) : + ∃ i, f i = ⨆ i, f i := + csSup_mem_of_not_isSuccLimit (Set.range_nonempty f) hf hf' + +lemma exists_eq_ciInf_of_not_isPredLimit + (hf : BddBelow (Set.range f)) (hf' : ¬ IsPredLimit (⨅ i, f i)) : + ∃ i, f i = ⨅ i, f i := + csInf_mem_of_not_isPredLimit (Set.range_nonempty f) hf hf' + +lemma IsLUB.mem_of_nonempty_of_not_isSuccLimit + (hs : IsLUB s x) (hne : s.Nonempty) (hx : ¬ IsSuccLimit x) : x ∈ s := + hs.csSup_eq hne ▸ csSup_mem_of_not_isSuccLimit hne hs.bddAbove (hs.csSup_eq hne ▸ hx) + +lemma IsGLB.mem_of_nonempty_of_not_isPredLimit + (hs : IsGLB s x) (hne : s.Nonempty) (hx : ¬ IsPredLimit x) : x ∈ s := + hs.csInf_eq hne ▸ csInf_mem_of_not_isPredLimit hne hs.bddBelow (hs.csInf_eq hne ▸ hx) + +lemma IsLUB.exists_of_nonempty_of_not_isSuccLimit + (hf : IsLUB (Set.range f) x) (hx : ¬ IsSuccLimit x) : + ∃ i, f i = x := hf.mem_of_nonempty_of_not_isSuccLimit (Set.range_nonempty f) hx + +lemma IsGLB.exists_of_nonempty_of_not_isPredLimit + (hf : IsGLB (Set.range f) x) (hx : ¬ IsPredLimit x) : + ∃ i, f i = x := hf.mem_of_nonempty_of_not_isPredLimit (Set.range_nonempty f) hx + +end ConditionallyCompleteLinearOrder + +section ConditionallyCompleteLinearOrderBot +variable [ConditionallyCompleteLinearOrderBot α] {f : ι → α} {s : Set α} {x : α} + +/-- See `csSup_mem_of_not_isSuccLimit` for the `ConditionallyCompleteLinearOrder` version. -/ +lemma csSup_mem_of_not_isSuccLimit' + (hbdd : BddAbove s) (hlim : ¬ IsSuccLimit (sSup s)) : + sSup s ∈ s := by + obtain (rfl|hs) := s.eq_empty_or_nonempty + · simp [isSuccLimit_bot] at hlim + · exact csSup_mem_of_not_isSuccLimit hs hbdd hlim + +/-- See `exists_eq_ciSup_of_not_isSuccLimit` for the +`ConditionallyCompleteLinearOrder` version. -/ +lemma exists_eq_ciSup_of_not_isSuccLimit' + (hf : BddAbove (Set.range f)) (hf' : ¬ IsSuccLimit (⨆ i, f i)) : + ∃ i, f i = ⨆ i, f i := + csSup_mem_of_not_isSuccLimit' hf hf' + +lemma IsLUB.mem_of_not_isSuccLimit (hs : IsLUB s x) (hx : ¬ IsSuccLimit x) : + x ∈ s := by + obtain (rfl|hs') := s.eq_empty_or_nonempty + · simp [show x = ⊥ by simpa using hs, isSuccLimit_bot] at hx + · exact hs.mem_of_nonempty_of_not_isSuccLimit hs' hx + +lemma IsLUB.exists_of_not_isSuccLimit (hf : IsLUB (Set.range f) x) (hx : ¬ IsSuccLimit x) : + ∃ i, f i = x := hf.mem_of_not_isSuccLimit hx + +end ConditionallyCompleteLinearOrderBot + +section CompleteLinearOrder +variable [CompleteLinearOrder α] {s : Set α} {f : ι → α} {x : α} + +lemma sSup_mem_of_not_isSuccLimit (hlim : ¬ IsSuccLimit (sSup s)) : + sSup s ∈ s := by + obtain ⟨y, hy⟩ := not_forall_not.mp hlim + obtain ⟨i, his, hi⟩ := lt_sSup_iff.mp hy.lt + exact eq_of_le_of_not_lt (le_sSup his) (hy.2 hi) ▸ his + +lemma sInf_mem_of_not_isPredLimit (hlim : ¬ IsPredLimit (sInf s)) : + sInf s ∈ s := by + obtain ⟨y, hy⟩ := not_forall_not.mp hlim + obtain ⟨i, his, hi⟩ := sInf_lt_iff.mp hy.lt + exact eq_of_le_of_not_lt (sInf_le his) (hy.2 · hi) ▸ his + +lemma exists_eq_iSup_of_not_isSuccLimit (hf : ¬ IsSuccLimit (⨆ i, f i)) : + ∃ i, f i = ⨆ i, f i := + sSup_mem_of_not_isSuccLimit hf + +lemma exists_eq_iInf_of_not_isPredLimit (hf : ¬ IsPredLimit (⨅ i, f i)) : + ∃ i, f i = ⨅ i, f i := + sInf_mem_of_not_isPredLimit hf + +lemma IsGLB.mem_of_not_isPredLimit (hs : IsGLB s x) (hx : ¬ IsPredLimit x) : + x ∈ s := + hs.sInf_eq ▸ sInf_mem_of_not_isPredLimit (hs.sInf_eq ▸ hx) + +lemma IsGLB.exists_of_not_isPredLimit (hf : IsGLB (Set.range f) x) (hx : ¬ IsPredLimit x) : + ∃ i, f i = x := hf.mem_of_not_isPredLimit hx + +end CompleteLinearOrder diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 2cd01ce9b2b7b..ef94f9d48f0e0 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -431,7 +431,7 @@ theorem strong_law_aux1 {c : ℝ} (c_one : 1 < c) {ε : ℝ} (εpos : 0 < ε) : set Y := fun n : ℕ => truncation (X n) n set S := fun n => ∑ i in range n, Y i with hS let u : ℕ → ℕ := fun n => ⌊c ^ n⌋₊ - have u_mono : Monotone u := fun i j hij => Nat.floor_mono (pow_le_pow c_one.le hij) + have u_mono : Monotone u := fun i j hij => Nat.floor_mono (pow_le_pow_right c_one.le hij) have I1 : ∀ K, ∑ j in range K, ((j : ℝ) ^ 2)⁻¹ * Var[Y j] ≤ 2 * 𝔼[X 0] := by intro K calc diff --git a/Mathlib/RingTheory/Artinian.lean b/Mathlib/RingTheory/Artinian.lean index f4bc31b001c3d..45ecbf69526b7 100644 --- a/Mathlib/RingTheory/Artinian.lean +++ b/Mathlib/RingTheory/Artinian.lean @@ -427,7 +427,7 @@ variable {R : Type*} [CommRing R] [IsArtinianRing R] theorem isNilpotent_jacobson_bot : IsNilpotent (Ideal.jacobson (⊥ : Ideal R)) := by let Jac := Ideal.jacobson (⊥ : Ideal R) - let f : ℕ →o (Ideal R)ᵒᵈ := ⟨fun n => Jac ^ n, fun _ _ h => Ideal.pow_le_pow h⟩ + let f : ℕ →o (Ideal R)ᵒᵈ := ⟨fun n => Jac ^ n, fun _ _ h => Ideal.pow_le_pow_right h⟩ obtain ⟨n, hn⟩ : ∃ n, ∀ m, n ≤ m → Jac ^ n = Jac ^ m := IsArtinian.monotone_stabilizes f refine' ⟨n, _⟩ let J : Ideal R := annihilator (Jac ^ n) diff --git a/Mathlib/RingTheory/DedekindDomain/Different.lean b/Mathlib/RingTheory/DedekindDomain/Different.lean index 7b66f8c57ab53..145c346ce267b 100644 --- a/Mathlib/RingTheory/DedekindDomain/Different.lean +++ b/Mathlib/RingTheory/DedekindDomain/Different.lean @@ -13,9 +13,9 @@ import Mathlib.RingTheory.DedekindDomain.IntegralClosure ## Main definition - `Submodule.traceDual`: The dual `L`-sub `B`-module under the trace form. - `FractionalIdeal.dual`: The dual fractional ideal under the trace form. +- `differentIdeal`: The different ideal of an extension of integral domains. ## TODO -- Define the relative different ideal - Show properties of the different ideal -/ @@ -153,6 +153,35 @@ lemma isIntegral_discr_mul_of_mem_traceDual variable (A K) +lemma map_equiv_traceDual [NoZeroSMulDivisors A B] (I : Submodule B (FractionRing B)) : + (traceDual A (FractionRing A) I).map (FractionRing.algEquiv B L) = + traceDual A K (I.map (FractionRing.algEquiv B L)) := by + show Submodule.map (FractionRing.algEquiv B L).toLinearEquiv.toLinearMap _ = + traceDual A K (I.map (FractionRing.algEquiv B L).toLinearEquiv.toLinearMap) + rw [Submodule.map_equiv_eq_comap_symm, Submodule.map_equiv_eq_comap_symm] + ext x + simp only [AlgEquiv.toLinearEquiv_symm, AlgEquiv.toLinearEquiv_toLinearMap, + traceDual, traceForm_apply, Submodule.mem_comap, AlgEquiv.toLinearMap_apply, + Submodule.mem_mk, AddSubmonoid.mem_mk, AddSubsemigroup.mem_mk, Set.mem_setOf_eq] + apply (FractionRing.algEquiv B L).forall_congr + simp only [restrictScalars_mem, traceForm_apply, AlgEquiv.toEquiv_eq_coe, + EquivLike.coe_coe, mem_comap, AlgEquiv.toLinearMap_apply, AlgEquiv.symm_apply_apply] + refine fun {y} ↦ (forall_congr' $ fun hy ↦ ?_) + rw [Algebra.trace_eq_of_equiv_equiv (FractionRing.algEquiv A K).toRingEquiv + (FractionRing.algEquiv B L).toRingEquiv] + swap + · apply IsLocalization.ringHom_ext (M := A⁰); ext + simp only [AlgEquiv.toRingEquiv_eq_coe, AlgEquiv.toRingEquiv_toRingHom, RingHom.coe_comp, + RingHom.coe_coe, Function.comp_apply, AlgEquiv.commutes, ← IsScalarTower.algebraMap_apply] + rw [IsScalarTower.algebraMap_apply A B (FractionRing B), AlgEquiv.commutes, + ← IsScalarTower.algebraMap_apply] + simp only [AlgEquiv.toRingEquiv_eq_coe, _root_.map_mul, AlgEquiv.coe_ringEquiv, + AlgEquiv.apply_symm_apply] + show (FractionRing.algEquiv A K).symm _ ∈ (algebraMap A (FractionRing A)).range ↔ _ + rw [← (FractionRing.algEquiv A K).symm.toAlgHom.comp_algebraMap, ← RingHom.map_range, + AlgEquiv.toAlgHom_eq_coe, AlgEquiv.coe_ringHom_commutes, Subring.mem_map_equiv] + simp + open scoped Classical namespace FractionalIdeal @@ -176,11 +205,13 @@ def dual (I : FractionalIdeal B⁰ L) : · ext w; exact (IsIntegralClosure.isIntegral_iff (A := B)).symm · rw [Algebra.smul_def, RingHom.map_mul, hy, ← Algebra.smul_def]⟩ -variable [IsDedekindDomain B] {A K} {I J : FractionalIdeal B⁰ L} (hI : I ≠ 0) (hJ : J ≠ 0) +variable [IsDedekindDomain B] {I J : FractionalIdeal B⁰ L} (hI : I ≠ 0) (hJ : J ≠ 0) lemma coe_dual : (dual A K I : Submodule B L) = Iᵛ := by rw [dual, dif_neg hI]; rfl +variable (B L) + @[simp] lemma coe_dual_one : (dual A K (1 : FractionalIdeal B⁰ L) : Submodule B L) = 1ᵛ := by @@ -191,10 +222,14 @@ lemma coe_dual_one : lemma dual_zero : dual A K (0 : FractionalIdeal B⁰ L) = 0 := by rw [dual, dif_pos rfl] +variable {A K L B} + lemma mem_dual {x} : x ∈ dual A K I ↔ ∀ a ∈ I, traceForm K L x a ∈ (algebraMap A K).range := by rw [dual, dif_neg hI]; rfl +variable (A K) + lemma dual_ne_zero : dual A K I ≠ 0 := by obtain ⟨b, hb, hb'⟩ := I.prop @@ -213,13 +248,18 @@ lemma dual_ne_zero : exact IsIntegralClosure.isIntegral_iff (A := B) · exact (Algebra.smul_def _ _).symm +variable {A K} + @[simp] lemma dual_eq_zero_iff : - dual A K I = 0 ↔ I = 0 := ⟨not_imp_not.mp dual_ne_zero, fun e ↦ e.symm ▸ dual_zero⟩ + dual A K I = 0 ↔ I = 0 := + ⟨not_imp_not.mp (dual_ne_zero A K), fun e ↦ e.symm ▸ dual_zero A K L B⟩ lemma dual_ne_zero_iff : dual A K I ≠ 0 ↔ I ≠ 0 := dual_eq_zero_iff.not +variable (A K) + lemma le_dual_inv_aux (hIJ : I * J ≤ 1) : J ≤ dual A K I := by rw [dual, dif_neg hI] @@ -233,27 +273,27 @@ lemma le_dual_inv_aux (hIJ : I * J ≤ 1) : lemma one_le_dual_one : 1 ≤ dual A K (1 : FractionalIdeal B⁰ L) := - le_dual_inv_aux one_ne_zero (by rw [one_mul]) + le_dual_inv_aux A K one_ne_zero (by rw [one_mul]) lemma le_dual_iff : I ≤ dual A K J ↔ I * J ≤ dual A K 1 := by by_cases hI : I = 0 · simp [hI, zero_le] - rw [← coe_le_coe, ← coe_le_coe, coe_mul, coe_dual hJ, coe_dual_one, le_traceDual] + rw [← coe_le_coe, ← coe_le_coe, coe_mul, coe_dual A K hJ, coe_dual_one, le_traceDual] variable (I) lemma inv_le_dual : I⁻¹ ≤ dual A K I := - if hI : I = 0 then by simp [hI] else le_dual_inv_aux hI (le_of_eq (mul_inv_cancel hI)) + if hI : I = 0 then by simp [hI] else le_dual_inv_aux A K hI (le_of_eq (mul_inv_cancel hI)) lemma dual_inv_le : (dual A K I)⁻¹ ≤ I := by by_cases hI : I = 0; · simp [hI] convert mul_right_mono ((dual A K I)⁻¹) - (mul_left_mono I (inv_le_dual (A := A) (K := K) I)) using 1 + (mul_left_mono I (inv_le_dual A K I)) using 1 · simp only [mul_inv_cancel hI, one_mul] - · simp only [mul_inv_cancel (dual_ne_zero (hI := hI)), mul_assoc, mul_one] + · simp only [mul_inv_cancel (dual_ne_zero A K (hI := hI)), mul_assoc, mul_one] lemma dual_eq_mul_inv : dual A K I = dual A K 1 * I⁻¹ := by @@ -261,14 +301,14 @@ lemma dual_eq_mul_inv : apply le_antisymm · suffices : dual A K I * I ≤ dual A K 1 · convert mul_right_mono I⁻¹ this using 1; simp only [mul_inv_cancel hI, mul_one, mul_assoc] - rw [← le_dual_iff hI] - rw [le_dual_iff hI, mul_assoc, inv_mul_cancel hI, mul_one] + rw [← le_dual_iff A K hI] + rw [le_dual_iff A K hI, mul_assoc, inv_mul_cancel hI, mul_one] variable {I} lemma dual_div_dual : dual A K J / dual A K I = I / J := by - rw [dual_eq_mul_inv J, dual_eq_mul_inv I, mul_div_mul_comm, div_self, one_mul] + rw [dual_eq_mul_inv A K J, dual_eq_mul_inv A K I, mul_div_mul_comm, div_self, one_mul] exact inv_div_inv J I simp only [ne_eq, dual_eq_zero_iff, one_ne_zero, not_false_eq_true] @@ -278,7 +318,7 @@ lemma dual_mul_self : lemma self_mul_dual : I * dual A K I = dual A K 1 := by - rw [mul_comm, dual_mul_self hI] + rw [mul_comm, dual_mul_self A K hI] lemma dual_inv : dual A K I⁻¹ = dual A K 1 * I := by rw [dual_eq_mul_inv, inv_inv] @@ -288,25 +328,107 @@ variable (I) @[simp] lemma dual_dual : dual A K (dual A K I) = I := by - rw [dual_eq_mul_inv, dual_eq_mul_inv (I := I), mul_inv, inv_inv, ← mul_assoc, mul_inv_cancel, + rw [dual_eq_mul_inv, dual_eq_mul_inv A K (I := I), mul_inv, inv_inv, ← mul_assoc, mul_inv_cancel, one_mul] rw [dual_ne_zero_iff] exact one_ne_zero -lemma dual_involutive : - Function.Involutive (dual A K : FractionalIdeal B⁰ L → FractionalIdeal B⁰ L) := dual_dual - -lemma dual_injective : - Function.Injective (dual A K : FractionalIdeal B⁰ L → FractionalIdeal B⁰ L) := - dual_involutive.injective - variable {I} @[simp] lemma dual_le_dual : dual A K I ≤ dual A K J ↔ J ≤ I := by - nth_rewrite 2 [← dual_dual (A := A) (K := K) I] - rw [le_dual_iff hJ, le_dual_iff (I := J), mul_comm] - rwa [dual_ne_zero_iff] + nth_rewrite 2 [← dual_dual A K I] + rw [le_dual_iff A K hJ, le_dual_iff A K (I := J) (by rwa [dual_ne_zero_iff]), mul_comm] + +variable {A K} + +lemma dual_involutive : + Function.Involutive (dual A K : FractionalIdeal B⁰ L → FractionalIdeal B⁰ L) := dual_dual A K + +lemma dual_injective : + Function.Injective (dual A K : FractionalIdeal B⁰ L → FractionalIdeal B⁰ L) := + dual_involutive.injective end FractionalIdeal + +variable (B) +variable [IsDedekindDomain B] + +/-- The different ideal of an extension of integral domains `B/A` is the inverse of the dual of `A` +as an ideal of `B`. See `coeIdeal_differentIdeal` and `coeSubmodule_differentIdeal`. -/ +def differentIdeal [NoZeroSMulDivisors A B] : Ideal B := + (1 / Submodule.traceDual A (FractionRing A) 1 : Submodule B (FractionRing B)).comap + (Algebra.linearMap B (FractionRing B)) + +lemma coeSubmodule_differentIdeal_fractionRing + [NoZeroSMulDivisors A B] (hAB : Algebra.IsIntegral A B) + [IsSeparable (FractionRing A) (FractionRing B)] + [FiniteDimensional (FractionRing A) (FractionRing B)] : + coeSubmodule (FractionRing B) (differentIdeal A B) = + 1 / Submodule.traceDual A (FractionRing A) 1 := by + have : IsIntegralClosure B A (FractionRing B) := + IsIntegralClosure.of_isIntegrallyClosed _ _ _ hAB + rw [coeSubmodule, differentIdeal, Submodule.map_comap_eq, inf_eq_right] + have := FractionalIdeal.dual_inv_le (A := A) (K := FractionRing A) + (1 : FractionalIdeal B⁰ (FractionRing B)) + have : _ ≤ ((1 : FractionalIdeal B⁰ (FractionRing B)) : Submodule B (FractionRing B)) := this + simp only [← one_div, FractionalIdeal.val_eq_coe] at this + rw [FractionalIdeal.coe_div (FractionalIdeal.dual_ne_zero _ _ _), + FractionalIdeal.coe_dual] at this + simpa only [FractionalIdeal.coe_one] using this + · exact one_ne_zero + · exact one_ne_zero + +lemma coeSubmodule_differentIdeal [NoZeroSMulDivisors A B] : + coeSubmodule L (differentIdeal A B) = 1 / Submodule.traceDual A K 1 := by + have : (FractionRing.algEquiv B L).toLinearEquiv.comp (Algebra.linearMap B (FractionRing B)) = + Algebra.linearMap B L := by ext; simp + rw [coeSubmodule, ← this] + have H : RingHom.comp (algebraMap (FractionRing A) (FractionRing B)) + ↑(FractionRing.algEquiv A K).symm.toRingEquiv = + RingHom.comp ↑(FractionRing.algEquiv B L).symm.toRingEquiv (algebraMap K L) + · apply IsLocalization.ringHom_ext A⁰ + ext + simp only [AlgEquiv.toRingEquiv_eq_coe, RingHom.coe_comp, RingHom.coe_coe, + AlgEquiv.coe_ringEquiv, Function.comp_apply, AlgEquiv.commutes, + ← IsScalarTower.algebraMap_apply] + rw [IsScalarTower.algebraMap_apply A B L, AlgEquiv.commutes, ← IsScalarTower.algebraMap_apply] + have : IsSeparable (FractionRing A) (FractionRing B) := IsSeparable.of_equiv_equiv _ _ H + have : FiniteDimensional (FractionRing A) (FractionRing B) := Module.Finite.of_equiv_equiv _ _ H + simp only [AlgEquiv.toLinearEquiv_toLinearMap, Submodule.map_comp] + rw [← coeSubmodule, coeSubmodule_differentIdeal_fractionRing _ _ + (IsIntegralClosure.isIntegral_algebra _ L), + Submodule.map_div, ← AlgEquiv.toAlgHom_toLinearMap, Submodule.map_one] + congr 1 + refine (map_equiv_traceDual A K _).trans ?_ + congr 1 + ext + simp + +variable (L) + +lemma coeIdeal_differentIdeal [NoZeroSMulDivisors A B] : + ↑(differentIdeal A B) = (FractionalIdeal.dual A K (1 : FractionalIdeal B⁰ L))⁻¹ := by + apply FractionalIdeal.coeToSubmodule_injective + simp only [FractionalIdeal.coe_div + (FractionalIdeal.dual_ne_zero _ _ (@one_ne_zero (FractionalIdeal B⁰ L) _ _ _)), + FractionalIdeal.coe_coeIdeal, coeSubmodule_differentIdeal A K, inv_eq_one_div, + FractionalIdeal.coe_dual_one, FractionalIdeal.coe_one] + +variable {A K B L} + +lemma differentialIdeal_le_fractionalIdeal_iff + {I : FractionalIdeal B⁰ L} (hI : I ≠ 0) [NoZeroSMulDivisors A B] : + differentIdeal A B ≤ I ↔ (((I⁻¹ : _) : Submodule B L).restrictScalars A).map + ((Algebra.trace K L).restrictScalars A) ≤ 1 := by + rw [coeIdeal_differentIdeal A K L B, FractionalIdeal.inv_le_comm (by simp) hI, + ← FractionalIdeal.coe_le_coe, FractionalIdeal.coe_dual_one] + refine le_traceDual_iff_map_le_one.trans ?_ + simp + +lemma differentialIdeal_le_iff {I : Ideal B} (hI : I ≠ ⊥) [NoZeroSMulDivisors A B] : + differentIdeal A B ≤ I ↔ (((I⁻¹ : FractionalIdeal B⁰ L) : Submodule B L).restrictScalars A).map + ((Algebra.trace K L).restrictScalars A) ≤ 1 := + (FractionalIdeal.coeIdeal_le_coeIdeal _).symm.trans + (differentialIdeal_le_fractionalIdeal_iff (I := (I : FractionalIdeal B⁰ L)) (by simpa)) diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 630b3c8ea701f..9c16f152210c5 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -369,9 +369,7 @@ theorem exists_multiset_prod_cons_le_and_prod_not_le [IsDedekindDomain A] (hNF : wellFounded_lt.has_min {Z | (Z.map PrimeSpectrum.asIdeal).prod ≤ I ∧ (Z.map PrimeSpectrum.asIdeal).prod ≠ ⊥} ⟨Z₀, hZ₀.1, hZ₀.2⟩ - have hZM : Multiset.prod (Z.map PrimeSpectrum.asIdeal) ≤ M := le_trans hZI hIM - have hZ0 : Z ≠ 0 := by rintro rfl; simp [hM.ne_top] at hZM - obtain ⟨_, hPZ', hPM⟩ := (hM.isPrime.multiset_prod_le (mt Multiset.map_eq_zero.mp hZ0)).mp hZM + obtain ⟨_, hPZ', hPM⟩ := hM.isPrime.multiset_prod_le.mp (hZI.trans hIM) -- Then in fact there is a `P ∈ Z` with `P ≤ M`. obtain ⟨P, hPZ, rfl⟩ := Multiset.mem_map.mp hPZ' classical @@ -722,22 +720,22 @@ theorem Ideal.isPrime_iff_bot_or_prime {P : Ideal A} : IsPrime P ↔ P = ⊥ ∨ hp.elim (fun h => h.symm ▸ Ideal.bot_prime) Ideal.isPrime_of_prime⟩ #align ideal.is_prime_iff_bot_or_prime Ideal.isPrime_iff_bot_or_prime -theorem Ideal.strictAnti_pow (I : Ideal A) (hI0 : I ≠ ⊥) (hI1 : I ≠ ⊤) : +theorem Ideal.pow_right_strictAnti (I : Ideal A) (hI0 : I ≠ ⊥) (hI1 : I ≠ ⊤) : StrictAnti (I ^ · : ℕ → Ideal A) := strictAnti_nat_of_succ_lt fun e => Ideal.dvdNotUnit_iff_lt.mp ⟨pow_ne_zero _ hI0, I, mt isUnit_iff.mp hI1, pow_succ' I e⟩ -#align ideal.strict_anti_pow Ideal.strictAnti_pow +#align ideal.strict_anti_pow Ideal.pow_right_strictAnti theorem Ideal.pow_lt_self (I : Ideal A) (hI0 : I ≠ ⊥) (hI1 : I ≠ ⊤) (e : ℕ) (he : 2 ≤ e) : I ^ e < I := by - convert I.strictAnti_pow hI0 hI1 he + convert I.pow_right_strictAnti hI0 hI1 he dsimp only rw [pow_one] #align ideal.pow_lt_self Ideal.pow_lt_self theorem Ideal.exists_mem_pow_not_mem_pow_succ (I : Ideal A) (hI0 : I ≠ ⊥) (hI1 : I ≠ ⊤) (e : ℕ) : ∃ x ∈ I ^ e, x ∉ I ^ (e + 1) := - SetLike.exists_of_lt (I.strictAnti_pow hI0 hI1 e.lt_succ_self) + SetLike.exists_of_lt (I.pow_right_strictAnti hI0 hI1 e.lt_succ_self) #align ideal.exists_mem_pow_not_mem_pow_succ Ideal.exists_mem_pow_not_mem_pow_succ open UniqueFactorizationMonoid @@ -760,7 +758,7 @@ theorem Ideal.eq_prime_pow_of_succ_lt_of_le {P I : Ideal A} [P_prime : P.IsPrime theorem Ideal.pow_succ_lt_pow {P : Ideal A} [P_prime : P.IsPrime] (hP : P ≠ ⊥) (i : ℕ) : P ^ (i + 1) < P ^ i := - lt_of_le_of_ne (Ideal.pow_le_pow (Nat.le_succ _)) + lt_of_le_of_ne (Ideal.pow_le_pow_right (Nat.le_succ _)) (mt (pow_eq_pow_iff hP (mt Ideal.isUnit_iff.mp P_prime.ne_top)).mp i.succ_ne_self) #align ideal.pow_succ_lt_pow Ideal.pow_succ_lt_pow @@ -770,9 +768,18 @@ theorem Associates.le_singleton_iff (x : A) (n : ℕ) (I : Ideal A) : Ideal.dvd_span_singleton] #align associates.le_singleton_iff Associates.le_singleton_iff +variable {K} + +lemma FractionalIdeal.le_inv_comm {I J : FractionalIdeal A⁰ K} (hI : I ≠ 0) (hJ : J ≠ 0) : + I ≤ J⁻¹ ↔ J ≤ I⁻¹ := by + rw [inv_eq, inv_eq, le_div_iff_mul_le hI, le_div_iff_mul_le hJ, mul_comm] + +lemma FractionalIdeal.inv_le_comm {I J : FractionalIdeal A⁰ K} (hI : I ≠ 0) (hJ : J ≠ 0) : + I⁻¹ ≤ J ↔ J⁻¹ ≤ I := by + simpa using le_inv_comm (A := A) (K := K) (inv_ne_zero hI) (inv_ne_zero hJ) + open FractionalIdeal -variable {K} /-- Strengthening of `IsLocalization.exist_integer_multiples`: Let `J ≠ ⊤` be an ideal in a Dedekind domain `A`, and `f ≠ 0` a finite collection diff --git a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean index 24f2f23df8ec1..f62909baf08d2 100644 --- a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean +++ b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean @@ -64,12 +64,13 @@ variable [Algebra K L] [Algebra A L] [IsScalarTower A K L] variable [Algebra C L] [IsIntegralClosure C A L] [Algebra A C] [IsScalarTower A C L] -/- If `L` is a separable extension of `K = Frac(A)` and `L` has no zero smul divisors by `A`, +/- If `L` is an algebraic extension of `K = Frac(A)` and `L` has no zero smul divisors by `A`, then `L` is the localization of the integral closure `C` of `A` in `L` at `A⁰`. -/ -theorem IsIntegralClosure.isLocalization [IsSeparable K L] [NoZeroSMulDivisors A L] : +theorem IsIntegralClosure.isLocalization (hKL : Algebra.IsAlgebraic K L) : IsLocalization (Algebra.algebraMapSubmonoid C A⁰) L := by haveI : IsDomain C := (IsIntegralClosure.equiv A C L (integralClosure A L)).toMulEquiv.isDomain (integralClosure A L) + haveI : NoZeroSMulDivisors A L := NoZeroSMulDivisors.trans A K L haveI : NoZeroSMulDivisors A C := IsIntegralClosure.noZeroSMulDivisors A L refine' ⟨_, fun z => _, fun {x y} h => ⟨1, _⟩⟩ · rintro ⟨_, x, hx, rfl⟩ @@ -77,13 +78,17 @@ theorem IsIntegralClosure.isLocalization [IsSeparable K L] [NoZeroSMulDivisors A Subtype.coe_mk, map_ne_zero_iff _ (NoZeroSMulDivisors.algebraMap_injective A C)] exact mem_nonZeroDivisors_iff_ne_zero.mp hx · obtain ⟨m, hm⟩ := - IsIntegral.exists_multiple_integral_of_isLocalization A⁰ z (IsSeparable.isIntegral K z) + IsIntegral.exists_multiple_integral_of_isLocalization A⁰ z (hKL z).isIntegral obtain ⟨x, hx⟩ : ∃ x, algebraMap C L x = m • z := IsIntegralClosure.isIntegral_iff.mp hm refine' ⟨⟨x, algebraMap A C m, m, SetLike.coe_mem m, rfl⟩, _⟩ rw [Subtype.coe_mk, ← IsScalarTower.algebraMap_apply, hx, mul_comm, Submonoid.smul_def, smul_def] · simp only [IsIntegralClosure.algebraMap_injective C A L h] -#align is_integral_closure.is_localization IsIntegralClosure.isLocalization + +theorem IsIntegralClosure.isLocalization_of_isSeparable [IsSeparable K L] : + IsLocalization (Algebra.algebraMapSubmonoid C A⁰) L := + IsIntegralClosure.isLocalization A K L C (IsSeparable.isAlgebraic _ _) +#align is_integral_closure.is_localization IsIntegralClosure.isLocalization_of_isSeparable variable [FiniteDimensional K L] @@ -210,7 +215,7 @@ theorem IsIntegralClosure.rank [IsPrincipalIdealRing A] [NoZeroSMulDivisors A L] haveI : Module.Free A C := IsIntegralClosure.module_free A K L C haveI : IsNoetherian A C := IsIntegralClosure.isNoetherian A K L C haveI : IsLocalization (Algebra.algebraMapSubmonoid C A⁰) L := - IsIntegralClosure.isLocalization A K L C + IsIntegralClosure.isLocalization A K L C (Algebra.IsIntegral.of_finite _ _).isAlgebraic let b := Basis.localizationLocalization K A⁰ L (Module.Free.chooseBasis A C) rw [FiniteDimensional.finrank_eq_card_chooseBasisIndex, FiniteDimensional.finrank_eq_card_basis b] #align is_integral_closure.rank IsIntegralClosure.rank diff --git a/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean b/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean index 3bb772cd9389a..b12a797640f6d 100644 --- a/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean +++ b/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean @@ -81,7 +81,7 @@ theorem exists_maximalIdeal_pow_eq_of_principal [IsNoetherianRing R] [LocalRing by_cases hs₃ : s = 0; · rw [hs₃]; exact zero_mem _ obtain ⟨n, u, rfl⟩ := H' s hs₃ (le_maximalIdeal hI' hs₁) rw [mul_comm, Ideal.unit_mul_mem_iff_mem _ u.isUnit] at hs₁ ⊢ - apply Ideal.pow_le_pow (Nat.find_min' this hs₁) + apply Ideal.pow_le_pow_right (Nat.find_min' this hs₁) apply Ideal.pow_mem_pow exact (H _).mpr (dvd_refl _) · rw [hx, Ideal.span_singleton_pow, Ideal.span_le, Set.singleton_subset_iff] @@ -246,7 +246,7 @@ theorem DiscreteValuationRing.TFAE [IsNoetherianRing R] [LocalRing R] [IsDomain obtain ⟨n, rfl⟩ := H I hI obtain ⟨m, rfl⟩ := H J hJ cases' le_total m n with h' h' - · left; exact Ideal.pow_le_pow h' - · right; exact Ideal.pow_le_pow h' + · left; exact Ideal.pow_le_pow_right h' + · right; exact Ideal.pow_le_pow_right h' tfae_finish #align discrete_valuation_ring.tfae DiscreteValuationRing.TFAE diff --git a/Mathlib/RingTheory/Etale.lean b/Mathlib/RingTheory/Etale.lean index a1e42956e7aa9..d6ba0d8dbd727 100644 --- a/Mathlib/RingTheory/Etale.lean +++ b/Mathlib/RingTheory/Etale.lean @@ -338,7 +338,7 @@ theorem FormallySmooth.of_split [FormallySmooth R P] (g : A →ₐ[R] P ⧸ (Rin have : _ = i (f x) := (FormallySmooth.mk_lift I ⟨2, hI⟩ (i.comp f) x : _) rwa [hx, map_zero, ← Ideal.Quotient.mk_eq_mk, Submodule.Quotient.mk_eq_zero] at this intro x hx - have := (Ideal.pow_mono this 2).trans (Ideal.le_comap_pow _ 2) hx + have := (Ideal.pow_right_mono this 2).trans (Ideal.le_comap_pow _ 2) hx rwa [hI] at this have : i.comp f.kerSquareLift = (Ideal.Quotient.mkₐ R _).comp l := by apply AlgHom.coe_ringHom_injective diff --git a/Mathlib/RingTheory/Finiteness.lean b/Mathlib/RingTheory/Finiteness.lean index fa2d1ece3dfb4..38eaa127af764 100644 --- a/Mathlib/RingTheory/Finiteness.lean +++ b/Mathlib/RingTheory/Finiteness.lean @@ -527,9 +527,9 @@ theorem exists_radical_pow_le_of_fg {R : Type*} [CommSemiring R] (I : Ideal R) ( rw [← Ideal.add_eq_sup, add_pow, Ideal.sum_eq_sup, Finset.sup_le_iff] refine' fun i _ => Ideal.mul_le_right.trans _ obtain h | h := le_or_lt n i - · apply Ideal.mul_le_right.trans ((Ideal.pow_le_pow h).trans hn) + · apply Ideal.mul_le_right.trans ((Ideal.pow_le_pow_right h).trans hn) · apply Ideal.mul_le_left.trans - refine' (Ideal.pow_le_pow _).trans hm + refine' (Ideal.pow_le_pow_right _).trans hm rw [add_comm, Nat.add_sub_assoc h.le] apply Nat.le_add_right #align ideal.exists_radical_pow_le_of_fg Ideal.exists_radical_pow_le_of_fg @@ -588,6 +588,11 @@ theorem of_surjective [hM : Finite R M] (f : M →ₗ[R] N) (hf : Surjective f) exact hM.1.map f⟩ #align module.finite.of_surjective Module.Finite.of_surjective +instance quotient (R) {A M} [Semiring R] [AddCommGroup M] [Ring A] [Module A M] [Module R M] + [SMul R A] [IsScalarTower R A M] [Finite R M] + (N : Submodule A M) : Finite R (M ⧸ N) := + Module.Finite.of_surjective (N.mkQ.restrictScalars R) N.mkQ_surjective + /-- The range of a linear map from a finite module is finite. -/ instance range [Finite R M] (f : M →ₗ[R] N) : Finite R (LinearMap.range f) := of_surjective f.rangeRestrict fun ⟨_, y, hy⟩ => ⟨y, Subtype.ext hy⟩ @@ -652,6 +657,19 @@ theorem trans {R : Type*} (A M : Type*) [CommSemiring R] [Semiring A] [Algebra R Submodule.restrictScalars_top]⟩⟩ #align module.finite.trans Module.Finite.trans +lemma of_equiv_equiv {A₁ B₁ A₂ B₂ : Type*} [CommRing A₁] [CommRing B₁] + [CommRing A₂] [CommRing B₂] [Algebra A₁ B₁] [Algebra A₂ B₂] (e₁ : A₁ ≃+* A₂) (e₂ : B₁ ≃+* B₂) + (he : RingHom.comp (algebraMap A₂ B₂) ↑e₁ = RingHom.comp ↑e₂ (algebraMap A₁ B₁)) + [Module.Finite A₁ B₁] : Module.Finite A₂ B₂ := by + letI := e₁.toRingHom.toAlgebra + letI := ((algebraMap A₁ B₁).comp e₁.symm.toRingHom).toAlgebra + haveI : IsScalarTower A₁ A₂ B₁ := IsScalarTower.of_algebraMap_eq + (fun x ↦ by simp [RingHom.algebraMap_toAlgebra]) + let e : B₁ ≃ₐ[A₂] B₂ := { e₂ with commutes' := fun r ↦ by simpa [RingHom.algebraMap_toAlgebra] + using FunLike.congr_fun he.symm (e₁.symm r) } + haveI := Module.Finite.of_restrictScalars_finite A₁ A₂ B₁ + exact Module.Finite.equiv e.toLinearEquiv + end Algebra end Finite diff --git a/Mathlib/RingTheory/Henselian.lean b/Mathlib/RingTheory/Henselian.lean index 203d6bd8676cd..b3f04b7021ba1 100644 --- a/Mathlib/RingTheory/Henselian.lean +++ b/Mathlib/RingTheory/Henselian.lean @@ -227,7 +227,7 @@ instance (priority := 100) IsAdicComplete.henselianRing (R : Type*) [CommRing R] simp only [Finset.mem_Ico] rintro i ⟨h2i, _⟩ have aux : n + 2 ≤ i * (n + 1) := by trans 2 * (n + 1) <;> nlinarith only [h2i] - refine' Ideal.mul_mem_left _ _ (Ideal.pow_le_pow aux _) + refine' Ideal.mul_mem_left _ _ (Ideal.pow_le_pow_right aux _) rw [pow_mul'] refine' Ideal.pow_mem_pow ((Ideal.neg_mem_iff _).2 <| Ideal.mul_mem_right _ _ ih) _ -- we are now in the position to show that `c : ℕ → R` is a Cauchy sequence @@ -242,7 +242,7 @@ instance (priority := 100) IsAdicComplete.henselianRing (R : Type*) [CommRing R] refine' ih.add _ symm rw [SModEq.zero, Ideal.neg_mem_iff] - refine' Ideal.mul_mem_right _ _ (Ideal.pow_le_pow _ (hfcI _)) + refine' Ideal.mul_mem_right _ _ (Ideal.pow_le_pow_right _ (hfcI _)) rw [add_assoc] exact le_self_add -- hence the sequence converges to some limit point `a`, which is the `a` we are looking for @@ -255,7 +255,7 @@ instance (priority := 100) IsAdicComplete.henselianRing (R : Type*) [CommRing R] rw [← Ideal.one_eq_top, Ideal.smul_eq_mul, mul_one] at ha ⊢ refine' (ha.symm.eval f).trans _ rw [SModEq.zero] - exact Ideal.pow_le_pow le_self_add (hfcI _) + exact Ideal.pow_le_pow_right le_self_add (hfcI _) · show a - a₀ ∈ I specialize ha 1 rw [hc, pow_one, ← Ideal.one_eq_top, Ideal.smul_eq_mul, mul_one, sub_eq_add_neg] at ha diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index b48bbab50c5b6..17e7d6ced29fa 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -790,24 +790,24 @@ theorem sup_mul : (I ⊔ J) * K = I * K ⊔ J * K := variable {I J K} -theorem pow_le_pow {m n : ℕ} (h : m ≤ n) : I ^ n ≤ I ^ m := by +theorem pow_le_pow_right {m n : ℕ} (h : m ≤ n) : I ^ n ≤ I ^ m := by cases' Nat.exists_eq_add_of_le h with k hk rw [hk, pow_add] exact le_trans mul_le_inf inf_le_left -#align ideal.pow_le_pow Ideal.pow_le_pow +#align ideal.pow_le_pow_right Ideal.pow_le_pow_right theorem pow_le_self {n : ℕ} (hn : n ≠ 0) : I ^ n ≤ I := calc - I ^ n ≤ I ^ 1 := pow_le_pow (Nat.pos_of_ne_zero hn) + I ^ n ≤ I ^ 1 := pow_le_pow_right (Nat.pos_of_ne_zero hn) _ = I := pow_one _ #align ideal.pow_le_self Ideal.pow_le_self -theorem pow_mono {I J : Ideal R} (e : I ≤ J) (n : ℕ) : I ^ n ≤ J ^ n := by +theorem pow_right_mono {I J : Ideal R} (e : I ≤ J) (n : ℕ) : I ^ n ≤ J ^ n := by induction' n with _ hn · rw [pow_zero, pow_zero] · rw [pow_succ, pow_succ] exact Ideal.mul_mono e hn -#align ideal.pow_mono Ideal.pow_mono +#align ideal.pow_right_mono Ideal.pow_right_mono theorem mul_eq_bot {R : Type*} [CommSemiring R] [NoZeroDivisors R] {I J : Ideal R} : I * J = ⊥ ↔ I = ⊥ ∨ J = ⊥ := @@ -1089,71 +1089,39 @@ theorem radical_pow (n : ℕ) (H : n > 0) : radical (I ^ n) = radical I := H #align ideal.radical_pow Ideal.radical_pow -theorem IsPrime.mul_le {I J P : Ideal R} (hp : IsPrime P) : I * J ≤ P ↔ I ≤ P ∨ J ≤ P := - ⟨fun h => - or_iff_not_imp_left.2 fun hip _ hj => - let ⟨_, hi, hip⟩ := Set.not_subset.1 hip - (hp.mem_or_mem <| h <| mul_mem_mul hi hj).resolve_left hip, - fun h => - Or.casesOn h (le_trans <| le_trans mul_le_inf inf_le_left) - (le_trans <| le_trans mul_le_inf inf_le_right)⟩ +theorem IsPrime.mul_le {I J P : Ideal R} (hp : IsPrime P) : I * J ≤ P ↔ I ≤ P ∨ J ≤ P := by + rw [or_comm, Ideal.mul_le] + simp_rw [hp.mul_mem_iff_mem_or_mem, SetLike.le_def, ← forall_or_left, or_comm, forall_or_left] #align ideal.is_prime.mul_le Ideal.IsPrime.mul_le theorem IsPrime.inf_le {I J P : Ideal R} (hp : IsPrime P) : I ⊓ J ≤ P ↔ I ≤ P ∨ J ≤ P := - ⟨fun h => hp.mul_le.1 <| le_trans mul_le_inf h, fun h => - Or.casesOn h (le_trans inf_le_left) (le_trans inf_le_right)⟩ + ⟨fun h ↦ hp.mul_le.1 <| mul_le_inf.trans h, fun h ↦ h.elim inf_le_left.trans inf_le_right.trans⟩ #align ideal.is_prime.inf_le Ideal.IsPrime.inf_le -theorem IsPrime.multiset_prod_le {s : Multiset (Ideal R)} {P : Ideal R} (hp : IsPrime P) - (hne : s ≠ 0) : s.prod ≤ P ↔ ∃ I ∈ s, I ≤ P := by - suffices s.prod ≤ P → ∃ I ∈ s, I ≤ P from - ⟨this, fun ⟨i, his, hip⟩ => le_trans multiset_prod_le_inf <| le_trans (Multiset.inf_le his) hip⟩ - classical - obtain ⟨b, hb⟩ : ∃ b, b ∈ s := Multiset.exists_mem_of_ne_zero hne - obtain ⟨t, rfl⟩ : ∃ t, s = b ::ₘ t - exact ⟨s.erase b, (Multiset.cons_erase hb).symm⟩ - refine' t.induction_on _ _ - · simp only [exists_prop, Multiset.cons_zero, Multiset.prod_singleton, Multiset.mem_singleton, - exists_eq_left, imp_self] - intro a s ih h - rw [Multiset.cons_swap, Multiset.prod_cons, hp.mul_le] at h - rw [Multiset.cons_swap] - cases' h with h h - · exact ⟨a, Multiset.mem_cons_self a _, h⟩ - obtain ⟨I, hI, ih⟩ : ∃ I ∈ b ::ₘ s, I ≤ P := ih h - exact ⟨I, Multiset.mem_cons_of_mem hI, ih⟩ +theorem IsPrime.multiset_prod_le {s : Multiset (Ideal R)} {P : Ideal R} (hp : IsPrime P) : + s.prod ≤ P ↔ ∃ I ∈ s, I ≤ P := + s.induction_on (by simp [hp.ne_top]) fun I s ih ↦ by simp [hp.mul_le, ih] #align ideal.is_prime.multiset_prod_le Ideal.IsPrime.multiset_prod_le theorem IsPrime.multiset_prod_map_le {s : Multiset ι} (f : ι → Ideal R) {P : Ideal R} - (hp : IsPrime P) (hne : s ≠ 0) : (s.map f).prod ≤ P ↔ ∃ i ∈ s, f i ≤ P := by - rw [hp.multiset_prod_le (mt Multiset.map_eq_zero.mp hne)] - simp_rw [Multiset.mem_map, exists_exists_and_eq_and] + (hp : IsPrime P) : (s.map f).prod ≤ P ↔ ∃ i ∈ s, f i ≤ P := by + simp_rw [hp.multiset_prod_le, Multiset.mem_map, exists_exists_and_eq_and] #align ideal.is_prime.multiset_prod_map_le Ideal.IsPrime.multiset_prod_map_le -theorem IsPrime.prod_le {s : Finset ι} {f : ι → Ideal R} {P : Ideal R} (hp : IsPrime P) - (hne : s.Nonempty) : s.prod f ≤ P ↔ ∃ i ∈ s, f i ≤ P := - hp.multiset_prod_map_le f (mt Finset.val_eq_zero.mp hne.ne_empty) +theorem IsPrime.prod_le {s : Finset ι} {f : ι → Ideal R} {P : Ideal R} (hp : IsPrime P) : + s.prod f ≤ P ↔ ∃ i ∈ s, f i ≤ P := + hp.multiset_prod_map_le f #align ideal.is_prime.prod_le Ideal.IsPrime.prod_le -theorem IsPrime.inf_le' {s : Finset ι} {f : ι → Ideal R} {P : Ideal R} (hp : IsPrime P) - (hsne : s.Nonempty) : s.inf f ≤ P ↔ ∃ i ∈ s, f i ≤ P := - ⟨fun h => (hp.prod_le hsne).1 <| le_trans prod_le_inf h, fun ⟨_, his, hip⟩ => - le_trans (Finset.inf_le his) hip⟩ +theorem IsPrime.inf_le' {s : Finset ι} {f : ι → Ideal R} {P : Ideal R} (hp : IsPrime P) : + s.inf f ≤ P ↔ ∃ i ∈ s, f i ≤ P := + ⟨fun h ↦ hp.prod_le.1 <| prod_le_inf.trans h, fun ⟨_, his, hip⟩ ↦ (Finset.inf_le his).trans hip⟩ #align ideal.is_prime.inf_le' Ideal.IsPrime.inf_le' -- Porting note: needed to add explicit coercions (· : Set R). theorem subset_union {R : Type u} [Ring R] {I J K : Ideal R} : (I : Set R) ⊆ J ∪ K ↔ I ≤ J ∨ I ≤ K := - ⟨fun h => - or_iff_not_imp_left.2 fun hij s hsi => - let ⟨r, hri, hrj⟩ := Set.not_subset.1 hij - by_contradiction fun hsk => - Or.casesOn (h <| I.add_mem hri hsi) - (fun hj => hrj <| add_sub_cancel r s ▸ J.sub_mem hj ((h hsi).resolve_right hsk)) fun hk => - hsk <| add_sub_cancel' r s ▸ K.sub_mem hk ((h hri).resolve_left hrj), - fun h => - Or.casesOn h (fun h => Set.Subset.trans h <| Set.subset_union_left (J : Set R) K) fun h => - Set.Subset.trans h <| Set.subset_union_right (J : Set R) K ⟩ + AddSubgroupClass.subset_union #align ideal.subset_union Ideal.subset_union theorem subset_union_prime' {R : Type u} [CommRing R] {s : Finset ι} {f : ι → Ideal R} {a b : ι} @@ -1233,10 +1201,7 @@ theorem subset_union_prime' {R : Type u} [CommRing R] {s : Finset ι} {f : ι by_cases Hi : I ≤ f i · exact Or.inr (Or.inr ⟨i, Finset.mem_insert_self i t, Hi⟩) have : ¬I ⊓ f a ⊓ f b ⊓ t.inf f ≤ f i := by - rcases t.eq_empty_or_nonempty with (rfl | hsne) - · rw [Finset.inf_empty, inf_top_eq, hp.1.inf_le, hp.1.inf_le, not_or, not_or] - exact ⟨⟨Hi, Ha⟩, Hb⟩ - simp only [hp.1.inf_le, hp.1.inf_le' hsne, not_or] + simp only [hp.1.inf_le, hp.1.inf_le', not_or] exact ⟨⟨⟨Hi, Ha⟩, Hb⟩, Ht⟩ rcases Set.not_subset.1 this with ⟨r, ⟨⟨⟨hrI, hra⟩, hrb⟩, hr⟩, hri⟩ by_cases HI : (I : Set R) ⊆ f a ∪ f b ∪ ⋃ j ∈ (↑t : Set ι), f j diff --git a/Mathlib/RingTheory/IntegralClosure.lean b/Mathlib/RingTheory/IntegralClosure.lean index 86743a62a9475..249bd9e14d3ce 100644 --- a/Mathlib/RingTheory/IntegralClosure.lean +++ b/Mathlib/RingTheory/IntegralClosure.lean @@ -170,22 +170,29 @@ theorem isIntegral_iff_isIntegral_closure_finite {r : B} : exact hsr.of_subring _ #align is_integral_iff_is_integral_closure_finite isIntegral_iff_isIntegral_closure_finite -theorem IsIntegral.fg_adjoin_singleton {x : B} (hx : IsIntegral R x) : - (Algebra.adjoin R {x}).toSubmodule.FG := by - rcases hx with ⟨f, hfm, hfx⟩ - use (Finset.range <| f.natDegree + 1).image (x ^ ·) +theorem Submodule.span_range_natDegree_eq_adjoin {R A} [CommRing R] [Semiring A] [Algebra R A] + {x : A} {f : R[X]} (hf : f.Monic) (hfx : aeval x f = 0) : + span R (Finset.image (x ^ ·) (Finset.range (natDegree f))) = + Subalgebra.toSubmodule (Algebra.adjoin R {x}) := by + nontriviality A + have hf1 : f ≠ 1 := by rintro rfl; simp [one_ne_zero' A] at hfx refine (span_le.mpr fun s hs ↦ ?_).antisymm fun r hr ↦ ?_ · rcases Finset.mem_image.1 hs with ⟨k, -, rfl⟩ exact (Algebra.adjoin R {x}).pow_mem (Algebra.subset_adjoin rfl) k rw [Subalgebra.mem_toSubmodule, Algebra.adjoin_singleton_eq_range_aeval] at hr rcases (aeval x).mem_range.mp hr with ⟨p, rfl⟩ - rw [← modByMonic_add_div p hfm, map_add, map_mul, aeval_def x f, hfx, + rw [← modByMonic_add_div p hf, map_add, map_mul, hfx, zero_mul, add_zero, ← sum_C_mul_X_pow_eq (p %ₘ f), aeval_def, eval₂_sum, sum_def] refine sum_mem fun k hkq ↦ ?_ rw [C_mul_X_pow_eq_monomial, eval₂_monomial, ← Algebra.smul_def] - exact smul_mem _ _ (subset_span <| Finset.mem_image_of_mem _ <| Finset.mem_range_succ_iff.mpr <| - (le_natDegree_of_mem_supp _ hkq).trans <| natDegree_modByMonic_le p hfm) -#align fg_adjoin_singleton_of_integral IsIntegral.fg_adjoin_singleton + exact smul_mem _ _ (subset_span <| Finset.mem_image_of_mem _ <| Finset.mem_range.mpr <| + (le_natDegree_of_mem_supp _ hkq).trans_lt <| natDegree_modByMonic_lt p hf hf1) + +theorem IsIntegral.fg_adjoin_singleton {x : B} (hx : IsIntegral R x) : + (Algebra.adjoin R {x}).toSubmodule.FG := by + rcases hx with ⟨f, hfm, hfx⟩ + use (Finset.range <| f.natDegree).image (x ^ ·) + exact span_range_natDegree_eq_adjoin hfm (by rwa [aeval_def]) theorem fg_adjoin_of_finite {s : Set A} (hfs : s.Finite) (his : ∀ x ∈ s, IsIntegral R x) : (Algebra.adjoin R s).toSubmodule.FG := diff --git a/Mathlib/RingTheory/IntegralRestrict.lean b/Mathlib/RingTheory/IntegralRestrict.lean new file mode 100644 index 0000000000000..6b8c14ce86206 --- /dev/null +++ b/Mathlib/RingTheory/IntegralRestrict.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2023 Andrew Yang, Patrick Lutz. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.IntegrallyClosed +import Mathlib.RingTheory.Norm +import Mathlib.RingTheory.DedekindDomain.IntegralClosure +/-! +# Restriction of various maps between fields to integrally closed subrings. + +In this file, we assume `A` is an integrally closed domain; `K` is the fraction ring of `A`; +`L` is a finite (separable) extension of `K`; `B` is the integral closure of `A` in `L`. +We call this the AKLB setup. + +## Main definition +- `galRestrict`: The restriction `Aut(L/K) → Aut(B/A)` as an `MulEquiv` in an AKLB setup. + +## TODO +Define the restriction of norms and traces. + +-/ +open BigOperators nonZeroDivisors + +variable (A K L B : Type*) [CommRing A] [CommRing B] [Algebra A B] [Field K] [Field L] + [Algebra A K] [IsFractionRing A K] [Algebra B L] + [Algebra K L] [Algebra A L] [IsScalarTower A B L] [IsScalarTower A K L] + [IsIntegralClosure B A L] [FiniteDimensional K L] + +/-- The lift `End(B/A) → End(L/K)` in an ALKB setup. +This is inverse to the restriction. See `galRestrictHom`. -/ +noncomputable +def galLift (σ : B →ₐ[A] B) : L →ₐ[K] L := + haveI := (IsFractionRing.injective A K).isDomain + haveI := NoZeroSMulDivisors.trans A K L + haveI := IsIntegralClosure.isLocalization A K L B (Algebra.IsIntegral.of_finite _ _).isAlgebraic + haveI H : ∀ (y : Algebra.algebraMapSubmonoid B A⁰), + IsUnit (((algebraMap B L).comp σ) (y : B)) := by + rintro ⟨_, x, hx, rfl⟩ + simpa only [RingHom.coe_comp, RingHom.coe_coe, Function.comp_apply, AlgHom.commutes, + isUnit_iff_ne_zero, ne_eq, map_eq_zero_iff _ (NoZeroSMulDivisors.algebraMap_injective _ _), + ← IsScalarTower.algebraMap_apply] using nonZeroDivisors.ne_zero hx + haveI H_eq : (IsLocalization.lift (S := L) H).comp (algebraMap K L) = (algebraMap K L) := by + apply IsLocalization.ringHom_ext A⁰ + ext + simp only [RingHom.coe_comp, Function.comp_apply, ← IsScalarTower.algebraMap_apply A K L, + IsScalarTower.algebraMap_apply A B L, IsLocalization.lift_eq, + RingHom.coe_coe, AlgHom.commutes] + { IsLocalization.lift (S := L) H with commutes' := FunLike.congr_fun H_eq } + +/-- The restriction `End(L/K) → End(B/A)` in an AKLB setup. +Also see `galRestrict` for the `AlgEquiv` version. -/ +noncomputable +def galRestrictHom : (L →ₐ[K] L) ≃* (B →ₐ[A] B) where + toFun := fun f ↦ (IsIntegralClosure.equiv A (integralClosure A L) L B).toAlgHom.comp + (((f.restrictScalars A).comp (IsScalarTower.toAlgHom A B L)).codRestrict + (integralClosure A L) (fun x ↦ IsIntegral.map _ (IsIntegralClosure.isIntegral A L x))) + map_mul' := by + intros σ₁ σ₂ + ext x + apply (IsIntegralClosure.equiv A (integralClosure A L) L B).symm.injective + ext + dsimp + simp only [AlgEquiv.symm_apply_apply, AlgHom.coe_codRestrict, AlgHom.coe_comp, + AlgHom.coe_restrictScalars', IsScalarTower.coe_toAlgHom', Function.comp_apply, + AlgHom.mul_apply, IsIntegralClosure.algebraMap_equiv, Subalgebra.algebraMap_eq] + rfl + invFun := galLift A K L B + left_inv σ := + have := (IsFractionRing.injective A K).isDomain + have := IsIntegralClosure.isLocalization A K L B (Algebra.IsIntegral.of_finite _ _).isAlgebraic + AlgHom.coe_ringHom_injective <| IsLocalization.ringHom_ext (Algebra.algebraMapSubmonoid B A⁰) + <| RingHom.ext fun x ↦ by simp [Subalgebra.algebraMap_eq, galLift] + right_inv σ := + have := (IsFractionRing.injective A K).isDomain + have := IsIntegralClosure.isLocalization A K L B (Algebra.IsIntegral.of_finite _ _).isAlgebraic + AlgHom.ext fun x ↦ + IsIntegralClosure.algebraMap_injective B A L (by simp [Subalgebra.algebraMap_eq, galLift]) + +@[simp] +lemma algebraMap_galRestrictHom_apply (σ : L →ₐ[K] L) (x : B) : + algebraMap B L (galRestrictHom A K L B σ x) = σ (algebraMap B L x) := by + simp [galRestrictHom, Subalgebra.algebraMap_eq] + +@[simp, nolint unusedHavesSuffices] -- false positive from unfolding galRestrictHom +lemma galRestrictHom_symm_algebraMap_apply (σ : B →ₐ[A] B) (x : B) : + (galRestrictHom A K L B).symm σ (algebraMap B L x) = algebraMap B L (σ x) := by + have := (IsFractionRing.injective A K).isDomain + have := IsIntegralClosure.isLocalization A K L B (Algebra.IsIntegral.of_finite _ _).isAlgebraic + simp [galRestrictHom, galLift, Subalgebra.algebraMap_eq] + +/-- The restriction `Aut(L/K) → Aut(B/A)` in an AKLB setup. -/ +noncomputable +def galRestrict : (L ≃ₐ[K] L) ≃* (B ≃ₐ[A] B) := + (AlgEquiv.algHomUnitsEquiv K L).symm.trans + ((Units.mapEquiv <| galRestrictHom A K L B).trans (AlgEquiv.algHomUnitsEquiv A B)) + +variable {K L} + +lemma coe_galRestrict_apply (σ : L ≃ₐ[K] L) : + (galRestrict A K L B σ : B →ₐ[A] B) = galRestrictHom A K L B σ := rfl + +variable {B} + +lemma galRestrict_apply (σ : L ≃ₐ[K] L) (x : B) : + galRestrict A K L B σ x = galRestrictHom A K L B σ x := rfl + +lemma algebraMap_galRestrict_apply (σ : L ≃ₐ[K] L) (x : B) : + algebraMap B L (galRestrict A K L B σ x) = σ (algebraMap B L x) := + algebraMap_galRestrictHom_apply A K L B σ.toAlgHom x + +variable (K L B) + +lemma prod_galRestrict_eq_norm [IsGalois K L] [IsIntegrallyClosed A] (x : B) : + (∏ σ : L ≃ₐ[K] L, galRestrict A K L B σ x) = + algebraMap A B (IsIntegralClosure.mk' (R := A) A (Algebra.norm K <| algebraMap B L x) + (Algebra.isIntegral_norm K (IsIntegralClosure.isIntegral A L x).algebraMap)) := by + apply IsIntegralClosure.algebraMap_injective B A L + rw [← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_eq A K L] + simp only [map_prod, algebraMap_galRestrict_apply, IsIntegralClosure.algebraMap_mk', + Algebra.norm_eq_prod_automorphisms, AlgHom.coe_coe, RingHom.coe_comp, Function.comp_apply] diff --git a/Mathlib/RingTheory/IntegrallyClosed.lean b/Mathlib/RingTheory/IntegrallyClosed.lean index fcb3d0700c6e0..2c2ec88419b13 100644 --- a/Mathlib/RingTheory/IntegrallyClosed.lean +++ b/Mathlib/RingTheory/IntegrallyClosed.lean @@ -79,7 +79,7 @@ end Iff namespace IsIntegrallyClosed -variable {R : Type*} [CommRing R] [id : IsDomain R] [iic : IsIntegrallyClosed R] +variable {R S : Type*} [CommRing R] [CommRing S] [id : IsDomain R] [iic : IsIntegrallyClosed R] variable {K : Type*} [CommRing K] [Algebra R K] [ifr : IsFractionRing R K] @@ -101,7 +101,17 @@ theorem exists_algebraMap_eq_of_pow_mem_subalgebra {K : Type*} [CommRing K] [Alg exists_algebraMap_eq_of_isIntegral_pow hn <| isIntegral_iff.mpr ⟨⟨x ^ n, hx⟩, rfl⟩ #align is_integrally_closed.exists_algebra_map_eq_of_pow_mem_subalgebra IsIntegrallyClosed.exists_algebraMap_eq_of_pow_mem_subalgebra -variable (K) +variable (R S K) + +lemma _root_.IsIntegralClosure.of_isIntegrallyClosed + [Algebra S R] [Algebra S K] [IsScalarTower S R K] (hRS : Algebra.IsIntegral S R) : + IsIntegralClosure R S K := by + refine ⟨IsLocalization.injective _ le_rfl, fun {x} ↦ + ⟨fun hx ↦ IsIntegralClosure.isIntegral_iff.mp (IsIntegral.tower_top (A := R) hx), ?_⟩⟩ + rintro ⟨y, rfl⟩ + exact IsIntegral.map (IsScalarTower.toAlgHom S R K) (hRS y) + +variable {R} theorem integralClosure_eq_bot_iff : integralClosure R K = ⊥ ↔ IsIntegrallyClosed R := by refine' eq_bot_iff.trans _ diff --git a/Mathlib/RingTheory/LittleWedderburn.lean b/Mathlib/RingTheory/LittleWedderburn.lean index 18a29f8e752a8..ab7138df571cc 100644 --- a/Mathlib/RingTheory/LittleWedderburn.lean +++ b/Mathlib/RingTheory/LittleWedderburn.lean @@ -128,7 +128,7 @@ private theorem center_eq_top [Finite D] (hD : InductionHyp D) : Subring.center rw [← aux, ← aux, ← eval_mul] refine (evalRingHom ↑q).map_dvd (X_pow_sub_one_mul_cyclotomic_dvd_X_pow_sub_one_of_dvd ℤ ?_) refine Nat.mem_properDivisors.mpr ⟨⟨_, (finrank_mul_finrank Z Zx D).symm⟩, ?_⟩ - rw [← (Nat.pow_right_strictMono hq).lt_iff_lt, ← card_D, ← card_Zx] + rw [← pow_lt_pow_iff_right hq, ← card_D, ← card_Zx] obtain ⟨b, -, hb⟩ := SetLike.exists_of_lt hZx.lt_top refine card_lt_of_injective_of_not_mem _ Subtype.val_injective (?_ : b ∉ _) rintro ⟨b, rfl⟩ diff --git a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean index f5bb57ffcafd9..d25cc3befe232 100644 --- a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean @@ -494,6 +494,19 @@ instance center.instNonUnitalCommSemiring : NonUnitalCommSemiring (center R) := { Subsemigroup.center.commSemigroup, NonUnitalSubsemiringClass.toNonUnitalNonAssocSemiring (center R) with } +/-- A point-free means of proving membership in the center, for a non-associative ring. + +This can be helpful when working with types that have ext lemmas for `R →+ R`. -/ +lemma _root_.Set.mem_center_iff_addMonoidHom (a : R) : + a ∈ Set.center R ↔ + AddMonoidHom.mulLeft a = .mulRight a ∧ + AddMonoidHom.compr₂ .mul (.mulLeft a) = .comp .mul (.mulLeft a) ∧ + AddMonoidHom.comp .mul (.mulRight a) = .compl₂ .mul (.mulLeft a) ∧ + AddMonoidHom.compr₂ .mul (.mulRight a) = .compl₂ .mul (.mulRight a) := by + rw [Set.mem_center_iff, isMulCentral_iff] + simp_rw [FunLike.ext_iff] + rfl + end NonUnitalNonAssocSemiring section NonUnitalSemiring diff --git a/Mathlib/RingTheory/Perfection.lean b/Mathlib/RingTheory/Perfection.lean index 7cf7b3a8d2d53..c436dda16be99 100644 --- a/Mathlib/RingTheory/Perfection.lean +++ b/Mathlib/RingTheory/Perfection.lean @@ -600,8 +600,8 @@ theorem valAux_add (f g : PreTilt K v O hv p) : rw [valAux_eq hm, valAux_eq hn, valAux_eq hk, RingHom.map_add] cases' le_max_iff.1 (ModP.preVal_add (coeff _ _ (max (max m n) k) f) (coeff _ _ (max (max m n) k) g)) with h h - · exact le_max_of_le_left (pow_le_pow_of_le_left' h _) - · exact le_max_of_le_right (pow_le_pow_of_le_left' h _) + · exact le_max_of_le_left (pow_le_pow_left' h _) + · exact le_max_of_le_right (pow_le_pow_left' h _) #align pre_tilt.val_aux_add PreTilt.valAux_add variable (K v O hv p) diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index 873224e280a7c..5b210a31edb35 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -180,6 +180,72 @@ theorem eval_eq_sum_degreeLTEquiv {n : ℕ} {p : R[X]} (hp : p ∈ degreeLT R n) exact (sum_fin _ (by simp_rw [zero_mul, forall_const]) (mem_degreeLT.mp hp)).symm #align polynomial.eval_eq_sum_degree_lt_equiv Polynomial.eval_eq_sum_degreeLTEquiv +theorem degreeLT_succ_eq_degreeLE {n : ℕ} : degreeLT R (n + 1) = degreeLE R n := by + ext x + by_cases x_zero : x = 0 + · simp_rw [x_zero, Submodule.zero_mem] + · rw [mem_degreeLT, mem_degreeLE, ← natDegree_lt_iff_degree_lt (by rwa [ne_eq]), + ← natDegree_le_iff_degree_le, Nat.lt_succ] + +/-- For every polynomial `p` in the span of a set `s : Set R[X]`, there exists a polynomial of + `p' ∈ s` with higher degree. See also `Polynomial.exists_degree_le_of_mem_span_of_finite`. -/ +theorem exists_degree_le_of_mem_span {s : Set R[X]} {p : R[X]} + (hs : s.Nonempty) (hp : p ∈ Submodule.span R s) : + ∃ p' ∈ s, degree p ≤ degree p' := by + by_contra! h + by_cases hp_zero : p = 0 + · rw [hp_zero, degree_zero] at h + rcases hs with ⟨x,hx⟩ + exact not_lt_bot (h x hx) + · have : p ∈ degreeLT R (natDegree p) := by + refine (Submodule.span_le.mpr fun p' p'_mem => ?_) hp + rw [SetLike.mem_coe, mem_degreeLT, Nat.cast_withBot] + exact lt_of_lt_of_le (h p' p'_mem) degree_le_natDegree + rwa [mem_degreeLT, Nat.cast_withBot, degree_eq_natDegree hp_zero, + Nat.cast_withBot, lt_self_iff_false] at this + +/-- A stronger version of `Polynomial.exists_degree_le_of_mem_span` under the assumption that the + set `s : R[X]` is finite. There exists a polynomial `p' ∈ s` whose degree dominates the degree of + every element of `p ∈ span R s`-/ +theorem exists_degree_le_of_mem_span_of_finite {s : Set R[X]} (s_fin : s.Finite) (hs : s.Nonempty) : + ∃ p' ∈ s, ∀ (p : R[X]), p ∈ Submodule.span R s → degree p ≤ degree p' := by + rcases Set.Finite.exists_maximal_wrt degree s s_fin hs with ⟨a, has, hmax⟩ + refine ⟨a, has, fun p hp => ?_⟩ + rcases exists_degree_le_of_mem_span hs hp with ⟨p', hp'⟩ + have p'max := hmax p' hp'.left + by_cases h : degree a ≤ degree p' + · rw [← p'max h] at hp'; exact hp'.right + · exact le_trans hp'.right (not_le.mp h).le + +/-- The span of every finite set of polynomials is contained in a `degreeLE n` for some `n`. -/ +theorem span_le_degreeLE_of_finite {s : Set R[X]} (s_fin : s.Finite) : + ∃ n : ℕ, Submodule.span R s ≤ degreeLE R n := by + by_cases s_emp : s.Nonempty + · rcases exists_degree_le_of_mem_span_of_finite s_fin s_emp with ⟨p', _, hp'max⟩ + exact ⟨natDegree p', fun p hp => mem_degreeLE.mpr ((hp'max _ hp).trans degree_le_natDegree)⟩ + · rw [Set.not_nonempty_iff_eq_empty] at s_emp + rw [s_emp, Submodule.span_empty] + exact ⟨0, bot_le⟩ + +/-- The span of every finite set of polynomials is contained in a `degreeLT n` for some `n`. -/ +theorem span_of_finite_le_degreeLT {s : Set R[X]} (s_fin : s.Finite) : + ∃ n : ℕ, Submodule.span R s ≤ degreeLT R n := by + rcases span_le_degreeLE_of_finite s_fin with ⟨n, _⟩ + exact ⟨n + 1, by rwa [degreeLT_succ_eq_degreeLE]⟩ + +/-- If `R` is a nontrivial ring, the polynomials `R[X]` are not finite as an `R`-module. When `R` is +a field, this is equivalent to `R[X]` being an infinite-dimensional vector space over `R`. -/ +theorem not_finite [Nontrivial R] : ¬ Module.Finite R R[X] := by + rw [Module.finite_def, Submodule.fg_def] + push_neg + intro s hs contra + rcases span_le_degreeLE_of_finite hs with ⟨n,hn⟩ + have : ((X : R[X]) ^ (n + 1)) ∈ Polynomial.degreeLE R ↑n := by + rw [contra] at hn + exact hn Submodule.mem_top + rw [mem_degreeLE, degree_X_pow, Nat.cast_le, add_le_iff_nonpos_right, nonpos_iff_eq_zero] at this + exact one_ne_zero this + /-- The finset of nonzero coefficients of a polynomial. -/ def frange (p : R[X]) : Finset R := letI := Classical.decEq R @@ -239,7 +305,7 @@ theorem Monic.geom_sum {P : R[X]} (hP : P.Monic) (hdeg : 0 < P.natDegree) {n : · simp only [Finset.mem_range, degree_eq_natDegree (hP.pow _).ne_zero] simp only [Nat.cast_lt, hP.natDegree_pow] intro k - exact nsmul_lt_nsmul hdeg + exact nsmul_lt_nsmul_left hdeg · rw [bot_lt_iff_ne_bot, Ne.def, degree_eq_bot] exact (hP.pow _).ne_zero #align polynomial.monic.geom_sum Polynomial.Monic.geom_sum @@ -632,7 +698,7 @@ theorem _root_.Polynomial.coeff_prod_mem_ideal_pow_tsub {ι : Type*} (s : Finset apply sum_mem rintro ⟨i, j⟩ e obtain rfl : i + j = k := mem_antidiagonal.mp e - apply Ideal.pow_le_pow add_tsub_add_le_tsub_add_tsub + apply Ideal.pow_le_pow_right add_tsub_add_le_tsub_add_tsub rw [pow_add] exact Ideal.mul_mem_mul (h _ (Finset.mem_insert.mpr <| Or.inl rfl) _) diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Expand.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Expand.lean index f6b4eaa205b7c..1f76d7e29433f 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Expand.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Expand.lean @@ -175,7 +175,7 @@ theorem isRoot_cyclotomic_prime_pow_mul_iff_of_charP {m k p : ℕ} {R : Type*} [ rw [cyclotomic_mul_prime_pow_eq R (NeZero.not_char_dvd R p m) hk, IsRoot.def, eval_pow, h, zero_pow] simp only [tsub_pos_iff_lt] - apply pow_strictMono_right hp.out.one_lt (Nat.pred_lt hk.ne') + apply pow_right_strictMono hp.out.one_lt (Nat.pred_lt hk.ne') #align polynomial.is_root_cyclotomic_prime_pow_mul_iff_of_char_p Polynomial.isRoot_cyclotomic_prime_pow_mul_iff_of_charP end CharP diff --git a/Mathlib/RingTheory/QuotientNilpotent.lean b/Mathlib/RingTheory/QuotientNilpotent.lean index 3d1e9248e9155..437e6dc96931b 100644 --- a/Mathlib/RingTheory/QuotientNilpotent.lean +++ b/Mathlib/RingTheory/QuotientNilpotent.lean @@ -46,7 +46,7 @@ theorem Ideal.IsNilpotent.induction_on (hI : IsNilpotent I) apply h₂ (I ^ 2) _ (Ideal.pow_le_self two_ne_zero) · apply H n.succ _ (I ^ 2) · rw [← pow_mul, eq_bot_iff, ← hI, Nat.succ_eq_add_one, Nat.succ_eq_add_one] - apply Ideal.pow_le_pow (by linarith) + apply Ideal.pow_le_pow_right (by linarith) · exact n.succ.lt_succ_self · apply h₁ rw [← Ideal.map_pow, Ideal.map_quotient_self] diff --git a/Mathlib/RingTheory/TensorProduct.lean b/Mathlib/RingTheory/TensorProduct.lean index 012abb6ecb3b5..0bda602a6d98c 100644 --- a/Mathlib/RingTheory/TensorProduct.lean +++ b/Mathlib/RingTheory/TensorProduct.lean @@ -699,7 +699,10 @@ theorem lift_comp_includeRight (f : A →ₐ[S] C) (g : B →ₐ[R] C) (hfg : Pairs of algebra morphisms that commute are equivalent to algebra morphisms from the tensor product. -This is `Algebra.TensorProduct.lift` as an equivalence. -/ +This is `Algebra.TensorProduct.lift` as an equivalence. + +See also `GradedTensorProduct.liftEquiv` for an alternative commutativity requirement for graded +algebra. -/ @[simps] def liftEquiv [IsScalarTower R S A] [IsScalarTower R S C] : {fg : (A →ₐ[S] C) × (B →ₐ[R] C) // ∀ x y, Commute (fg.1 x) (fg.2 y)} diff --git a/Mathlib/RingTheory/Trace.lean b/Mathlib/RingTheory/Trace.lean index e97e24b66b0da..183407924e4e7 100644 --- a/Mathlib/RingTheory/Trace.lean +++ b/Mathlib/RingTheory/Trace.lean @@ -13,6 +13,7 @@ import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure import Mathlib.FieldTheory.PrimitiveElement import Mathlib.FieldTheory.Galois import Mathlib.RingTheory.PowerBasis +import Mathlib.FieldTheory.Minpoly.MinpolyDiv #align_import ring_theory.trace from "leanprover-community/mathlib"@"3e068ece210655b7b9a9477c3aff38a492400aa1" @@ -43,6 +44,8 @@ the roots of the minimal polynomial of `s` over `R`. algebraically closed field * `traceForm_nondegenerate`: the trace form over a separable extension is a nondegenerate bilinear form +* `traceForm_dualBasis_powerBasis_eq`: The dual basis of a powerbasis `{1, x, x²...}` under the + trace form is `aᵢ / f'(x)`, with `f` being the minpoly of `x` and `f / (X - x) = ∑ aᵢxⁱ`. ## Implementation notes @@ -660,4 +663,32 @@ theorem Algebra.trace_surjective [FiniteDimensional K L] [IsSeparable K L] : rw [LinearMap.range_eq_bot] exact Algebra.trace_ne_zero K L +variable {K L} + +/-- +The dual basis of a powerbasis `{1, x, x²...}` under the trace form is `aᵢ / f'(x)`, +with `f` being the minimal polynomial of `x` and `f / (X - x) = ∑ aᵢxⁱ`. +-/ +lemma traceForm_dualBasis_powerBasis_eq [FiniteDimensional K L] [IsSeparable K L] + (pb : PowerBasis K L) (i) : + (Algebra.traceForm K L).dualBasis (traceForm_nondegenerate K L) pb.basis i = + (minpolyDiv K pb.gen).coeff i / aeval pb.gen (derivative <| minpoly K pb.gen) := by + classical + apply ((Algebra.traceForm K L).toDual (traceForm_nondegenerate K L)).injective + apply pb.basis.ext + intro j + simp only [BilinForm.toDual_def, BilinForm.apply_dualBasis_left] + apply (algebraMap K (AlgebraicClosure K)).injective + have := congr_arg (coeff · i) (sum_smul_minpolyDiv_eq_X_pow (AlgebraicClosure K) + pb.adjoin_gen_eq_top (r := j) (pb.finrank.symm ▸ j.prop)) + simp only [AlgEquiv.toAlgHom_eq_coe, Polynomial.map_smul, map_div₀, + map_pow, RingHom.coe_coe, AlgHom.coe_coe, finset_sum_coeff, coeff_smul, coeff_map, smul_eq_mul, + coeff_X_pow, ← Fin.ext_iff, @eq_comm _ i] at this + rw [PowerBasis.coe_basis, Algebra.traceForm_apply, RingHom.map_ite_one_zero, + ← this, trace_eq_sum_embeddings (E := AlgebraicClosure K)] + apply Finset.sum_congr rfl + intro σ _ + simp only [_root_.map_mul, map_div₀, map_pow] + ring + end DetNeZero diff --git a/Mathlib/RingTheory/Valuation/Integral.lean b/Mathlib/RingTheory/Valuation/Integral.lean index e0aef90c67948..e8b7395cf51fa 100644 --- a/Mathlib/RingTheory/Valuation/Integral.lean +++ b/Mathlib/RingTheory/Valuation/Integral.lean @@ -41,8 +41,8 @@ theorem mem_of_integral {x : R} (hx : IsIntegral O x) : x ∈ v.integer := rw [eval₂_mul, eval₂_pow, eval₂_C, eval₂_X, v.map_mul, v.map_pow, ← one_mul (v x ^ p.natDegree)] cases' (hv.2 <| p.coeff i).lt_or_eq with hvpi hvpi - · exact mul_lt_mul₀ hvpi (pow_lt_pow₀ hvx <| Finset.mem_range.1 hi) - · erw [hvpi]; rw [one_mul, one_mul]; exact pow_lt_pow₀ hvx (Finset.mem_range.1 hi) + · exact mul_lt_mul₀ hvpi (pow_lt_pow_right₀ hvx <| Finset.mem_range.1 hi) + · erw [hvpi]; rw [one_mul, one_mul]; exact pow_lt_pow_right₀ hvx (Finset.mem_range.1 hi) #align valuation.integers.mem_of_integral Valuation.Integers.mem_of_integral protected theorem integralClosure : integralClosure O R = ⊥ := diff --git a/Mathlib/RingTheory/WittVector/Frobenius.lean b/Mathlib/RingTheory/WittVector/Frobenius.lean index f7c15d5f29db0..3d0f384daa5ef 100644 --- a/Mathlib/RingTheory/WittVector/Frobenius.lean +++ b/Mathlib/RingTheory/WittVector/Frobenius.lean @@ -138,7 +138,7 @@ theorem map_frobeniusPoly.key₂ {n i j : ℕ} (hi : i ≤ n) (hj : j < p ^ (n - add_assoc, tsub_right_comm, add_comm i, tsub_add_cancel_of_le (le_tsub_of_add_le_right ((le_tsub_iff_left hi).mp h₁))] have hle : p ^ m ≤ j + 1 := h ▸ Nat.le_of_dvd j.succ_pos (multiplicity.pow_multiplicity_dvd _) - exact ⟨(pow_le_pow_iff hp.1.one_lt).1 (hle.trans hj), + exact ⟨(pow_le_pow_iff_right hp.1.one_lt).1 (hle.trans hj), Nat.le_of_lt_succ ((Nat.lt_pow_self hp.1.one_lt m).trans_le hle)⟩ #align witt_vector.map_frobenius_poly.key₂ WittVector.map_frobeniusPoly.key₂ diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 946dc432dad93..8c3e3617af161 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -10,7 +10,7 @@ import Mathlib.Data.Nat.PartENat import Mathlib.Data.Set.Countable import Mathlib.Logic.Small.Basic import Mathlib.Order.ConditionallyCompleteLattice.Basic -import Mathlib.Order.SuccPred.Limit +import Mathlib.Order.SuccPred.CompleteLinearOrder import Mathlib.SetTheory.Cardinal.SchroederBernstein import Mathlib.Tactic.Positivity import Mathlib.Tactic.PPWithUniv @@ -379,6 +379,7 @@ instance : Zero Cardinal.{u} := instance : Inhabited Cardinal.{u} := ⟨0⟩ +@[simp] theorem mk_eq_zero (α : Type u) [IsEmpty α] : #α = 0 := (Equiv.equivOfIsEmpty α (ULift (Fin 0))).cardinal_eq #align cardinal.mk_eq_zero Cardinal.mk_eq_zero @@ -1011,6 +1012,22 @@ protected theorem iSup_of_empty {ι} (f : ι → Cardinal) [IsEmpty ι] : iSup f ciSup_of_empty f #align cardinal.supr_of_empty Cardinal.iSup_of_empty +lemma exists_eq_of_iSup_eq_of_not_isSuccLimit + {ι : Type u} (f : ι → Cardinal.{max u v}) (ω : Cardinal.{max u v}) + (hω : ¬ Order.IsSuccLimit ω) + (h : ⨆ i : ι, f i = ω) : ∃ i, f i = ω := + IsLUB.exists_of_not_isSuccLimit (h ▸ isLUB_csSup' (bddAbove_range.{u, v} f)) hω + +lemma exists_eq_of_iSup_eq_of_not_isLimit + {ι : Type u} [hι : Nonempty ι] (f : ι → Cardinal.{max u v}) (ω : Cardinal.{max u v}) + (hω : ¬ ω.IsLimit) + (h : ⨆ i : ι, f i = ω) : ∃ i, f i = ω := by + refine (not_and_or.mp hω).elim (fun e ↦ ⟨hι.some, ?_⟩) + (Cardinal.exists_eq_of_iSup_eq_of_not_isSuccLimit.{u, v} f ω · h) + cases not_not.mp e + rw [← le_zero_iff] at h ⊢ + exact (le_ciSup (Cardinal.bddAbove_range.{u, v} f) _).trans h + -- Portin note: simpNF is not happy with universe levels. @[simp, nolint simpNF] theorem lift_mk_shrink (α : Type u) [Small.{v} α] : @@ -1449,13 +1466,18 @@ theorem isLimit_aleph0 : IsLimit ℵ₀ := ⟨aleph0_ne_zero, isSuccLimit_aleph0⟩ #align cardinal.is_limit_aleph_0 Cardinal.isLimit_aleph0 +lemma not_isLimit_natCast : (n : ℕ) → ¬ IsLimit (n : Cardinal.{u}) + | 0, e => e.1 rfl + | Nat.succ n, e => Order.not_isSuccLimit_succ _ (nat_succ n ▸ e.2) + theorem IsLimit.aleph0_le {c : Cardinal} (h : IsLimit c) : ℵ₀ ≤ c := by by_contra! h' - rcases lt_aleph0.1 h' with ⟨_ | n, rfl⟩ - · exact h.ne_zero.irrefl - · rw [nat_succ] at h - exact not_isSuccLimit_succ _ h.isSuccLimit -#align cardinal.is_limit.aleph_0_le Cardinal.IsLimit.aleph0_le + rcases lt_aleph0.1 h' with ⟨n, rfl⟩ + exact not_isLimit_natCast n h + +lemma exists_eq_natCast_of_iSup_eq {ι : Type u} [Nonempty ι] (f : ι → Cardinal.{max u v}) + (n : ℕ) (h : ⨆ i, f i = n) : ∃ i, f i = n := + exists_eq_of_iSup_eq_of_not_isLimit.{u, v} f _ (not_isLimit_natCast n) h @[simp] theorem range_natCast : range ((↑) : ℕ → Cardinal) = Iio ℵ₀ := diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 43a6de2953e97..b5b10eacecd20 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -923,8 +923,7 @@ theorem mk_bounded_subset {α : Type*} (h : ∀ x < #α, (2^x) < #α) {r : α theorem mk_subset_mk_lt_cof {α : Type*} (h : ∀ x < #α, (2^x) < #α) : #{ s : Set α // #s < cof (#α).ord } = #α := by rcases eq_or_ne #α 0 with (ha | ha) - · rw [ha] - simp [fun s => (Cardinal.zero_le s).not_lt] + · simp [ha] have h' : IsStrongLimit #α := ⟨ha, h⟩ rcases ord_eq α with ⟨r, wo, hr⟩ haveI := wo diff --git a/Mathlib/SetTheory/Cardinal/Finite.lean b/Mathlib/SetTheory/Cardinal/Finite.lean index 01af4796a066e..9c7c41fae7c48 100644 --- a/Mathlib/SetTheory/Cardinal/Finite.lean +++ b/Mathlib/SetTheory/Cardinal/Finite.lean @@ -42,7 +42,18 @@ theorem card_eq_fintype_card [Fintype α] : Nat.card α = Fintype.card α := mk_toNat_eq_card #align nat.card_eq_fintype_card Nat.card_eq_fintype_card -lemma card_eq_zero_of_isEmpty [IsEmpty α] : Nat.card α = 0 := by simp [Nat.card] +lemma card_eq_finsetCard (s : Finset α) : Nat.card s = s.card := by + simp only [Nat.card_eq_fintype_card, Fintype.card_coe] + +lemma card_eq_card_toFinset (s : Set α) [Fintype s] : Nat.card s = s.toFinset.card := by + simp only [← Nat.card_eq_finsetCard, s.mem_toFinset] + +lemma card_eq_card_finite_toFinset {s : Set α} (hs : s.Finite) : Nat.card s = hs.toFinset.card := by + simp only [← Nat.card_eq_finsetCard, hs.mem_toFinset] + +@[simp] theorem card_of_isEmpty [IsEmpty α] : Nat.card α = 0 := by simp [Nat.card] +#align nat.card_of_is_empty Nat.card_of_isEmpty + @[simp] lemma card_eq_zero_of_infinite [Infinite α] : Nat.card α = 0 := mk_toNat_of_infinite #align nat.card_eq_zero_of_infinite Nat.card_eq_zero_of_infinite @@ -84,9 +95,41 @@ theorem card_eq_of_equiv_fin {α : Type*} {n : ℕ} (f : α ≃ Fin n) : Nat.car simpa only [card_eq_fintype_card, Fintype.card_fin] using card_congr f #align nat.card_eq_of_equiv_fin Nat.card_eq_of_equiv_fin -lemma card_mono {s t : Set α} (ht : t.Finite) (h : s ⊆ t) : Nat.card s ≤ Nat.card t := +section Set +open Set +variable {s t : Set α} + +lemma card_mono (ht : t.Finite) (h : s ⊆ t) : Nat.card s ≤ Nat.card t := toNat_le_of_le_of_lt_aleph0 ht.lt_aleph0 <| mk_le_mk_of_subset h +lemma card_image_le (hs : s.Finite) : Nat.card (f '' s) ≤ Nat.card s := + have := hs.to_subtype; card_le_card_of_surjective (imageFactorization f s) surjective_onto_image + +lemma card_image_of_injOn (hf : s.InjOn f) : Nat.card (f '' s) = Nat.card s := by + classical + obtain hs | hs := s.finite_or_infinite + · have := hs.fintype + have := fintypeImage s f + simp_rw [Nat.card_eq_fintype_card, Set.card_image_of_inj_on hf] + · have := hs.to_subtype + have := (hs.image hf).to_subtype + simp [Nat.card_eq_zero_of_infinite] + +lemma card_image_of_injective (hf : Injective f) (s : Set α) : + Nat.card (f '' s) = Nat.card s := card_image_of_injOn $ hf.injOn _ + +lemma card_image_equiv (e : α ≃ β) : Nat.card (e '' s) = Nat.card s := + Nat.card_congr (e.image s).symm + +lemma card_preimage_of_injOn {s : Set β} (hf : (f ⁻¹' s).InjOn f) (hsf : s ⊆ range f) : + Nat.card (f ⁻¹' s) = Nat.card s := by + rw [← Nat.card_image_of_injOn hf, image_preimage_eq_iff.2 hsf] + +lemma card_preimage_of_injective {s : Set β} (hf : Injective f) (hsf : s ⊆ range f) : + Nat.card (f ⁻¹' s) = Nat.card s := card_preimage_of_injOn (hf.injOn _) hsf + +end Set + /-- If the cardinality is positive, that means it is a finite type, so there is an equivalence between `α` and `Fin (Nat.card α)`. See also `Finite.equivFin`. -/ def equivFinOfCardPos {α : Type*} (h : Nat.card α ≠ 0) : α ≃ Fin (Nat.card α) := by @@ -117,9 +160,6 @@ theorem card_eq_two_iff' (x : α) : Nat.card α = 2 ↔ ∃! y, y ≠ x := toNat_eq_ofNat.trans (mk_eq_two_iff' x) #align nat.card_eq_two_iff' Nat.card_eq_two_iff' -theorem card_of_isEmpty [IsEmpty α] : Nat.card α = 0 := by simp -#align nat.card_of_is_empty Nat.card_of_isEmpty - @[simp] theorem card_sum [Finite α] [Finite β] : Nat.card (α ⊕ β) = Nat.card α + Nat.card β := by have := Fintype.ofFinite α @@ -159,6 +199,16 @@ theorem card_zmod (n : ℕ) : Nat.card (ZMod n) = n := by end Nat +namespace Set + +lemma card_singleton_prod (a : α) (t : Set β) : Nat.card ({a} ×ˢ t) = Nat.card t := by + rw [singleton_prod, Nat.card_image_of_injective (Prod.mk.inj_left a)] + +lemma card_prod_singleton (s : Set α) (b : β) : Nat.card (s ×ˢ {b}) = Nat.card s := by + rw [prod_singleton, Nat.card_image_of_injective (Prod.mk.inj_right b)] + +end Set + namespace PartENat /-- `PartENat.card α` is the cardinality of `α` as an extended natural number. diff --git a/Mathlib/SetTheory/Game/Short.lean b/Mathlib/SetTheory/Game/Short.lean index 353e43f87ddaa..b44cac9bc0f95 100644 --- a/Mathlib/SetTheory/Game/Short.lean +++ b/Mathlib/SetTheory/Game/Short.lean @@ -169,8 +169,10 @@ theorem short_birthday (x : PGame.{u}) : [Short x] → x.birthday < Ordinal.omeg #align pgame.short_birthday SetTheory.PGame.short_birthday /-- This leads to infinite loops if made into an instance. -/ -def Short.ofIsEmpty {l r xL xR} [IsEmpty l] [IsEmpty r] : Short (PGame.mk l r xL xR) := - Short.mk isEmptyElim isEmptyElim +def Short.ofIsEmpty {l r xL xR} [IsEmpty l] [IsEmpty r] : Short (PGame.mk l r xL xR) := by + have : Fintype l := Fintype.ofIsEmpty + have : Fintype r := Fintype.ofIsEmpty + exact Short.mk isEmptyElim isEmptyElim #align pgame.short.of_is_empty SetTheory.PGame.Short.ofIsEmpty instance short0 : Short 0 := diff --git a/Mathlib/Tactic/ComputeDegree.lean b/Mathlib/Tactic/ComputeDegree.lean index 5ab9de971c149..1044c6b9db4df 100644 --- a/Mathlib/Tactic/ComputeDegree.lean +++ b/Mathlib/Tactic/ComputeDegree.lean @@ -116,6 +116,16 @@ theorem coeff_pow_of_natDegree_le_of_eq_ite' [Semiring R] {m n o : ℕ} {a : R} · exact natDegree_pow_le_of_le m ‹_› · exact Iff.mp ne_comm h +theorem natDegree_smul_le_of_le {n : ℕ} {a : R} {f : R[X]} (hf : natDegree f ≤ n) : + natDegree (a • f) ≤ n := + (natDegree_smul_le a f).trans hf + +theorem degree_smul_le_of_le {n : ℕ} {a : R} {f : R[X]} (hf : degree f ≤ n) : + degree (a • f) ≤ n := + (degree_smul_le a f).trans hf + +theorem coeff_smul {n : ℕ} {a : R} {f : R[X]} : (a • f).coeff n = a * f.coeff n := rfl + section congr_lemmas /-- The following two lemmas should be viewed as a hand-made "congr"-lemmas. @@ -320,6 +330,8 @@ def dispatchLemma π ``natDegree_monomial_le ``degree_monomial_le ``coeff_monomial | .inr ``Polynomial.C => π ``natDegree_C_le ``degree_C_le ``coeff_C + | .inr ``HSMul.hSMul => + π ``natDegree_smul_le_of_le ``degree_smul_le_of_le ``coeff_smul | _ => π ``le_rfl ``le_rfl ``rfl /-- `try_rfl mvs` takes as input a list of `MVarId`s, scans them partitioning them into two @@ -466,6 +478,19 @@ elab_rules : tactic | `(tactic| compute_degree $[!%$bang]?) => focus <| withMain throwError Lean.MessageData.joinSep (m!"The given degree is '{deg}'. However,\n" :: errors) "\n" +/-- `monicity` tries to solve a goal of the form `Monic f`. +It converts the goal into a goal of the form `natDegree f ≤ n` and one of the form `f.coeff n = 1` +and calls `compute_degree` on those two goals. + +The variant `monicity!` starts like `monicity`, but calls `compute_degree!` on the two side-goals. +-/ +macro (name := monicityMacro) "monicity" : tactic => + `(tactic| (apply monic_of_natDegree_le_of_coeff_eq_one <;> compute_degree)) + +@[inherit_doc monicityMacro] +macro "monicity!" : tactic => + `(tactic| (apply monic_of_natDegree_le_of_coeff_eq_one <;> compute_degree!)) + end Tactic end Mathlib.Tactic.ComputeDegree diff --git a/Mathlib/Tactic/GCongr.lean b/Mathlib/Tactic/GCongr.lean index 0e4a2cf7a58c9..00eb904655a0a 100644 --- a/Mathlib/Tactic/GCongr.lean +++ b/Mathlib/Tactic/GCongr.lean @@ -102,8 +102,8 @@ attribute [gcongr] /-! # ≤, ^ -/ attribute [gcongr] - pow_le_pow pow_le_pow' zpow_le_zpow zpow_le_of_le -- ff ^ tt - pow_le_pow_of_le_left pow_le_pow_of_le_left' zpow_le_zpow' -- tt ^ ff + pow_le_pow_right pow_le_pow_right' zpow_le_zpow zpow_le_of_le -- ff ^ tt + pow_le_pow_left pow_le_pow_left' zpow_le_zpow' -- tt ^ ff /-! # <, ^ -/ @@ -112,8 +112,8 @@ theorem zpow_lt_of_lt [LinearOrderedSemifield α] {a : α} {m n : ℤ} (hx : 1 < zpow_strictMono hx h attribute [gcongr] - pow_lt_pow pow_lt_pow' zpow_lt_zpow zpow_lt_of_lt -- ff ^ tt - pow_lt_pow_of_lt_left zpow_lt_zpow' -- tt ^ ff + pow_lt_pow_right pow_lt_pow_right' zpow_lt_zpow zpow_lt_of_lt -- ff ^ tt + pow_lt_pow_left zpow_lt_zpow' -- tt ^ ff /-! # coercions -/ diff --git a/Mathlib/Tactic/HelpCmd.lean b/Mathlib/Tactic/HelpCmd.lean index f706e7282f595..11a979baf8953 100644 --- a/Mathlib/Tactic/HelpCmd.lean +++ b/Mathlib/Tactic/HelpCmd.lean @@ -41,7 +41,9 @@ it will appear as a `(currently: true)` note next to the option. The form `#help option id` will show only options that begin with `id`. -/ -elab "#help " &"option" id:(ppSpace Parser.rawIdent)? : command => do +syntax withPosition("#help " colGt &"option" (colGt ppSpace Parser.rawIdent)?) : command + +private def elabHelpOption (id : Option Ident) : CommandElabM Unit := do let id := id.map (·.raw.getId.toString false) let mut decls : Lean.RBMap _ _ compare := {} for (name, decl) in show Lean.RBMap .. from ← getOptionDecls do @@ -69,6 +71,8 @@ elab "#help " &"option" id:(ppSpace Parser.rawIdent)? : command => do msg := msg ++ .nest 2 (f!"option {name} : {msg1}" ++ .line ++ decl.descr) ++ .line ++ .line logInfo msg +elab_rules : command | `(#help option $(id)?) => elabHelpOption id + /-- The command `#help attribute` (or the short form `#help attr`) shows all attributes that have been defined in the current environment. @@ -83,7 +87,10 @@ as the docstring will be displayed here. The form `#help attr id` will show only attributes that begin with `id`. -/ -elab "#help " (&"attr" <|> &"attribute") id:(ppSpace Parser.rawIdent)? : command => do +syntax withPosition("#help " colGt (&"attr" <|> &"attribute") + (colGt ppSpace Parser.rawIdent)?) : command + +private def elabHelpAttr (id : Option Ident) : CommandElabM Unit := do let id := id.map (·.raw.getId.toString false) let mut decls : Lean.RBMap _ _ compare := {} for (name, decl) in ← attributeMapRef.get do @@ -105,6 +112,10 @@ elab "#help " (&"attr" <|> &"attribute") id:(ppSpace Parser.rawIdent)? : command msg := msg ++ .nest 2 msg1 ++ .line ++ .line logInfo msg +elab_rules : command + | `(#help attr $(id)?) => elabHelpAttr id + | `(#help attribute $(id)?) => elabHelpAttr id + /-- Gets the initial string token in a parser description. For example, for a declaration like `syntax "bla" "baz" term : tactic`, it returns `some "bla"`. Returns `none` for syntax declarations that don't start with a string constant. -/ @@ -136,7 +147,9 @@ but you can click to go to the definition.) It also shows the doc string if avai The form `#help cats id` will show only syntax categories that begin with `id`. -/ -elab "#help " &"cats" id:(ppSpace Parser.rawIdent)? : command => do +syntax withPosition("#help " colGt &"cats" (colGt ppSpace Parser.rawIdent)?) : command + +private def elabHelpCats (id : Option Ident) : CommandElabM Unit := do let id := id.map (·.raw.getId.toString false) let mut decls : Lean.RBMap _ _ compare := {} for (name, cat) in (Parser.parserExtension.getState (← getEnv)).categories do @@ -158,6 +171,8 @@ elab "#help " &"cats" id:(ppSpace Parser.rawIdent)? : command => do msg := msg ++ .nest 2 msg1 ++ (.line ++ .line : Format) logInfo msg +elab_rules : command | `(#help cats $(id)?) => elabHelpCats id + /-- The command `#help cat C` shows all syntaxes that have been defined in syntax category `C` in the current environment. @@ -173,11 +188,11 @@ name of the syntax (which you can also click to go to the definition), and the d * The form `#help cat+ C` will also show information about any `macro`s and `elab`s associated to the listed syntaxes. -/ -elab "#help " &"cat" more:"+"? ppSpace catStx:ident - id:(ppSpace (Parser.rawIdent <|> str))? : command => do - let id := id.map fun id ↦ match id.raw with - | .ident _ _ v _ => v.toString false - | id => id.isStrLit?.get! +syntax withPosition("#help " colGt &"cat" "+"? colGt ident + (colGt ppSpace (Parser.rawIdent <|> str))?) : command + +private def elabHelpCat (more : Option Syntax) (catStx : Ident) (id : Option String) : + CommandElabM Unit := do let mut decls : Lean.RBMap _ _ compare := {} let mut rest : Lean.RBMap _ _ compare := {} let catName := catStx.getId.eraseMacroScopes @@ -231,30 +246,47 @@ elab "#help " &"cat" more:"+"? ppSpace catStx:ident msg ← addMsg k msg m!"syntax ... [{mkConst k}]" logInfo msg +elab_rules : command + | `(#help cat $[+%$more]? $cat) => elabHelpCat more cat none + | `(#help cat $[+%$more]? $cat $id:ident) => elabHelpCat more cat (id.getId.toString false) + | `(#help cat $[+%$more]? $cat $id:str) => elabHelpCat more cat id.getString + /-- The command `#help term` shows all term syntaxes that have been defined in the current environment. See `#help cat` for more information. -/ -macro "#help " tk:&"term" more:"+"? id:(ppSpace (Parser.rawIdent <|> str))? : command => - `(#help cat$[+%$more]? $(mkIdentFrom tk `term) $(id.map (⟨·.raw⟩))?) +syntax withPosition("#help " colGt &"term" "+"? + (colGt ppSpace (Parser.rawIdent <|> str))?) : command +macro_rules + | `(#help term%$tk $[+%$more]? $(id)?) => + `(#help cat$[+%$more]? $(mkIdentFrom tk `term) $(id)?) /-- The command `#help tactic` shows all tactics that have been defined in the current environment. See `#help cat` for more information. -/ -macro "#help " tk:&"tactic" more:"+"? id:(ppSpace (Parser.rawIdent <|> str))? : command => do - `(#help cat$[+%$more]? $(mkIdentFrom tk `tactic) $(id.map (⟨·.raw⟩))?) +syntax withPosition("#help " colGt &"tactic" "+"? + (colGt ppSpace (Parser.rawIdent <|> str))?) : command +macro_rules + | `(#help tactic%$tk $[+%$more]? $(id)?) => + `(#help cat$[+%$more]? $(mkIdentFrom tk `tactic) $(id)?) /-- The command `#help conv` shows all tactics that have been defined in the current environment. See `#help cat` for more information. -/ -macro "#help " tk:&"conv" more:"+"? id:(ppSpace (Parser.rawIdent <|> str))? : command => - `(#help cat$[+%$more]? $(mkIdentFrom tk `conv) $(id.map (⟨·.raw⟩))?) +syntax withPosition("#help " colGt &"conv" "+"? + (colGt ppSpace (Parser.rawIdent <|> str))?) : command +macro_rules + | `(#help conv%$tk $[+%$more]? $(id)?) => + `(#help cat$[+%$more]? $(mkIdentFrom tk `conv) $(id)?) /-- The command `#help command` shows all commands that have been defined in the current environment. See `#help cat` for more information. -/ -macro "#help " tk:&"command" more:"+"? id:(ppSpace (Parser.rawIdent <|> str))? : command => - `(#help cat$[+%$more]? $(mkIdentFrom tk `command) $(id.map (⟨·.raw⟩))?) +syntax withPosition("#help " colGt &"command" "+"? + (colGt ppSpace (Parser.rawIdent <|> str))?) : command +macro_rules + | `(#help command%$tk $[+%$more]? $(id)?) => + `(#help cat$[+%$more]? $(mkIdentFrom tk `command) $(id)?) diff --git a/Mathlib/Tactic/Linarith/Frontend.lean b/Mathlib/Tactic/Linarith/Frontend.lean index 04b01eadfbbe6..79a270de231f7 100644 --- a/Mathlib/Tactic/Linarith/Frontend.lean +++ b/Mathlib/Tactic/Linarith/Frontend.lean @@ -351,7 +351,7 @@ optional arguments: problems. * `transparency` controls how hard `linarith` will try to match atoms to each other. By default it will only unfold `reducible` definitions. -* If `split_hypotheses` is true, `linarith` will split conjunctions in the context into separate +* If `splitHypotheses` is true, `linarith` will split conjunctions in the context into separate hypotheses. * If `splitNe` is `true`, `linarith` will case split on disequality hypotheses. For a given `x ≠ y` hypothesis, `linarith` is run with both `x < y` and `x > y`, diff --git a/Mathlib/Tactic/ModCases.lean b/Mathlib/Tactic/ModCases.lean index 678b0feada813..97e2c05360ae9 100644 --- a/Mathlib/Tactic/ModCases.lean +++ b/Mathlib/Tactic/ModCases.lean @@ -7,14 +7,17 @@ import Mathlib.Data.Int.ModEq /-! # `mod_cases` tactic -The `mod_cases` tactic does case disjunction on `e % n`, where `e : ℤ`, to yield a number of -subgoals in which `e ≡ 0 [ZMOD n]`, ..., `e ≡ n-1 [ZMOD n]` are assumed. +The `mod_cases` tactic does case disjunction on `e % n`, where `e : ℤ` or `e : ℕ`, +to yield `n` new subgoals corresponding to the possible values of `e` modulo `n`. -/ set_option autoImplicit true namespace Mathlib.Tactic.ModCases -open Lean Meta Elab Tactic Term Qq Int +open Lean Meta Elab Tactic Term Qq + +namespace IntMod +open Int /-- `OnModCases n a lb p` represents a partial proof by cases that @@ -73,32 +76,119 @@ partial def proveOnModCases (n : Q(ℕ)) (a : Q(ℤ)) (b : Q(ℕ)) (p : Q(Sort u pure (q(onModCases_succ $b $g $pr), g.mvarId! :: acc) /-- -* The tactic `mod_cases h : e % 3` will perform a case disjunction on `e : ℤ` and yield subgoals - containing the assumptions `h : e ≡ 0 [ZMOD 3]`, `h : e ≡ 1 [ZMOD 3]`, `h : e ≡ 2 [ZMOD 3]` - respectively. +Int case of `mod_cases h : e % n`. +-/ +def modCases (h : TSyntax `Lean.binderIdent) (e : Q(ℤ)) (n : ℕ) : TacticM Unit := do + let ⟨u, p, g⟩ ← inferTypeQ (.mvar (← getMainGoal)) + have lit : Q(ℕ) := mkRawNatLit n + let p₁ : Nat.ble 1 $lit =Q true := ⟨⟩ + let (p₂, gs) ← proveOnModCases lit e (mkRawNatLit 0) p + let gs ← gs.mapM fun g => do + let (fvar, g) ← match h with + | `(binderIdent| $n:ident) => g.intro n.getId + | _ => g.intro `H + g.withContext <| (Expr.fvar fvar).addLocalVarInfoForBinderIdent h + pure g + g.mvarId!.assign q(onModCases_start $p $e $lit $p₁ $p₂) + replaceMainGoal gs + +end IntMod + +namespace NatMod + +/-- +`OnModCases n a lb p` represents a partial proof by cases that +there exists `0 ≤ m < n` such that `a ≡ m (mod n)`. +It asserts that if `∃ m, lb ≤ m < n ∧ a ≡ m (mod n)` holds, then `p` +(where `p` is the current goal). +-/ +def OnModCases (n : ℕ) (a : ℕ) (lb : ℕ) (p : Sort _) := +∀ m, lb ≤ m ∧ m < n ∧ a ≡ m [MOD n] → p + +/-- +The first theorem we apply says that `∃ m, 0 ≤ m < n ∧ a ≡ m (mod n)`. +The actual mathematical content of the proof is here. +-/ +@[inline] def onModCases_start (p : Sort _) (a : ℕ) (n : ℕ) (hn : Nat.ble 1 n = true) + (H : OnModCases n a (nat_lit 0) p) : p := + H (a % n) <| by + refine ⟨Nat.zero_le _, ?_, ?_⟩ + · exact Nat.mod_lt _ (Nat.le_of_ble_eq_true hn) + · rw [Nat.ModEq, Nat.mod_mod] + + +/-- +The end point is that once we have reduced to `∃ m, n ≤ m < n ∧ a ≡ m (mod n)` +there are no more cases to consider. +-/ +@[inline] def onModCases_stop (p : Sort _) (n : ℕ) (a : ℕ) : OnModCases n a n p := + fun _ h => (Nat.not_lt.2 h.1 h.2.1).elim + +/-- +The successor case decomposes `∃ m, b ≤ m < n ∧ a ≡ m (mod n)` into +`a ≡ b (mod n) ∨ ∃ m, b+1 ≤ m < n ∧ a ≡ m (mod n)`, +and the `a ≡ b (mod n) → p` case becomes a subgoal. +-/ +@[inline] def onModCases_succ {p : Sort _} {n : ℕ} {a : ℕ} (b : ℕ) + (h : a ≡ b [MOD n] → p) (H : OnModCases n a (Nat.add b 1) p) : + OnModCases n a b p := + fun z ⟨h₁, h₂⟩ => if e : b = z then h (e ▸ h₂.2) else H _ ⟨Nat.lt_of_le_of_ne h₁ e, h₂⟩ + +/-- +Proves an expression of the form `OnModCases n a b p` where `n` and `b` are raw nat literals +and `b ≤ n`. Returns the list of subgoals `?gi : a ≡ i [MOD n] → p`. +-/ +partial def proveOnModCases (n : Q(ℕ)) (a : Q(ℕ)) (b : Q(ℕ)) (p : Q(Sort u)) : + MetaM (Q(OnModCases $n $a $b $p) × List MVarId) := do + if n.natLit! ≤ b.natLit! then + pure ((q(onModCases_stop $p $n $a) : Expr), []) + else + let ty := q($a ≡ $b [MOD $n] → $p) + let g ← mkFreshExprMVarQ ty + let ((pr : Q(OnModCases $n $a (Nat.add $b 1) $p)), acc) ← + proveOnModCases n a (mkRawNatLit (b.natLit! + 1)) p + pure ((q(onModCases_succ $b $g $pr) : Expr), g.mvarId! :: acc) + +/-- +Nat case of `mod_cases h : e % n`. +-/ +def modCases (h : TSyntax `Lean.binderIdent) (e : Q(ℕ)) (n : ℕ) : TacticM Unit := do + let ⟨u, p, g⟩ ← inferTypeQ (.mvar (← getMainGoal)) + have lit : Q(ℕ) := mkRawNatLit n + let p₁ : Q(Nat.ble 1 $lit = true) := (q(Eq.refl true) : Expr) + let (p₂, gs) ← proveOnModCases lit e (mkRawNatLit 0) p + let gs ← gs.mapM fun g => do + let (fvar, g) ← match h with + | `(binderIdent| $n:ident) => g.intro n.getId + | _ => g.intro `H + g.withContext <| (Expr.fvar fvar).addLocalVarInfoForBinderIdent h + pure g + g.mvarId!.assign q(onModCases_start $p $e $lit $p₁ $p₂) + replaceMainGoal gs + +end NatMod + +/-- +* The tactic `mod_cases h : e % 3` will perform a case disjunction on `e`. + If `e : ℤ`, then it will yield subgoals containing the assumptions + `h : e ≡ 0 [ZMOD 3]`, `h : e ≡ 1 [ZMOD 3]`, `h : e ≡ 2 [ZMOD 3]` + respectively. If `e : ℕ` instead, then it works similarly, except with + `[MOD 3]` instead of `[ZMOD 3]`. * In general, `mod_cases h : e % n` works - when `n` is a positive numeral and `e` is an expression of type `ℤ`. + when `n` is a positive numeral and `e` is an expression of type `ℕ` or `ℤ`. * If `h` is omitted as in `mod_cases e % n`, it will be default-named `H`. -/ -syntax "mod_cases " (atomic(binderIdent " : "))? term:71 " % " num : tactic +syntax "mod_cases " (atomic(binderIdent ":"))? term:71 " % " num : tactic elab_rules : tactic | `(tactic| mod_cases $[$h :]? $e % $n) => do let n := n.getNat if n == 0 then Elab.throwUnsupportedSyntax - let g ← getMainGoal - g.withContext do - let ⟨u, p, g⟩ ← inferTypeQ (.mvar g) - let e : Q(ℤ) ← Tactic.elabTermEnsuringType e q(ℤ) let h := h.getD (← `(binderIdent| _)) - have lit : Q(ℕ) := mkRawNatLit n - let p₁ : Nat.ble 1 $lit =Q true := ⟨⟩ - let (p₂, gs) ← proveOnModCases lit e (mkRawNatLit 0) p - let gs ← gs.mapM fun g => do - let (fvar, g) ← match h with - | `(binderIdent| $n:ident) => g.intro n.getId - | _ => g.intro `H - g.withContext <| (Expr.fvar fvar).addLocalVarInfoForBinderIdent h - pure g - g.mvarId!.assign q(onModCases_start $p $e $lit $p₁ $p₂) - replaceMainGoal gs + withMainContext do + let e ← Tactic.elabTerm e none + let α : Q(Type) ← inferType e + match α with + | ~q(ℤ) => IntMod.modCases h e n + | ~q(ℕ) => NatMod.modCases h e n + | _ => throwError "mod_cases only works with Int and Nat" diff --git a/Mathlib/Tactic/Positivity/Basic.lean b/Mathlib/Tactic/Positivity/Basic.lean index 6840ce0ca5e46..7ef1b4b58a3f0 100644 --- a/Mathlib/Tactic/Positivity/Basic.lean +++ b/Mathlib/Tactic/Positivity/Basic.lean @@ -13,7 +13,7 @@ import Mathlib.Data.Rat.Cast.Order import Mathlib.Tactic.Positivity.Core import Mathlib.Tactic.HaveI import Mathlib.Algebra.GroupPower.Order -import Mathlib.Algebra.Order.Field.Basic +import Mathlib.Algebra.Order.Field.Power import Qq /-! @@ -238,25 +238,17 @@ where `a` and `b` are integers. -/ let ra ← core zα pα a; let rb ← core zα pα b guard <|← withDefault <| withNewMCtxDepth <| isDefEq f q(HDiv.hDiv (α := ℤ) (β := ℤ)) match ra, rb with - | .positive pa, .positive pb => - have pa' : Q(0 < $a) := pa - have pb' : Q(0 < $b) := pb + | .positive (pa : Q(0 < $a)), .positive (pb : Q(0 < $b)) => if pa == pb then -- Only attempts to prove `0 < a / a`, otherwise falls back to `0 ≤ a / b` - pure (.positive (q(int_div_self_pos $pa') : Expr)) + pure (.positive (q(int_div_self_pos $pa) : Expr)) else - pure (.nonnegative (q(int_div_nonneg_of_pos_of_pos $pa' $pb') : Expr)) - | .positive pa, .nonnegative pb => - have pa' : Q(0 < $a) := pa - have pb' : Q(0 ≤ $b) := pb - pure (.nonnegative (q(int_div_nonneg_of_pos_of_nonneg $pa' $pb') : Expr)) - | .nonnegative pa, .positive pb => - have pa' : Q(0 ≤ $a) := pa - have pb' : Q(0 < $b) := pb - pure (.nonnegative (q(int_div_nonneg_of_nonneg_of_pos $pa' $pb') : Expr)) - | .nonnegative pa, .nonnegative pb => - have pa' : Q(0 ≤ $a) := pa - have pb' : Q(0 ≤ $b) := pb - pure (.nonnegative (q(Int.ediv_nonneg $pa' $pb') : Expr)) + pure (.nonnegative (q(int_div_nonneg_of_pos_of_pos $pa $pb) : Expr)) + | .positive (pa : Q(0 < $a)), .nonnegative (pb : Q(0 ≤ $b)) => + pure (.nonnegative (q(int_div_nonneg_of_pos_of_nonneg $pa $pb) : Expr)) + | .nonnegative (pa : Q(0 ≤ $a)), .positive (pb : Q(0 < $b)) => + pure (.nonnegative (q(int_div_nonneg_of_nonneg_of_pos $pa $pb) : Expr)) + | .nonnegative (pa : Q(0 ≤ $a)), .nonnegative (pb : Q(0 ≤ $b)) => + pure (.nonnegative (q(Int.ediv_nonneg $pa $pb) : Expr)) | _, _ => pure .none section LinearOrderedSemifield @@ -376,13 +368,65 @@ def evalPow : PositivityExt where eval {u α} zα pα e := do | .nonzero pa => ofNonzero pa (← synthInstanceQ (_ : Q(Type u))) | .none => pure .none +/-- The `positivity` extension which identifies expressions of the form `a ^ (b : ℤ)`, +such that `positivity` successfully recognises both `a` and `b`. -/ +@[positivity (_ : α) ^ (_ : ℤ), Pow.pow _ (_ : ℤ)] +def evalZpow : PositivityExt where eval {u α} zα pα e := do + let .app (.app _ (a : Q($α))) (b : Q(ℤ)) ← withReducible (whnf e) | throwError "not ^" + let result ← catchNone do + let _a ← synthInstanceQ q(LinearOrderedField $α) + assumeInstancesCommute + match ← whnfR b with + | .app (.app (.app (.const `OfNat.ofNat _) _) (.lit (Literal.natVal n))) _ => + guard (n % 2 = 0) + have m : Q(ℕ) := mkRawNatLit (n / 2) + haveI' : $b =Q $m + $m := ⟨⟩ -- b = bit0 m + haveI' : $e =Q $a ^ $b := ⟨⟩ + pure (by exact .nonnegative q(zpow_bit0_nonneg $a $m)) + | .app (.app (.app (.const `Neg.neg _) _) _) b' => + let b' ← whnfR b' + let .true := b'.isAppOfArity ``OfNat.ofNat 3 | throwError "not a ^ -n where n is a literal" + let some n := (b'.getRevArg! 1).natLit? | throwError "not a ^ -n where n is a literal" + guard (n % 2 = 0) + have m : Q(ℕ) := mkRawNatLit (n / 2) + haveI' : $b =Q (-$m) + (-$m) := ⟨⟩ -- b = bit0 (-m) + haveI' : $e =Q $a ^ $b := ⟨⟩ + pure (by exact .nonnegative q(zpow_bit0_nonneg $a (-$m))) + | _ => throwError "not a ^ n where n is a literal or a negated literal" + orElse result do + let ra ← core zα pα a + let ofNonneg (pa : Q(0 ≤ $a)) (_oα : Q(LinearOrderedSemifield $α)) : + MetaM (Strictness zα pα e) := do + haveI' : $e =Q $a ^ $b := ⟨⟩ + assumeInstancesCommute + pure (by exact .nonnegative (q(zpow_nonneg $pa $b))) + let ofNonzero (pa : Q($a ≠ 0)) (_oα : Q(GroupWithZero $α)) : MetaM (Strictness zα pα e) := do + haveI' : $e =Q $a ^ $b := ⟨⟩ + let _a ← synthInstanceQ q(GroupWithZero $α) + assumeInstancesCommute + pure (.nonzero (by exact q(zpow_ne_zero $b $pa))) + match ra with + | .positive pa => + try + let _a ← synthInstanceQ (q(LinearOrderedSemifield $α) : Q(Type u)) + haveI' : $e =Q $a ^ $b := ⟨⟩ + assumeInstancesCommute + pure (by exact .positive (q(zpow_pos_of_pos $pa $b))) + catch e : Exception => + trace[Tactic.positivity.failure] "{e.toMessageData}" + let oα ← synthInstanceQ q(LinearOrderedSemifield $α) + orElse (← catchNone (ofNonneg q(le_of_lt $pa) oα)) (ofNonzero q(ne_of_gt $pa) oα) + | .nonnegative pa => ofNonneg pa (← synthInstanceQ (_ : Q(Type u))) + | .nonzero pa => ofNonzero pa (← synthInstanceQ (_ : Q(Type u))) + | .none => pure .none + private theorem abs_pos_of_ne_zero {α : Type*} [AddGroup α] [LinearOrder α] [CovariantClass α α (·+·) (·≤·)] {a : α} : a ≠ 0 → 0 < |a| := abs_pos.mpr /-- The `positivity` extension which identifies expressions of the form `|a|`. -/ @[positivity |(_ : α)|] -def evalAbs : PositivityExt where eval {_ _α} zα pα e := do - let (.app _ (a : Q($_α))) ← withReducible (whnf e) | throwError "not |·|" +def evalAbs : PositivityExt where eval {u} (α : Q(Type u)) zα pα (e : Q($α)) := do + let ~q(@Abs.abs _ (_) $a) := e | throwError "not |·|" try match ← core zα pα a with | .positive pa => @@ -404,8 +448,8 @@ Since the output type of `Int.natAbs` is `ℕ`, the nonnegative case is handled `positivity` tactic. -/ @[positivity Int.natAbs _] -def evalNatAbs : PositivityExt where eval {_u _α} _zα _pα e := do - let (.app _ (a : Q(Int))) ← withReducible (whnf e) | throwError "not Int.natAbs" +def evalNatAbs : PositivityExt where eval {_u _α} _zα _pα (e : Q(ℕ)) := do + let ~q(Int.natAbs $a) := e | throwError "not Int.natAbs" let zα' : Q(Zero Int) := q(inferInstance) let pα' : Q(PartialOrder Int) := q(inferInstance) let ra ← core zα' pα' a @@ -421,14 +465,13 @@ def evalNatAbs : PositivityExt where eval {_u _α} _zα _pα e := do @[positivity Nat.cast _] def evalNatCast : PositivityExt where eval {u α} _zα _pα e := do - let (.app _ (a : Q(Nat))) ← withReducible (whnf e) | throwError "not Nat.cast" + let ~q(@Nat.cast _ (_) ($a : ℕ)) := e | throwError "not Nat.cast" let zα' : Q(Zero Nat) := q(inferInstance) let pα' : Q(PartialOrder Nat) := q(inferInstance) - let ra ← core zα' pα' a - let _oα ← synthInstanceQ q(OrderedSemiring $α) + let (_oα : Q(OrderedSemiring $α)) ← synthInstanceQ q(OrderedSemiring $α) haveI' : $e =Q Nat.cast $a := ⟨⟩ assumeInstancesCommute - match ra with + match ← core zα' pα' a with | .positive pa => let _nt ← synthInstanceQ q(Nontrivial $α) pure (.positive q(Nat.cast_pos.mpr $pa)) @@ -437,7 +480,7 @@ def evalNatCast : PositivityExt where eval {u α} _zα _pα e := do @[positivity Int.cast _] def evalIntCast : PositivityExt where eval {u α} _zα _pα e := do - let (.app _ (a : Q(Int))) ← withReducible (whnf e) | throwError "not Int.cast" + let ~q(@Int.cast _ (_) ($a : ℤ)) := e | throwError "not Int.cast" let zα' : Q(Zero Int) := q(inferInstance) let pα' : Q(PartialOrder Int) := q(inferInstance) let ra ← core zα' pα' a @@ -466,8 +509,7 @@ def evalIntCast : PositivityExt where eval {u α} _zα _pα e := do /-- Extension for Rat.cast. -/ @[positivity Rat.cast _] def evalRatCast : PositivityExt where eval {u α} _zα _pα e := do - let _rα : Q(RatCast $α) ← synthInstanceQ (q(RatCast $α)) - let ~q(Rat.cast ($a : ℚ)) := e | throwError "not Rat.cast" + let ~q(@Rat.cast _ (_) ($a : ℚ)) := e | throwError "not Rat.cast" let zα' : Q(Zero ℚ) := q(inferInstance) let pα' : Q(PartialOrder ℚ) := q(inferInstance) match ← core zα' pα' a with @@ -485,8 +527,8 @@ def evalRatCast : PositivityExt where eval {u α} _zα _pα e := do /-- Extension for Nat.succ. -/ @[positivity Nat.succ _] -def evalNatSucc : PositivityExt where eval {_u _α} _zα _pα e := do - let (.app _ (a : Q(Nat))) ← withReducible (whnf e) | throwError "not Nat.succ" +def evalNatSucc : PositivityExt where eval {_u _α} _zα _pα (e : Q(ℕ)) := do + let ~q(Nat.succ $a) := e | throwError "not Nat.succ" pure (.positive (q(Nat.succ_pos $a) : Expr)) /-- Extension for Nat.factorial. -/ diff --git a/Mathlib/Testing/SlimCheck/Sampleable.lean b/Mathlib/Testing/SlimCheck/Sampleable.lean index b580a2d1f7525..97b2fb3892a85 100644 --- a/Mathlib/Testing/SlimCheck/Sampleable.lean +++ b/Mathlib/Testing/SlimCheck/Sampleable.lean @@ -288,18 +288,20 @@ def printSamples {t : Type} [Repr t] (g : Gen t) : IO PUnit := do open Lean Meta Qq /-- Create a `Gen α` expression from the argument of `#sample` -/ -def mkGenerator (e : Expr) : MetaM (Expr × Expr) := do - let t ← inferType e - match t.getAppFnArgs with - | (`Gen, #[t]) => do - let repr_inst ← synthInstance (← mkAppM ``Repr #[t]) - pure (repr_inst, e) - | _ => do - let sampleableExt_inst ← synthInstance (mkApp (← mkConstWithFreshMVarLevels ``SampleableExt) e) - let repr_inst ← mkAppOptM ``SampleableExt.proxyRepr #[e, sampleableExt_inst] - let gen ← mkAppOptM ``SampleableExt.sample #[none, sampleableExt_inst] - pure (repr_inst, gen) - +def mkGenerator (e : Expr) : MetaM (Σ (u : Level) (α : Q(Type $u)), Q(Repr $α) × Q(Gen $α)) := do + match ← inferTypeQ e with + | ⟨.succ u, ~q(Gen $α), gen⟩ => + let repr_inst ← synthInstanceQ (q(Repr $α) : Q(Type $u)) + pure ⟨u, α, repr_inst, gen⟩ + | ⟨.succ u, ~q(Sort u), α⟩ => do + let v ← mkFreshLevelMVar + let _sampleableExt_inst ← synthInstanceQ (q(SampleableExt.{u,v} $α) : Q(Sort (max u (v + 2)))) + let v ← instantiateLevelMVars v + let repr_inst := q(SampleableExt.proxyRepr (α := $α)) + let gen := q(SampleableExt.sample (α := $α)) + pure ⟨v, q(SampleableExt.proxy $α), repr_inst, gen⟩ + | ⟨_u, t, _e⟩ => + throwError "Must be a Sort u` or a `Gen α`, got value of type{indentExpr t}" open Elab /-- @@ -339,8 +341,10 @@ values of type `type` using an increasing size parameter. elab "#sample " e:term : command => Command.runTermElabM fun _ => do let e ← Elab.Term.elabTermAndSynthesize e none - let (repr_inst, gen) ← mkGenerator e - let printSamples ← mkAppOptM ``printSamples #[none, repr_inst, gen] + let g ← mkGenerator e + -- `printSamples` only works in `Type 0` (see the porting note) + let ⟨0, α, _, gen⟩ := g | throwError "Cannot sample from {g.1} due to its universe" + let printSamples := q(printSamples (t := $α) $gen) let code ← unsafe evalExpr (IO PUnit) q(IO PUnit) printSamples _ ← code diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean index 53dcb7c6685f3..cee7485e5e9d8 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean @@ -267,7 +267,7 @@ theorem Finite.of_summable_const [LinearOrderedAddCommGroup α] [TopologicalSpac simpa using sum_le_hasSum s (fun a _ => hb.le) hf.hasSum obtain ⟨n, hn⟩ := Archimedean.arch (∑' _ : ι, b) hb have : ∀ s : Finset ι, s.card ≤ n := fun s => by - simpa [nsmul_le_nsmul_iff hb] using (H s).trans hn + simpa [nsmul_le_nsmul_iff_left hb] using (H s).trans hn have : Fintype ι := fintypeOfFinsetCardLe n this infer_instance diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 72e581dbfc63a..6d5ac30602e06 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -714,6 +714,9 @@ theorem one_apply (x : M₁) : (1 : M₁ →L[R₁] M₁) x = x := rfl #align continuous_linear_map.one_apply ContinuousLinearMap.one_apply +instance [Nontrivial M₁] : Nontrivial (M₁ →L[R₁] M₁) := + ⟨0, 1, fun e ↦ have ⟨x, hx⟩ := exists_ne (0 : M₁); hx (by simpa using FunLike.congr_fun e.symm x)⟩ + section Add variable [ContinuousAdd M₂] diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index afb64a2a6e7f9..fda352aed6b75 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -133,7 +133,7 @@ theorem LinearMap.continuous_of_isClosed_ker (l : E →ₗ[𝕜] 𝕜) Continuous l := by -- `l` is either constant or surjective. If it is constant, the result is trivial. by_cases H : finrank 𝕜 (LinearMap.range l) = 0 - · rw [finrank_eq_zero, LinearMap.range_eq_bot] at H + · rw [Submodule.finrank_eq_zero, LinearMap.range_eq_bot] at H rw [H] exact continuous_zero · -- In the case where `l` is surjective, we factor it as `φ : (E ⧸ l.ker) ≃ₗ[𝕜] 𝕜`. Note that @@ -225,7 +225,7 @@ private theorem continuous_equivFun_basis_aux [ht2 : T2Space E] {ι : Type v} [F have H₂ : ∀ f : E →ₗ[𝕜] 𝕜, Continuous f := by intro f by_cases H : finrank 𝕜 (LinearMap.range f) = 0 - · rw [finrank_eq_zero, LinearMap.range_eq_bot] at H + · rw [Submodule.finrank_eq_zero, LinearMap.range_eq_bot] at H rw [H] exact continuous_zero · have : finrank 𝕜 (LinearMap.ker f) = n := by diff --git a/Mathlib/Topology/Algebra/Nonarchimedean/AdicTopology.lean b/Mathlib/Topology/Algebra/Nonarchimedean/AdicTopology.lean index df8353733e861..af02d73ea693a 100644 --- a/Mathlib/Topology/Algebra/Nonarchimedean/AdicTopology.lean +++ b/Mathlib/Topology/Algebra/Nonarchimedean/AdicTopology.lean @@ -57,7 +57,7 @@ theorem adic_basis (I : Ideal R) : SubmodulesRingBasis fun n : ℕ => (I ^ n • suffices ∀ i j : ℕ, ∃ k, I ^ k ≤ I ^ i ∧ I ^ k ≤ I ^ j by simpa only [smul_eq_mul, mul_top, Algebra.id.map_eq_id, map_id, le_inf_iff] using this intro i j - exact ⟨max i j, pow_le_pow (le_max_left i j), pow_le_pow (le_max_right i j)⟩ + exact ⟨max i j, pow_le_pow_right (le_max_left i j), pow_le_pow_right (le_max_right i j)⟩ leftMul := by suffices ∀ (a : R) (i : ℕ), ∃ j : ℕ, a • I ^ j ≤ I ^ i by simpa only [smul_top_eq_map, Algebra.id.map_eq_id, map_id] using this @@ -119,8 +119,8 @@ theorem adic_module_basis : { inter := fun i j => ⟨max i j, le_inf_iff.mpr - ⟨smul_mono_left <| pow_le_pow (le_max_left i j), - smul_mono_left <| pow_le_pow (le_max_right i j)⟩⟩ + ⟨smul_mono_left <| pow_le_pow_right (le_max_left i j), + smul_mono_left <| pow_le_pow_right (le_max_right i j)⟩⟩ smul := fun m i => ⟨(I ^ i • ⊤ : Ideal R), ⟨i, by simp⟩, fun a a_in => by replace a_in : a ∈ I ^ i := by simpa [(I ^ i).mul_top] using a_in @@ -201,7 +201,7 @@ theorem is_ideal_adic_pow {J : Ideal R} (h : IsAdic J) {n : ℕ} (hn : 0 < n) : · exfalso exact Nat.not_succ_le_zero 0 hn rw [← pow_mul, Nat.succ_mul] - apply Ideal.pow_le_pow + apply Ideal.pow_le_pow_right apply Nat.le_add_left #align is_ideal_adic_pow is_ideal_adic_pow diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index b81a2a4d7c5bb..5c2b5736e149c 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -709,6 +709,11 @@ theorem closure_diff_interior (s : Set α) : closure s \ interior s = frontier s rfl #align closure_diff_interior closure_diff_interior +/-- Interior and frontier are disjoint. -/ +lemma disjoint_interior_frontier : Disjoint (interior s) (frontier s) := by + rw [disjoint_iff_inter_eq_empty, ← closure_diff_interior, diff_eq, + ← inter_assoc, inter_comm, ← inter_assoc, compl_inter_self, empty_inter] + @[simp] theorem closure_diff_frontier (s : Set α) : closure s \ frontier s = interior s := by rw [frontier, diff_diff_right_self, inter_eq_self_of_subset_right interior_subset_closure] diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 5ce81050e5fba..e681390b4ad7e 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -1650,6 +1650,16 @@ end Sigma section ULift +theorem ULift.isOpen_iff [TopologicalSpace α] {s : Set (ULift.{v} α)} : + IsOpen s ↔ IsOpen (ULift.up ⁻¹' s) := by + unfold ULift.topologicalSpace + erw [← Equiv.ulift.coinduced_symm] + rfl + +theorem ULift.isClosed_iff [TopologicalSpace α] {s : Set (ULift.{v} α)} : + IsClosed s ↔ IsClosed (ULift.up ⁻¹' s) := by + rw [← isOpen_compl_iff, ← isOpen_compl_iff, isOpen_iff, preimage_compl] + @[continuity] theorem continuous_uLift_down [TopologicalSpace α] : Continuous (ULift.down : ULift.{v, u} α → α) := continuous_induced_dom diff --git a/Mathlib/Topology/EMetricSpace/Paracompact.lean b/Mathlib/Topology/EMetricSpace/Paracompact.lean index ba1e0e2ac41ea..a5228884ea43b 100644 --- a/Mathlib/Topology/EMetricSpace/Paracompact.lean +++ b/Mathlib/Topology/EMetricSpace/Paracompact.lean @@ -43,7 +43,7 @@ instance (priority := 100) instParacompactSpace [PseudoEMetricSpace α] : Paraco have pow_pos : ∀ k : ℕ, (0 : ℝ≥0∞) < 2⁻¹ ^ k := fun k => ENNReal.pow_pos (ENNReal.inv_pos.2 ENNReal.two_ne_top) _ have hpow_le : ∀ {m n : ℕ}, m ≤ n → (2⁻¹ : ℝ≥0∞) ^ n ≤ 2⁻¹ ^ m := @fun m n h => - pow_le_pow_of_le_one' (ENNReal.inv_le_one.2 ENNReal.one_lt_two.le) h + pow_le_pow_right_of_le_one' (ENNReal.inv_le_one.2 ENNReal.one_lt_two.le) h have h2pow : ∀ n : ℕ, 2 * (2⁻¹ : ℝ≥0∞) ^ (n + 1) = 2⁻¹ ^ n := fun n => by simp [pow_succ, ← mul_assoc, ENNReal.mul_inv_cancel two_ne_zero two_ne_top] -- Consider an open covering `S : Set (Set α)` diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index 54a38c1c0cbfa..7d7c46b4aa2ef 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -69,6 +69,10 @@ instance : CoeFun (X ≃ₜ Y) fun _ ↦ X → Y := ⟨FunLike.coe⟩ rfl #align homeomorph.homeomorph_mk_coe Homeomorph.homeomorph_mk_coe +/-- The unique homeomorphism between two empty types. -/ +protected def empty [IsEmpty X] [IsEmpty Y] : X ≃ₜ Y where + __ := Equiv.equivOfIsEmpty X Y + /-- Inverse of a homeomorphism. -/ protected def symm (h : X ≃ₜ Y) : Y ≃ₜ X where continuous_toFun := h.continuous_invFun diff --git a/Mathlib/Topology/Instances/NNReal.lean b/Mathlib/Topology/Instances/NNReal.lean index b8cbdac0d0802..8bd81f84d22a7 100644 --- a/Mathlib/Topology/Instances/NNReal.lean +++ b/Mathlib/Topology/Instances/NNReal.lean @@ -256,7 +256,7 @@ nonrec theorem tendsto_tsum_compl_atTop_zero {α : Type*} (f : α → ℝ≥0) : /-- `x ↦ x ^ n` as an order isomorphism of `ℝ≥0`. -/ def powOrderIso (n : ℕ) (hn : n ≠ 0) : ℝ≥0 ≃o ℝ≥0 := StrictMono.orderIsoOfSurjective (fun x ↦ x ^ n) (fun x y h => - strictMonoOn_pow hn.bot_lt (zero_le x) (zero_le y) h) <| + pow_left_strictMonoOn hn (zero_le x) (zero_le y) h) <| (continuous_id.pow _).surjective (tendsto_pow_atTop hn) <| by simpa [OrderBot.atBot_eq, pos_iff_ne_zero] #align nnreal.pow_order_iso NNReal.powOrderIso diff --git a/Mathlib/Topology/IsLocalHomeomorph.lean b/Mathlib/Topology/IsLocalHomeomorph.lean index 4348b9b62e049..b6445d9ad5987 100644 --- a/Mathlib/Topology/IsLocalHomeomorph.lean +++ b/Mathlib/Topology/IsLocalHomeomorph.lean @@ -15,12 +15,19 @@ This file defines local homeomorphisms. ## Main definitions -* `IsLocalHomeomorph`: A function `f : X → Y` satisfies `IsLocalHomeomorph` if for each - point `x : X`, the restriction of `f` to some open neighborhood `U` of `x` gives a homeomorphism +For a function `f : X → Y ` between topological spaces, we say +* `IsLocalHomeomorphOn f s` if `f` is a local homeomorphism around each point of `s`: for each + `x : X`, the restriction of `f` to some open neighborhood `U` of `x` gives a homeomorphism between `U` and an open subset of `Y`. +* `IsLocalHomeomorph f`: `f` is a local homeomorphism, i.e. it's a local homeomorphism on `univ`. + +Note that `IsLocalHomeomorph` is a global condition. This is in contrast to +`PartialHomeomorph`, which is a homeomorphism between specific open subsets. + +## Main results +* local homeomorphisms are locally injective open maps +* more! - Note that `IsLocalHomeomorph` is a global condition. This is in contrast to - `PartialHomeomorph`, which is a homeomorphism between specific open subsets. -/ @@ -56,24 +63,24 @@ namespace IsLocalHomeomorphOn /-- Proves that `f` satisfies `IsLocalHomeomorphOn f s`. The condition `h` is weaker than the definition of `IsLocalHomeomorphOn f s`, since it only requires `e : PartialHomeomorph X Y` to agree with `f` on its source `e.source`, as opposed to on the whole space `X`. -/ -theorem mk (h : ∀ x ∈ s, ∃ e : PartialHomeomorph X Y, x ∈ e.source ∧ ∀ y ∈ e.source, f y = e y) : +theorem mk (h : ∀ x ∈ s, ∃ e : PartialHomeomorph X Y, x ∈ e.source ∧ Set.EqOn f e e.source) : IsLocalHomeomorphOn f s := by intro x hx obtain ⟨e, hx, he⟩ := h x hx exact ⟨{ e with toFun := f - map_source' := fun x hx => by rw [he x hx]; exact e.map_source' hx - left_inv' := fun x hx => by rw [he x hx]; exact e.left_inv' hx - right_inv' := fun y hy => by rw [he _ (e.map_target' hy)]; exact e.right_inv' hy + map_source' := fun _x hx ↦ by rw [he hx]; exact e.map_source' hx + left_inv' := fun _x hx ↦ by rw [he hx]; exact e.left_inv' hx + right_inv' := fun _y hy ↦ by rw [he (e.map_target' hy)]; exact e.right_inv' hy continuousOn_toFun := (continuousOn_congr he).mpr e.continuousOn_toFun }, hx, rfl⟩ #align is_locally_homeomorph_on.mk IsLocalHomeomorphOn.mk variable {g f s t} -theorem mono {t : Set X} (hf : IsLocalHomeomorphOn f t) (hst : s ⊆ t) : - IsLocalHomeomorphOn f s := fun x hx ↦ hf x (hst hx) +theorem mono {t : Set X} (hf : IsLocalHomeomorphOn f t) (hst : s ⊆ t) : IsLocalHomeomorphOn f s := + fun x hx ↦ hf x (hst hx) theorem of_comp_left (hgf : IsLocalHomeomorphOn (g ∘ f) s) (hg : IsLocalHomeomorphOn g (f '' s)) (cont : ∀ x ∈ s, ContinuousAt f x) : IsLocalHomeomorphOn f s := mk f s fun x hx ↦ by @@ -84,8 +91,7 @@ theorem of_comp_left (hgf : IsLocalHomeomorphOn (g ∘ f) s) (hg : IsLocalHomeom fun y hy ↦ ?_⟩ change f y = g.symm (gf y) have : f y ∈ g.source := by apply interior_subset hy.1.2 - rw [← he, g.eq_symm_apply this (by apply g.map_source this)] - rfl + rw [← he, g.eq_symm_apply this (by apply g.map_source this), Function.comp_apply] theorem of_comp_right (hgf : IsLocalHomeomorphOn (g ∘ f) s) (hf : IsLocalHomeomorphOn f s) : IsLocalHomeomorphOn g (f '' s) := mk g _ <| by @@ -108,7 +114,7 @@ protected theorem continuousAt (hf : IsLocalHomeomorphOn f s) {x : X} (hx : x #align is_locally_homeomorph_on.continuous_at IsLocalHomeomorphOn.continuousAt protected theorem continuousOn (hf : IsLocalHomeomorphOn f s) : ContinuousOn f s := - ContinuousAt.continuousOn fun _x => hf.continuousAt + ContinuousAt.continuousOn fun _x ↦ hf.continuousAt #align is_locally_homeomorph_on.continuous_on IsLocalHomeomorphOn.continuousOn protected theorem comp (hg : IsLocalHomeomorphOn g t) (hf : IsLocalHomeomorphOn f s) @@ -157,10 +163,10 @@ namespace IsLocalHomeomorph /-- Proves that `f` satisfies `IsLocalHomeomorph f`. The condition `h` is weaker than the definition of `IsLocalHomeomorph f`, since it only requires `e : PartialHomeomorph X Y` to agree with `f` on its source `e.source`, as opposed to on the whole space `X`. -/ -theorem mk (h : ∀ x : X, ∃ e : PartialHomeomorph X Y, x ∈ e.source ∧ ∀ y ∈ e.source, f y = e y) : +theorem mk (h : ∀ x : X, ∃ e : PartialHomeomorph X Y, x ∈ e.source ∧ Set.EqOn f e e.source) : IsLocalHomeomorph f := isLocalHomeomorph_iff_isLocalHomeomorphOn_univ.mpr - (IsLocalHomeomorphOn.mk f Set.univ fun x _hx => h x) + (IsLocalHomeomorphOn.mk f Set.univ fun x _hx ↦ h x) #align is_locally_homeomorph.mk IsLocalHomeomorph.mk variable {g f} @@ -177,20 +183,24 @@ theorem map_nhds_eq (hf : IsLocalHomeomorph f) (x : X) : (𝓝 x).map f = 𝓝 ( hf.isLocalHomeomorphOn.map_nhds_eq (Set.mem_univ x) #align is_locally_homeomorph.map_nhds_eq IsLocalHomeomorph.map_nhds_eq +/-- A local homeomorphism is continuous. -/ protected theorem continuous (hf : IsLocalHomeomorph f) : Continuous f := continuous_iff_continuousOn_univ.mpr hf.isLocalHomeomorphOn.continuousOn #align is_locally_homeomorph.continuous IsLocalHomeomorph.continuous +/-- A local homeomorphism is an open map. -/ protected theorem isOpenMap (hf : IsLocalHomeomorph f) : IsOpenMap f := - IsOpenMap.of_nhds_le fun x => ge_of_eq (hf.map_nhds_eq x) + IsOpenMap.of_nhds_le fun x ↦ ge_of_eq (hf.map_nhds_eq x) #align is_locally_homeomorph.is_open_map IsLocalHomeomorph.isOpenMap +/-- The composition of local homeomorphisms is a local homeomorphism. -/ protected theorem comp (hg : IsLocalHomeomorph g) (hf : IsLocalHomeomorph f) : IsLocalHomeomorph (g ∘ f) := isLocalHomeomorph_iff_isLocalHomeomorphOn_univ.mpr (hg.isLocalHomeomorphOn.comp hf.isLocalHomeomorphOn (Set.univ.mapsTo_univ f)) #align is_locally_homeomorph.comp IsLocalHomeomorph.comp +/-- An injective local homeomorphism is an open embedding. -/ theorem openEmbedding_of_injective (hf : IsLocalHomeomorph f) (hi : f.Injective) : OpenEmbedding f := openEmbedding_of_continuous_injective_open hf.continuous hi hf.isOpenMap diff --git a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean index e932af2bae0aa..52960eda285c9 100644 --- a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean +++ b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean @@ -224,7 +224,7 @@ theorem _root_.IsOpen.exists_iUnion_isClosed {U : Set α} (hU : IsOpen U) : show Monotone F · intro m n hmn x hx simp only [mem_Ici, mem_preimage] at hx ⊢ - apply le_trans (pow_le_pow_of_le_one' a_lt_one.le hmn) hx + apply le_trans (pow_le_pow_right_of_le_one' a_lt_one.le hmn) hx #align is_open.exists_Union_is_closed IsOpen.exists_iUnion_isClosed theorem _root_.IsCompact.exists_infEdist_eq_edist (hs : IsCompact s) (hne : s.Nonempty) (x : α) : diff --git a/Mathlib/Topology/MetricSpace/PiNat.lean b/Mathlib/Topology/MetricSpace/PiNat.lean index 88af496faa80c..40c96261c1ec9 100644 --- a/Mathlib/Topology/MetricSpace/PiNat.lean +++ b/Mathlib/Topology/MetricSpace/PiNat.lean @@ -56,7 +56,7 @@ open Classical Topology Filter open TopologicalSpace Set Metric Filter Function -attribute [local simp] pow_le_pow_iff one_lt_two inv_le_inv zero_le_two zero_lt_two +attribute [local simp] pow_le_pow_iff_right one_lt_two inv_le_inv zero_le_two zero_lt_two variable {E : ℕ → Type*} @@ -295,7 +295,7 @@ theorem dist_triangle_nonarch (x y z : ∀ n, E n) : dist x z ≤ max (dist x y) rcases eq_or_ne y z with (rfl | hyz) · simp simp only [dist_eq_of_ne, hxz, hxy, hyz, inv_le_inv, one_div, inv_pow, zero_lt_two, Ne.def, - not_false_iff, le_max_iff, pow_le_pow_iff, one_lt_two, pow_pos, + not_false_iff, le_max_iff, pow_le_pow_iff_right, one_lt_two, pow_pos, min_le_iff.1 (min_firstDiff_le x y z hxz)] #align pi_nat.dist_triangle_nonarch PiNat.dist_triangle_nonarch @@ -328,7 +328,7 @@ theorem apply_eq_of_dist_lt {x y : ∀ n, E n} {n : ℕ} (h : dist x y < (1 / 2) rcases eq_or_ne x y with (rfl | hne) · rfl have : n < firstDiff x y := by - simpa [dist_eq_of_ne hne, inv_lt_inv, pow_lt_pow_iff, one_lt_two] using h + simpa [dist_eq_of_ne hne, inv_lt_inv, pow_lt_pow_iff_right, one_lt_two] using h exact apply_eq_of_lt_firstDiff (hi.trans_lt this) #align pi_nat.apply_eq_of_dist_lt PiNat.apply_eq_of_dist_lt diff --git a/Mathlib/Topology/MetricSpace/PseudoMetric.lean b/Mathlib/Topology/MetricSpace/PseudoMetric.lean index ad480a35f4372..745d0b44e1553 100644 --- a/Mathlib/Topology/MetricSpace/PseudoMetric.lean +++ b/Mathlib/Topology/MetricSpace/PseudoMetric.lean @@ -1564,7 +1564,8 @@ theorem NNReal.dist_eq (a b : ℝ≥0) : dist a b = |(a : ℝ) - b| := rfl theorem NNReal.nndist_eq (a b : ℝ≥0) : nndist a b = max (a - b) (b - a) := eq_of_forall_ge_iff fun _ => by - simp only [← NNReal.coe_le_coe, coe_nndist, dist_eq, max_le_iff, abs_sub_le_iff, + simp only [max_le_iff, tsub_le_iff_right (α := ℝ≥0)] + simp only [← NNReal.coe_le_coe, coe_nndist, dist_eq, abs_sub_le_iff, tsub_le_iff_right, NNReal.coe_add] #align nnreal.nndist_eq NNReal.nndist_eq diff --git a/Mathlib/Topology/Metrizable/Uniformity.lean b/Mathlib/Topology/Metrizable/Uniformity.lean index 787d9b2e19fc5..a92d51d8e47da 100644 --- a/Mathlib/Topology/Metrizable/Uniformity.lean +++ b/Mathlib/Topology/Metrizable/Uniformity.lean @@ -226,7 +226,7 @@ protected theorem UniformSpace.metrizable_uniformity (X : Type*) [UniformSpace X intro x y n dsimp only [] split_ifs with h - · rw [(strictAnti_pow hr.1 hr.2).le_iff_le, Nat.find_le_iff] + · rw [(pow_right_strictAnti hr.1 hr.2).le_iff_le, Nat.find_le_iff] exact ⟨fun ⟨m, hmn, hm⟩ hn => hm (hB.antitone hmn hn), fun h => ⟨n, le_rfl, h⟩⟩ · push_neg at h simp only [h, not_true, (pow_pos hr.1 _).not_le] diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index bccc397e77bd7..d828fd9ce59af 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -362,6 +362,9 @@ instance Pi.instT0Space {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace ⟨fun _ _ h => funext fun i => (h.map (continuous_apply i)).eq⟩ #align pi.t0_space Pi.instT0Space +instance ULift.instT0Space [T0Space X] : T0Space (ULift X) := + embedding_uLift_down.t0Space + theorem T0Space.of_cover (h : ∀ x y, Inseparable x y → ∃ s : Set X, x ∈ s ∧ y ∈ s ∧ T0Space s) : T0Space X := by refine' ⟨fun x y hxy => _⟩ @@ -627,6 +630,9 @@ instance {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, T1Space (∀ i, X i) := ⟨fun f => univ_pi_singleton f ▸ isClosed_set_pi fun _ _ => isClosed_singleton⟩ +instance ULift.instT1Space [T1Space X] : T1Space (ULift X) := + embedding_uLift_down.t1Space + -- see Note [lower instance priority] instance (priority := 100) T1Space.t0Space [T1Space X] : T0Space X := ⟨fun _ _ h => h.specializes.eq⟩ @@ -1234,6 +1240,9 @@ theorem Embedding.t2Space [TopologicalSpace Y] [T2Space Y] {f : X → Y} (hf : E .of_injective_continuous hf.inj hf.continuous #align embedding.t2_space Embedding.t2Space +instance ULift.instT2Space [T2Space X] : T2Space (ULift X) := + embedding_uLift_down.t2Space + instance [T2Space X] [TopologicalSpace Y] [T2Space Y] : T2Space (X ⊕ Y) := by constructor @@ -1837,6 +1846,9 @@ instance Subtype.t3Space [T3Space X] {p : X → Prop} : T3Space (Subtype p) := embedding_subtype_val.t3Space #align subtype.t3_space Subtype.t3Space +instance ULift.instT3Space [T3Space X] : T3Space (ULift X) := + embedding_uLift_down.t3Space + instance [TopologicalSpace Y] [T3Space X] [T3Space Y] : T3Space (X × Y) := ⟨⟩ instance {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, T3Space (X i)] : @@ -1974,6 +1986,9 @@ protected theorem ClosedEmbedding.t4Space [TopologicalSpace Y] [T4Space Y] {f : toNormalSpace := hf.normalSpace #align closed_embedding.normal_space ClosedEmbedding.t4Space +instance ULift.instT4Space [T4Space X] : T4Space (ULift X) := + ULift.closedEmbedding_down.t4Space + namespace SeparationQuotient /-- The `SeparationQuotient` of a normal space is a normal space. -/ @@ -2020,6 +2035,9 @@ theorem Embedding.t5Space [TopologicalSpace Y] [T5Space Y] {e : X → Y} (he : E instance [T5Space X] {p : X → Prop} : T5Space { x // p x } := embedding_subtype_val.t5Space +instance ULift.instT5Space [T5Space X] : T5Space (ULift X) := + embedding_uLift_down.t5Space + -- see Note [lower instance priority] /-- A `T₅` space is a `T₄` space. -/ instance (priority := 100) T5Space.toT4Space [T5Space X] : T4Space X where diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index 81b3234fa8e6a..238c5bc0e8d80 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -102,6 +102,13 @@ theorem tsupport_smul_subset_left {M α} [TopologicalSpace X] [Zero M] [Zero α] closure_mono <| support_smul_subset_left f g #align tsupport_smul_subset_left tsupport_smul_subset_left +@[to_additive] +theorem mulTSupport_mul [TopologicalSpace X] [Monoid α] {f g : X → α} : + (mulTSupport fun x ↦ f x * g x) ⊆ mulTSupport f ∪ mulTSupport g := + closure_minimal + ((mulSupport_mul f g).trans (union_subset_union (subset_mulTSupport _) (subset_mulTSupport _))) + (isClosed_closure.union isClosed_closure) + section variable [TopologicalSpace α] [TopologicalSpace α'] diff --git a/Mathlib/Topology/UnitInterval.lean b/Mathlib/Topology/UnitInterval.lean index 531578c12b74d..44797fad16522 100644 --- a/Mathlib/Topology/UnitInterval.lean +++ b/Mathlib/Topology/UnitInterval.lean @@ -226,10 +226,10 @@ lemma addNSMul_eq_right [Archimedean α] (hδ : 0 < δ) : obtain ⟨m, hm⟩ := Archimedean.arch (b - a) hδ refine ⟨m, fun n hn ↦ ?_⟩ rw [addNSMul, coe_projIcc, add_comm, min_eq_left_iff.mpr, max_eq_right h] - exact sub_le_iff_le_add.mp (hm.trans <| nsmul_le_nsmul hδ.le hn) + exact sub_le_iff_le_add.mp (hm.trans <| nsmul_le_nsmul_left hδ.le hn) lemma monotone_addNSMul (hδ : 0 ≤ δ) : Monotone (addNSMul h δ) := - fun _ _ hnm ↦ monotone_projIcc h <| (add_le_add_iff_left _).mpr (nsmul_le_nsmul hδ hnm) + fun _ _ hnm ↦ monotone_projIcc h <| (add_le_add_iff_left _).mpr (nsmul_le_nsmul_left hδ hnm) lemma abs_sub_addNSMul_le (hδ : 0 ≤ δ) {t : Icc a b} (n : ℕ) (ht : t ∈ Icc (addNSMul h δ n) (addNSMul h δ (n+1))) : diff --git a/Mathlib/Util/AssertExists.lean b/Mathlib/Util/AssertExists.lean index 8c1e855c80e9e..4c901b2d562da 100644 --- a/Mathlib/Util/AssertExists.lean +++ b/Mathlib/Util/AssertExists.lean @@ -49,4 +49,7 @@ You should *not* delete the `assert_not_exists` statement without careful discus -/ elab "assert_not_exists " n:ident : command => do let [] ← try resolveGlobalConstWithInfos n catch _ => return - | throw <| Exception.error n m!"Declaration {n} is not allowed to exist in this file" + | throw <| .error n <| m!"Declaration {n} is not allowed to be imported by this file.\n\n{ + ""}These invariants are maintained by `assert_not_exists` statements, { + ""}and exist in order to ensure that \"complicated\" parts of the library { + ""}are not accidentally introduced as dependencies of \"simple\" parts of the library." diff --git a/lake-manifest.json b/lake-manifest.json index cd58ceb1d04d3..a90df0d688464 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -4,7 +4,7 @@ [{"url": "https://github.com/leanprover/std4", "type": "git", "subDir": null, - "rev": "6b4cf96c89e53cfcd73350bbcd90333a051ff4f0", + "rev": "9dd24a3493cceefa2bede383f21e4ef548990b68", "name": "std", "manifestFile": "lake-manifest.json", "inputRev": "main", diff --git a/test/ComputeDegree.lean b/test/ComputeDegree.lean index 3cf30c9a91e55..3f4d209c9650f 100644 --- a/test/ComputeDegree.lean +++ b/test/ComputeDegree.lean @@ -154,6 +154,9 @@ example (h : n ≤ 5) : natDegree (monomial n (5 + n : R)) ≤ 5 := by compute_d -- Expr.fvar example {f : R[X]} : natDegree f ≤ natDegree f := by compute_degree +example {R} [Semiring R] [Nontrivial R] : Monic (1 * X ^ 5 + X ^ 6 * monomial 10 1 : R[X]) := by + monicity! + end native_mathlib4_tests section tests_from_mathlib3 @@ -219,4 +222,13 @@ example {F} [Ring F] {a : F} : natDegree (X ^ 3 + C a * X ^ 10 : F[X]) ≤ 10 := example [Semiring R] : natDegree (7 * X : R[X]) ≤ 1 := by compute_degree +example [Semiring R] {a : R} : natDegree (a • X ^ 5 : R[X]) ≤ 5 := by + compute_degree + +example [Semiring R] {a : R} (a0 : a ≠ 0) : natDegree (a • X ^ 5 + X : R[X]) = 5 := by + compute_degree! + +example [Semiring R] {a : R} (a0 : a ≠ 0) : degree (a • X ^ 5 + X ^ 2 : R[X]) = 5 := by + compute_degree! + end tests_from_mathlib3 diff --git a/test/help_cmd.lean b/test/help_cmd.lean new file mode 100644 index 0000000000000..710dd1ac512b5 --- /dev/null +++ b/test/help_cmd.lean @@ -0,0 +1,6 @@ +import Mathlib.Tactic.HelpCmd +import Std.Tactic.GuardMsgs + +#guard_msgs(error, drop info) in +#help tactic +def foo := 1 diff --git a/test/linarith.lean b/test/linarith.lean index cb89b23fcd5f0..4a67985cb6734 100644 --- a/test/linarith.lean +++ b/test/linarith.lean @@ -192,7 +192,7 @@ by linarith (config := {exfalso := false}) example (x y : Rat) (h : 6 + ((x + 4) * x + (6 + 3 * y) * y) = 3 ∧ (x + 4) * x ≥ 0 ∧ (6 + 3 * y) * y ≥ 0) : False := by fail_if_success - linarith (config := {split_hypotheses := false}) + linarith (config := {splitHypotheses := false}) linarith example (h : 1 < 0) (g : ¬ 37 < 42) (k : True) (l : (-7 : ℤ) < 5) : 3 < 7 := by diff --git a/test/mod_cases.lean b/test/mod_cases.lean index f7393fc8e86db..d7d8292197519 100644 --- a/test/mod_cases.lean +++ b/test/mod_cases.lean @@ -7,6 +7,12 @@ example (n : ℤ) : 3 ∣ n ^ 3 - n := by · guard_hyp H :ₛ n ≡ 1 [ZMOD 3]; guard_target = 3 ∣ n ^ 3 - n; exact test_sorry · guard_hyp H :ₛ n ≡ 2 [ZMOD 3]; guard_target = 3 ∣ n ^ 3 - n; exact test_sorry +example (n : ℕ) : 3 ∣ n ^ 3 + n := by + mod_cases n % 3 + · guard_hyp H :~ n ≡ 0 [MOD 3]; guard_target = 3 ∣ n ^ 3 + n; exact test_sorry + · guard_hyp H :~ n ≡ 1 [MOD 3]; guard_target = 3 ∣ n ^ 3 + n; exact test_sorry + · guard_hyp H :~ n ≡ 2 [MOD 3]; guard_target = 3 ∣ n ^ 3 + n; exact test_sorry + -- test case for https://github.com/leanprover-community/mathlib4/issues/1851 example (n : ℕ) (z : ℤ) : n = n := by induction n with diff --git a/test/positivity.lean b/test/positivity.lean index 028cdba4cb6eb..bd5451621f0dd 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -179,9 +179,10 @@ example [OrderedSemiring α] {a : α} {n : ℕ} (ha : 0 ≤ a) : 0 ≤ a ^ n := example [StrictOrderedSemiring α] {a : α} {n : ℕ} (ha : 0 < a) : 0 < a ^ n := by positivity example [LinearOrderedSemifield α] (a : α) : 0 < a ^ (0 : ℤ) := by positivity --- example [LinearOrderedField α] (a : α) (n : ℤ) : 0 ≤ a ^ bit0 n := by positivity --- example [LinearOrderedSemifield α] {a : α} {n : ℤ} (ha : 0 ≤ a) : 0 ≤ a ^ n := by positivity --- example [LinearOrderedSemifield α] {a : α} {n : ℤ} (ha : 0 < a) : 0 < a ^ n := by positivity +example [LinearOrderedField α] (a : α) : 0 ≤ a ^ (18 : ℤ) := by positivity +example [LinearOrderedField α] (a : α) : 0 ≤ a ^ (-34 : ℤ) := by positivity +example [LinearOrderedSemifield α] {a : α} {n : ℤ} (ha : 0 ≤ a) : 0 ≤ a ^ n := by positivity +example [LinearOrderedSemifield α] {a : α} {n : ℤ} (ha : 0 < a) : 0 < a ^ n := by positivity -- example {a b : Cardinal.{u}} (ha : 0 < a) : 0 < a ^ b := by positivity -- example {a b : Ordinal.{u}} (ha : 0 < a) : 0 < a ^ b := by positivity