diff --git a/Archive/Wiedijk100Theorems/BallotProblem.lean b/Archive/Wiedijk100Theorems/BallotProblem.lean index 5d22c4ad829b2..8fbd5b2ecf7d0 100644 --- a/Archive/Wiedijk100Theorems/BallotProblem.lean +++ b/Archive/Wiedijk100Theorems/BallotProblem.lean @@ -235,7 +235,7 @@ theorem first_vote_neg (p q : ℕ) (h : 0 < p + q) : have := condCount_compl {l : List ℤ | l.headI = 1}ᶜ (countedSequence_finite p q) (countedSequence_nonempty p q) rw [compl_compl, first_vote_pos _ _ h] at this - rw [← ENNReal.sub_eq_of_add_eq _ this, ENNReal.eq_div_iff, ENNReal.mul_sub, mul_one, + rw [ENNReal.eq_sub_of_add_eq _ this, ENNReal.eq_div_iff, ENNReal.mul_sub, mul_one, ENNReal.mul_div_cancel', ENNReal.add_sub_cancel_left] all_goals simp_all [ENNReal.div_eq_top] diff --git a/Mathlib.lean b/Mathlib.lean index f7b1e13424d66..267dd1c310b92 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1097,6 +1097,7 @@ import Mathlib.Analysis.Convex.Cone.Extension import Mathlib.Analysis.Convex.Cone.InnerDual import Mathlib.Analysis.Convex.Cone.Pointed import Mathlib.Analysis.Convex.Cone.Proper +import Mathlib.Analysis.Convex.Continuous import Mathlib.Analysis.Convex.Contractible import Mathlib.Analysis.Convex.Deriv import Mathlib.Analysis.Convex.EGauge @@ -2518,6 +2519,7 @@ import Mathlib.Data.Real.GoldenRatio import Mathlib.Data.Real.Hyperreal import Mathlib.Data.Real.Irrational import Mathlib.Data.Real.Pi.Bounds +import Mathlib.Data.Real.Pi.Irrational import Mathlib.Data.Real.Pi.Leibniz import Mathlib.Data.Real.Pi.Wallis import Mathlib.Data.Real.Pointwise @@ -4067,8 +4069,10 @@ import Mathlib.RingTheory.TensorProduct.MvPolynomial import Mathlib.RingTheory.Trace.Basic import Mathlib.RingTheory.Trace.Defs import Mathlib.RingTheory.TwoSidedIdeal.Basic +import Mathlib.RingTheory.TwoSidedIdeal.BigOperators import Mathlib.RingTheory.TwoSidedIdeal.Instances import Mathlib.RingTheory.TwoSidedIdeal.Lattice +import Mathlib.RingTheory.TwoSidedIdeal.Operations import Mathlib.RingTheory.UniqueFactorizationDomain import Mathlib.RingTheory.Unramified.Basic import Mathlib.RingTheory.Unramified.Derivations diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index 29d91b4ab7439..a7a16de103ee9 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -185,8 +185,7 @@ theorem End_algebraMap_isUnit_inv_apply_eq_iff {x : R} mpr := fun H => H.symm ▸ by apply_fun ⇑h.unit.val using ((Module.End_isUnit_iff _).mp h).injective - erw [End_isUnit_apply_inv_apply_of_isUnit] - rfl } + simpa using End_isUnit_apply_inv_apply_of_isUnit h (x • m') } theorem End_algebraMap_isUnit_inv_apply_eq_iff' {x : R} (h : IsUnit (algebraMap R (Module.End S M) x)) (m m' : M) : @@ -195,8 +194,7 @@ theorem End_algebraMap_isUnit_inv_apply_eq_iff' {x : R} mpr := fun H => H.symm ▸ by apply_fun (↑h.unit : M → M) using ((Module.End_isUnit_iff _).mp h).injective - erw [End_isUnit_apply_inv_apply_of_isUnit] - rfl } + simpa using End_isUnit_apply_inv_apply_of_isUnit h (x • m') |>.symm } end diff --git a/Mathlib/Algebra/CharP/Defs.lean b/Mathlib/Algebra/CharP/Defs.lean index d470a6086987e..34b448a97de44 100644 --- a/Mathlib/Algebra/CharP/Defs.lean +++ b/Mathlib/Algebra/CharP/Defs.lean @@ -11,6 +11,7 @@ import Mathlib.Data.Nat.Cast.Prod import Mathlib.Data.Nat.Find import Mathlib.Data.Nat.Prime.Defs import Mathlib.Data.ULift +import Mathlib.Tactic.NormNum.Basic /-! # Characteristic of semirings @@ -102,10 +103,10 @@ lemma intCast_injOn_Ico [IsRightCancelAdd R] : InjOn (Int.cast : ℤ → R) (Ico lemma intCast_eq_zero_iff (a : ℤ) : (a : R) = 0 ↔ (p : ℤ) ∣ a := by rcases lt_trichotomy a 0 with (h | rfl | h) - · rw [← neg_eq_zero, ← Int.cast_neg, ← dvd_neg] + · rw [← neg_eq_zero, ← Int.cast_neg, ← Int.dvd_neg] lift -a to ℕ using neg_nonneg.mpr (le_of_lt h) with b rw [Int.cast_natCast, CharP.cast_eq_zero_iff R p, Int.natCast_dvd_natCast] - · simp only [Int.cast_zero, eq_self_iff_true, dvd_zero] + · simp only [Int.cast_zero, eq_self_iff_true, Int.dvd_zero] · lift a to ℕ using le_of_lt h with b rw [Int.cast_natCast, CharP.cast_eq_zero_iff R p, Int.natCast_dvd_natCast] diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index f495bba039807..febcba6ed781c 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -452,7 +452,7 @@ theorem map_pow [Monoid G] [Monoid H] [MonoidHomClass F G H] (f : F) (a : G) : | n + 1 => by rw [pow_succ, pow_succ, map_mul, map_pow f a n] @[to_additive (attr := simp)] -lemma map_comp_pow [Group G] [DivisionMonoid H] [MonoidHomClass F G H] (f : F) (g : ι → G) (n : ℕ) : +lemma map_comp_pow [Monoid G] [Monoid H] [MonoidHomClass F G H] (f : F) (g : ι → G) (n : ℕ) : f ∘ (g ^ n) = f ∘ g ^ n := by ext; simp @[to_additive] diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index 0d13d5ae72a71..2df0f53da5529 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -69,7 +69,7 @@ Definitions in the file: * `MonoidHom.ker f` : the kernel of a group homomorphism `f` is the subgroup of elements `x : G` such that `f x = 1` -* `MonoidHom.eq_locus f g` : given group homomorphisms `f`, `g`, the elements of `G` such that +* `MonoidHom.eqLocus f g` : given group homomorphisms `f`, `g`, the elements of `G` such that `f x = g x` form a subgroup of `G` ## Implementation notes diff --git a/Mathlib/Algebra/Lie/Subalgebra.lean b/Mathlib/Algebra/Lie/Subalgebra.lean index 649208ace2362..a1d215735203f 100644 --- a/Mathlib/Algebra/Lie/Subalgebra.lean +++ b/Mathlib/Algebra/Lie/Subalgebra.lean @@ -224,7 +224,7 @@ variable [Module R M] `L`, we may regard `M` as a Lie module of `L'` by restriction. -/ instance lieModule [LieModule R L M] : LieModule R L' M where smul_lie t x m := by - rw [coe_bracket_of_module]; erw [smul_lie]; simp only [coe_bracket_of_module] + rw [coe_bracket_of_module, Submodule.coe_smul_of_tower, smul_lie, coe_bracket_of_module] lie_smul t x m := by simp only [coe_bracket_of_module, lie_smul] /-- An `L`-equivariant map of Lie modules `M → N` is `L'`-equivariant for any Lie subalgebra diff --git a/Mathlib/Algebra/ModEq.lean b/Mathlib/Algebra/ModEq.lean index 8c4bcc0fd4def..a275aad2c7254 100644 --- a/Mathlib/Algebra/ModEq.lean +++ b/Mathlib/Algebra/ModEq.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies -/ import Mathlib.Data.Int.ModEq import Mathlib.Algebra.Field.Basic +import Mathlib.Algebra.Order.Ring.Int import Mathlib.GroupTheory.QuotientGroup.Basic /-! diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index 3d210451cac46..502f5ddd8ebdf 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -122,7 +122,7 @@ protected theorem smul_def (f : M ≃ₗ[R] M) (a : M) : f • a = f a := /-- `LinearEquiv.applyDistribMulAction` is faithful. -/ instance apply_faithfulSMul : FaithfulSMul (M ≃ₗ[R] M) M := - ⟨@fun _ _ ↦ LinearEquiv.ext⟩ + ⟨LinearEquiv.ext⟩ instance apply_smulCommClass [SMul S R] [SMul S M] [IsScalarTower S R M] : SMulCommClass S (M ≃ₗ[R] M) M where diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index e442944e3cb7b..e2c0d89223846 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -933,9 +933,8 @@ end Actions section RestrictScalarsAsLinearMap -variable {R S M N : Type*} [Semiring R] [Semiring S] [AddCommGroup M] [AddCommGroup N] [Module R M] - [Module R N] [Module S M] [Module S N] - [LinearMap.CompatibleSMul M N R S] +variable {R S M N P : Type*} [Semiring R] [Semiring S] [AddCommMonoid M] [AddCommMonoid N] + [Module R M] [Module R N] [Module S M] [Module S N] [CompatibleSMul M N R S] variable (R S M N) in @[simp] @@ -948,7 +947,9 @@ theorem restrictScalars_add (f g : M →ₗ[S] N) : rfl @[simp] -theorem restrictScalars_neg (f : M →ₗ[S] N) : (-f).restrictScalars R = -f.restrictScalars R := +theorem restrictScalars_neg {M N : Type*} [AddCommGroup M] [AddCommGroup N] + [Module R M] [Module R N] [Module S M] [Module S N] [CompatibleSMul M N R S] + (f : M →ₗ[S] N) : (-f).restrictScalars R = -f.restrictScalars R := rfl variable {R₁ : Type*} [Semiring R₁] [Module R₁ N] [SMulCommClass S R₁ N] [SMulCommClass R R₁ N] @@ -958,6 +959,18 @@ theorem restrictScalars_smul (c : R₁) (f : M →ₗ[S] N) : (c • f).restrictScalars R = c • f.restrictScalars R := rfl +@[simp] +lemma restrictScalars_comp [AddCommMonoid P] [Module S P] [Module R P] + [CompatibleSMul N P R S] [CompatibleSMul M P R S] (f : N →ₗ[S] P) (g : M →ₗ[S] N) : + (f ∘ₗ g).restrictScalars R = f.restrictScalars R ∘ₗ g.restrictScalars R := by + rfl + +@[simp] +lemma restrictScalars_trans {T : Type*} [CommSemiring T] [Module T M] [Module T N] + [CompatibleSMul M N S T] [CompatibleSMul M N R T] (f : M →ₗ[T] N) : + (f.restrictScalars S).restrictScalars R = f.restrictScalars R := + rfl + variable (S M N R R₁) /-- `LinearMap.restrictScalars` as a `LinearMap`. -/ diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index da687de7c7942..722c1bdf7b69a 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -1357,6 +1357,11 @@ theorem comp_aeval {B : Type*} [CommSemiring B] [Algebra R B] (φ : S₁ →ₐ[ ext i simp +lemma comp_aeval_apply {B : Type*} [CommSemiring B] [Algebra R B] (φ : S₁ →ₐ[R] B) + (p : MvPolynomial σ R) : + φ (aeval f p) = aeval (fun i ↦ φ (f i)) p := by + rw [← comp_aeval, AlgHom.coe_comp, comp_apply] + @[simp] theorem map_aeval {B : Type*} [CommSemiring B] (g : σ → S₁) (φ : S₁ →+* B) (p : MvPolynomial σ R) : φ (aeval g p) = eval₂Hom (φ.comp (algebraMap R S₁)) (fun i => φ (g i)) p := by @@ -1537,6 +1542,17 @@ theorem eval_mem {p : MvPolynomial σ S} {s : subS} (hs : ∀ i ∈ p.support, p end EvalMem +variable {S T : Type*} [CommSemiring S] [Algebra R S] [CommSemiring T] [Algebra R T] [Algebra S T] + [IsScalarTower R S T] + +lemma aeval_sum_elim {σ τ : Type*} (p : MvPolynomial (σ ⊕ τ) R) (f : τ → S) (g : σ → T) : + (aeval (Sum.elim g (algebraMap S T ∘ f))) p = + (aeval g) ((aeval (Sum.elim X (C ∘ f))) p) := by + induction' p using MvPolynomial.induction_on with r p q hp hq p i h + · simp [← IsScalarTower.algebraMap_apply] + · simp [hp, hq] + · cases i <;> simp [h] + end CommSemiring end MvPolynomial diff --git a/Mathlib/Algebra/MvPolynomial/PDeriv.lean b/Mathlib/Algebra/MvPolynomial/PDeriv.lean index a4755788c2c2a..b7aa67a2aad48 100644 --- a/Mathlib/Algebra/MvPolynomial/PDeriv.lean +++ b/Mathlib/Algebra/MvPolynomial/PDeriv.lean @@ -115,6 +115,27 @@ theorem pderiv_map {S} [CommSemiring S] {φ : R →+* S} {f : MvPolynomial σ R} · simp [eq] · simp [eq, h] +lemma pderiv_rename {τ : Type*} [DecidableEq τ] [DecidableEq σ] {f : σ → τ} + (hf : Function.Injective f) (x : σ) (p : MvPolynomial σ R) : + pderiv (f x) (rename f p) = rename f (pderiv x p) := by + induction' p using MvPolynomial.induction_on with a p q hp hq p a h + · simp + · simp [hp, hq] + · simp only [map_mul, MvPolynomial.rename_X, Derivation.leibniz, MvPolynomial.pderiv_X, + Pi.single_apply, hf.eq_iff, smul_eq_mul, mul_ite, mul_one, mul_zero, h, map_add, add_left_inj] + split_ifs <;> simp + +lemma aeval_sum_elim_pderiv_inl {S τ : Type*} [CommRing S] [Algebra R S] + (p : MvPolynomial (σ ⊕ τ) R) (f : τ → S) (j : σ) : + aeval (Sum.elim X (C ∘ f)) ((pderiv (Sum.inl j)) p) = + (pderiv j) ((aeval (Sum.elim X (C ∘ f))) p) := by + classical + induction' p using MvPolynomial.induction_on with a p q hp hq p q h + · simp + · simp [hp, hq] + · simp only [Derivation.leibniz, pderiv_X, smul_eq_mul, map_add, map_mul, aeval_X, h] + cases q <;> simp [Pi.single_apply] + end PDeriv end MvPolynomial diff --git a/Mathlib/Algebra/Order/BigOperators/Expect.lean b/Mathlib/Algebra/Order/BigOperators/Expect.lean index 30064a4bcb627..a449b9ce36435 100644 --- a/Mathlib/Algebra/Order/BigOperators/Expect.lean +++ b/Mathlib/Algebra/Order/BigOperators/Expect.lean @@ -130,8 +130,7 @@ end LinearOrderedAddCommMonoid section LinearOrderedAddCommGroup variable [LinearOrderedAddCommGroup α] [Module ℚ≥0 α] [PosSMulMono ℚ≥0 α] --- TODO: Norm version -lemma abs_expect_le_expect_abs (s : Finset ι) (f : ι → α) : |𝔼 i ∈ s, f i| ≤ 𝔼 i ∈ s, |f i| := +lemma abs_expect_le (s : Finset ι) (f : ι → α) : |𝔼 i ∈ s, f i| ≤ 𝔼 i ∈ s, |f i| := le_expect_of_subadditive abs_zero abs_add (fun _ ↦ abs_nnqsmul _) end LinearOrderedAddCommGroup diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index 2f1547de189c6..c7a0670ed2b05 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -221,6 +221,16 @@ theorem abs_sum_of_nonneg' {G : Type*} [LinearOrderedAddCommGroup G] {f : ι → (hf : ∀ i, 0 ≤ f i) : |∑ i ∈ s, f i| = ∑ i ∈ s, f i := by rw [abs_of_nonneg (Finset.sum_nonneg' hf)] +section CommMonoid +variable [CommMonoid α] [LE α] [CovariantClass α α (· * ·) (· ≤ ·)] {s : Finset ι} {f : ι → α} + +@[to_additive (attr := simp)] +lemma mulLECancellable_prod : + MulLECancellable (∏ i ∈ s, f i) ↔ ∀ ⦃i⦄, i ∈ s → MulLECancellable (f i) := by + induction' s using Finset.cons_induction with i s hi ih <;> simp [*] + +end CommMonoid + section Pigeonhole variable [DecidableEq β] diff --git a/Mathlib/Algebra/Order/Chebyshev.lean b/Mathlib/Algebra/Order/Chebyshev.lean index ebbe5643ee00c..439d79b0c4e2b 100644 --- a/Mathlib/Algebra/Order/Chebyshev.lean +++ b/Mathlib/Algebra/Order/Chebyshev.lean @@ -4,9 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mantas Bakšys, Yaël Dillies -/ import Mathlib.Algebra.Order.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Monovary import Mathlib.Algebra.Order.Rearrangement -import Mathlib.Algebra.Order.Ring.Basic import Mathlib.GroupTheory.Perm.Cycle.Basic +import Mathlib.Tactic.GCongr +import Mathlib.Tactic.Positivity.Basic +import Mathlib.Tactic.Positivity.Finset /-! # Chebyshev's sum inequality @@ -44,28 +47,26 @@ variable {ι α β : Type*} section SMul - -variable [LinearOrderedRing α] [LinearOrderedAddCommGroup β] [Module α β] [OrderedSMul α β] - {s : Finset ι} {σ : Perm ι} {f : ι → α} {g : ι → β} +variable [LinearOrderedSemiring α] [ExistsAddOfLE α] [LinearOrderedCancelAddCommMonoid β] + [Module α β] [OrderedSMul α β] {s : Finset ι} {σ : Perm ι} {f : ι → α} {g : ι → β} /-- **Chebyshev's Sum Inequality**: When `f` and `g` monovary together (eg they are both monotone/antitone), the scalar product of their sum is less than the size of the set times their scalar product. -/ theorem MonovaryOn.sum_smul_sum_le_card_smul_sum (hfg : MonovaryOn f g s) : - ((∑ i ∈ s, f i) • ∑ i ∈ s, g i) ≤ s.card • ∑ i ∈ s, f i • g i := by + (∑ i ∈ s, f i) • ∑ i ∈ s, g i ≤ s.card • ∑ i ∈ s, f i • g i := by classical - obtain ⟨σ, hσ, hs⟩ := s.countable_toSet.exists_cycleOn - rw [← card_range s.card, sum_smul_sum_eq_sum_perm hσ] - exact - sum_le_card_nsmul _ _ _ fun n _ => - hfg.sum_smul_comp_perm_le_sum_smul fun x hx => hs fun h => hx <| IsFixedPt.perm_pow h _ + obtain ⟨σ, hσ, hs⟩ := s.countable_toSet.exists_cycleOn + rw [← card_range s.card, sum_smul_sum_eq_sum_perm hσ] + exact sum_le_card_nsmul _ _ _ fun n _ ↦ + hfg.sum_smul_comp_perm_le_sum_smul fun x hx ↦ hs fun h ↦ hx <| IsFixedPt.perm_pow h _ /-- **Chebyshev's Sum Inequality**: When `f` and `g` antivary together (eg one is monotone, the other is antitone), the scalar product of their sum is less than the size of the set times their scalar product. -/ theorem AntivaryOn.card_smul_sum_le_sum_smul_sum (hfg : AntivaryOn f g s) : - (s.card • ∑ i ∈ s, f i • g i) ≤ (∑ i ∈ s, f i) • ∑ i ∈ s, g i := by - exact hfg.dual_right.sum_smul_sum_le_card_smul_sum + s.card • ∑ i ∈ s, f i • g i ≤ (∑ i ∈ s, f i) • ∑ i ∈ s, g i := + hfg.dual_right.sum_smul_sum_le_card_smul_sum variable [Fintype ι] @@ -73,15 +74,15 @@ variable [Fintype ι] monotone/antitone), the scalar product of their sum is less than the size of the set times their scalar product. -/ theorem Monovary.sum_smul_sum_le_card_smul_sum (hfg : Monovary f g) : - ((∑ i, f i) • ∑ i, g i) ≤ Fintype.card ι • ∑ i, f i • g i := + (∑ i, f i) • ∑ i, g i ≤ Fintype.card ι • ∑ i, f i • g i := (hfg.monovaryOn _).sum_smul_sum_le_card_smul_sum /-- **Chebyshev's Sum Inequality**: When `f` and `g` antivary together (eg one is monotone, the other is antitone), the scalar product of their sum is less than the size of the set times their scalar product. -/ theorem Antivary.card_smul_sum_le_sum_smul_sum (hfg : Antivary f g) : - (Fintype.card ι • ∑ i, f i • g i) ≤ (∑ i, f i) • ∑ i, g i := by - exact (hfg.dual_right.monovaryOn _).sum_smul_sum_le_card_smul_sum + Fintype.card ι • ∑ i, f i • g i ≤ (∑ i, f i) • ∑ i, g i := + (hfg.dual_right.monovaryOn _).sum_smul_sum_le_card_smul_sum end SMul @@ -93,14 +94,13 @@ Special cases of the above when scalar multiplication is actually multiplication section Mul - -variable [LinearOrderedRing α] {s : Finset ι} {σ : Perm ι} {f g : ι → α} +variable [LinearOrderedSemiring α] [ExistsAddOfLE α] {s : Finset ι} {σ : Perm ι} {f g : ι → α} /-- **Chebyshev's Sum Inequality**: When `f` and `g` monovary together (eg they are both monotone/antitone), the product of their sum is less than the size of the set times their scalar product. -/ theorem MonovaryOn.sum_mul_sum_le_card_mul_sum (hfg : MonovaryOn f g s) : - ((∑ i ∈ s, f i) * ∑ i ∈ s, g i) ≤ s.card * ∑ i ∈ s, f i * g i := by + (∑ i ∈ s, f i) * ∑ i ∈ s, g i ≤ s.card * ∑ i ∈ s, f i * g i := by rw [← nsmul_eq_mul] exact hfg.sum_smul_sum_le_card_smul_sum @@ -108,10 +108,26 @@ theorem MonovaryOn.sum_mul_sum_le_card_mul_sum (hfg : MonovaryOn f g s) : other is antitone), the product of their sum is greater than the size of the set times their scalar product. -/ theorem AntivaryOn.card_mul_sum_le_sum_mul_sum (hfg : AntivaryOn f g s) : - ((s.card : α) * ∑ i ∈ s, f i * g i) ≤ (∑ i ∈ s, f i) * ∑ i ∈ s, g i := by + (s.card : α) * ∑ i ∈ s, f i * g i ≤ (∑ i ∈ s, f i) * ∑ i ∈ s, g i := by rw [← nsmul_eq_mul] exact hfg.card_smul_sum_le_sum_smul_sum +/-- Special case of **Jensen's inequality** for sums of powers. -/ +lemma pow_sum_le_card_mul_sum_pow (hf : ∀ i ∈ s, 0 ≤ f i) : + ∀ n, (∑ i ∈ s, f i) ^ (n + 1) ≤ (s.card : α) ^ n * ∑ i ∈ s, f i ^ (n + 1) + | 0 => by simp + | n + 1 => + calc + _ = (∑ i ∈ s, f i) ^ (n + 1) * ∑ i ∈ s, f i := by rw [pow_succ] + _ ≤ (s.card ^ n * ∑ i ∈ s, f i ^ (n + 1)) * ∑ i ∈ s, f i := by + gcongr + exacts [sum_nonneg hf, pow_sum_le_card_mul_sum_pow hf _] + _ = s.card ^ n * ((∑ i ∈ s, f i ^ (n + 1)) * ∑ i ∈ s, f i) := by rw [mul_assoc] + _ ≤ s.card ^ n * (s.card * ∑ i ∈ s, f i ^ (n + 1) * f i) := by + gcongr _ * ?_ + exact ((monovaryOn_self ..).pow_left₀ hf _).sum_mul_sum_le_card_mul_sum + _ = _ := by simp_rw [← mul_assoc, ← pow_succ] + /-- Special case of **Chebyshev's Sum Inequality** or the **Cauchy-Schwarz Inequality**: The square of the sum is less than the size of the set times the sum of the squares. -/ theorem sq_sum_le_card_mul_sum_sq : (∑ i ∈ s, f i) ^ 2 ≤ s.card * ∑ i ∈ s, f i ^ 2 := by @@ -124,25 +140,32 @@ variable [Fintype ι] monotone/antitone), the product of their sum is less than the size of the set times their scalar product. -/ theorem Monovary.sum_mul_sum_le_card_mul_sum (hfg : Monovary f g) : - ((∑ i, f i) * ∑ i, g i) ≤ Fintype.card ι * ∑ i, f i * g i := + (∑ i, f i) * ∑ i, g i ≤ Fintype.card ι * ∑ i, f i * g i := (hfg.monovaryOn _).sum_mul_sum_le_card_mul_sum /-- **Chebyshev's Sum Inequality**: When `f` and `g` antivary together (eg one is monotone, the other is antitone), the product of their sum is less than the size of the set times their scalar product. -/ theorem Antivary.card_mul_sum_le_sum_mul_sum (hfg : Antivary f g) : - ((Fintype.card ι : α) * ∑ i, f i * g i) ≤ (∑ i, f i) * ∑ i, g i := + Fintype.card ι * ∑ i, f i * g i ≤ (∑ i, f i) * ∑ i, g i := (hfg.antivaryOn _).card_mul_sum_le_sum_mul_sum end Mul -variable [LinearOrderedField α] {s : Finset ι} {f : ι → α} +variable [LinearOrderedSemifield α] [ExistsAddOfLE α] {s : Finset ι} {f : ι → α} + +/-- Special case of **Jensen's inequality** for sums of powers. -/ +lemma pow_sum_div_card_le_sum_pow (hf : ∀ i ∈ s, 0 ≤ f i) (n : ℕ) : + (∑ i ∈ s, f i) ^ (n + 1) / s.card ^ n ≤ ∑ i ∈ s, f i ^ (n + 1) := by + obtain rfl | hs := s.eq_empty_or_nonempty + · simp + rw [div_le_iff₀' (by positivity)] + exact pow_sum_le_card_mul_sum_pow hf _ theorem sum_div_card_sq_le_sum_sq_div_card : ((∑ i ∈ s, f i) / s.card) ^ 2 ≤ (∑ i ∈ s, f i ^ 2) / s.card := by obtain rfl | hs := s.eq_empty_or_nonempty · simp - rw [← card_pos, ← @Nat.cast_pos α] at hs - rw [div_pow, div_le_div_iff (sq_pos_of_ne_zero hs.ne') hs, sq (s.card : α), mul_left_comm, ← - mul_assoc] - exact mul_le_mul_of_nonneg_right sq_sum_le_card_mul_sum_sq hs.le + rw [div_pow, div_le_div_iff (by positivity) (by positivity), sq (s.card : α), mul_left_comm, + ← mul_assoc] + exact mul_le_mul_of_nonneg_right sq_sum_le_card_mul_sum_sq (by positivity) diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Int.lean b/Mathlib/Algebra/Order/Group/Unbundled/Int.lean index 773e4ce5f46eb..96c6eed502379 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Int.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Int.lean @@ -82,6 +82,15 @@ theorem abs_le_one_iff {a : ℤ} : |a| ≤ 1 ↔ a = 0 ∨ a = 1 ∨ a = -1 := b theorem one_le_abs {z : ℤ} (h₀ : z ≠ 0) : 1 ≤ |z| := add_one_le_iff.mpr (abs_pos.mpr h₀) +lemma eq_zero_of_abs_lt_dvd {m x : ℤ} (h1 : m ∣ x) (h2 : |x| < m) : x = 0 := by + by_contra h + have := Int.natAbs_le_of_dvd_ne_zero h1 h + rw [Int.abs_eq_natAbs] at h2 + omega + +lemma abs_sub_lt_of_lt_lt {m a b : ℕ} (ha : a < m) (hb : b < m) : |(b : ℤ) - a| < m := by + rw [abs_lt]; omega + /-! #### `/` -/ theorem ediv_eq_zero_of_lt_abs {a b : ℤ} (H1 : 0 ≤ a) (H2 : a < |b|) : a / b = 0 := diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean index d6509bd4ccb64..f33502550e3da 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean @@ -1289,7 +1289,7 @@ theorem Contravariant.MulLECancellable [Mul α] [LE α] [ContravariantClass α MulLECancellable a := fun _ _ => le_of_mul_le_mul_left' -@[to_additive] +@[to_additive (attr := simp)] theorem mulLECancellable_one [Monoid α] [LE α] : MulLECancellable (1 : α) := fun a b => by simpa only [one_mul] using id @@ -1350,4 +1350,20 @@ protected theorem mul_le_iff_le_one_left [MulOneClass α] [i : @Std.Commutative [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} (ha : MulLECancellable a) : b * a ≤ a ↔ b ≤ 1 := by rw [i.comm, ha.mul_le_iff_le_one_right] +@[to_additive] lemma mul [Semigroup α] {a b : α} (ha : MulLECancellable a) + (hb : MulLECancellable b) : MulLECancellable (a * b) := + fun c d hcd ↦ hb <| ha <| by rwa [← mul_assoc, ← mul_assoc] + +@[to_additive] lemma of_mul_right [Semigroup α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} + (h : MulLECancellable (a * b)) : MulLECancellable b := + fun c d hcd ↦ h <| by rw [mul_assoc, mul_assoc]; exact mul_le_mul_left' hcd _ + +@[to_additive] lemma of_mul_left [CommSemigroup α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} + (h : MulLECancellable (a * b)) : MulLECancellable a := (mul_comm a b ▸ h).of_mul_right + end MulLECancellable + +@[to_additive (attr := simp)] +lemma mulLECancellable_mul [LE α] [CommSemigroup α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} : + MulLECancellable (a * b) ↔ MulLECancellable a ∧ MulLECancellable b := + ⟨fun h ↦ ⟨h.of_mul_left, h.of_mul_right⟩, fun h ↦ h.1.mul h.2⟩ diff --git a/Mathlib/Algebra/Order/Rearrangement.lean b/Mathlib/Algebra/Order/Rearrangement.lean index f032d3104e240..f69acca4d99d3 100644 --- a/Mathlib/Algebra/Order/Rearrangement.lean +++ b/Mathlib/Algebra/Order/Rearrangement.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mantas Bakšys -/ import Mathlib.Algebra.BigOperators.Group.Finset -import Mathlib.Algebra.Order.Group.Instances import Mathlib.Algebra.Order.Module.OrderedSMul import Mathlib.Algebra.Order.Module.Synonym import Mathlib.Data.Prod.Lex @@ -41,23 +40,30 @@ convenience. The case for `Monotone`/`Antitone` pairs of functions over a `LinearOrder` is not deduced in this file because it is easily deducible from the `Monovary` API. + +## TODO + +Add equality cases for when the permute function is injective. This comes from the following fact: +If `Monovary f g`, `Injective g` and `σ` is a permutation, then `Monovary f (g ∘ σ) ↔ σ = 1`. -/ open Equiv Equiv.Perm Finset Function OrderDual -variable {ι α β : Type*} +variable {ι α β : Type*} [LinearOrderedSemiring α] [ExistsAddOfLE α] + [LinearOrderedCancelAddCommMonoid β] [Module α β] /-! ### Scalar multiplication versions -/ - section SMul -variable [LinearOrderedRing α] [LinearOrderedAddCommGroup β] [Module α β] [OrderedSMul α β] - {s : Finset ι} {σ : Perm ι} {f : ι → α} {g : ι → β} +/-! #### Weak rearrangement inequality -/ + +section weak_inequality +variable [PosSMulMono α β] {s : Finset ι} {σ : Perm ι} {f : ι → α} {g : ι → β} /-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when -`f` and `g` monovary together. Stated by permuting the entries of `g`. -/ +`f` and `g` monovary together on `s`. Stated by permuting the entries of `g`. -/ theorem MonovaryOn.sum_smul_comp_perm_le_sum_smul (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g (σ i) ≤ ∑ i ∈ s, f i • g i := by classical @@ -106,9 +112,62 @@ theorem MonovaryOn.sum_smul_comp_perm_le_sum_smul (hfg : MonovaryOn f g s) rintro rfl exact has hx.2 +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when +`f` and `g` antivary together on `s`. Stated by permuting the entries of `g`. -/ +theorem AntivaryOn.sum_smul_le_sum_smul_comp_perm (hfg : AntivaryOn f g s) + (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g i ≤ ∑ i ∈ s, f i • g (σ i) := + hfg.dual_right.sum_smul_comp_perm_le_sum_smul hσ + +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when +`f` and `g` monovary together on `s`. Stated by permuting the entries of `f`. -/ +theorem MonovaryOn.sum_comp_perm_smul_le_sum_smul (hfg : MonovaryOn f g s) + (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) • g i ≤ ∑ i ∈ s, f i • g i := by + convert hfg.sum_smul_comp_perm_le_sum_smul + (show { x | σ⁻¹ x ≠ x } ⊆ s by simp only [set_support_inv_eq, hσ]) using 1 + exact σ.sum_comp' s (fun i j ↦ f i • g j) hσ + +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when +`f` and `g` antivary together on `s`. Stated by permuting the entries of `f`. -/ +theorem AntivaryOn.sum_smul_le_sum_comp_perm_smul (hfg : AntivaryOn f g s) + (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g i ≤ ∑ i ∈ s, f (σ i) • g i := + hfg.dual_right.sum_comp_perm_smul_le_sum_smul hσ + +variable [Fintype ι] + +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when +`f` and `g` monovary together. Stated by permuting the entries of `g`. -/ +theorem Monovary.sum_smul_comp_perm_le_sum_smul (hfg : Monovary f g) : + ∑ i, f i • g (σ i) ≤ ∑ i, f i • g i := + (hfg.monovaryOn _).sum_smul_comp_perm_le_sum_smul fun _ _ ↦ mem_univ _ + +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when +`f` and `g` antivary together. Stated by permuting the entries of `g`. -/ +theorem Antivary.sum_smul_le_sum_smul_comp_perm (hfg : Antivary f g) : + ∑ i, f i • g i ≤ ∑ i, f i • g (σ i) := + (hfg.antivaryOn _).sum_smul_le_sum_smul_comp_perm fun _ _ ↦ mem_univ _ + +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when +`f` and `g` monovary together. Stated by permuting the entries of `f`. -/ +theorem Monovary.sum_comp_perm_smul_le_sum_smul (hfg : Monovary f g) : + ∑ i, f (σ i) • g i ≤ ∑ i, f i • g i := + (hfg.monovaryOn _).sum_comp_perm_smul_le_sum_smul fun _ _ ↦ mem_univ _ + +/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when +`f` and `g` antivary together. Stated by permuting the entries of `f`. -/ +theorem Antivary.sum_smul_le_sum_comp_perm_smul (hfg : Antivary f g) : + ∑ i, f i • g i ≤ ∑ i, f (σ i) • g i := + (hfg.antivaryOn _).sum_smul_le_sum_comp_perm_smul fun _ _ ↦ mem_univ _ + +end weak_inequality + +/-! #### Equality case of the rearrangement inequality -/ + +section equality_case +variable [PosSMulStrictMono α β] {s : Finset ι} {σ : Perm ι} {f : ι → α} {g : ι → β} + /-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which monovary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` monovary -together. Stated by permuting the entries of `g`. -/ +`g`, which monovary together on `s`, is unchanged by a permutation if and only if `f` and `g ∘ σ` +monovary together on `s`. Stated by permuting the entries of `g`. -/ theorem MonovaryOn.sum_smul_comp_perm_eq_sum_smul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g (σ i) = ∑ i ∈ s, f i • g i ↔ MonovaryOn f (g ∘ σ) s := by @@ -134,26 +193,17 @@ theorem MonovaryOn.sum_smul_comp_perm_eq_sum_smul_iff (hfg : MonovaryOn f g s) · convert h.sum_smul_comp_perm_le_sum_smul ((set_support_inv_eq _).subset.trans hσ) using 1 simp_rw [Function.comp_apply, apply_inv_self] -/-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of -`f` and `g`, which monovary together, is strictly decreased by a permutation if and only if -`f` and `g ∘ σ` do not monovary together. Stated by permuting the entries of `g`. -/ -theorem MonovaryOn.sum_smul_comp_perm_lt_sum_smul_iff (hfg : MonovaryOn f g s) +/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and +`g`, which antivary together on `s`, is unchanged by a permutation if and only if `f` and `g ∘ σ` +antivary together on `s`. Stated by permuting the entries of `g`. -/ +theorem AntivaryOn.sum_smul_comp_perm_eq_sum_smul_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : - ∑ i ∈ s, f i • g (σ i) < ∑ i ∈ s, f i • g i ↔ ¬MonovaryOn f (g ∘ σ) s := by - simp [← hfg.sum_smul_comp_perm_eq_sum_smul_iff hσ, lt_iff_le_and_ne, - hfg.sum_smul_comp_perm_le_sum_smul hσ] - -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when -`f` and `g` monovary together. Stated by permuting the entries of `f`. -/ -theorem MonovaryOn.sum_comp_perm_smul_le_sum_smul (hfg : MonovaryOn f g s) - (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) • g i ≤ ∑ i ∈ s, f i • g i := by - convert hfg.sum_smul_comp_perm_le_sum_smul - (show { x | σ⁻¹ x ≠ x } ⊆ s by simp only [set_support_inv_eq, hσ]) using 1 - exact σ.sum_comp' s (fun i j ↦ f i • g j) hσ + ∑ i ∈ s, f i • g (σ i) = ∑ i ∈ s, f i • g i ↔ AntivaryOn f (g ∘ σ) s := + (hfg.dual_right.sum_smul_comp_perm_eq_sum_smul_iff hσ).trans monovaryOn_toDual_right /-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which monovary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` monovary -together. Stated by permuting the entries of `f`. -/ +`g`, which monovary together on `s`, is unchanged by a permutation if and only if `f ∘ σ` and `g` +monovary together on `s`. Stated by permuting the entries of `f`. -/ theorem MonovaryOn.sum_comp_perm_smul_eq_sum_smul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) • g i = ∑ i ∈ s, f i • g i ↔ MonovaryOn (f ∘ σ) g s := by @@ -171,63 +221,93 @@ theorem MonovaryOn.sum_comp_perm_smul_eq_sum_smul_iff (hfg : MonovaryOn f g s) · rw [σ.symm.eq_preimage_iff_image_eq] exact Set.image_perm hσinv -/-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of -`f` and `g`, which monovary together, is strictly decreased by a permutation if and only if -`f ∘ σ` and `g` do not monovary together. Stated by permuting the entries of `f`. -/ -theorem MonovaryOn.sum_comp_perm_smul_lt_sum_smul_iff (hfg : MonovaryOn f g s) +/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and +`g`, which antivary together on `s`, is unchanged by a permutation if and only if `f ∘ σ` and `g` +antivary together on `s`. Stated by permuting the entries of `f`. -/ +theorem AntivaryOn.sum_comp_perm_smul_eq_sum_smul_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : - ∑ i ∈ s, f (σ i) • g i < ∑ i ∈ s, f i • g i ↔ ¬MonovaryOn (f ∘ σ) g s := by - simp [← hfg.sum_comp_perm_smul_eq_sum_smul_iff hσ, lt_iff_le_and_ne, - hfg.sum_comp_perm_smul_le_sum_smul hσ] + ∑ i ∈ s, f (σ i) • g i = ∑ i ∈ s, f i • g i ↔ AntivaryOn (f ∘ σ) g s := + (hfg.dual_right.sum_comp_perm_smul_eq_sum_smul_iff hσ).trans monovaryOn_toDual_right -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when -`f` and `g` antivary together. Stated by permuting the entries of `g`. -/ -theorem AntivaryOn.sum_smul_le_sum_smul_comp_perm (hfg : AntivaryOn f g s) - (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g i ≤ ∑ i ∈ s, f i • g (σ i) := - hfg.dual_right.sum_smul_comp_perm_le_sum_smul hσ +@[deprecated (since := "2024-06-25")] +alias AntivaryOn.sum_smul_eq_sum_comp_perm_smul_iff := AntivaryOn.sum_comp_perm_smul_eq_sum_smul_iff + +variable [Fintype ι] + +/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and +`g`, which monovary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` monovary +together. Stated by permuting the entries of `g`. -/ +theorem Monovary.sum_smul_comp_perm_eq_sum_smul_iff (hfg : Monovary f g) : + ∑ i, f i • g (σ i) = ∑ i, f i • g i ↔ Monovary f (g ∘ σ) := by + simp [(hfg.monovaryOn _).sum_smul_comp_perm_eq_sum_smul_iff fun _ _ ↦ mem_univ _] + +/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and +`g`, which monovary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` monovary +together. Stated by permuting the entries of `g`. -/ +theorem Monovary.sum_comp_perm_smul_eq_sum_smul_iff (hfg : Monovary f g) : + ∑ i, f (σ i) • g i = ∑ i, f i • g i ↔ Monovary (f ∘ σ) g := by + simp [(hfg.monovaryOn _).sum_comp_perm_smul_eq_sum_smul_iff fun _ _ ↦ mem_univ _] /-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g`, which antivary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` antivary together. Stated by permuting the entries of `g`. -/ -theorem AntivaryOn.sum_smul_comp_perm_eq_sum_smul_iff (hfg : AntivaryOn f g s) - (hσ : {x | σ x ≠ x} ⊆ s) : - ∑ i ∈ s, f i • g (σ i) = ∑ i ∈ s, f i • g i ↔ AntivaryOn f (g ∘ σ) s := - (hfg.dual_right.sum_smul_comp_perm_eq_sum_smul_iff hσ).trans monovaryOn_toDual_right +theorem Antivary.sum_smul_comp_perm_eq_sum_smul_iff (hfg : Antivary f g) : + ∑ i, f i • g (σ i) = ∑ i, f i • g i ↔ Antivary f (g ∘ σ) := by + simp [(hfg.antivaryOn _).sum_smul_comp_perm_eq_sum_smul_iff fun _ _ ↦ mem_univ _] @[deprecated (since := "2024-06-25")] -alias AntivaryOn.sum_smul_eq_sum_smul_comp_perm_iff := AntivaryOn.sum_smul_comp_perm_eq_sum_smul_iff +alias Antivary.sum_smul_eq_sum_smul_comp_perm_iff := Antivary.sum_smul_comp_perm_eq_sum_smul_iff + +/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and +`g`, which antivary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` antivary +together. Stated by permuting the entries of `f`. -/ +theorem Antivary.sum_comp_perm_smul_eq_sum_smul_iff (hfg : Antivary f g) : + ∑ i, f (σ i) • g i = ∑ i, f i • g i ↔ Antivary (f ∘ σ) g := by + simp [(hfg.antivaryOn _).sum_comp_perm_smul_eq_sum_smul_iff fun _ _ ↦ mem_univ _] + +@[deprecated (since := "2024-06-25")] +alias Antivary.sum_smul_eq_sum_comp_perm_smul_iff := Antivary.sum_comp_perm_smul_eq_sum_smul_iff + +end equality_case + +/-! #### Strict rearrangement inequality -/ + +section strict_inequality +variable [PosSMulStrictMono α β] {s : Finset ι} {σ : Perm ι} {f : ι → α} {g : ι → β} /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of -`f` and `g`, which antivary together, is strictly decreased by a permutation if and only if -`f` and `g ∘ σ` do not antivary together. Stated by permuting the entries of `g`. -/ +`f` and `g`, which monovary together on `s`, is strictly decreased by a permutation if and only if +`f` and `g ∘ σ` do not monovary together on `s`. Stated by permuting the entries of `g`. -/ +theorem MonovaryOn.sum_smul_comp_perm_lt_sum_smul_iff (hfg : MonovaryOn f g s) + (hσ : {x | σ x ≠ x} ⊆ s) : + ∑ i ∈ s, f i • g (σ i) < ∑ i ∈ s, f i • g i ↔ ¬MonovaryOn f (g ∘ σ) s := by + simp [← hfg.sum_smul_comp_perm_eq_sum_smul_iff hσ, lt_iff_le_and_ne, + hfg.sum_smul_comp_perm_le_sum_smul hσ] + +/-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of +`f` and `g`, which antivary together on `s`, is strictly decreased by a permutation if and only if +`f` and `g ∘ σ` do not antivary together on `s`. Stated by permuting the entries of `g`. -/ theorem AntivaryOn.sum_smul_lt_sum_smul_comp_perm_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g i < ∑ i ∈ s, f i • g (σ i) ↔ ¬AntivaryOn f (g ∘ σ) s := by simp [← hfg.sum_smul_comp_perm_eq_sum_smul_iff hσ, lt_iff_le_and_ne, eq_comm, hfg.sum_smul_le_sum_smul_comp_perm hσ] -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when -`f` and `g` antivary together. Stated by permuting the entries of `f`. -/ -theorem AntivaryOn.sum_smul_le_sum_comp_perm_smul (hfg : AntivaryOn f g s) - (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g i ≤ ∑ i ∈ s, f (σ i) • g i := by - convert hfg.sum_smul_le_sum_smul_comp_perm - (show { x | σ⁻¹ x ≠ x } ⊆ s by simp only [set_support_inv_eq, hσ]) using 1 - exact σ.sum_comp' s (fun i j ↦ f i • g j) hσ - -/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which antivary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` antivary -together. Stated by permuting the entries of `f`. -/ -theorem AntivaryOn.sum_comp_perm_smul_eq_sum_smul_iff (hfg : AntivaryOn f g s) +/-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of +`f` and `g`, which monovary together on `s`, is strictly decreased by a permutation if and only if +`f ∘ σ` and `g` do not monovary together on `s`. Stated by permuting the entries of `f`. -/ +theorem MonovaryOn.sum_comp_perm_smul_lt_sum_smul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : - ∑ i ∈ s, f (σ i) • g i = ∑ i ∈ s, f i • g i ↔ AntivaryOn (f ∘ σ) g s := - (hfg.dual_right.sum_comp_perm_smul_eq_sum_smul_iff hσ).trans monovaryOn_toDual_right + ∑ i ∈ s, f (σ i) • g i < ∑ i ∈ s, f i • g i ↔ ¬MonovaryOn (f ∘ σ) g s := by + simp [← hfg.sum_comp_perm_smul_eq_sum_smul_iff hσ, lt_iff_le_and_ne, + hfg.sum_comp_perm_smul_le_sum_smul hσ] @[deprecated (since := "2024-06-25")] -alias AntivaryOn.sum_smul_eq_sum_comp_perm_smul_iff := AntivaryOn.sum_comp_perm_smul_eq_sum_smul_iff +alias AntivaryOn.sum_smul_eq_sum_smul_comp_perm_iff := AntivaryOn.sum_smul_comp_perm_eq_sum_smul_iff /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of -`f` and `g`, which antivary together, is strictly decreased by a permutation if and only if -`f ∘ σ` and `g` do not antivary together. Stated by permuting the entries of `f`. -/ +`f` and `g`, which antivary together on `s`, is strictly decreased by a permutation if and only if +`f ∘ σ` and `g` do not antivary together on `s`. Stated by permuting the entries of `f`. -/ theorem AntivaryOn.sum_smul_lt_sum_comp_perm_smul_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g i < ∑ i ∈ s, f (σ i) • g i ↔ ¬AntivaryOn (f ∘ σ) g s := by @@ -236,19 +316,6 @@ theorem AntivaryOn.sum_smul_lt_sum_comp_perm_smul_iff (hfg : AntivaryOn f g s) variable [Fintype ι] -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when -`f` and `g` monovary together. Stated by permuting the entries of `g`. -/ -theorem Monovary.sum_smul_comp_perm_le_sum_smul (hfg : Monovary f g) : - ∑ i, f i • g (σ i) ≤ ∑ i, f i • g i := - (hfg.monovaryOn _).sum_smul_comp_perm_le_sum_smul fun _ _ ↦ mem_univ _ - -/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which monovary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` monovary -together. Stated by permuting the entries of `g`. -/ -theorem Monovary.sum_smul_comp_perm_eq_sum_smul_iff (hfg : Monovary f g) : - ∑ i, f i • g (σ i) = ∑ i, f i • g i ↔ Monovary f (g ∘ σ) := by - simp [(hfg.monovaryOn _).sum_smul_comp_perm_eq_sum_smul_iff fun _ _ ↦ mem_univ _] - /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g`, which monovary together, is strictly decreased by a permutation if and only if `f` and `g ∘ σ` do not monovary together. Stated by permuting the entries of `g`. -/ @@ -256,19 +323,6 @@ theorem Monovary.sum_smul_comp_perm_lt_sum_smul_iff (hfg : Monovary f g) : ∑ i, f i • g (σ i) < ∑ i, f i • g i ↔ ¬Monovary f (g ∘ σ) := by simp [(hfg.monovaryOn _).sum_smul_comp_perm_lt_sum_smul_iff fun _ _ ↦ mem_univ _] -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is maximized when -`f` and `g` monovary together. Stated by permuting the entries of `f`. -/ -theorem Monovary.sum_comp_perm_smul_le_sum_smul (hfg : Monovary f g) : - ∑ i, f (σ i) • g i ≤ ∑ i, f i • g i := - (hfg.monovaryOn _).sum_comp_perm_smul_le_sum_smul fun _ _ ↦ mem_univ _ - -/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which monovary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` monovary -together. Stated by permuting the entries of `g`. -/ -theorem Monovary.sum_comp_perm_smul_eq_sum_smul_iff (hfg : Monovary f g) : - ∑ i, f (σ i) • g i = ∑ i, f i • g i ↔ Monovary (f ∘ σ) g := by - simp [(hfg.monovaryOn _).sum_comp_perm_smul_eq_sum_smul_iff fun _ _ ↦ mem_univ _] - /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g`, which monovary together, is strictly decreased by a permutation if and only if `f` and `g ∘ σ` do not monovary together. Stated by permuting the entries of `g`. -/ @@ -276,22 +330,6 @@ theorem Monovary.sum_comp_perm_smul_lt_sum_smul_iff (hfg : Monovary f g) : ∑ i, f (σ i) • g i < ∑ i, f i • g i ↔ ¬Monovary (f ∘ σ) g := by simp [(hfg.monovaryOn _).sum_comp_perm_smul_lt_sum_smul_iff fun _ _ ↦ mem_univ _] -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when -`f` and `g` antivary together. Stated by permuting the entries of `g`. -/ -theorem Antivary.sum_smul_le_sum_smul_comp_perm (hfg : Antivary f g) : - ∑ i, f i • g i ≤ ∑ i, f i • g (σ i) := - (hfg.antivaryOn _).sum_smul_le_sum_smul_comp_perm fun _ _ ↦ mem_univ _ - -/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which antivary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` antivary -together. Stated by permuting the entries of `g`. -/ -theorem Antivary.sum_smul_comp_perm_eq_sum_smul_iff (hfg : Antivary f g) : - ∑ i, f i • g (σ i) = ∑ i, f i • g i ↔ Antivary f (g ∘ σ) := by - simp [(hfg.antivaryOn _).sum_smul_comp_perm_eq_sum_smul_iff fun _ _ ↦ mem_univ _] - -@[deprecated (since := "2024-06-25")] -alias Antivary.sum_smul_eq_sum_smul_comp_perm_iff := Antivary.sum_smul_comp_perm_eq_sum_smul_iff - /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g`, which antivary together, is strictly decreased by a permutation if and only if `f` and `g ∘ σ` do not antivary together. Stated by permuting the entries of `g`. -/ @@ -299,22 +337,6 @@ theorem Antivary.sum_smul_lt_sum_smul_comp_perm_iff (hfg : Antivary f g) : ∑ i, f i • g i < ∑ i, f i • g (σ i) ↔ ¬Antivary f (g ∘ σ) := by simp [(hfg.antivaryOn _).sum_smul_lt_sum_smul_comp_perm_iff fun _ _ ↦ mem_univ _] -/-- **Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g` is minimized when -`f` and `g` antivary together. Stated by permuting the entries of `f`. -/ -theorem Antivary.sum_smul_le_sum_comp_perm_smul (hfg : Antivary f g) : - ∑ i, f i • g i ≤ ∑ i, f (σ i) • g i := - (hfg.antivaryOn _).sum_smul_le_sum_comp_perm_smul fun _ _ ↦ mem_univ _ - -/-- **Equality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and -`g`, which antivary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` antivary -together. Stated by permuting the entries of `f`. -/ -theorem Antivary.sum_comp_perm_smul_eq_sum_smul_iff (hfg : Antivary f g) : - ∑ i, f (σ i) • g i = ∑ i, f i • g i ↔ Antivary (f ∘ σ) g := by - simp [(hfg.antivaryOn _).sum_comp_perm_smul_eq_sum_smul_iff fun _ _ ↦ mem_univ _] - -@[deprecated (since := "2024-06-25")] -alias Antivary.sum_smul_eq_sum_comp_perm_smul_iff := Antivary.sum_comp_perm_smul_eq_sum_smul_iff - /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of `f` and `g`, which antivary together, is strictly decreased by a permutation if and only if `f ∘ σ` and `g` do not antivary together. Stated by permuting the entries of `f`. -/ @@ -322,6 +344,7 @@ theorem Antivary.sum_smul_lt_sum_comp_perm_smul_iff (hfg : Antivary f g) : ∑ i, f i • g i < ∑ i, f (σ i) • g i ↔ ¬Antivary (f ∘ σ) g := by simp [(hfg.antivaryOn _).sum_smul_lt_sum_comp_perm_smul_iff fun _ _ ↦ mem_univ _] +end strict_inequality end SMul /-! @@ -330,87 +353,84 @@ end SMul Special cases of the above when scalar multiplication is actually multiplication. -/ - section Mul - - -variable [LinearOrderedRing α] {s : Finset ι} {σ : Perm ι} {f g : ι → α} +variable {s : Finset ι} {σ : Perm ι} {f g : ι → α} /-- **Rearrangement Inequality**: Pointwise multiplication of `f` and `g` is maximized when `f` and -`g` monovary together. Stated by permuting the entries of `g`. -/ -theorem MonovaryOn.sum_mul_comp_perm_le_sum_mul (hfg : MonovaryOn f g s) - (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g (σ i) ≤ ∑ i ∈ s, f i * g i := +`g` monovary together on `s`. Stated by permuting the entries of `g`. -/ +theorem MonovaryOn.sum_mul_comp_perm_le_sum_mul (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : + ∑ i ∈ s, f i * g (σ i) ≤ ∑ i ∈ s, f i * g i := hfg.sum_smul_comp_perm_le_sum_smul hσ /-- **Equality case of the Rearrangement Inequality**: Pointwise multiplication of `f` and `g`, -which monovary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` monovary -together. Stated by permuting the entries of `g`. -/ +which monovary together on `s`, is unchanged by a permutation if and only if `f` and `g ∘ σ` +monovary together on `s`. Stated by permuting the entries of `g`. -/ theorem MonovaryOn.sum_mul_comp_perm_eq_sum_mul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g (σ i) = ∑ i ∈ s, f i * g i ↔ MonovaryOn f (g ∘ σ) s := hfg.sum_smul_comp_perm_eq_sum_smul_iff hσ /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise scalar multiplication of -`f` and `g`, which monovary together, is strictly decreased by a permutation if and only if -`f` and `g ∘ σ` do not monovary together. Stated by permuting the entries of `g`. -/ +`f` and `g`, which monovary together on `s`, is strictly decreased by a permutation if and only if +`f` and `g ∘ σ` do not monovary together on `s`. Stated by permuting the entries of `g`. -/ theorem MonovaryOn.sum_mul_comp_perm_lt_sum_mul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i • g (σ i) < ∑ i ∈ s, f i • g i ↔ ¬MonovaryOn f (g ∘ σ) s := hfg.sum_smul_comp_perm_lt_sum_smul_iff hσ /-- **Rearrangement Inequality**: Pointwise multiplication of `f` and `g` is maximized when `f` and -`g` monovary together. Stated by permuting the entries of `f`. -/ +`g` monovary together on `s`. Stated by permuting the entries of `f`. -/ theorem MonovaryOn.sum_comp_perm_mul_le_sum_mul (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) * g i ≤ ∑ i ∈ s, f i * g i := hfg.sum_comp_perm_smul_le_sum_smul hσ /-- **Equality case of the Rearrangement Inequality**: Pointwise multiplication of `f` and `g`, -which monovary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` monovary -together. Stated by permuting the entries of `f`. -/ +which monovary together on `s`, is unchanged by a permutation if and only if `f ∘ σ` and `g` +monovary together on `s`. Stated by permuting the entries of `f`. -/ theorem MonovaryOn.sum_comp_perm_mul_eq_sum_mul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) * g i = ∑ i ∈ s, f i * g i ↔ MonovaryOn (f ∘ σ) g s := hfg.sum_comp_perm_smul_eq_sum_smul_iff hσ /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise multiplication of -`f` and `g`, which monovary together, is strictly decreased by a permutation if and only if -`f ∘ σ` and `g` do not monovary together. Stated by permuting the entries of `f`. -/ +`f` and `g`, which monovary together on `s`, is strictly decreased by a permutation if and only if +`f ∘ σ` and `g` do not monovary together on `s`. Stated by permuting the entries of `f`. -/ theorem MonovaryOn.sum_comp_perm_mul_lt_sum_mul_iff (hfg : MonovaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) * g i < ∑ i ∈ s, f i * g i ↔ ¬MonovaryOn (f ∘ σ) g s := hfg.sum_comp_perm_smul_lt_sum_smul_iff hσ /-- **Rearrangement Inequality**: Pointwise multiplication of `f` and `g` is minimized when `f` and -`g` antivary together. Stated by permuting the entries of `g`. -/ +`g` antivary together on `s`. Stated by permuting the entries of `g`. -/ theorem AntivaryOn.sum_mul_le_sum_mul_comp_perm (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g i ≤ ∑ i ∈ s, f i * g (σ i) := hfg.sum_smul_le_sum_smul_comp_perm hσ /-- **Equality case of the Rearrangement Inequality**: Pointwise multiplication of `f` and `g`, -which antivary together, is unchanged by a permutation if and only if `f` and `g ∘ σ` antivary -together. Stated by permuting the entries of `g`. -/ +which antivary together on `s`, is unchanged by a permutation if and only if `f` and `g ∘ σ` +antivary together on `s`. Stated by permuting the entries of `g`. -/ theorem AntivaryOn.sum_mul_eq_sum_mul_comp_perm_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g (σ i) = ∑ i ∈ s, f i * g i ↔ AntivaryOn f (g ∘ σ) s := hfg.sum_smul_comp_perm_eq_sum_smul_iff hσ /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise multiplication of -`f` and `g`, which antivary together, is strictly decreased by a permutation if and only if -`f` and `g ∘ σ` do not antivary together. Stated by permuting the entries of `g`. -/ +`f` and `g`, which antivary together on `s`, is strictly decreased by a permutation if and only if +`f` and `g ∘ σ` do not antivary together on `s`. Stated by permuting the entries of `g`. -/ theorem AntivaryOn.sum_mul_lt_sum_mul_comp_perm_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g i < ∑ i ∈ s, f i * g (σ i) ↔ ¬AntivaryOn f (g ∘ σ) s := hfg.sum_smul_lt_sum_smul_comp_perm_iff hσ /-- **Rearrangement Inequality**: Pointwise multiplication of `f` and `g` is minimized when `f` and -`g` antivary together. Stated by permuting the entries of `f`. -/ +`g` antivary together on `s`. Stated by permuting the entries of `f`. -/ theorem AntivaryOn.sum_mul_le_sum_comp_perm_mul (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g i ≤ ∑ i ∈ s, f (σ i) * g i := hfg.sum_smul_le_sum_comp_perm_smul hσ /-- **Equality case of the Rearrangement Inequality**: Pointwise multiplication of `f` and `g`, -which antivary together, is unchanged by a permutation if and only if `f ∘ σ` and `g` antivary -together. Stated by permuting the entries of `f`. -/ +which antivary together on `s`, is unchanged by a permutation if and only if `f ∘ σ` and `g` +antivary together on `s`. Stated by permuting the entries of `f`. -/ theorem AntivaryOn.sum_comp_perm_mul_eq_sum_mul_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f (σ i) * g i = ∑ i ∈ s, f i * g i ↔ AntivaryOn (f ∘ σ) g s := @@ -420,8 +440,8 @@ theorem AntivaryOn.sum_comp_perm_mul_eq_sum_mul_iff (hfg : AntivaryOn f g s) alias AntivaryOn.sum_mul_eq_sum_comp_perm_mul_iff := AntivaryOn.sum_comp_perm_mul_eq_sum_mul_iff /-- **Strict inequality case of the Rearrangement Inequality**: Pointwise multiplication of -`f` and `g`, which antivary together, is strictly decreased by a permutation if and only if -`f ∘ σ` and `g` do not antivary together. Stated by permuting the entries of `f`. -/ +`f` and `g`, which antivary together on `s`, is strictly decreased by a permutation if and only if +`f ∘ σ` and `g` do not antivary together on `s`. Stated by permuting the entries of `f`. -/ theorem AntivaryOn.sum_mul_lt_sum_comp_perm_mul_iff (hfg : AntivaryOn f g s) (hσ : {x | σ x ≠ x} ⊆ s) : ∑ i ∈ s, f i * g i < ∑ i ∈ s, f (σ i) * g i ↔ ¬AntivaryOn (f ∘ σ) g s := diff --git a/Mathlib/Algebra/Order/Sub/Defs.lean b/Mathlib/Algebra/Order/Sub/Defs.lean index 08e4b08db1ebc..55d910166545c 100644 --- a/Mathlib/Algebra/Order/Sub/Defs.lean +++ b/Mathlib/Algebra/Order/Sub/Defs.lean @@ -243,17 +243,39 @@ theorem tsub_right_comm : a - b - c = a - c - b := by namespace AddLECancellable +/-- See `AddLECancellable.tsub_eq_of_eq_add'` for a version assuming that `a = c + b` itself is +cancellable rather than `b`. -/ protected theorem tsub_eq_of_eq_add (hb : AddLECancellable b) (h : a = c + b) : a - b = c := le_antisymm (tsub_le_iff_right.mpr h.le) <| by rw [h] exact hb.le_add_tsub +/-- Weaker version of `AddLECancellable.tsub_eq_of_eq_add` assuming that `a = c + b` itself is +cancellable rather than `b`. -/ +protected lemma tsub_eq_of_eq_add' [CovariantClass α α (· + ·) (· ≤ ·)] (ha : AddLECancellable a) + (h : a = c + b) : a - b = c := (h ▸ ha).of_add_right.tsub_eq_of_eq_add h + +/-- See `AddLECancellable.eq_tsub_of_add_eq'` for a version assuming that `b = a + c` itself is +cancellable rather than `c`. -/ protected theorem eq_tsub_of_add_eq (hc : AddLECancellable c) (h : a + c = b) : a = b - c := (hc.tsub_eq_of_eq_add h.symm).symm +/-- Weaker version of `AddLECancellable.eq_tsub_of_add_eq` assuming that `b = a + c` itself is +cancellable rather than `c`. -/ +protected lemma eq_tsub_of_add_eq' [CovariantClass α α (· + ·) (· ≤ ·)] (hb : AddLECancellable b) + (h : a + c = b) : a = b - c := (hb.tsub_eq_of_eq_add' h.symm).symm + +/-- See `AddLECancellable.tsub_eq_of_eq_add_rev'` for a version assuming that `a = b + c` itself is +cancellable rather than `b`. -/ protected theorem tsub_eq_of_eq_add_rev (hb : AddLECancellable b) (h : a = b + c) : a - b = c := hb.tsub_eq_of_eq_add <| by rw [add_comm, h] +/-- Weaker version of `AddLECancellable.tsub_eq_of_eq_add_rev` assuming that `a = b + c` itself is +cancellable rather than `b`. -/ +protected lemma tsub_eq_of_eq_add_rev' [CovariantClass α α (· + ·) (· ≤ ·)] + (ha : AddLECancellable a) (h : a = b + c) : a - b = c := + ha.tsub_eq_of_eq_add' <| by rw [add_comm, h] + @[simp] protected theorem add_tsub_cancel_right (hb : AddLECancellable b) : a + b - b = a := hb.tsub_eq_of_eq_add <| by rw [add_comm] diff --git a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean index 1d9ab3832d4fb..3a3fe3f22832f 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean @@ -237,6 +237,12 @@ theorem natDegree_natCast (n : ℕ) : natDegree (n : R[X]) = 0 := by @[deprecated (since := "2024-04-17")] alias natDegree_nat_cast := natDegree_natCast +-- See note [no_index around OfNat.ofNat] +@[simp] +theorem natDegree_ofNat (n : ℕ) [Nat.AtLeastTwo n] : + natDegree (no_index (OfNat.ofNat n : R[X])) = 0 := + natDegree_natCast _ + theorem degree_natCast_le (n : ℕ) : degree (n : R[X]) ≤ 0 := degree_le_of_natDegree_le (by simp) @[deprecated (since := "2024-04-17")] diff --git a/Mathlib/Algebra/Polynomial/Splits.lean b/Mathlib/Algebra/Polynomial/Splits.lean index 955c85eca5bf7..3b0f053ebb0ea 100644 --- a/Mathlib/Algebra/Polynomial/Splits.lean +++ b/Mathlib/Algebra/Polynomial/Splits.lean @@ -437,7 +437,7 @@ theorem aeval_root_derivative_of_splits [Algebra K L] [DecidableEq L] {P : K[X]} rw [eval_multiset_prod_X_sub_C_derivative hr] /-- If `P` is a monic polynomial that splits, then `coeff P 0` equals the product of the roots. -/ -theorem prod_roots_eq_coeff_zero_of_monic_of_split {P : K[X]} (hmo : P.Monic) +theorem prod_roots_eq_coeff_zero_of_monic_of_splits {P : K[X]} (hmo : P.Monic) (hP : P.Splits (RingHom.id K)) : coeff P 0 = (-1) ^ P.natDegree * P.roots.prod := by nth_rw 1 [eq_prod_roots_of_monic_of_splits_id hmo hP] rw [coeff_zero_eq_eval_zero, eval_multiset_prod, Multiset.map_map] @@ -449,6 +449,9 @@ theorem prod_roots_eq_coeff_zero_of_monic_of_split {P : K[X]} (hmo : P.Monic) rw [neg_eq_neg_one_mul] simp only [splits_iff_card_roots.1 hP, neg_mul, one_mul, Multiset.prod_map_neg] +@[deprecated (since := "2024-10-01")] +alias prod_roots_eq_coeff_zero_of_monic_of_split := prod_roots_eq_coeff_zero_of_monic_of_splits + /-- If `P` is a monic polynomial that splits, then `P.nextCoeff` equals the sum of the roots. -/ theorem sum_roots_eq_nextCoeff_of_monic_of_split {P : K[X]} (hmo : P.Monic) (hP : P.Splits (RingHom.id K)) : P.nextCoeff = -P.roots.sum := by diff --git a/Mathlib/Algebra/TrivSqZeroExt.lean b/Mathlib/Algebra/TrivSqZeroExt.lean index 0e976107f9917..4eb552d919fc2 100644 --- a/Mathlib/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Algebra/TrivSqZeroExt.lean @@ -690,7 +690,9 @@ section Inv variable {R : Type u} {M : Type v} variable [Neg M] [Inv R] [SMul Rᵐᵒᵖ M] [SMul R M] -/-- Inversion of the trivial-square-zero extension, sending $r + m$ to $r^{-1} - r^{-1}mr^{-1}$. -/ +/-- Inversion of the trivial-square-zero extension, sending $r + m$ to $r^{-1} - r^{-1}mr^{-1}$. + +Strictly this is only a _two_-sided inverse when the left and right actions associate. -/ instance instInv : Inv (tsze R M) := ⟨fun b => (b.1⁻¹, -(b.1⁻¹ •> b.2 <• b.1⁻¹))⟩ @@ -702,6 +704,75 @@ instance instInv : Inv (tsze R M) := end Inv +/-! This section is heavily inspired by analogous results about matrices. -/ +section Invertible +variable {R : Type u} {M : Type v} +variable [AddCommGroup M] [Semiring R] [Module Rᵐᵒᵖ M] [Module R M] + +/-- `x.fst : R` is invertible when `x : tzre R M` is. -/ +abbrev invertibleFstOfInvertible (x : tsze R M) [Invertible x] : Invertible x.fst where + invOf := (⅟x).fst + invOf_mul_self := by rw [← fst_mul, invOf_mul_self, fst_one] + mul_invOf_self := by rw [← fst_mul, mul_invOf_self, fst_one] + +theorem fst_invOf (x : tsze R M) [Invertible x] [Invertible x.fst] : (⅟x).fst = ⅟(x.fst) := by + letI := invertibleFstOfInvertible x + convert (rfl : _ = ⅟ x.fst) + +theorem mul_left_eq_one (r : R) (x : tsze R M) (h : r * x.fst = 1) : + (inl r + inr (-((r •> x.snd) <• r))) * x = 1 := by + ext <;> dsimp + · rw [add_zero, h] + · rw [add_zero, zero_add, smul_neg, op_smul_op_smul, h, op_one, one_smul, + add_neg_cancel] + +theorem mul_right_eq_one (x : tsze R M) (r : R) (h : x.fst * r = 1) : + x * (inl r + inr (-(r •> (x.snd <• r)))) = 1 := by + ext <;> dsimp + · rw [add_zero, h] + · rw [add_zero, zero_add, smul_neg, smul_smul, h, one_smul, neg_add_cancel] + +variable [SMulCommClass R Rᵐᵒᵖ M] + +/-- `x : tzre R M` is invertible when `x.fst : R` is. -/ +abbrev invertibleOfInvertibleFst (x : tsze R M) [Invertible x.fst] : Invertible x where + invOf := (⅟x.fst, -(⅟x.fst •> x.snd <• ⅟x.fst)) + invOf_mul_self := by + convert mul_left_eq_one _ _ (invOf_mul_self x.fst) + ext <;> simp + mul_invOf_self := by + convert mul_right_eq_one _ _ (mul_invOf_self x.fst) + ext <;> simp [smul_comm] + +theorem snd_invOf (x : tsze R M) [Invertible x] [Invertible x.fst] : + (⅟x).snd = -(⅟x.fst •> x.snd <• ⅟x.fst) := by + letI := invertibleOfInvertibleFst x + convert congr_arg (TrivSqZeroExt.snd (R := R) (M := M)) (_ : _ = ⅟ x) + convert rfl + +/-- Together `TrivSqZeroExt.detInvertibleOfInvertible` and `TrivSqZeroExt.invertibleOfDetInvertible` +form an equivalence, although both sides of the equiv are subsingleton anyway. -/ +@[simps] +def invertibleEquivInvertibleFst (x : tsze R M) : Invertible x ≃ Invertible x.fst where + toFun _ := invertibleFstOfInvertible x + invFun _ := invertibleOfInvertibleFst x + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ + +/-- When lowered to a prop, `Matrix.invertibleEquivInvertibleFst` forms an `iff`. -/ +theorem isUnit_iff_isUnit_fst {x : tsze R M} : IsUnit x ↔ IsUnit x.fst := by + simp only [← nonempty_invertible_iff_isUnit, (invertibleEquivInvertibleFst x).nonempty_congr] + +@[simp] +theorem isUnit_inl_iff {r : R} : IsUnit (inl r : tsze R M) ↔ IsUnit r := by + rw [isUnit_iff_isUnit_fst, fst_inl] + +@[simp] +theorem isUnit_inr_iff {m : M} : IsUnit (inr m : tsze R M) ↔ Subsingleton R := by + simp_rw [isUnit_iff_isUnit_fst, fst_inr, isUnit_zero_iff, subsingleton_iff_zero_eq_one] + +end Invertible + section DivisionSemiring variable {R : Type u} {M : Type v} variable [DivisionSemiring R] [AddCommGroup M] [Module Rᵐᵒᵖ M] [Module R M] @@ -727,18 +798,19 @@ protected theorem inv_one : (1 : tsze R M)⁻¹ = (1 : tsze R M) := by rw [← inl_one, TrivSqZeroExt.inv_inl, inv_one] protected theorem inv_mul_cancel {x : tsze R M} (hx : fst x ≠ 0) : x⁻¹ * x = 1 := by - ext - · rw [fst_mul, fst_inv, inv_mul_cancel₀ hx, fst_one] - · rw [snd_mul, snd_inv, snd_one, smul_neg, op_smul_op_smul, inv_mul_cancel₀ hx, op_one, one_smul, - fst_inv, add_neg_cancel] + convert mul_left_eq_one _ _ (_root_.inv_mul_cancel₀ hx) using 2 + ext <;> simp variable [SMulCommClass R Rᵐᵒᵖ M] +@[simp] theorem invOf_eq_inv (x : tsze R M) [Invertible x] : ⅟x = x⁻¹ := by + letI := invertibleFstOfInvertible x + ext <;> simp [fst_invOf, snd_invOf] + protected theorem mul_inv_cancel {x : tsze R M} (hx : fst x ≠ 0) : x * x⁻¹ = 1 := by - ext - · rw [fst_mul, fst_inv, fst_one, mul_inv_cancel₀ hx] - · rw [snd_mul, snd_inv, snd_one, smul_neg, smul_comm, smul_smul, mul_inv_cancel₀ hx, one_smul, - fst_inv, neg_add_cancel] + have : Invertible x.fst := Units.invertible (.mk0 _ hx) + have := invertibleOfInvertibleFst x + rw [← invOf_eq_inv, mul_invOf_self] protected theorem mul_inv_rev (a b : tsze R M) : (a * b)⁻¹ = b⁻¹ * a⁻¹ := by @@ -763,6 +835,10 @@ protected theorem inv_inv {x : tsze R M} (hx : fst x ≠ 0) : x⁻¹⁻¹ = x := rw [fst_inv] apply inv_ne_zero hx +@[simp] +theorem isUnit_inv_iff {x : tsze R M} : IsUnit x⁻¹ ↔ IsUnit x := by + simp_rw [isUnit_iff_isUnit_fst, fst_inv, isUnit_iff_ne_zero, ne_eq, inv_eq_zero] + end DivisionSemiring section DivisionRing diff --git a/Mathlib/Analysis/Convex/Continuous.lean b/Mathlib/Analysis/Convex/Continuous.lean new file mode 100644 index 0000000000000..746b4cf9313a6 --- /dev/null +++ b/Mathlib/Analysis/Convex/Continuous.lean @@ -0,0 +1,232 @@ +/- +Copyright (c) 2023 Yaël Dillies, Zichen Wang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Zichen Wang +-/ +import Mathlib.Analysis.Convex.Normed + +/-! +# Convex functions are continuous + +This file proves that a convex function from a finite dimensional real normed space to `ℝ` is +continuous. +-/ + +open FiniteDimensional Metric Set List Bornology +open scoped Topology + +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] + {C : Set E} {f : E → ℝ} {x₀ : E} {ε r r' M : ℝ} + +lemma ConvexOn.lipschitzOnWith_of_abs_le (hf : ConvexOn ℝ (ball x₀ r) f) (hε : 0 < ε) + (hM : ∀ a, dist a x₀ < r → |f a| ≤ M) : + LipschitzOnWith (2 * M / ε).toNNReal f (ball x₀ (r - ε)) := by + set K := 2 * M / ε with hK + have oneside {x y : E} (hx : x ∈ ball x₀ (r - ε)) (hy : y ∈ ball x₀ (r - ε)) : + f x - f y ≤ K * ‖x - y‖ := by + obtain rfl | hxy := eq_or_ne x y + · simp + have hx₀r : ball x₀ (r - ε) ⊆ ball x₀ r := ball_subset_ball <| by linarith + have hx' : x ∈ ball x₀ r := hx₀r hx + have hy' : y ∈ ball x₀ r := hx₀r hy + let z := x + (ε / ‖x - y‖) • (x - y) + replace hxy : 0 < ‖x - y‖ := by rwa [norm_sub_pos_iff] + have hz : z ∈ ball x₀ r := mem_ball_iff_norm.2 <| by + calc + _ = ‖(x - x₀) + (ε / ‖x - y‖) • (x - y)‖ := by simp only [z, add_sub_right_comm] + _ ≤ ‖x - x₀‖ + ‖(ε / ‖x - y‖) • (x - y)‖ := norm_add_le .. + _ < r - ε + ε := + add_lt_add_of_lt_of_le (mem_ball_iff_norm.1 hx) <| by + simp [norm_smul, abs_of_nonneg, hε.le, hxy.ne'] + _ = r := by simp + let a := ε / (ε + ‖x - y‖) + let b := ‖x - y‖ / (ε + ‖x - y‖) + have hab : a + b = 1 := by field_simp [a, b] + have hxyz : x = a • y + b • z := by + calc + x = a • x + b • x := by rw [Convex.combo_self hab] + _ = a • y + b • z := by simp [z, a, b, smul_smul, hxy.ne', smul_sub]; abel + rw [hK, mul_comm, ← mul_div_assoc, le_div_iff₀' hε] + calc + ε * (f x - f y) ≤ ‖x - y‖ * (f z - f x) := by + rw [mul_sub, mul_sub, sub_le_sub_iff, ← add_mul] + have h := hf.2 hy' hz (by positivity) (by positivity) hab + field_simp [← hxyz, a, b, ← mul_div_right_comm] at h + rwa [← le_div_iff₀' (by positivity), add_comm (_ * _)] + _ ≤ _ := by + rw [sub_eq_add_neg (f _), two_mul] + gcongr + · exact (le_abs_self _).trans <| hM _ hz + · exact (neg_le_abs _).trans <| hM _ hx' + refine .of_dist_le' fun x hx y hy ↦ ?_ + simp_rw [dist_eq_norm_sub, Real.norm_eq_abs, abs_sub_le_iff] + exact ⟨oneside hx hy, norm_sub_rev x _ ▸ oneside hy hx⟩ + +lemma ConcaveOn.lipschitzOnWith_of_abs_le (hf : ConcaveOn ℝ (ball x₀ r) f) (hε : 0 < ε) + (hM : ∀ a, dist a x₀ < r → |f a| ≤ M) : + LipschitzOnWith (2 * M / ε).toNNReal f (ball x₀ (r - ε)) := by + simpa using hf.neg.lipschitzOnWith_of_abs_le hε <| by simpa using hM + +lemma ConvexOn.exists_lipschitzOnWith_of_isBounded (hf : ConvexOn ℝ (ball x₀ r) f) (hr : r' < r) + (hf' : IsBounded (f '' ball x₀ r)) : ∃ K, LipschitzOnWith K f (ball x₀ r') := by + rw [isBounded_iff_subset_ball 0] at hf' + simp only [Set.subset_def, mem_image, mem_ball, dist_zero_right, Real.norm_eq_abs, + forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] at hf' + obtain ⟨M, hM⟩ := hf' + rw [← sub_sub_cancel r r'] + exact ⟨_, hf.lipschitzOnWith_of_abs_le (sub_pos.2 hr) fun a ha ↦ (hM a ha).le⟩ + +lemma ConcaveOn.exists_lipschitzOnWith_of_isBounded (hf : ConcaveOn ℝ (ball x₀ r) f) (hr : r' < r) + (hf' : IsBounded (f '' ball x₀ r)) : ∃ K, LipschitzOnWith K f (ball x₀ r') := by + replace hf' : IsBounded ((-f) '' ball x₀ r) := by convert hf'.neg; ext; simp [neg_eq_iff_eq_neg] + simpa using hf.neg.exists_lipschitzOnWith_of_isBounded hr hf' + +lemma ConvexOn.isBoundedUnder_abs (hf : ConvexOn ℝ C f) {x₀ : E} (hC : C ∈ 𝓝 x₀) : + (𝓝 x₀).IsBoundedUnder (· ≤ ·) |f| ↔ (𝓝 x₀).IsBoundedUnder (· ≤ ·) f := by + refine ⟨fun h ↦ h.mono_le <| .of_forall fun x ↦ le_abs_self _, ?_⟩ + rintro ⟨r, hr⟩ + refine ⟨|r| + 2 * |f x₀|, ?_⟩ + have : (𝓝 x₀).Tendsto (fun y => 2 • x₀ - y) (𝓝 x₀) := + tendsto_nhds_nhds.2 (⟨·, ·, by simp [two_nsmul, dist_comm]⟩) + simp only [Filter.eventually_map, Pi.abs_apply, abs_le'] at hr ⊢ + filter_upwards [this.eventually_mem hC, hC, hr, this.eventually hr] with y hx hx' hfr hfr' + refine ⟨hfr.trans <| (le_abs_self _).trans <| by simp, ?_⟩ + rw [← sub_le_iff_le_add, neg_sub_comm, sub_le_iff_le_add', ← abs_two, ← abs_mul] + calc + -|2 * f x₀| ≤ 2 * f x₀ := neg_abs_le _ + _ ≤ f y + f (2 • x₀ - y) := by + have := hf.2 hx' hx (by positivity) (by positivity) (add_halves _) + simp only [one_div, ← Nat.cast_smul_eq_nsmul ℝ, Nat.cast_ofNat, smul_sub, ne_eq, + OfNat.ofNat_ne_zero, not_false_eq_true, inv_smul_smul₀, add_sub_cancel, smul_eq_mul] at this + cancel_denoms at this + rwa [← Nat.cast_two, Nat.cast_smul_eq_nsmul] at this + _ ≤ f y + |r| := by gcongr; exact hfr'.trans (le_abs_self _) + +lemma ConcaveOn.isBoundedUnder_abs (hf : ConcaveOn ℝ C f) {x₀ : E} (hC : C ∈ 𝓝 x₀) : + (𝓝 x₀).IsBoundedUnder (· ≤ ·) |f| ↔ (𝓝 x₀).IsBoundedUnder (· ≥ ·) f := by + simpa [Pi.neg_def, Pi.abs_def] using hf.neg.isBoundedUnder_abs hC + +lemma ConvexOn.continuousOn_tfae (hC : IsOpen C) (hC' : C.Nonempty) (hf : ConvexOn ℝ C f) : TFAE [ + LocallyLipschitzOn C f, + ContinuousOn f C, + ∃ x₀ ∈ C, ContinuousAt f x₀, + ∃ x₀ ∈ C, (𝓝 x₀).IsBoundedUnder (· ≤ ·) f, + ∀ ⦃x₀⦄, x₀ ∈ C → (𝓝 x₀).IsBoundedUnder (· ≤ ·) f, + ∀ ⦃x₀⦄, x₀ ∈ C → (𝓝 x₀).IsBoundedUnder (· ≤ ·) |f|] := by + tfae_have 1 → 2 + · exact LocallyLipschitzOn.continuousOn + tfae_have 2 → 3 + · obtain ⟨x₀, hx₀⟩ := hC' + exact fun h ↦ ⟨x₀, hx₀, h.continuousAt <| hC.mem_nhds hx₀⟩ + tfae_have 3 → 4 + · rintro ⟨x₀, hx₀, h⟩ + exact ⟨x₀, hx₀, f x₀ + 1, by simpa using h.eventually (eventually_le_nhds (by simp))⟩ + tfae_have 4 → 5 + · rintro ⟨x₀, hx₀, r, hr⟩ x hx + have : ∀ᶠ δ in 𝓝 (0 : ℝ), (1 - δ)⁻¹ • x - (δ / (1 - δ)) • x₀ ∈ C := by + have h : ContinuousAt (fun δ : ℝ ↦ (1 - δ)⁻¹ • x - (δ / (1 - δ)) • x₀) 0 := by + fun_prop (disch := norm_num) + exact h (by simpa using hC.mem_nhds hx) + obtain ⟨δ, hδ₀, hy, hδ₁⟩ := (this.and <| eventually_lt_nhds zero_lt_one).exists_gt + set y := (1 - δ)⁻¹ • x - (δ / (1 - δ)) • x₀ + refine ⟨max r (f y), ?_⟩ + simp only [Filter.eventually_map, Pi.abs_apply] at hr ⊢ + obtain ⟨ε, hε, hr⟩ := Metric.eventually_nhds_iff.1 <| hr.and (hC.eventually_mem hx₀) + refine Metric.eventually_nhds_iff.2 ⟨ε * δ, by positivity, fun z hz ↦ ?_⟩ + have hx₀' : δ⁻¹ • (x - y) + y = x₀ := MulAction.injective₀ (sub_ne_zero.2 hδ₁.ne') <| by + simp [y, smul_sub, smul_smul, hδ₀.ne', div_eq_mul_inv, sub_ne_zero.2 hδ₁.ne', mul_left_comm, + sub_mul, sub_smul] + let w := δ⁻¹ • (z - y) + y + have hwyz : δ • w + (1 - δ) • y = z := by simp [w, hδ₀.ne', sub_smul] + have hw : dist w x₀ < ε := by + simpa [w, ← hx₀', dist_smul₀, abs_of_nonneg, hδ₀.le, inv_mul_lt_iff', hδ₀] + calc + f z ≤ max (f w) (f y) := + hf.le_max_of_mem_segment (hr hw).2 hy ⟨_, _, hδ₀.le, sub_nonneg.2 hδ₁.le, by simp, hwyz⟩ + _ ≤ max r (f y) := by gcongr; exact (hr hw).1 + tfae_have 6 ↔ 5 + · exact forall₂_congr fun x₀ hx₀ ↦ hf.isBoundedUnder_abs (hC.mem_nhds hx₀) + tfae_have 6 → 1 + · rintro h x hx + obtain ⟨r, hr⟩ := h hx + obtain ⟨ε, hε, hεD⟩ := Metric.mem_nhds_iff.1 <| Filter.inter_mem (hC.mem_nhds hx) hr + simp only [preimage_setOf_eq, Pi.abs_apply, subset_inter_iff, hC.nhdsWithin_eq hx] at hεD ⊢ + obtain ⟨K, hK⟩ := exists_lipschitzOnWith_of_isBounded (hf.subset hεD.1 (convex_ball ..)) + (half_lt_self hε) <| isBounded_iff_forall_norm_le.2 ⟨r, by simpa using hεD.2⟩ + exact ⟨K, _, ball_mem_nhds _ (by simpa), hK⟩ + tfae_finish + +lemma ConcaveOn.continuousOn_tfae (hC : IsOpen C) (hC' : C.Nonempty) (hf : ConcaveOn ℝ C f) : TFAE [ + LocallyLipschitzOn C f, + ContinuousOn f C, + ∃ x₀ ∈ C, ContinuousAt f x₀, + ∃ x₀ ∈ C, (𝓝 x₀).IsBoundedUnder (· ≥ ·) f, + ∀ ⦃x₀⦄, x₀ ∈ C → (𝓝 x₀).IsBoundedUnder (· ≥ ·) f, + ∀ ⦃x₀⦄, x₀ ∈ C → (𝓝 x₀).IsBoundedUnder (· ≤ ·) |f|] := by + have := hf.neg.continuousOn_tfae hC hC' + simp at this + convert this using 8 <;> exact (Equiv.neg ℝ).exists_congr (by simp) + +lemma ConvexOn.locallyLipschitzOn_iff_continuousOn (hC : IsOpen C) (hf : ConvexOn ℝ C f) : + LocallyLipschitzOn C f ↔ ContinuousOn f C := by + obtain rfl | hC' := C.eq_empty_or_nonempty + · simp + · exact (hf.continuousOn_tfae hC hC').out 0 1 + +lemma ConcaveOn.locallyLipschitzOn_iff_continuousOn (hC : IsOpen C) (hf : ConcaveOn ℝ C f) : + LocallyLipschitzOn C f ↔ ContinuousOn f C := by + simpa using hf.neg.locallyLipschitzOn_iff_continuousOn hC + +variable [FiniteDimensional ℝ E] + +protected lemma ConvexOn.locallyLipschitzOn (hC : IsOpen C) (hf : ConvexOn ℝ C f) : + LocallyLipschitzOn C f := by + obtain rfl | ⟨x₀, hx₀⟩ := C.eq_empty_or_nonempty + · simp + · obtain ⟨b, hx₀b, hbC⟩ := exists_mem_interior_convexHull_affineBasis (hC.mem_nhds hx₀) + refine ((hf.continuousOn_tfae hC ⟨x₀, hx₀⟩).out 3 0).mp ?_ + refine ⟨x₀, hx₀, BddAbove.isBoundedUnder (IsOpen.mem_nhds isOpen_interior hx₀b) ?_⟩ + exact (hf.bddAbove_convexHull ((subset_convexHull ..).trans hbC) + ((finite_range _).image _).bddAbove).mono (by gcongr; exact interior_subset) + +protected lemma ConcaveOn.locallyLipschitzOn (hC : IsOpen C) (hf : ConcaveOn ℝ C f) : + LocallyLipschitzOn C f := by simpa using hf.neg.locallyLipschitzOn hC + +protected lemma ConvexOn.continuousOn (hC : IsOpen C) (hf : ConvexOn ℝ C f) : + ContinuousOn f C := (hf.locallyLipschitzOn hC).continuousOn + +protected lemma ConcaveOn.continuousOn (hC : IsOpen C) (hf : ConcaveOn ℝ C f) : + ContinuousOn f C := (hf.locallyLipschitzOn hC).continuousOn + +lemma ConvexOn.locallyLipschitzOn_interior (hf : ConvexOn ℝ C f) : + LocallyLipschitzOn (interior C) f := + (hf.subset interior_subset hf.1.interior).locallyLipschitzOn isOpen_interior + +lemma ConcaveOn.locallyLipschitzOn_interior (hf : ConcaveOn ℝ C f) : + LocallyLipschitzOn (interior C) f := + (hf.subset interior_subset hf.1.interior).locallyLipschitzOn isOpen_interior + +lemma ConvexOn.continuousOn_interior (hf : ConvexOn ℝ C f) : ContinuousOn f (interior C) := + hf.locallyLipschitzOn_interior.continuousOn + +lemma ConcaveOn.continuousOn_interior (hf : ConcaveOn ℝ C f) : ContinuousOn f (interior C) := + hf.locallyLipschitzOn_interior.continuousOn + +protected lemma ConvexOn.locallyLipschitz (hf : ConvexOn ℝ univ f) : LocallyLipschitz f := by + simpa using hf.locallyLipschitzOn_interior + +protected lemma ConcaveOn.locallyLipschitz (hf : ConcaveOn ℝ univ f) : LocallyLipschitz f := by + simpa using hf.locallyLipschitzOn_interior + +-- Commented out since `intrinsicInterior` is not imported (but should be once these are proved) +-- proof_wanted ConvexOn.locallyLipschitzOn_intrinsicInterior (hf : ConvexOn ℝ C f) : +-- ContinuousOn f (intrinsicInterior ℝ C) + +-- proof_wanted ConcaveOn.locallyLipschitzOn_intrinsicInterior (hf : ConcaveOn ℝ C f) : +-- ContinuousOn f (intrinsicInterior ℝ C) + +-- proof_wanted ConvexOn.continuousOn_intrinsicInterior (hf : ConvexOn ℝ C f) : +-- ContinuousOn f (intrinsicInterior ℝ C) + +-- proof_wanted ConcaveOn.continuousOn_intrinsicInterior (hf : ConcaveOn ℝ C f) : +-- ContinuousOn f (intrinsicInterior ℝ C) diff --git a/Mathlib/Analysis/Convex/Intrinsic.lean b/Mathlib/Analysis/Convex/Intrinsic.lean index b244bd60d24fb..82cc1b1b0da2d 100644 --- a/Mathlib/Analysis/Convex/Intrinsic.lean +++ b/Mathlib/Analysis/Convex/Intrinsic.lean @@ -96,24 +96,6 @@ theorem intrinsicFrontier_subset_intrinsicClosure : intrinsicFrontier 𝕜 s ⊆ theorem subset_intrinsicClosure : s ⊆ intrinsicClosure 𝕜 s := fun x hx => ⟨⟨x, subset_affineSpan _ _ hx⟩, subset_closure hx, rfl⟩ -lemma intrinsicInterior_eq_interior_of_span (hs : affineSpan 𝕜 s = ⊤) : - intrinsicInterior 𝕜 s = interior s := by - set f : affineSpan 𝕜 s ≃ₜ P := .trans (.setCongr (congr_arg SetLike.coe hs)) (.Set.univ _) - change f '' interior (f ⁻¹' s) = interior s - rw [f.image_interior, f.image_preimage] - -lemma intrinsicFrontier_eq_frontier_of_span (hs : affineSpan 𝕜 s = ⊤) : - intrinsicFrontier 𝕜 s = frontier s := by - set f : affineSpan 𝕜 s ≃ₜ P := .trans (.setCongr (congr_arg SetLike.coe hs)) (.Set.univ _) - change f '' frontier (f ⁻¹' s) = frontier s - rw [f.image_frontier, f.image_preimage] - -lemma intrinsicClosure_eq_closure_of_span (hs : affineSpan 𝕜 s = ⊤) : - intrinsicClosure 𝕜 s = closure s := by - set f : affineSpan 𝕜 s ≃ₜ P := .trans (.setCongr (congr_arg SetLike.coe hs)) (.Set.univ _) - change f '' closure (f ⁻¹' s) = closure s - rw [f.image_closure, f.image_preimage] - @[simp] theorem intrinsicInterior_empty : intrinsicInterior 𝕜 (∅ : Set P) = ∅ := by simp [intrinsicInterior] @@ -123,15 +105,6 @@ theorem intrinsicFrontier_empty : intrinsicFrontier 𝕜 (∅ : Set P) = ∅ := @[simp] theorem intrinsicClosure_empty : intrinsicClosure 𝕜 (∅ : Set P) = ∅ := by simp [intrinsicClosure] -@[simp] lemma intrinsicInterior_univ : intrinsicInterior 𝕜 (univ : Set P) = univ := by - simp [intrinsicInterior] - -@[simp] lemma intrinsicFrontier_univ : intrinsicFrontier 𝕜 (univ : Set P) = ∅ := by - simp [intrinsicFrontier] - -@[simp] lemma intrinsicClosure_univ : intrinsicClosure 𝕜 (univ : Set P) = univ := by - simp [intrinsicClosure] - @[simp] theorem intrinsicClosure_nonempty : (intrinsicClosure 𝕜 s).Nonempty ↔ s.Nonempty := ⟨by simp_rw [nonempty_iff_ne_empty]; rintro h rfl; exact h intrinsicClosure_empty, diff --git a/Mathlib/Analysis/MeanInequalitiesPow.lean b/Mathlib/Analysis/MeanInequalitiesPow.lean index daa489d39347c..fdee6d8fdbe46 100644 --- a/Mathlib/Analysis/MeanInequalitiesPow.lean +++ b/Mathlib/Analysis/MeanInequalitiesPow.lean @@ -61,23 +61,6 @@ theorem pow_arith_mean_le_arith_mean_pow_of_even (w z : ι → ℝ) (hw : ∀ i (∑ i ∈ s, w i * z i) ^ n ≤ ∑ i ∈ s, w i * z i ^ n := hn.convexOn_pow.map_sum_le hw hw' fun _ _ => Set.mem_univ _ -/-- Specific case of Jensen's inequality for sums of powers -/ -theorem pow_sum_div_card_le_sum_pow {f : ι → ℝ} (n : ℕ) (hf : ∀ a ∈ s, 0 ≤ f a) : - (∑ x ∈ s, f x) ^ (n + 1) / (s.card : ℝ) ^ n ≤ ∑ x ∈ s, f x ^ (n + 1) := by - rcases s.eq_empty_or_nonempty with (rfl | hs) - · simp_rw [Finset.sum_empty, zero_pow n.succ_ne_zero, zero_div]; rfl - · have hs0 : 0 < (s.card : ℝ) := Nat.cast_pos.2 hs.card_pos - suffices (∑ x ∈ s, f x / s.card) ^ (n + 1) ≤ ∑ x ∈ s, f x ^ (n + 1) / s.card by - rwa [← Finset.sum_div, ← Finset.sum_div, div_pow, pow_succ (s.card : ℝ), ← div_div, - div_le_iff₀ hs0, div_mul, div_self hs0.ne', div_one] at this - have := - @ConvexOn.map_sum_le ℝ ℝ ℝ ι _ _ _ _ _ _ (Set.Ici 0) (fun x => x ^ (n + 1)) s - (fun _ => 1 / s.card) ((↑) ∘ f) (convexOn_pow (n + 1)) ?_ ?_ fun i hi => - Set.mem_Ici.2 (hf i hi) - · simpa only [inv_mul_eq_div, one_div, Algebra.id.smul_eq_mul] using this - · simp only [one_div, inv_nonneg, Nat.cast_nonneg, imp_true_iff] - · simpa only [one_div, Finset.sum_const, nsmul_eq_mul] using mul_inv_cancel₀ hs0.ne' - theorem zpow_arith_mean_le_arith_mean_zpow (w z : ι → ℝ) (hw : ∀ i ∈ s, 0 ≤ w i) (hw' : ∑ i ∈ s, w i = 1) (hz : ∀ i ∈ s, 0 < z i) (m : ℤ) : (∑ i ∈ s, w i * z i) ^ m ≤ ∑ i ∈ s, w i * z i ^ m := @@ -111,11 +94,6 @@ theorem pow_arith_mean_le_arith_mean_pow (w z : ι → ℝ≥0) (hw' : ∑ i ∈ Real.pow_arith_mean_le_arith_mean_pow s _ _ (fun i _ => (w i).coe_nonneg) (mod_cast hw') (fun i _ => (z i).coe_nonneg) n -theorem pow_sum_div_card_le_sum_pow (f : ι → ℝ≥0) (n : ℕ) : - (∑ x ∈ s, f x) ^ (n + 1) / (s.card : ℝ) ^ n ≤ ∑ x ∈ s, f x ^ (n + 1) := by - simpa only [← NNReal.coe_le_coe, NNReal.coe_sum, Nonneg.coe_div, NNReal.coe_pow] using - @Real.pow_sum_div_card_le_sum_pow ι s (((↑) : ℝ≥0 → ℝ) ∘ f) n fun _ _ => NNReal.coe_nonneg _ - /-- Weighted generalized mean inequality, version for sums over finite sets, with `ℝ≥0`-valued functions and real exponents. -/ theorem rpow_arith_mean_le_arith_mean_rpow (w z : ι → ℝ≥0) (hw' : ∑ i ∈ s, w i = 1) {p : ℝ} diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index bde983756cf20..1eba1978859f4 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -1129,7 +1129,7 @@ theorem nnnorm_prod_le (s : Finset ι) (f : ι → E) : ‖∏ a ∈ s, f a‖ @[to_additive] theorem nnnorm_prod_le_of_le (s : Finset ι) {f : ι → E} {n : ι → ℝ≥0} (h : ∀ b ∈ s, ‖f b‖₊ ≤ n b) : ‖∏ b ∈ s, f b‖₊ ≤ ∑ b ∈ s, n b := - (norm_prod_le_of_le s h).trans_eq NNReal.coe_sum.symm + (norm_prod_le_of_le s h).trans_eq (NNReal.coe_sum ..).symm namespace Real diff --git a/Mathlib/Analysis/Normed/Group/Constructions.lean b/Mathlib/Analysis/Normed/Group/Constructions.lean index a43e9e963c9ee..c9ee9cc5d178f 100644 --- a/Mathlib/Analysis/Normed/Group/Constructions.lean +++ b/Mathlib/Analysis/Normed/Group/Constructions.lean @@ -362,7 +362,7 @@ lemma Pi.sum_norm_apply_le_norm' : ∑ i, ‖f i‖ ≤ Fintype.card ι • ‖f @[to_additive Pi.sum_nnnorm_apply_le_nnnorm "The $L^1$ norm is less than the $L^\\infty$ norm scaled by the cardinality."] lemma Pi.sum_nnnorm_apply_le_nnnorm' : ∑ i, ‖f i‖₊ ≤ Fintype.card ι • ‖f‖₊ := - NNReal.coe_sum.trans_le <| Pi.sum_norm_apply_le_norm' _ + (NNReal.coe_sum ..).trans_le <| Pi.sum_norm_apply_le_norm' _ end SeminormedGroup diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 869aa9fe0a463..82077c49fead5 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ import Mathlib.Algebra.Algebra.Field +import Mathlib.Algebra.Order.BigOperators.Expect import Mathlib.Algebra.Order.Star.Basic import Mathlib.Analysis.CStarAlgebra.Basic import Mathlib.Analysis.Normed.Operator.ContinuousLinearMap @@ -39,7 +40,7 @@ their counterparts in `Mathlib/Analysis/Complex/Basic.lean` (which causes linter A few lemmas requiring heavier imports are in `Mathlib/Data/RCLike/Lemmas.lean`. -/ -open scoped ComplexConjugate +open scoped BigOperators ComplexConjugate section @@ -234,6 +235,10 @@ theorem norm_ofReal (r : ℝ) : ‖(r : K)‖ = |r| := instance (priority := 100) charZero_rclike : CharZero K := (RingHom.charZero_iff (algebraMap ℝ K).injective).1 inferInstance +@[rclike_simps, norm_cast] +lemma ofReal_expect {α : Type*} (s : Finset α) (f : α → ℝ) : 𝔼 i ∈ s, f i = 𝔼 i ∈ s, (f i : K) := + map_expect (algebraMap ..) .. + /-! ### The imaginary unit, `I` -/ /-- The imaginary unit. -/ @@ -605,13 +610,23 @@ variable (K) in lemma nnnorm_nsmul [NormedAddCommGroup E] [NormedSpace K E] (n : ℕ) (x : E) : ‖n • x‖₊ = n • ‖x‖₊ := by simpa [Nat.cast_smul_eq_nsmul] using nnnorm_smul (n : K) x +section NormedField +variable [NormedField E] [CharZero E] [NormedSpace K E] +include K + variable (K) in -lemma norm_nnqsmul [NormedField E] [CharZero E] [NormedSpace K E] (q : ℚ≥0) (x : E) : - ‖q • x‖ = q • ‖x‖ := by simpa [NNRat.cast_smul_eq_nnqsmul] using norm_smul (q : K) x +lemma norm_nnqsmul (q : ℚ≥0) (x : E) : ‖q • x‖ = q • ‖x‖ := by + simpa [NNRat.cast_smul_eq_nnqsmul] using norm_smul (q : K) x variable (K) in -lemma nnnorm_nnqsmul [NormedField E] [CharZero E] [NormedSpace K E] (q : ℚ≥0) (x : E) : - ‖q • x‖₊ = q • ‖x‖₊ := by simpa [NNRat.cast_smul_eq_nnqsmul] using nnnorm_smul (q : K) x +lemma nnnorm_nnqsmul (q : ℚ≥0) (x : E) : ‖q • x‖₊ = q • ‖x‖₊ := by + simpa [NNRat.cast_smul_eq_nnqsmul] using nnnorm_smul (q : K) x + +@[bound] +lemma norm_expect_le {ι : Type*} {s : Finset ι} {f : ι → E} : ‖𝔼 i ∈ s, f i‖ ≤ 𝔼 i ∈ s, ‖f i‖ := + Finset.le_expect_of_subadditive norm_zero norm_add_le fun _ _ ↦ by rw [norm_nnqsmul K] + +end NormedField theorem mul_self_norm (z : K) : ‖z‖ * ‖z‖ = normSq z := by rw [normSq_eq_def', sq] diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean index c12393d6b86df..86b3f777ce844 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean @@ -553,8 +553,7 @@ lemma integral_rpow_mul_exp_neg_mul_Ioi {a r : ℝ} (ha : 0 < a) (hr : 0 < r) : convert integral_cpow_mul_exp_neg_mul_Ioi (by rwa [ofReal_re] : 0 < (a : ℂ).re) hr refine _root_.integral_ofReal.symm.trans <| setIntegral_congr measurableSet_Ioi (fun t ht ↦ ?_) norm_cast - rw [← ofReal_cpow (le_of_lt ht), RCLike.ofReal_mul] - rfl + simp_rw [← ofReal_cpow ht.le, RCLike.ofReal_mul, coe_algebraMap] open Lean.Meta Qq Mathlib.Meta.Positivity in /-- The `positivity` extension which identifies expressions of the form `Gamma a`. -/ diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean index e0595d589e11d..e8e17232197f1 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean @@ -173,11 +173,9 @@ theorem lt_tan {x : ℝ} (h1 : 0 < x) (h2 : x < π / 2) : x < tan x := by let U := Ico 0 (π / 2) have intU : interior U = Ioo 0 (π / 2) := interior_Ico have half_pi_pos : 0 < π / 2 := div_pos pi_pos two_pos - have cos_pos : ∀ {y : ℝ}, y ∈ U → 0 < cos y := by - intro y hy + have cos_pos {y : ℝ} (hy : y ∈ U) : 0 < cos y := by exact cos_pos_of_mem_Ioo (Ico_subset_Ioo_left (neg_lt_zero.mpr half_pi_pos) hy) - have sin_pos : ∀ {y : ℝ}, y ∈ interior U → 0 < sin y := by - intro y hy + have sin_pos {y : ℝ} (hy : y ∈ interior U) : 0 < sin y := by rw [intU] at hy exact sin_pos_of_mem_Ioo (Ioo_subset_Ioo_right (div_le_self pi_pos.le one_le_two) hy) have tan_cts_U : ContinuousOn tan U := by @@ -186,8 +184,7 @@ theorem lt_tan {x : ℝ} (h1 : 0 < x) (h2 : x < π / 2) : x < tan x := by simp only [mem_setOf_eq] exact (cos_pos hz).ne' have tan_minus_id_cts : ContinuousOn (fun y : ℝ => tan y - y) U := tan_cts_U.sub continuousOn_id - have deriv_pos : ∀ y : ℝ, y ∈ interior U → 0 < deriv (fun y' : ℝ => tan y' - y') y := by - intro y hy + have deriv_pos (y : ℝ) (hy : y ∈ interior U) : 0 < deriv (fun y' : ℝ => tan y' - y') y := by have := cos_pos (interior_subset hy) simp only [deriv_tan_sub_id y this.ne', one_div, gt_iff_lt, sub_pos] norm_cast diff --git a/Mathlib/Combinatorics/Additive/FreimanHom.lean b/Mathlib/Combinatorics/Additive/FreimanHom.lean index dee5a50bf940d..01cf0ef4b01c3 100644 --- a/Mathlib/Combinatorics/Additive/FreimanHom.lean +++ b/Mathlib/Combinatorics/Additive/FreimanHom.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.Order.BigOperators.Group.Multiset +import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Data.ZMod.Defs /-! diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 6483469eea3bc..bc2017ba047a1 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -442,6 +442,15 @@ lemma re_ofNat (n : ℕ) [n.AtLeastTwo] : (no_index (OfNat.ofNat n) : ℂ).re = @[simp, norm_cast] lemma ratCast_re (q : ℚ) : (q : ℂ).re = q := rfl @[simp, norm_cast] lemma ratCast_im (q : ℚ) : (q : ℂ).im = 0 := rfl +lemma re_nsmul (n : ℕ) (z : ℂ) : (n • z).re = n • z.re := smul_re .. +lemma im_nsmul (n : ℕ) (z : ℂ) : (n • z).im = n • z.im := smul_im .. +lemma re_zsmul (n : ℤ) (z : ℂ) : (n • z).re = n • z.re := smul_re .. +lemma im_zsmul (n : ℤ) (z : ℂ) : (n • z).im = n • z.im := smul_im .. +@[simp] lemma re_nnqsmul (q : ℚ≥0) (z : ℂ) : (q • z).re = q • z.re := smul_re .. +@[simp] lemma im_nnqsmul (q : ℚ≥0) (z : ℂ) : (q • z).im = q • z.im := smul_im .. +@[simp] lemma re_qsmul (q : ℚ) (z : ℂ) : (q • z).re = q • z.re := smul_re .. +@[simp] lemma im_qsmul (q : ℚ) (z : ℂ) : (q • z).im = q • z.im := smul_im .. + @[deprecated (since := "2024-04-17")] alias rat_cast_im := ratCast_im @@ -628,6 +637,28 @@ def ofReal : ℝ →+* ℂ where theorem ofReal_eq_coe (r : ℝ) : ofReal r = r := rfl +variable {α : Type*} + +@[simp] lemma ofReal_comp_add (f g : α → ℝ) : ofReal' ∘ (f + g) = ofReal' ∘ f + ofReal' ∘ g := + map_comp_add ofReal .. + +@[simp] lemma ofReal_comp_sub (f g : α → ℝ) : ofReal' ∘ (f - g) = ofReal' ∘ f - ofReal' ∘ g := + map_comp_sub ofReal .. + +@[simp] lemma ofReal_comp_neg (f : α → ℝ) : ofReal' ∘ (-f) = -(ofReal' ∘ f) := map_comp_neg ofReal _ + +lemma ofReal_comp_nsmul (n : ℕ) (f : α → ℝ) : ofReal' ∘ (n • f) = n • (ofReal' ∘ f) := + map_comp_nsmul ofReal .. + +lemma ofReal_comp_zsmul (n : ℤ) (f : α → ℝ) : ofReal' ∘ (n • f) = n • (ofReal' ∘ f) := + map_comp_zsmul ofReal .. + +@[simp] lemma ofReal_comp_mul (f g : α → ℝ) : ofReal' ∘ (f * g) = ofReal' ∘ f * ofReal' ∘ g := + map_comp_mul ofReal .. + +@[simp] lemma ofReal_comp_pow (f : α → ℝ) (n : ℕ) : ofReal' ∘ (f ^ n) = (ofReal' ∘ f) ^ n := + map_comp_pow ofReal .. + @[simp] theorem I_sq : I ^ 2 = -1 := by rw [sq, I_mul_I] diff --git a/Mathlib/Data/Complex/BigOperators.lean b/Mathlib/Data/Complex/BigOperators.lean index 43c85d7750bfa..cd9cb5f344fa0 100644 --- a/Mathlib/Data/Complex/BigOperators.lean +++ b/Mathlib/Data/Complex/BigOperators.lean @@ -3,14 +3,15 @@ Copyright (c) 2017 Kevin Buzzard. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Buzzard, Mario Carneiro -/ -import Mathlib.Algebra.BigOperators.Group.Finset +import Mathlib.Algebra.BigOperators.Expect import Mathlib.Data.Complex.Basic /-! # Finite sums and products of complex numbers - -/ +open scoped BigOperators + namespace Complex variable {α : Type*} (s : Finset α) @@ -23,12 +24,24 @@ theorem ofReal_prod (f : α → ℝ) : ((∏ i ∈ s, f i : ℝ) : ℂ) = ∏ i theorem ofReal_sum (f : α → ℝ) : ((∑ i ∈ s, f i : ℝ) : ℂ) = ∑ i ∈ s, (f i : ℂ) := map_sum ofReal _ _ +@[simp, norm_cast] +lemma ofReal_expect (f : α → ℝ) : (𝔼 i ∈ s, f i : ℝ) = 𝔼 i ∈ s, (f i : ℂ) := + map_expect ofReal .. + @[simp] theorem re_sum (f : α → ℂ) : (∑ i ∈ s, f i).re = ∑ i ∈ s, (f i).re := map_sum reAddGroupHom f s +@[simp] +lemma re_expect (f : α → ℂ) : (𝔼 i ∈ s, f i).re = 𝔼 i ∈ s, (f i).re := + map_expect (LinearMap.mk reAddGroupHom.toAddHom (by simp)) f s + @[simp] theorem im_sum (f : α → ℂ) : (∑ i ∈ s, f i).im = ∑ i ∈ s, (f i).im := map_sum imAddGroupHom f s +@[simp] +lemma im_expect (f : α → ℂ) : (𝔼 i ∈ s, f i).im = 𝔼 i ∈ s, (f i).im := + map_expect (LinearMap.mk imAddGroupHom.toAddHom (by simp)) f s + end Complex diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index d0e3217ac181c..1be528c876575 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -438,9 +438,7 @@ protected theorem half_lt_self (hz : a ≠ 0) (ht : a ≠ ∞) : a / 2 < a := by protected theorem half_le_self : a / 2 ≤ a := le_add_self.trans_eq <| ENNReal.add_halves _ -theorem sub_half (h : a ≠ ∞) : a - a / 2 = a / 2 := by - lift a to ℝ≥0 using h - exact sub_eq_of_add_eq (mul_ne_top coe_ne_top <| by simp) (ENNReal.add_halves a) +theorem sub_half (h : a ≠ ∞) : a - a / 2 = a / 2 := ENNReal.sub_eq_of_eq_add' h a.add_halves.symm @[simp] theorem one_sub_inv_two : (1 : ℝ≥0∞) - 2⁻¹ = 2⁻¹ := by diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index 9f0b1dbbbddc3..bf0335bc473f4 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -314,15 +314,37 @@ theorem natCast_sub (m n : ℕ) : ↑(m - n) = (m - n : ℝ≥0∞) := by @[deprecated (since := "2024-04-17")] alias nat_cast_sub := natCast_sub +/-- See `ENNReal.sub_eq_of_eq_add'` for a version assuming that `a = c + b` itself is finite rather +than `b`. -/ protected theorem sub_eq_of_eq_add (hb : b ≠ ∞) : a = c + b → a - b = c := (cancel_of_ne hb).tsub_eq_of_eq_add +/-- Weaker version of `ENNReal.sub_eq_of_eq_add` assuming that `a = c + b` itself is finite rather +han `b`. -/ +protected lemma sub_eq_of_eq_add' (ha : a ≠ ∞) : a = c + b → a - b = c := + (cancel_of_ne ha).tsub_eq_of_eq_add' + +/-- See `ENNReal.eq_sub_of_add_eq'` for a version assuming that `b = a + c` itself is finite rather +than `c`. -/ protected theorem eq_sub_of_add_eq (hc : c ≠ ∞) : a + c = b → a = b - c := (cancel_of_ne hc).eq_tsub_of_add_eq +/-- Weaker version of `ENNReal.eq_sub_of_add_eq` assuming that `b = a + c` itself is finite rather +than `c`. -/ +protected lemma eq_sub_of_add_eq' (hb : b ≠ ∞) : a + c = b → a = b - c := + (cancel_of_ne hb).eq_tsub_of_add_eq' + +/-- See `ENNReal.sub_eq_of_eq_add_rev'` for a version assuming that `a = b + c` itself is finite +rather than `b`. -/ protected theorem sub_eq_of_eq_add_rev (hb : b ≠ ∞) : a = b + c → a - b = c := (cancel_of_ne hb).tsub_eq_of_eq_add_rev +/-- Weaker version of `ENNReal.sub_eq_of_eq_add_rev` assuming that `a = b + c` itself is finite +rather than `b`. -/ +protected lemma sub_eq_of_eq_add_rev' (ha : a ≠ ∞) : a = b + c → a - b = c := + (cancel_of_ne ha).tsub_eq_of_eq_add_rev' + +@[deprecated ENNReal.sub_eq_of_eq_add (since := "2024-09-30")] theorem sub_eq_of_add_eq (hb : b ≠ ∞) (hc : a + b = c) : c - b = a := ENNReal.sub_eq_of_eq_add hb hc.symm @@ -338,7 +360,7 @@ protected theorem sub_add_eq_add_sub (hab : b ≤ a) (b_ne_top : b ≠ ∞) : a - b + c = a + c - b := by by_cases c_top : c = ∞ · simpa [c_top] using ENNReal.eq_sub_of_add_eq b_ne_top rfl - refine (sub_eq_of_add_eq b_ne_top ?_).symm + refine ENNReal.eq_sub_of_add_eq b_ne_top ?_ simp only [add_assoc, add_comm c b] simpa only [← add_assoc] using (add_left_inj c_top).mpr <| tsub_add_cancel_of_le hab diff --git a/Mathlib/Data/ENNReal/Real.lean b/Mathlib/Data/ENNReal/Real.lean index 4b6df926ef130..972d93ec15a0c 100644 --- a/Mathlib/Data/ENNReal/Real.lean +++ b/Mathlib/Data/ENNReal/Real.lean @@ -235,7 +235,7 @@ lemma ofNat_le_ofReal {n : ℕ} [n.AtLeastTwo] {p : ℝ} : no_index (OfNat.ofNat n) ≤ ENNReal.ofReal p ↔ OfNat.ofNat n ≤ p := natCast_le_ofReal (NeZero.ne n) -@[simp] +@[simp, norm_cast] lemma ofReal_le_natCast {r : ℝ} {n : ℕ} : ENNReal.ofReal r ≤ n ↔ r ≤ n := coe_le_coe.trans Real.toNNReal_le_natCast diff --git a/Mathlib/Data/Finsupp/Weight.lean b/Mathlib/Data/Finsupp/Weight.lean index 2ed68e2f237b0..ed767ee615ac8 100644 --- a/Mathlib/Data/Finsupp/Weight.lean +++ b/Mathlib/Data/Finsupp/Weight.lean @@ -157,10 +157,9 @@ theorem weight_eq_zero_iff_eq_zero ext s simp only [Finsupp.coe_zero, Pi.zero_apply] by_contra hs - apply NonTorsionWeight.ne_zero w _ - -- TODO: this is subtle, but I'm not sure what to best say in a comment... - · rw [← nonpos_iff_eq_zero, ← h] - exact le_weight_of_ne_zero' w hs + apply NonTorsionWeight.ne_zero w s + rw [← nonpos_iff_eq_zero, ← h] + exact le_weight_of_ne_zero' w hs · intro h rw [h, map_zero] diff --git a/Mathlib/Data/Int/ModEq.lean b/Mathlib/Data/Int/ModEq.lean index 2f973a2e81d79..d171639d57b4d 100644 --- a/Mathlib/Data/Int/ModEq.lean +++ b/Mathlib/Data/Int/ModEq.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import Mathlib.Data.Nat.ModEq -import Mathlib.Tactic.Abel -import Mathlib.Tactic.GCongr.CoreAttrs /-! @@ -92,8 +90,7 @@ theorem mod_modEq (a n) : a % n ≡ a [ZMOD n] := @[simp] theorem neg_modEq_neg : -a ≡ -b [ZMOD n] ↔ a ≡ b [ZMOD n] := by --- Porting note: Restore old proof once #3309 is through - simp [-sub_neg_eq_add, neg_sub_neg, modEq_iff_dvd, dvd_sub_comm] + simp only [modEq_iff_dvd, (by omega : -b - -a = -(b - a)), Int.dvd_neg] @[simp] theorem modEq_neg : a ≡ b [ZMOD -n] ↔ a ≡ b [ZMOD n] := by simp [modEq_iff_dvd] @@ -105,9 +102,9 @@ protected theorem of_dvd (d : m ∣ n) (h : a ≡ b [ZMOD n]) : a ≡ b [ZMOD m] protected theorem mul_left' (h : a ≡ b [ZMOD n]) : c * a ≡ c * b [ZMOD c * n] := by obtain hc | rfl | hc := lt_trichotomy c 0 - · rw [← neg_modEq_neg, ← modEq_neg, ← neg_mul, ← neg_mul, ← neg_mul] + · rw [← neg_modEq_neg, ← modEq_neg, ← Int.neg_mul, ← Int.neg_mul, ← Int.neg_mul] simp only [ModEq, mul_emod_mul_of_pos _ _ (neg_pos.2 hc), h.eq] - · simp only [zero_mul, ModEq.rfl] + · simp only [Int.zero_mul, ModEq.rfl] · simp only [ModEq, mul_emod_mul_of_pos _ _ hc, h.eq] protected theorem mul_right' (h : a ≡ b [ZMOD n]) : a * c ≡ b * c [ZMOD n * c] := by @@ -115,7 +112,7 @@ protected theorem mul_right' (h : a ≡ b [ZMOD n]) : a * c ≡ b * c [ZMOD n * @[gcongr] protected theorem add (h₁ : a ≡ b [ZMOD n]) (h₂ : c ≡ d [ZMOD n]) : a + c ≡ b + d [ZMOD n] := - modEq_iff_dvd.2 <| by convert dvd_add h₁.dvd h₂.dvd using 1; abel + modEq_iff_dvd.2 <| by convert Int.dvd_add h₁.dvd h₂.dvd using 1; omega @[gcongr] protected theorem add_left (c : ℤ) (h : a ≡ b [ZMOD n]) : c + a ≡ c + b [ZMOD n] := ModEq.rfl.add h @@ -125,10 +122,10 @@ protected theorem add (h₁ : a ≡ b [ZMOD n]) (h₂ : c ≡ d [ZMOD n]) : a + protected theorem add_left_cancel (h₁ : a ≡ b [ZMOD n]) (h₂ : a + c ≡ b + d [ZMOD n]) : c ≡ d [ZMOD n] := - have : d - c = b + d - (a + c) - (b - a) := by abel + have : d - c = b + d - (a + c) - (b - a) := by omega modEq_iff_dvd.2 <| by rw [this] - exact dvd_sub h₂.dvd h₁.dvd + exact Int.dvd_sub h₂.dvd h₁.dvd protected theorem add_left_cancel' (c : ℤ) (h : c + a ≡ c + b [ZMOD n]) : a ≡ b [ZMOD n] := ModEq.rfl.add_left_cancel h @@ -183,7 +180,7 @@ theorem cancel_right_div_gcd (hm : 0 < m) (h : a * c ≡ b * c [ZMOD m]) : rw [modEq_iff_dvd] at h ⊢ -- Porting note: removed `show` due to leanprover-community/mathlib4#3305 refine Int.dvd_of_dvd_mul_right_of_gcd_one (?_ : m / d ∣ c / d * (b - a)) ?_ - · rw [mul_comm, ← Int.mul_ediv_assoc (b - a) gcd_dvd_right, sub_mul] + · rw [mul_comm, ← Int.mul_ediv_assoc (b - a) gcd_dvd_right, Int.sub_mul] exact Int.ediv_dvd_ediv gcd_dvd_left h · rw [gcd_div gcd_dvd_left gcd_dvd_right, natAbs_ofNat, Nat.div_self (gcd_pos_of_ne_zero_left c hm.ne')] @@ -233,7 +230,7 @@ theorem modEq_add_fac {a b n : ℤ} (c : ℤ) (ha : a ≡ b [ZMOD n]) : a + n * _ ≡ b [ZMOD n] := by rw [add_zero] theorem modEq_sub_fac {a b n : ℤ} (c : ℤ) (ha : a ≡ b [ZMOD n]) : a - n * c ≡ b [ZMOD n] := by - convert Int.modEq_add_fac (-c) ha using 1; rw [mul_neg, sub_eq_add_neg] + convert Int.modEq_add_fac (-c) ha using 1; rw [Int.mul_neg, sub_eq_add_neg] theorem modEq_add_fac_self {a t n : ℤ} : a + n * t ≡ a [ZMOD n] := modEq_add_fac _ ModEq.rfl diff --git a/Mathlib/Data/Int/Order/Lemmas.lean b/Mathlib/Data/Int/Order/Lemmas.lean index e7bc9e148fc42..d664a5b284aca 100644 --- a/Mathlib/Data/Int/Order/Lemmas.lean +++ b/Mathlib/Data/Int/Order/Lemmas.lean @@ -34,15 +34,4 @@ theorem natAbs_le_iff_mul_self_le {a b : ℤ} : a.natAbs ≤ b.natAbs ↔ a * a rw [← abs_le_iff_mul_self_le, abs_eq_natAbs, abs_eq_natAbs] exact Int.ofNat_le.symm -/-! ### units -/ - - -theorem eq_zero_of_abs_lt_dvd {m x : ℤ} (h1 : m ∣ x) (h2 : |x| < m) : x = 0 := by - obtain rfl | hm := eq_or_ne m 0 - · exact Int.zero_dvd.1 h1 - rcases h1 with ⟨d, rfl⟩ - apply mul_eq_zero_of_right - rw [← abs_lt_one_iff, ← mul_lt_iff_lt_one_right (abs_pos.mpr hm), ← abs_mul] - exact lt_of_lt_of_le h2 (le_abs_self m) - end Int diff --git a/Mathlib/Data/Matrix/Basic.lean b/Mathlib/Data/Matrix/Basic.lean index dfb96ff464262..6bcfb3a9df82c 100644 --- a/Mathlib/Data/Matrix/Basic.lean +++ b/Mathlib/Data/Matrix/Basic.lean @@ -492,6 +492,11 @@ theorem diagonal_conjTranspose [AddMonoid α] [StarAddMonoid α] (v : n → α) rw [conjTranspose, diagonal_transpose, diagonal_map (star_zero _)] rfl +theorem diagonal_unique [Unique m] [DecidableEq m] [Zero α] (d : m → α) : + diagonal d = of fun _ _ => d default := by + ext i j + rw [Subsingleton.elim i default, Subsingleton.elim j default, diagonal_apply_eq _ _, of_apply] + section One variable [Zero α] [One α] diff --git a/Mathlib/Data/NNReal/Basic.lean b/Mathlib/Data/NNReal/Basic.lean index df15ccc645e66..acc5c21138221 100644 --- a/Mathlib/Data/NNReal/Basic.lean +++ b/Mathlib/Data/NNReal/Basic.lean @@ -3,16 +3,13 @@ Copyright (c) 2018 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ -import Mathlib.Algebra.Algebra.Defs +import Mathlib.Algebra.BigOperators.Expect import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Algebra.Order.Field.Canonical.Basic -import Mathlib.Algebra.Order.Nonneg.Field import Mathlib.Algebra.Order.Nonneg.Floor +import Mathlib.Algebra.Ring.Regular import Mathlib.Data.Real.Pointwise import Mathlib.Order.ConditionallyCompleteLattice.Group -import Mathlib.Tactic.Bound.Attribute -import Mathlib.Tactic.GCongr.CoreAttrs -import Mathlib.Algebra.Ring.Regular /-! # Nonnegative real numbers @@ -55,6 +52,7 @@ This file defines `ℝ≥0` as a localized notation for `NNReal`. assert_not_exists Star open Function +open scoped BigOperators -- to ensure these instances are computable /-- Nonnegative real numbers. -/ @@ -278,22 +276,26 @@ theorem coe_multiset_sum (s : Multiset ℝ≥0) : ((s.sum : ℝ≥0) : ℝ) = (s theorem coe_multiset_prod (s : Multiset ℝ≥0) : ((s.prod : ℝ≥0) : ℝ) = (s.map (↑)).prod := map_multiset_prod toRealHom s +variable {ι : Type*} {s : Finset ι} {f : ι → ℝ} + @[simp, norm_cast] -theorem coe_sum {α} {s : Finset α} {f : α → ℝ≥0} : ↑(∑ a ∈ s, f a) = ∑ a ∈ s, (f a : ℝ) := +theorem coe_sum (s : Finset ι) (f : ι → ℝ≥0) : ∑ i ∈ s, f i = ∑ i ∈ s, (f i : ℝ) := map_sum toRealHom _ _ -theorem _root_.Real.toNNReal_sum_of_nonneg {α} {s : Finset α} {f : α → ℝ} - (hf : ∀ a, a ∈ s → 0 ≤ f a) : +@[simp, norm_cast] +lemma coe_expect (s : Finset ι) (f : ι → ℝ≥0) : 𝔼 i ∈ s, f i = 𝔼 i ∈ s, (f i : ℝ) := + map_expect toRealHom .. + +theorem _root_.Real.toNNReal_sum_of_nonneg (hf : ∀ i ∈ s, 0 ≤ f i) : Real.toNNReal (∑ a ∈ s, f a) = ∑ a ∈ s, Real.toNNReal (f a) := by rw [← coe_inj, NNReal.coe_sum, Real.coe_toNNReal _ (Finset.sum_nonneg hf)] exact Finset.sum_congr rfl fun x hxs => by rw [Real.coe_toNNReal _ (hf x hxs)] @[simp, norm_cast] -theorem coe_prod {α} {s : Finset α} {f : α → ℝ≥0} : ↑(∏ a ∈ s, f a) = ∏ a ∈ s, (f a : ℝ) := +theorem coe_prod (s : Finset ι) (f : ι → ℝ≥0) : ↑(∏ a ∈ s, f a) = ∏ a ∈ s, (f a : ℝ) := map_prod toRealHom _ _ -theorem _root_.Real.toNNReal_prod_of_nonneg {α} {s : Finset α} {f : α → ℝ} - (hf : ∀ a, a ∈ s → 0 ≤ f a) : +theorem _root_.Real.toNNReal_prod_of_nonneg (hf : ∀ a, a ∈ s → 0 ≤ f a) : Real.toNNReal (∏ a ∈ s, f a) = ∏ a ∈ s, Real.toNNReal (f a) := by rw [← coe_inj, NNReal.coe_prod, Real.coe_toNNReal _ (Finset.prod_nonneg hf)] exact Finset.prod_congr rfl fun x hxs => by rw [Real.coe_toNNReal _ (hf x hxs)] diff --git a/Mathlib/Data/Nat/ModEq.lean b/Mathlib/Data/Nat/ModEq.lean index 5877782fc1102..29ff2fed4da9d 100644 --- a/Mathlib/Data/Nat/ModEq.lean +++ b/Mathlib/Data/Nat/ModEq.lean @@ -3,10 +3,8 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Algebra.Ring.Regular +import Mathlib.Algebra.Order.Group.Unbundled.Int import Mathlib.Data.Int.GCD -import Mathlib.Data.Int.Order.Lemmas -import Mathlib.Tactic.NormNum.Basic /-! # Congruences modulo a natural number @@ -24,6 +22,7 @@ and proves basic properties about it such as the Chinese Remainder Theorem ModEq, congruence, mod, MOD, modulo -/ +assert_not_exists OrderedAddCommMonoid assert_not_exists Function.support namespace Nat @@ -37,8 +36,8 @@ notation:50 a " ≡ " b " [MOD " n "]" => ModEq n a b variable {m n a b c d : ℕ} --- Porting note: This instance should be derivable automatically -instance : Decidable (ModEq n a b) := decEq (a % n) (b % n) +-- Since `ModEq` is semi-reducible, we need to provide the decidable instance manually +instance : Decidable (ModEq n a b) := inferInstanceAs <| Decidable (a % n = b % n) namespace ModEq @@ -91,7 +90,7 @@ theorem mod_modEq (a n) : a % n ≡ a [MOD n] := namespace ModEq lemma of_dvd (d : m ∣ n) (h : a ≡ b [MOD n]) : a ≡ b [MOD m] := - modEq_of_dvd <| d.natCast.trans h.dvd + modEq_of_dvd <| Int.ofNat_dvd.mpr d |>.trans h.dvd protected theorem mul_left' (c : ℕ) (h : a ≡ b [MOD n]) : c * a ≡ c * b [MOD c * n] := by unfold ModEq at *; rw [mul_mod_mul_left, mul_mod_mul_left, h] @@ -122,7 +121,7 @@ protected theorem pow (m : ℕ) (h : a ≡ b [MOD n]) : a ^ m ≡ b ^ m [MOD n] @[gcongr] protected theorem add (h₁ : a ≡ b [MOD n]) (h₂ : c ≡ d [MOD n]) : a + c ≡ b + d [MOD n] := by rw [modEq_iff_dvd, Int.ofNat_add, Int.ofNat_add, add_sub_add_comm] - exact dvd_add h₁.dvd h₂.dvd + exact Int.dvd_add h₁.dvd h₂.dvd @[gcongr] protected theorem add_left (c : ℕ) (h : a ≡ b [MOD n]) : c + a ≡ c + b [MOD n] := @@ -136,7 +135,7 @@ protected theorem add_left_cancel (h₁ : a ≡ b [MOD n]) (h₂ : a + c ≡ b + c ≡ d [MOD n] := by simp only [modEq_iff_dvd, Int.ofNat_add] at * rw [add_sub_add_comm] at h₂ - convert _root_.dvd_sub h₂ h₁ using 1 + convert Int.dvd_sub h₂ h₁ using 1 rw [add_sub_cancel_left] protected theorem add_left_cancel' (c : ℕ) (h : c + a ≡ c + b [MOD n]) : a ≡ b [MOD n] := @@ -155,7 +154,8 @@ protected theorem add_right_cancel' (c : ℕ) (h : a + c ≡ b + c [MOD n]) : a For cancelling left multiplication in the modulus, see `Nat.ModEq.of_mul_left`. -/ protected theorem mul_left_cancel' {a b c m : ℕ} (hc : c ≠ 0) : c * a ≡ c * b [MOD c * m] → a ≡ b [MOD m] := by - simp [modEq_iff_dvd, ← mul_sub, mul_dvd_mul_iff_left (by simp [hc] : (c : ℤ) ≠ 0)] + simp only [modEq_iff_dvd, Int.natCast_mul, ← Int.mul_sub] + exact fun h => (Int.dvd_of_mul_dvd_mul_left (Int.ofNat_ne_zero.mpr hc) h) protected theorem mul_left_cancel_iff' {a b c m : ℕ} (hc : c ≠ 0) : c * a ≡ c * b [MOD c * m] ↔ a ≡ b [MOD m] := @@ -166,7 +166,8 @@ protected theorem mul_left_cancel_iff' {a b c m : ℕ} (hc : c ≠ 0) : For cancelling right multiplication in the modulus, see `Nat.ModEq.of_mul_right`. -/ protected theorem mul_right_cancel' {a b c m : ℕ} (hc : c ≠ 0) : a * c ≡ b * c [MOD m * c] → a ≡ b [MOD m] := by - simp [modEq_iff_dvd, ← sub_mul, mul_dvd_mul_iff_right (by simp [hc] : (c : ℤ) ≠ 0)] + simp only [modEq_iff_dvd, Int.natCast_mul, ← Int.sub_mul] + exact fun h => (Int.dvd_of_mul_dvd_mul_right (Int.ofNat_ne_zero.mpr hc) h) protected theorem mul_right_cancel_iff' {a b c m : ℕ} (hc : c ≠ 0) : a * c ≡ b * c [MOD m * c] ↔ a ≡ b [MOD m] := @@ -204,10 +205,10 @@ namespace ModEq theorem le_of_lt_add (h1 : a ≡ b [MOD m]) (h2 : a < b + m) : a ≤ b := (le_total a b).elim id fun h3 => Nat.le_of_sub_eq_zero - (eq_zero_of_dvd_of_lt ((modEq_iff_dvd' h3).mp h1.symm) ((tsub_lt_iff_left h3).mpr h2)) + (eq_zero_of_dvd_of_lt ((modEq_iff_dvd' h3).mp h1.symm) (by omega)) theorem add_le_of_lt (h1 : a ≡ b [MOD m]) (h2 : a < b) : a + m ≤ b := - le_of_lt_add (add_modEq_right.trans h1) (add_lt_add_right h2 m) + le_of_lt_add (add_modEq_right.trans h1) (by omega) theorem dvd_iff (h : a ≡ b [MOD m]) (hdm : d ∣ m) : d ∣ a ↔ d ∣ b := by simp only [← modEq_zero_iff_dvd] @@ -227,9 +228,7 @@ lemma eq_of_abs_lt (h : a ≡ b [MOD m]) (h2 : |(b : ℤ) - a| < m) : a = b := b exact Int.eq_zero_of_abs_lt_dvd h.dvd h2 lemma eq_of_lt_of_lt (h : a ≡ b [MOD m]) (ha : a < m) (hb : b < m) : a = b := - h.eq_of_abs_lt <| abs_sub_lt_iff.2 - ⟨(sub_le_self _ <| Int.natCast_nonneg _).trans_lt <| Int.ofNat_lt.2 hb, - (sub_le_self _ <| Int.natCast_nonneg _).trans_lt <| Int.ofNat_lt.2 ha⟩ + h.eq_of_abs_lt <| Int.abs_sub_lt_of_lt_lt ha hb /-- To cancel a common factor `c` from a `ModEq` we must divide the modulus `m` by `gcd m c` -/ lemma cancel_left_div_gcd (hm : 0 < m) (h : c * a ≡ c * b [MOD m]) : a ≡ b [MOD m / gcd m c] := by @@ -241,7 +240,7 @@ lemma cancel_left_div_gcd (hm : 0 < m) (h : c * a ≡ c * b [MOD m]) : a ≡ b · show (m / d : ℤ) ∣ c / d * (b - a) rw [mul_comm, ← Int.mul_ediv_assoc (b - a) (Int.natCast_dvd_natCast.mpr hcd), mul_comm] apply Int.ediv_dvd_ediv (Int.natCast_dvd_natCast.mpr hmd) - rw [mul_sub] + rw [Int.mul_sub] exact modEq_iff_dvd.mp h · show Int.gcd (m / d) (c / d) = 1 simp only [← Int.natCast_div, Int.gcd_natCast_natCast (m / d) (c / d), gcd_div hmd hcd, @@ -299,18 +298,18 @@ def chineseRemainder' (h : a ≡ b [MOD gcd n m]) : { k // k ≡ a [MOD n] ∧ k have hcoedvd : ∀ t, (gcd n m : ℤ) ∣ t * (b - a) := fun t => h.dvd.mul_left _ have := gcd_eq_gcd_ab n m constructor <;> rw [Int.emod_def, ← sub_add] <;> - refine dvd_add ?_ (dvd_mul_of_dvd_left ?_ _) <;> + refine Int.dvd_add ?_ (dvd_mul_of_dvd_left ?_ _) <;> try norm_cast · rw [← sub_eq_iff_eq_add'] at this - rw [← this, sub_mul, ← add_sub_assoc, add_comm, add_sub_assoc, ← mul_sub, + rw [← this, Int.sub_mul, ← add_sub_assoc, add_comm, add_sub_assoc, ← Int.mul_sub, Int.add_ediv_of_dvd_left, Int.mul_ediv_cancel_left _ hnonzero, - Int.mul_ediv_assoc _ h.dvd, ← sub_sub, sub_self, zero_sub, dvd_neg, mul_assoc] + Int.mul_ediv_assoc _ h.dvd, ← sub_sub, sub_self, zero_sub, Int.dvd_neg, mul_assoc] · exact dvd_mul_right _ _ norm_cast exact dvd_mul_right _ _ · exact dvd_lcm_left n m · rw [← sub_eq_iff_eq_add] at this - rw [← this, sub_mul, sub_add, ← mul_sub, Int.sub_ediv_of_dvd, + rw [← this, Int.sub_mul, sub_add, ← Int.mul_sub, Int.sub_ediv_of_dvd, Int.mul_ediv_cancel_left _ hnonzero, Int.mul_ediv_assoc _ h.dvd, ← sub_add, sub_self, zero_add, mul_assoc] · exact dvd_mul_right _ _ @@ -407,7 +406,7 @@ protected theorem add_div_of_dvd_right {a b c : ℕ} (hca : c ∣ a) : (a + b) / add_div_eq_of_add_mod_lt (by rw [Nat.mod_eq_zero_of_dvd hca, zero_add] - exact Nat.mod_lt _ (pos_iff_ne_zero.mpr h)) + exact Nat.mod_lt _ (zero_lt_of_ne_zero h)) protected theorem add_div_of_dvd_left {a b c : ℕ} (hca : c ∣ b) : (a + b) / c = a / c + b / c := by rwa [add_comm, Nat.add_div_of_dvd_right, add_comm] @@ -430,27 +429,24 @@ theorem odd_mul_odd {n m : ℕ} : n % 2 = 1 → m % 2 = 1 → n * m % 2 = 1 := b theorem odd_mul_odd_div_two {m n : ℕ} (hm1 : m % 2 = 1) (hn1 : n % 2 = 1) : m * n / 2 = m * (n / 2) + m / 2 := - have hm0 : 0 < m := Nat.pos_of_ne_zero fun h => by simp_all have hn0 : 0 < n := Nat.pos_of_ne_zero fun h => by simp_all mul_right_injective₀ two_ne_zero <| by dsimp rw [mul_add, two_mul_odd_div_two hm1, mul_left_comm, two_mul_odd_div_two hn1, - two_mul_odd_div_two (Nat.odd_mul_odd hm1 hn1), mul_tsub, mul_one, ← - add_tsub_assoc_of_le (succ_le_of_lt hm0), - tsub_add_cancel_of_le (le_mul_of_one_le_right (Nat.zero_le _) hn0)] + two_mul_odd_div_two (Nat.odd_mul_odd hm1 hn1), Nat.mul_sub, mul_one, ← + Nat.add_sub_assoc (by omega), Nat.sub_add_cancel (Nat.le_mul_of_pos_right m hn0)] theorem odd_of_mod_four_eq_one {n : ℕ} : n % 4 = 1 → n % 2 = 1 := by - simpa [ModEq, show 2 * 2 = 4 by norm_num] using @ModEq.of_mul_left 2 n 1 2 + simpa [ModEq] using @ModEq.of_mul_left 2 n 1 2 theorem odd_of_mod_four_eq_three {n : ℕ} : n % 4 = 3 → n % 2 = 1 := by - simpa [ModEq, show 2 * 2 = 4 by norm_num, show 3 % 4 = 3 by norm_num] using - @ModEq.of_mul_left 2 n 3 2 + simpa [ModEq] using @ModEq.of_mul_left 2 n 3 2 /-- A natural number is odd iff it has residue `1` or `3` mod `4`-/ theorem odd_mod_four_iff {n : ℕ} : n % 2 = 1 ↔ n % 4 = 1 ∨ n % 4 = 3 := have help : ∀ m : ℕ, m < 4 → m % 2 = 1 → m = 1 ∨ m = 3 := by decide ⟨fun hn => - help (n % 4) (mod_lt n (by norm_num)) <| (mod_mod_of_dvd n (by decide : 2 ∣ 4)).trans hn, + help (n % 4) (mod_lt n (by omega)) <| (mod_mod_of_dvd n (by decide : 2 ∣ 4)).trans hn, fun h => Or.elim h odd_of_mod_four_eq_one odd_of_mod_four_eq_three⟩ lemma mod_eq_of_modEq {a b n} (h : a ≡ b [MOD n]) (hb : b < n) : a % n = b := diff --git a/Mathlib/Data/Real/Pi/Irrational.lean b/Mathlib/Data/Real/Pi/Irrational.lean new file mode 100644 index 0000000000000..8736c6f964f66 --- /dev/null +++ b/Mathlib/Data/Real/Pi/Irrational.lean @@ -0,0 +1,305 @@ +/- +Copyright (c) 2022 Bhavik Mehta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Bhavik Mehta +-/ +import Mathlib.Analysis.SpecialFunctions.Integrals +import Mathlib.Data.Real.Irrational + +/-! +# `Real.pi` is irrational + +The main result of this file is `irrational_pi`. + +The proof is adapted from https://en.wikipedia.org/wiki/Proof_that_%CF%80_is_irrational#Cartwright's_proof. + +The proof idea is as follows. +* Define a sequence of integrals `I n θ = ∫ x in (-1)..1, (1 - x ^ 2) ^ n * cos (x * θ)`. +* Give a recursion formula for `I (n + 2) θ * θ ^ 2` in terms of `I n θ` and `I (n + 1) θ`. + Note we do not find it helpful to define `J` as in the above proof, and instead work directly + with `I`. +* Define polynomials with integer coefficients `sinPoly n` and `cosPoly n` such that + `I n θ * θ ^ (2 * n + 1) = n ! * (sinPoly n θ * sin θ + cosPoly n θ * cos θ)`. + Note that in the informal proof, these polynomials are not defined explicitly, but we find it + useful to define them by recursion. +* Show that both these polynomials have degree bounded by `n`. +* Show that `0 < I n (π / 2) ≤ 2` for all `n`. +* Now we can finish: if `π / 2` is rational, write it as `a / b` with `a, b > 0`. Then + `b ^ (2 * n + 1) * sinPoly n (a / b)` is a positive integer by the degree bound. But it is equal + to `a ^ (2 * n + 1) / n ! * I n (π / 2) ≤ 2 * a * (2 * n + 1) / n !`, which converges to 0 as + `n → ∞`. + +-/ + +noncomputable section + +open intervalIntegral MeasureTheory.MeasureSpace Set Polynomial Real +open scoped Nat + +/-- The sequence of integrals used for Cartwright's proof of irrationality of `π`. -/ +private def I (n : ℕ) (θ : ℝ) : ℝ := ∫ x in (-1)..1, (1 - x ^ 2) ^ n * cos (x * θ) + +variable {n : ℕ} {θ : ℝ} + +private lemma I_zero : I 0 θ * θ = 2 * sin θ := by + rw [mul_comm, I] + simp [mul_integral_comp_mul_right, two_mul] + +/-- +Auxiliary for the proof that `π` is irrational. +While it is most natural to give the recursive formula for `I (n + 2) θ`, as well as give the second +base case of `I 1 θ`, it is in fact more convenient to give the recursive formula for `I (n + 1) θ` +in terms of `I n θ` and `I (n - 1) θ` (note the natural subtraction!). +Despite the usually inconvenient subtraction, this in fact allows deducing both of the above facts +with significantly fewer analysis computations. +In addition, note the `0 ^ n` on the right hand side - this is intentional, and again allows +combining the proof of the "usual" recursion formula and the base case `I 1 θ`. +-/ +private lemma recursion' (n : ℕ) : + I (n + 1) θ * θ ^ 2 = - (2 * 2 * ((n + 1) * (0 ^ n * cos θ))) + + 2 * (n + 1) * (2 * n + 1) * I n θ - 4 * (n + 1) * n * I (n - 1) θ := by + rw [I] + let f (x : ℝ) : ℝ := 1 - x ^ 2 + let u₁ (x : ℝ) : ℝ := f x ^ (n + 1) + let u₁' (x : ℝ) : ℝ := - (2 * (n + 1) * x * f x ^ n) + let v₁ (x : ℝ) : ℝ := sin (x * θ) + let v₁' (x : ℝ) : ℝ := cos (x * θ) * θ + let u₂ (x : ℝ) : ℝ := x * (f x) ^ n + let u₂' (x : ℝ) : ℝ := (f x) ^ n - 2 * n * x ^ 2 * (f x) ^ (n - 1) + let v₂ (x : ℝ) : ℝ := cos (x * θ) + let v₂' (x : ℝ) : ℝ := -sin (x * θ) * θ + have hfd : Continuous f := by fun_prop + have hu₁d : Continuous u₁' := by fun_prop + have hv₁d : Continuous v₁' := by fun_prop + have hu₂d : Continuous u₂' := by fun_prop + have hv₂d : Continuous v₂' := by fun_prop + have hu₁_eval_one : u₁ 1 = 0 := by simp only [u₁, f]; simp + have hu₁_eval_neg_one : u₁ (-1) = 0 := by simp only [u₁, f]; simp + have t : u₂ 1 * v₂ 1 - u₂ (-1) * v₂ (-1) = 2 * (0 ^ n * cos θ) := by simp [u₂, v₂, f, ← two_mul] + have hf (x) : HasDerivAt f (- 2 * x) x := by + convert (hasDerivAt_pow 2 x).const_sub 1 using 1 + simp + have hu₁ (x) : HasDerivAt u₁ (u₁' x) x := by + convert (hf x).pow _ using 1 + simp only [Nat.add_succ_sub_one, u₁', Nat.cast_add_one] + ring + have hv₁ (x) : HasDerivAt v₁ (v₁' x) x := (hasDerivAt_mul_const θ).sin + have hu₂ (x) : HasDerivAt u₂ (u₂' x) x := by + convert (hasDerivAt_id' x).mul ((hf x).pow _) using 1 + simp only [u₂'] + ring + have hv₂ (x) : HasDerivAt v₂ (v₂' x) x := (hasDerivAt_mul_const θ).cos + convert_to (∫ (x : ℝ) in (-1)..1, u₁ x * v₁' x) * θ = _ using 1 + · simp_rw [u₁, v₁', ← intervalIntegral.integral_mul_const, sq θ, mul_assoc] + rw [integral_mul_deriv_eq_deriv_mul (fun x _ => hu₁ x) (fun x _ => hv₁ x) + (hu₁d.intervalIntegrable _ _) (hv₁d.intervalIntegrable _ _), hu₁_eval_one, hu₁_eval_neg_one, + zero_mul, zero_mul, sub_zero, zero_sub, ← integral_neg, ← integral_mul_const] + convert_to ((-2 : ℝ) * (n + 1)) * ∫ (x : ℝ) in (-1)..1, (u₂ x * v₂' x) = _ using 1 + · rw [← integral_const_mul] + congr 1 with x + dsimp [u₁', v₁, u₂, v₂'] + ring + rw [integral_mul_deriv_eq_deriv_mul (fun x _ => hu₂ x) (fun x _ => hv₂ x) + (hu₂d.intervalIntegrable _ _) (hv₂d.intervalIntegrable _ _), + mul_sub, t, neg_mul, neg_mul, neg_mul, sub_neg_eq_add] + have (x) : u₂' x = (2 * n + 1) * f x ^ n - 2 * n * f x ^ (n - 1) := by + cases n with + | zero => simp [u₂'] + | succ n => ring! + simp_rw [this, sub_mul, mul_assoc _ _ (v₂ _)] + have : Continuous v₂ := by fun_prop + rw [mul_mul_mul_comm, integral_sub, mul_sub, add_sub_assoc] + · congr 1 + simp_rw [integral_const_mul] + ring! + all_goals exact Continuous.intervalIntegrable (by fun_prop) _ _ + +/-- +Auxiliary for the proof that `π` is irrational. +The recursive formula for `I (n + 2) θ * θ ^ 2` in terms of `I n θ` and `I (n + 1) θ`. +-/ +private lemma recursion (n : ℕ) : + I (n + 2) θ * θ ^ 2 = + 2 * (n + 2) * (2 * n + 3) * I (n + 1) θ - 4 * (n + 2) * (n + 1) * I n θ := by + rw [recursion' (n + 1)] + simp + ring! + +/-- +Auxiliary for the proof that `π` is irrational. +The second base case for the induction on `n`, giving an explicit formula for `I 1 θ`. +-/ +private lemma I_one : I 1 θ * θ ^ 3 = 4 * sin θ - 4 * θ * cos θ := by + rw [_root_.pow_succ, ← mul_assoc, recursion' 0, sub_mul, add_mul, mul_assoc _ (I 0 θ), I_zero] + ring + +/-- +Auxiliary for the proof that `π` is irrational. +The first of the two integer-coefficient polynomials that describe the behaviour of the +sequence of integrals `I`. +While not given in the informal proof, these are easy to deduce from the recursion formulae. +-/ +private def sinPoly : ℕ → ℤ[X] + | 0 => C 2 + | 1 => C 4 + | (n+2) => ((2 : ℤ) * (2 * n + 3)) • sinPoly (n + 1) + monomial 2 (-4) * sinPoly n + +/-- +Auxiliary for the proof that `π` is irrational. +The second of the two integer-coefficient polynomials that describe the behaviour of the +sequence of integrals `I`. +While not given in the informal proof, these are easy to deduce from the recursion formulae. +-/ +private def cosPoly : ℕ → ℤ[X] + | 0 => 0 + | 1 => monomial 1 (-4) + | (n+2) => ((2 : ℤ) * (2 * n + 3)) • cosPoly (n + 1) + monomial 2 (-4) * cosPoly n + +/-- +Auxiliary for the proof that `π` is irrational. +Prove a degree bound for `sinPoly n` by induction. Note this is where we find the value in an +explicit description of `sinPoly`. +-/ +private lemma sinPoly_natDegree_le : ∀ n : ℕ, (sinPoly n).natDegree ≤ n + | 0 => by simp [sinPoly] + | 1 => by simp only [natDegree_C, mul_one, zero_le', sinPoly] + | n + 2 => by + rw [sinPoly] + refine natDegree_add_le_of_degree_le ((natDegree_smul_le _ _).trans ?_) ?_ + · exact (sinPoly_natDegree_le (n + 1)).trans (by simp) + refine natDegree_mul_le.trans ?_ + simpa [add_comm 2] using sinPoly_natDegree_le n + +/-- +Auxiliary for the proof that `π` is irrational. +Prove a degree bound for `cosPoly n` by induction. Note this is where we find the value in an +explicit description of `cosPoly`. +-/ +private lemma cosPoly_natDegree_le : ∀ n : ℕ, (cosPoly n).natDegree ≤ n + | 0 => by simp [cosPoly] + | 1 => (natDegree_monomial_le _).trans (by simp) + | n + 2 => by + rw [cosPoly] + refine natDegree_add_le_of_degree_le ((natDegree_smul_le _ _).trans ?_) ?_ + · exact (cosPoly_natDegree_le (n + 1)).trans (by simp) + exact natDegree_mul_le.trans (by simp [add_comm 2, cosPoly_natDegree_le n]) + +/-- +Auxiliary for the proof that `π` is irrational. +The key lemma: the sequence of integrals `I` can be written as a linear combination of `sin` and +`cos`, with coefficients given by the polynomials `sinPoly` and `cosPoly`. +-/ +private lemma sinPoly_add_cosPoly_eval (θ : ℝ) : + ∀ n : ℕ, + I n θ * θ ^ (2 * n + 1) = n ! * ((sinPoly n).eval₂ (Int.castRingHom _) θ * sin θ + + (cosPoly n).eval₂ (Int.castRingHom _) θ * cos θ) + | 0 => by simp [sinPoly, cosPoly, I_zero] + | 1 => by simp [I_one, sinPoly, cosPoly, sub_eq_add_neg] + | n + 2 => by + calc I (n + 2) θ * θ ^ (2 * (n + 2) + 1) = I (n + 2) θ * θ ^ 2 * θ ^ (2 * n + 3) := by ring + _ = 2 * (n + 2) * (2 * n + 3) * (I (n + 1) θ * θ ^ (2 * (n + 1) + 1)) - + 4 * (n + 2) * (n + 1) * θ ^ 2 * (I n θ * θ ^ (2 * n + 1)) := by rw [recursion]; ring + _ = _ := by simp [sinPoly_add_cosPoly_eval, sinPoly, cosPoly, Nat.factorial_succ]; ring + +/-- +Auxiliary for the proof that `π` is irrational. +For a polynomial `p` with natural degree `≤ k` and integer coefficients, evaluating `p` at a +rational `a / b` gives a rational of the form `z / b ^ k`. +TODO: should this be moved elsewhere? It uses none of the pi-specific definitions. +-/ +private lemma is_integer {p : ℤ[X]} (a b : ℤ) {k : ℕ} (hp : p.natDegree ≤ k) : + ∃ z : ℤ, p.eval₂ (Int.castRingHom ℝ) (a / b) * b ^ k = z := by + rcases eq_or_ne b 0 with rfl | hb + · rcases k.eq_zero_or_pos with rfl | hk + · exact ⟨p.coeff 0, by simp⟩ + exact ⟨0, by simp [hk.ne']⟩ + refine ⟨∑ i in p.support, p.coeff i * a ^ i * b ^ (k - i), ?_⟩ + conv => lhs; rw [← sum_monomial_eq p] + rw [eval₂_sum, sum, Finset.sum_mul, Int.cast_sum] + simp only [eval₂_monomial, eq_intCast, div_pow, Int.cast_mul, Int.cast_pow] + refine Finset.sum_congr rfl (fun i hi => ?_) + have ik := (le_natDegree_of_mem_supp i hi).trans hp + rw [mul_assoc, div_mul_comm, ← Int.cast_pow, ← Int.cast_pow, ← Int.cast_pow, + ← pow_sub_mul_pow b ik, ← Int.cast_div_charZero, Int.mul_ediv_cancel _ (pow_ne_zero _ hb), + ← mul_assoc, mul_right_comm, ← Int.cast_pow] + exact dvd_mul_left _ _ + +open Filter + +/-- +Auxiliary for the proof that `π` is irrational. +The integrand in the definition of `I` is nonnegative and takes a positive value at least one point, +so the integral is positive. +-/ +private lemma I_pos : 0 < I n (π / 2) := by + refine integral_pos (by norm_num) (Continuous.continuousOn (by continuity)) ?_ ⟨0, by simp⟩ + refine fun x hx => mul_nonneg (pow_nonneg ?_ _) ?_ + · rw [sub_nonneg, sq_le_one_iff_abs_le_one, abs_le] + exact ⟨hx.1.le, hx.2⟩ + refine cos_nonneg_of_neg_pi_div_two_le_of_le ?_ ?_ <;> + nlinarith [hx.1, hx.2, pi_pos] + +/-- +Auxiliary for the proof that `π` is irrational. +The integrand in the definition of `I` is bounded by 1 and the interval has length 2, so the +integral is bounded above by `2`. +-/ +private lemma I_le (n : ℕ) : I n (π / 2) ≤ 2 := by + rw [← norm_of_nonneg I_pos.le] + refine (norm_integral_le_of_norm_le_const ?_).trans (show (1 : ℝ) * _ ≤ _ by norm_num) + intros x hx + simp only [uIoc_of_le, neg_le_self_iff, zero_le_one, mem_Ioc] at hx + rw [norm_eq_abs, abs_mul, abs_pow] + refine mul_le_one (pow_le_one _ (abs_nonneg _) ?_) (abs_nonneg _) (abs_cos_le_one _) + rw [abs_le] + constructor <;> nlinarith + +/-- +Auxiliary for the proof that `π` is irrational. +For any real `a`, we have that `a ^ (2n+1) / n!` tends to `0` as `n → ∞`. This is just a +reformulation of tendsto_pow_div_factorial_atTop, which asserts the same for `a ^ n / n!` +-/ +private lemma tendsto_pow_div_factorial_at_top_aux (a : ℝ) : + Tendsto (fun n => (a : ℝ) ^ (2 * n + 1) / n !) atTop (nhds 0) := by + rw [← mul_zero a] + refine ((tendsto_pow_div_factorial_atTop (a ^ 2)).const_mul a).congr (fun x => ?_) + rw [← pow_mul, mul_div_assoc', _root_.pow_succ'] + +/-- If `x` is rational, it can be written as `a / b` with `a : ℤ` and `b : ℕ` satisfying `b > 0`. -/ +private lemma not_irrational_exists_rep {x : ℝ} : + ¬Irrational x → ∃ (a : ℤ) (b : ℕ), 0 < b ∧ x = a / b := by + rw [Irrational, not_not, mem_range] + rintro ⟨q, rfl⟩ + exact ⟨q.num, q.den, q.pos, by exact_mod_cast (Rat.num_div_den _).symm⟩ + +@[simp] theorem irrational_pi : Irrational π := by + apply Irrational.of_div_nat 2 + rw [Nat.cast_two] + by_contra h' + obtain ⟨a, b, hb, h⟩ := not_irrational_exists_rep h' + have ha : (0 : ℝ) < a := by + have : 0 < (a : ℝ) / b := h ▸ pi_div_two_pos + rwa [lt_div_iff (by positivity), zero_mul] at this + have k (n : ℕ) : 0 < (a : ℝ) ^ (2 * n + 1) / n ! := by positivity + have j : ∀ᶠ n : ℕ in atTop, (a : ℝ) ^ (2 * n + 1) / n ! * I n (π / 2) < 1 := by + have := eventually_lt_of_tendsto_lt (show (0 : ℝ) < 1 / 2 by norm_num) + (tendsto_pow_div_factorial_at_top_aux a) + filter_upwards [this] with n hn + rw [lt_div_iff (zero_lt_two : (0 : ℝ) < 2)] at hn + exact hn.trans_le' (mul_le_mul_of_nonneg_left (I_le _) (by positivity)) + obtain ⟨n, hn⟩ := j.exists + have hn' : 0 < a ^ (2 * n + 1) / n ! * I n (π / 2) := mul_pos (k _) I_pos + obtain ⟨z, hz⟩ : ∃ z : ℤ, (sinPoly n).eval₂ (Int.castRingHom ℝ) (a / b) * b ^ (2 * n + 1) = z := + is_integer a b ((sinPoly_natDegree_le _).trans (by linarith)) + have e := sinPoly_add_cosPoly_eval (π / 2) n + rw [cos_pi_div_two, sin_pi_div_two, mul_zero, mul_one, add_zero] at e + have : a ^ (2 * n + 1) / n ! * I n (π / 2) = + eval₂ (Int.castRingHom ℝ) (π / 2) (sinPoly n) * b ^ (2 * n + 1) := by + nth_rw 2 [h] at e + field_simp at e ⊢ + linear_combination e + have : (0 : ℝ) < z ∧ (z : ℝ) < 1 := by simp [← hz, ← h, ← this, hn', hn] + norm_cast at this + omega + +end diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 9d317af2926e5..9b390e67bad19 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -1090,6 +1090,9 @@ theorem Injective.image_injective (hf : Injective f) : Injective (image f) := by intro s t h rw [← preimage_image_eq s hf, ← preimage_image_eq t hf, h] +lemma Injective.image_strictMono (inj : Function.Injective f) : StrictMono (image f) := + monotone_image.strictMono_of_injective inj.image_injective + theorem Surjective.preimage_subset_preimage_iff {s t : Set β} (hf : Surjective f) : f ⁻¹' s ⊆ f ⁻¹' t ↔ s ⊆ t := by apply Set.preimage_subset_preimage_iff diff --git a/Mathlib/Data/ZMod/Defs.lean b/Mathlib/Data/ZMod/Defs.lean index 324dec3252219..6977197f083cc 100644 --- a/Mathlib/Data/ZMod/Defs.lean +++ b/Mathlib/Data/ZMod/Defs.lean @@ -5,6 +5,7 @@ Authors: Eric Rodriguez -/ import Mathlib.Algebra.Group.Fin.Basic import Mathlib.Algebra.NeZero +import Mathlib.Algebra.Ring.Int import Mathlib.Data.Nat.ModEq import Mathlib.Data.Fintype.Card diff --git a/Mathlib/FieldTheory/Finiteness.lean b/Mathlib/FieldTheory/Finiteness.lean index 6c2e5621a9272..d6cc69e04c7ae 100644 --- a/Mathlib/FieldTheory/Finiteness.lean +++ b/Mathlib/FieldTheory/Finiteness.lean @@ -72,7 +72,7 @@ theorem coeSort_finsetBasisIndex [IsNoetherian K V] : /-- In a noetherian module over a division ring, there exists a finite basis. This is indexed by the `Finset` `IsNoetherian.finsetBasisIndex`. This is in contrast to the result `finite_basis_index (Basis.ofVectorSpace K V)`, -which provides a set and a `Set.finite`. +which provides a set and a `Set.Finite`. -/ noncomputable def finsetBasis [IsNoetherian K V] : Basis (finsetBasisIndex K V) K V := (Basis.ofVectorSpace K V).reindex (by rw [coeSort_finsetBasisIndex]) diff --git a/Mathlib/FieldTheory/Galois/Basic.lean b/Mathlib/FieldTheory/Galois/Basic.lean index 868454440fa9b..56ce629fc2bea 100644 --- a/Mathlib/FieldTheory/Galois/Basic.lean +++ b/Mathlib/FieldTheory/Galois/Basic.lean @@ -24,7 +24,7 @@ In this file we define Galois extensions as extensions which are both separable - `IntermediateField.fixingSubgroup_fixedField` : If `E/F` is finite dimensional (but not necessarily Galois) then `fixingSubgroup (fixedField H) = H` -- `IntermediateField.fixedField_fixingSubgroup`: If `E/F` is finite dimensional and Galois +- `IsGalois.fixedField_fixingSubgroup`: If `E/F` is finite dimensional and Galois then `fixedField (fixingSubgroup K) = K` Together, these two results prove the Galois correspondence. diff --git a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean index 3dba3c04792f4..0b4140f98c392 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean @@ -96,7 +96,7 @@ protected theorem ContMDiffAt.mfderiv {x₀ : N} (f : N → M → M') (g : N → rw [contMDiffAt_iff] at hf hg simp_rw [Function.comp_def, uncurry, extChartAt_prod, PartialEquiv.prod_coe_symm, ModelWithCorners.range_prod] at hf ⊢ - refine ContDiffWithinAt.fderivWithin ?_ hg.2 I.unique_diff hmn (mem_range_self _) ?_ + refine ContDiffWithinAt.fderivWithin ?_ hg.2 I.uniqueDiffOn hmn (mem_range_self _) ?_ · simp_rw [extChartAt_to_inv]; exact hf.2 · rw [← image_subset_iff] rintro _ ⟨x, -, rfl⟩ @@ -165,7 +165,7 @@ protected theorem ContMDiffAt.mfderiv {x₀ : N} (f : N → M → M') (g : N → PartialEquiv.mem_symm_trans_source _ (mem_extChartAt_source I' (f x₂ (g x₂))) h3x₂).differentiableWithinAt le_top have h3f := (h2x₂.mdifferentiableAt le_rfl).differentiableWithinAt_writtenInExtChartAt - refine fderivWithin.comp₃ _ hI' h3f hI ?_ ?_ ?_ ?_ (I.unique_diff _ <| mem_range_self _) + refine fderivWithin.comp₃ _ hI' h3f hI ?_ ?_ ?_ ?_ (I.uniqueDiffOn _ <| mem_range_self _) · exact fun x _ => mem_range_self _ · exact fun x _ => mem_range_self _ · simp_rw [writtenInExtChartAt, Function.comp_apply, @@ -577,7 +577,7 @@ theorem tangentMap_tangentBundle_pure [Is : SmoothManifoldWithCorners I M] (p : · simp · exact differentiableAt_id' · exact differentiableAt_const _ - · exact ModelWithCorners.unique_diff_at_image I + · exact ModelWithCorners.uniqueDiffWithinAt_image I · exact differentiableAt_id'.prod (differentiableAt_const _) simp (config := { unfoldPartialApp := true }) only [Bundle.zeroSection, tangentMap, mfderiv, A, if_pos, chartAt, FiberBundle.chartedSpace_chartAt, TangentBundle.trivializationAt_apply, diff --git a/Mathlib/Geometry/Manifold/Diffeomorph.lean b/Mathlib/Geometry/Manifold/Diffeomorph.lean index 2c6f481393275..07f8ab47bb49f 100644 --- a/Mathlib/Geometry/Manifold/Diffeomorph.lean +++ b/Mathlib/Geometry/Manifold/Diffeomorph.lean @@ -447,7 +447,7 @@ variable (I) (e : E ≃ₘ[𝕜] E') def transDiffeomorph (I : ModelWithCorners 𝕜 E H) (e : E ≃ₘ[𝕜] E') : ModelWithCorners 𝕜 E' H where toPartialEquiv := I.toPartialEquiv.trans e.toEquiv.toPartialEquiv source_eq := by simp - unique_diff' := by simp [range_comp e, I.unique_diff] + uniqueDiffOn' := by simp [range_comp e, I.uniqueDiffOn] continuous_toFun := e.continuous.comp I.continuous continuous_invFun := I.continuous_symm.comp e.symm.continuous diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index 5a704749a28a5..43cbf79926a58 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -136,7 +136,7 @@ def modelWithCornersEuclideanHalfSpace (n : ℕ) [NeZero n] : exact ⟨max_eq_left xprop, fun i _ => rfl⟩ right_inv' x hx := update_eq_iff.2 ⟨max_eq_left hx, fun i _ => rfl⟩ source_eq := rfl - unique_diff' := by + uniqueDiffOn' := by have : UniqueDiffOn ℝ _ := UniqueDiffOn.pi (Fin n) (fun _ => ℝ) _ _ fun i (_ : i ∈ ({0} : Set (Fin n))) => uniqueDiffOn_Ici 0 @@ -159,7 +159,7 @@ def modelWithCornersEuclideanQuadrant (n : ℕ) : left_inv' x _ := by ext i; simp only [Subtype.coe_mk, x.2 i, max_eq_left] right_inv' x hx := by ext1 i; simp only [hx i, max_eq_left] source_eq := rfl - unique_diff' := by + uniqueDiffOn' := by have this : UniqueDiffOn ℝ _ := UniqueDiffOn.univ_pi (Fin n) (fun _ => ℝ) _ fun _ => uniqueDiffOn_Ici 0 simpa only [pi_univ_Ici] using this diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean index f9b96b110f3fc..a5e16af8e340f 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean @@ -44,7 +44,7 @@ variable theorem uniqueMDiffWithinAt_univ : UniqueMDiffWithinAt I univ x := by unfold UniqueMDiffWithinAt simp only [preimage_univ, univ_inter] - exact I.unique_diff _ (mem_range_self _) + exact I.uniqueDiffOn _ (mem_range_self _) variable {I} diff --git a/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean b/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean index 75f5779ee5a46..5e0909b901206 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean @@ -324,7 +324,7 @@ theorem MDifferentiableAt.mfderiv_prod {f : M → M'} {g : M → M''} {x : M} classical simp_rw [mfderiv, if_pos (hf.prod_mk hg), if_pos hf, if_pos hg] exact hf.differentiableWithinAt_writtenInExtChartAt.fderivWithin_prod - hg.differentiableWithinAt_writtenInExtChartAt (I.unique_diff _ (mem_range_self _)) + hg.differentiableWithinAt_writtenInExtChartAt (I.uniqueDiffOn _ (mem_range_self _)) variable (I I' I'') diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 5ed443618c56c..9f6ab2d09bb29 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -138,7 +138,7 @@ structure ModelWithCorners (𝕜 : Type*) [NontriviallyNormedField 𝕜] (E : Ty [NormedAddCommGroup E] [NormedSpace 𝕜 E] (H : Type*) [TopologicalSpace H] extends PartialEquiv H E where source_eq : source = univ - unique_diff' : UniqueDiffOn 𝕜 toPartialEquiv.target + uniqueDiffOn' : UniqueDiffOn 𝕜 toPartialEquiv.target continuous_toFun : Continuous toFun := by continuity continuous_invFun : Continuous invFun := by continuity @@ -149,7 +149,7 @@ def modelWithCornersSelf (𝕜 : Type*) [NontriviallyNormedField 𝕜] (E : Type [NormedAddCommGroup E] [NormedSpace 𝕜 E] : ModelWithCorners 𝕜 E E where toPartialEquiv := PartialEquiv.refl E source_eq := rfl - unique_diff' := uniqueDiffOn_univ + uniqueDiffOn' := uniqueDiffOn_univ continuous_toFun := continuous_id continuous_invFun := continuous_id @@ -236,8 +236,11 @@ theorem target_eq : I.target = range (I : H → E) := by rw [← image_univ, ← I.source_eq] exact I.image_source_eq_target.symm -protected theorem unique_diff : UniqueDiffOn 𝕜 (range I) := - I.target_eq ▸ I.unique_diff' +protected theorem uniqueDiffOn : UniqueDiffOn 𝕜 (range I) := + I.target_eq ▸ I.uniqueDiffOn' + +@[deprecated (since := "2024-09-30")] +protected alias unique_diff := ModelWithCorners.uniqueDiffOn @[simp, mfld_simps] protected theorem left_inv (x : H) : I.symm (I x) = x := by refine I.left_inv' ?_; simp @@ -290,17 +293,26 @@ theorem symm_map_nhdsWithin_image {x : H} {s : Set H} : map I.symm (𝓝[I '' s] theorem symm_map_nhdsWithin_range (x : H) : map I.symm (𝓝[range I] I x) = 𝓝 x := by rw [← I.map_nhds_eq, map_map, I.symm_comp_self, map_id] -theorem unique_diff_preimage {s : Set H} (hs : IsOpen s) : +theorem uniqueDiffOn_preimage {s : Set H} (hs : IsOpen s) : UniqueDiffOn 𝕜 (I.symm ⁻¹' s ∩ range I) := by rw [inter_comm] - exact I.unique_diff.inter (hs.preimage I.continuous_invFun) + exact I.uniqueDiffOn.inter (hs.preimage I.continuous_invFun) + +@[deprecated (since := "2024-09-30")] +alias unique_diff_preimage := uniqueDiffOn_preimage -theorem unique_diff_preimage_source {β : Type*} [TopologicalSpace β] {e : PartialHomeomorph H β} : +theorem uniqueDiffOn_preimage_source {β : Type*} [TopologicalSpace β] {e : PartialHomeomorph H β} : UniqueDiffOn 𝕜 (I.symm ⁻¹' e.source ∩ range I) := - I.unique_diff_preimage e.open_source + I.uniqueDiffOn_preimage e.open_source + +@[deprecated (since := "2024-09-30")] +alias unique_diff_preimage_source := uniqueDiffOn_preimage_source + +theorem uniqueDiffWithinAt_image {x : H} : UniqueDiffWithinAt 𝕜 (range I) (I x) := + I.uniqueDiffOn _ (mem_range_self _) -theorem unique_diff_at_image {x : H} : UniqueDiffWithinAt 𝕜 (range I) (I x) := - I.unique_diff _ (mem_range_self _) +@[deprecated (since := "2024-09-30")] +alias unique_diff_at_image := uniqueDiffWithinAt_image theorem symm_continuousWithinAt_comp_right_iff {X} [TopologicalSpace X] {f : H → X} {s : Set H} {x : H} : @@ -369,7 +381,7 @@ def ModelWithCorners.prod {𝕜 : Type u} [NontriviallyNormedField 𝕜] {E : Ty invFun := fun x => (I.symm x.1, I'.symm x.2) source := { x | x.1 ∈ I.source ∧ x.2 ∈ I'.source } source_eq := by simp only [setOf_true, mfld_simps] - unique_diff' := I.unique_diff'.prod I'.unique_diff' + uniqueDiffOn' := I.uniqueDiffOn'.prod I'.uniqueDiffOn' continuous_toFun := I.continuous_toFun.prod_map I'.continuous_toFun continuous_invFun := I.continuous_invFun.prod_map I'.continuous_invFun } @@ -382,7 +394,7 @@ def ModelWithCorners.pi {𝕜 : Type u} [NontriviallyNormedField 𝕜] {ι : Typ ModelWithCorners 𝕜 (∀ i, E i) (ModelPi H) where toPartialEquiv := PartialEquiv.pi fun i => (I i).toPartialEquiv source_eq := by simp only [pi_univ, mfld_simps] - unique_diff' := UniqueDiffOn.pi ι E _ _ fun i _ => (I i).unique_diff' + uniqueDiffOn' := UniqueDiffOn.pi ι E _ _ fun i _ => (I i).uniqueDiffOn' continuous_toFun := continuous_pi fun i => (I i).continuous.comp (continuous_apply i) continuous_invFun := continuous_pi fun i => (I i).continuous_symm.comp (continuous_apply i) @@ -1057,7 +1069,7 @@ theorem extChartAt_target (x : M) : theorem uniqueDiffOn_extChartAt_target (x : M) : UniqueDiffOn 𝕜 (extChartAt I x).target := by rw [extChartAt_target] - exact I.unique_diff_preimage (chartAt H x).open_target + exact I.uniqueDiffOn_preimage (chartAt H x).open_target theorem uniqueDiffWithinAt_extChartAt_target (x : M) : UniqueDiffWithinAt 𝕜 (extChartAt I x).target (extChartAt I x x) := diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index fed71bc4299df..df7663e67c7ac 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -55,7 +55,7 @@ theorem contDiffOn_fderiv_coord_change (i j : atlas H M) : have h : ((i.1.extend I).symm ≫ j.1.extend I).source ⊆ range I := by rw [i.1.extend_coord_change_source]; apply image_subset_range intro x hx - refine (ContDiffWithinAt.fderivWithin_right ?_ I.unique_diff le_top <| h hx).mono h + refine (ContDiffWithinAt.fderivWithin_right ?_ I.uniqueDiffOn le_top <| h hx).mono h refine (PartialHomeomorph.contDiffOn_extend_coord_change I (subset_maximalAtlas I j.2) (subset_maximalAtlas I i.2) x hx).mono_of_mem ?_ exact i.1.extend_coord_change_source_mem_nhdsWithin j.1 I hx @@ -84,7 +84,7 @@ def tangentBundleCore : VectorBundleCore 𝕜 M E (atlas H M) where coordChange_self i x hx v := by simp only rw [Filter.EventuallyEq.fderivWithin_eq, fderivWithin_id', ContinuousLinearMap.id_apply] - · exact I.unique_diff_at_image + · exact I.uniqueDiffWithinAt_image · filter_upwards [i.1.extend_target_mem_nhdsWithin I hx] with y hy exact (i.1.extend I).right_inv hy · simp_rw [Function.comp_apply, i.1.extend_left_inv I hx] @@ -105,7 +105,7 @@ def tangentBundleCore : VectorBundleCore 𝕜 M E (atlas H M) where · exact (contDiffWithinAt_extend_coord_change' I (subset_maximalAtlas I j.2) (subset_maximalAtlas I i.2) hxj hxi).differentiableWithinAt le_top · intro x _; exact mem_range_self _ - · exact I.unique_diff_at_image + · exact I.uniqueDiffWithinAt_image · rw [Function.comp_apply, i.1.extend_left_inv I hxi] -- Porting note: moved to a separate `simp high` lemma b/c `simp` can simplify the LHS diff --git a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean index 988928ee6d456..b5db348da57a7 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean @@ -548,7 +548,7 @@ theorem mk'_spec' (x) (y : S) : f.toMap y * f.mk' x y = f.toMap x := by rw [mul_ @[to_additive] theorem eq_mk'_iff_mul_eq {x} {y : S} {z} : z = f.mk' x y ↔ z * f.toMap y = f.toMap x := - ⟨fun H ↦ by rw [H, mk'_spec], fun H ↦ by erw [mul_inv_right, H]⟩ + ⟨fun H ↦ by rw [H, mk'_spec], fun H ↦ by rw [mk', mul_inv_right, H]⟩ @[to_additive] theorem mk'_eq_iff_eq_mul {x} {y : S} {z} : f.mk' x y = z ↔ f.toMap x = z * f.toMap y := by @@ -773,9 +773,7 @@ theorem lift_comp : (f.lift hg).comp f.toMap = g := by ext; exact f.lift_eq hg _ @[to_additive (attr := simp)] theorem lift_of_comp (j : N →* P) : f.lift (f.isUnit_comp j) = j := by ext - rw [lift_spec] - show j _ = j _ * _ - erw [← j.map_mul, sec_spec'] + simp_rw [lift_spec, MonoidHom.comp_apply, ← j.map_mul, sec_spec'] @[to_additive] theorem epic_of_localizationMap {j k : N →* P} (h : ∀ a, j.comp f.toMap a = k.comp f.toMap a) : @@ -834,8 +832,8 @@ theorem lift_surjective_iff : obtain ⟨z, hz⟩ := H v obtain ⟨x, hx⟩ := f.surj z use x - rw [← hz, f.eq_mk'_iff_mul_eq.2 hx, lift_mk', mul_assoc, mul_comm _ (g ↑x.2)] - erw [IsUnit.mul_liftRight_inv (g.restrict S) hg, mul_one] + rw [← hz, f.eq_mk'_iff_mul_eq.2 hx, lift_mk', mul_assoc, mul_comm _ (g ↑x.2), + ← MonoidHom.restrict_apply, IsUnit.mul_liftRight_inv (g.restrict S) hg, mul_one] · intro H v obtain ⟨x, hx⟩ := H v use f.mk' x.1 x.2 @@ -1131,9 +1129,9 @@ of `AddCommMonoid`s, `k ∘ f` is a Localization map for `M` at `S`."] def ofMulEquivOfLocalizations (k : N ≃* P) : LocalizationMap S P := (k.toMonoidHom.comp f.toMap).toLocalizationMap (fun y ↦ isUnit_comp f k.toMonoidHom y) (fun v ↦ - let ⟨z, hz⟩ := k.toEquiv.surjective v + let ⟨z, hz⟩ := k.surjective v let ⟨x, hx⟩ := f.surj z - ⟨x, show v * k _ = k _ by rw [← hx, map_mul, ← hz]; rfl⟩) + ⟨x, show v * k _ = k _ by rw [← hx, map_mul, ← hz]⟩) fun x y ↦ (k.apply_eq_iff_eq.trans f.eq_iff_exists).1 @[to_additive (attr := simp)] @@ -1203,18 +1201,17 @@ def ofMulEquivOfDom {k : P ≃* M} (H : T.map k.toMonoidHom = S) : LocalizationM ⟨z, hz⟩) (fun z ↦ let ⟨x, hx⟩ := f.surj z - let ⟨v, hv⟩ := k.toEquiv.surjective x.1 - let ⟨w, hw⟩ := k.toEquiv.surjective x.2 - ⟨(v, ⟨w, H' ▸ show k w ∈ S from hw.symm ▸ x.2.2⟩), - show z * f.toMap (k.toEquiv w) = f.toMap (k.toEquiv v) by erw [hv, hw, hx]⟩) - fun x y ↦ - show f.toMap _ = f.toMap _ → _ by - erw [f.eq_iff_exists] - exact - fun ⟨c, hc⟩ ↦ - let ⟨d, hd⟩ := k.toEquiv.surjective c - ⟨⟨d, H' ▸ show k d ∈ S from hd.symm ▸ c.2⟩, by - erw [← hd, ← map_mul k, ← map_mul k] at hc; exact k.toEquiv.injective hc⟩ + let ⟨v, hv⟩ := k.surjective x.1 + let ⟨w, hw⟩ := k.surjective x.2 + ⟨(v, ⟨w, H' ▸ show k w ∈ S from hw.symm ▸ x.2.2⟩), by + simp_rw [MonoidHom.comp_apply, MulEquiv.toMonoidHom_eq_coe, MonoidHom.coe_coe, hv, hw, hx]⟩) + fun x y ↦ by + rw [MonoidHom.comp_apply, MonoidHom.comp_apply, MulEquiv.toMonoidHom_eq_coe, + MonoidHom.coe_coe, f.eq_iff_exists] + rintro ⟨c, hc⟩ + let ⟨d, hd⟩ := k.surjective c + refine ⟨⟨d, H' ▸ show k d ∈ S from hd.symm ▸ c.2⟩, ?_⟩ + rw [← hd, ← map_mul k, ← map_mul k] at hc; exact k.injective hc @[to_additive (attr := simp)] theorem ofMulEquivOfDom_apply {k : P ≃* M} (H : T.map k.toMonoidHom = S) (x) : diff --git a/Mathlib/Init/Algebra/Classes.lean b/Mathlib/Init/Algebra/Classes.lean index 98fbccbde3cf3..c7be944592e5c 100644 --- a/Mathlib/Init/Algebra/Classes.lean +++ b/Mathlib/Init/Algebra/Classes.lean @@ -57,7 +57,7 @@ instance (priority := 100) (α : Sort u) (lt : α → α → Prop) [IsStrictWeak section -variable {α : Sort u} {r : α → α → Prop} +variable {r : α → α → Prop} local infixl:50 " ≺ " => r @@ -83,7 +83,7 @@ namespace StrictWeakOrder section -variable {α : Sort u} {r : α → α → Prop} +variable {r : α → α → Prop} local infixl:50 " ≺ " => r diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index bba76de12bd95..ed68d3b7b7e44 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -591,8 +591,7 @@ theorem LinearIndependent.units_smul {v : ι → M} (hv : LinearIndependent R v) exact (hgs i hi).symm ▸ zero_smul _ _ · rw [← hsum, Finset.sum_congr rfl _] intros - erw [Pi.smul_apply, smul_assoc] - rfl + rw [Pi.smul_apply', smul_assoc, Units.smul_def] lemma LinearIndependent.eq_of_pair {x y : M} (h : LinearIndependent R ![x, y]) {s t s' t' : R} (h' : s • x + t • y = s' • x + t' • y) : s = s' ∧ t = t' := by diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/Eigs.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/Eigs.lean index ccd5c26310966..4eb330442505f 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/Eigs.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/Eigs.lean @@ -58,7 +58,7 @@ namespace Matrix theorem det_eq_prod_roots_charpoly_of_splits (hAps : A.charpoly.Splits (RingHom.id R)) : A.det = (Matrix.charpoly A).roots.prod := by rw [det_eq_sign_charpoly_coeff, ← charpoly_natDegree_eq_dim A, - Polynomial.prod_roots_eq_coeff_zero_of_monic_of_split A.charpoly_monic hAps, ← mul_assoc, + Polynomial.prod_roots_eq_coeff_zero_of_monic_of_splits A.charpoly_monic hAps, ← mul_assoc, ← pow_two, pow_right_comm, neg_one_sq, one_pow, one_mul] theorem trace_eq_sum_roots_charpoly_of_splits (hAps : A.charpoly.Splits (RingHom.id R)) : diff --git a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean index 1482f940cd82d..6ef3b1d0b480f 100644 --- a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean +++ b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean @@ -470,6 +470,10 @@ theorem nonsing_inv_nonsing_inv (h : IsUnit A.det) : A⁻¹⁻¹ = A := theorem isUnit_nonsing_inv_det_iff {A : Matrix n n α} : IsUnit A⁻¹.det ↔ IsUnit A.det := by rw [Matrix.det_nonsing_inv, isUnit_ring_inverse] +@[simp] +theorem isUnit_nonsing_inv_iff {A : Matrix n n α} : IsUnit A⁻¹ ↔ IsUnit A := by + simp_rw [isUnit_iff_isUnit_det, isUnit_nonsing_inv_det_iff] + -- `IsUnit.invertible` lifts the proposition `IsUnit A` to a constructive inverse of `A`. /-- A version of `Matrix.invertibleOfDetInvertible` with the inverse defeq to `A⁻¹` that is therefore noncomputable. -/ @@ -607,6 +611,24 @@ theorem inv_diagonal (v : n → α) : (diagonal v)⁻¹ = diagonal (Ring.inverse end Diagonal +/-- The inverse of a 1×1 or 0×0 matrix is always diagonal. + +While we could write this as `of fun _ _ => Ring.inverse (A default default)` on the RHS, this is +less useful because: + +* It wouldn't work for 0×0 matrices. +* More things are true about diagonal matrices than constant matrices, and so more lemmas exist. + +`Matrix.diagonal_unique` can be used to reach this form, while `Ring.inverse_eq_inv` can be used +to replace `Ring.inverse` with `⁻¹`. +-/ +@[simp] +theorem inv_subsingleton [Subsingleton m] [Fintype m] [DecidableEq m] (A : Matrix m m α) : + A⁻¹ = diagonal fun i => Ring.inverse (A i i) := by + rw [inv_def, adjugate_subsingleton, smul_one_eq_diagonal] + congr! with i + exact det_eq_elem_of_subsingleton _ _ + section Woodbury variable [Fintype m] [DecidableEq m] diff --git a/Mathlib/LinearAlgebra/Matrix/PosDef.lean b/Mathlib/LinearAlgebra/Matrix/PosDef.lean index 941af97a89e99..9343bc8070623 100644 --- a/Mathlib/LinearAlgebra/Matrix/PosDef.lean +++ b/Mathlib/LinearAlgebra/Matrix/PosDef.lean @@ -450,6 +450,25 @@ theorem det_pos [DecidableEq n] {M : Matrix n n 𝕜} (hM : M.PosDef) : 0 < det intro i _ simpa using hM.eigenvalues_pos i +theorem isUnit [DecidableEq n] {M : Matrix n n 𝕜} (hM : M.PosDef) : IsUnit M := + isUnit_iff_isUnit_det _ |>.2 <| hM.det_pos.ne'.isUnit + +protected theorem inv [DecidableEq n] {M : Matrix n n 𝕜} (hM : M.PosDef) : M⁻¹.PosDef := by + refine ⟨hM.isHermitian.inv, fun x hx => ?_⟩ + have := hM.2 (M⁻¹ *ᵥ x) ((Matrix.mulVec_injective_iff_isUnit.mpr ?_ |>.ne_iff' ?_).2 hx) + · let _inst := hM.isUnit.invertible + rwa [star_mulVec, mulVec_mulVec, Matrix.mul_inv_of_invertible, one_mulVec, + ← star_pos_iff, ← star_mulVec, ← star_dotProduct] at this + · simpa using hM.isUnit + · simp + +@[simp] +theorem _root_.Matrix.posDef_inv_iff [DecidableEq n] {M : Matrix n n 𝕜} : + M⁻¹.PosDef ↔ M.PosDef := + ⟨fun h => + letI := (Matrix.isUnit_nonsing_inv_iff.1 <| h.isUnit).invertible + Matrix.inv_inv_of_invertible M ▸ h.inv, (·.inv)⟩ + end PosDef end Matrix diff --git a/Mathlib/Logic/Function/FiberPartition.lean b/Mathlib/Logic/Function/FiberPartition.lean index e82895e5b975f..e1463b5d0e46b 100644 --- a/Mathlib/Logic/Function/FiberPartition.lean +++ b/Mathlib/Logic/Function/FiberPartition.lean @@ -62,7 +62,7 @@ lemma fiber_nonempty (f : Y → Z) (a : Fiber f) : Set.Nonempty a.val := by rw [mem_iff_eq_image, ← map_preimage_eq_image] lemma map_preimage_eq_image_map {W : Type*} (f : Y → Z) (g : Z → W) (a : Fiber (g ∘ f)) : - g (f a.preimage) = a.image := by rw [← map_preimage_eq_image]; rfl + g (f a.preimage) = a.image := by rw [← map_preimage_eq_image, comp_apply] lemma image_eq_image_mk (f : Y → Z) (g : X → Y) (a : Fiber (f ∘ g)) : a.image = (Fiber.mk f (g (a.preimage _))).image := by diff --git a/Mathlib/Logic/Godel/GodelBetaFunction.lean b/Mathlib/Logic/Godel/GodelBetaFunction.lean index 4fe1e162e2c5a..52ef094f4d92b 100644 --- a/Mathlib/Logic/Godel/GodelBetaFunction.lean +++ b/Mathlib/Logic/Godel/GodelBetaFunction.lean @@ -76,8 +76,7 @@ private lemma pairwise_coprime_coprimes (a : Fin m → ℕ) : Pairwise (Coprime have hja : j < supOfSeq a := lt_of_lt_of_le j.prop (le_step (le_max_left _ _)) exact coprime_mul_succ (Nat.succ_le_succ <| le_of_lt ltij) - (Nat.dvd_factorial - (by simp [Nat.succ_sub_succ, ltij]) + (Nat.dvd_factorial (by omega) (by simpa only [Nat.succ_sub_succ] using le_of_lt (lt_of_le_of_lt (sub_le j i) hja))) /-- Gödel's Beta Function. This is similar to `(Encodable.decodeList).get i`, but it is easier to diff --git a/Mathlib/Logic/Relator.lean b/Mathlib/Logic/Relator.lean index 83b0b7e2046b8..17ab1660850b2 100644 --- a/Mathlib/Logic/Relator.lean +++ b/Mathlib/Logic/Relator.lean @@ -116,7 +116,7 @@ lemma rel_eq {r : α → β → Prop} (hr : BiUnique r) : (r ⇒ r ⇒ (·↔·) open Function -variable {α : Type*} {r₁₁ : α → α → Prop} {r₁₂ : α → β → Prop} {r₂₁ : β → α → Prop} +variable {r₁₁ : α → α → Prop} {r₁₂ : α → β → Prop} {r₂₁ : β → α → Prop} {r₂₃ : β → γ → Prop} {r₁₃ : α → γ → Prop} namespace LeftTotal diff --git a/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean b/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean index db051f36a5bc3..0d8110dc76977 100644 --- a/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean @@ -127,10 +127,10 @@ theorem LocallyIntegrableOn.aestronglyMeasurable [SecondCountableTopology X] rw [this, aestronglyMeasurable_iUnion_iff] exact fun i : ℕ => (hu i).aestronglyMeasurable -/-- If `s` is either open, or closed, then `f` is locally integrable on `s` iff it is integrable on -every compact subset contained in `s`. -/ +/-- If `s` is locally closed (e.g. open or closed), then `f` is locally integrable on `s` iff it is +integrable on every compact subset contained in `s`. -/ theorem locallyIntegrableOn_iff [LocallyCompactSpace X] (hs : IsLocallyClosed s) : - LocallyIntegrableOn f s μ ↔ ∀ (k : Set X), k ⊆ s → (IsCompact k → IntegrableOn f k μ) := by + LocallyIntegrableOn f s μ ↔ ∀ (k : Set X), k ⊆ s → IsCompact k → IntegrableOn f k μ := by refine ⟨fun hf k hk ↦ hf.integrableOn_compact_subset hk, fun hf x hx ↦ ?_⟩ rcases hs with ⟨U, Z, hU, hZ, rfl⟩ rcases exists_compact_subset hU hx.1 with ⟨K, hK, hxK, hKU⟩ diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index fc342e5424d56..d7473d6e2d8df 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -1021,6 +1021,12 @@ theorem abs_integral_le_integral_abs (hab : a ≤ b) : |∫ x in a..b, f x ∂μ| ≤ ∫ x in a..b, |f x| ∂μ := by simpa only [← Real.norm_eq_abs] using norm_integral_le_integral_norm hab +lemma integral_pos (hab : a < b) + (hfc : ContinuousOn f (Icc a b)) (hle : ∀ x ∈ Ioc a b, 0 ≤ f x) (hlt : ∃ c ∈ Icc a b, 0 < f c) : + 0 < ∫ x in a..b, f x := + (integral_lt_integral_of_continuousOn_of_le_of_exists_lt hab + continuousOn_const hfc hle hlt).trans_eq' (by simp) + section Mono theorem integral_mono_interval {c d} (hca : c ≤ a) (hab : a ≤ b) (hbd : b ≤ d) diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index 4f9ff4c003ac1..54c09f7095a7a 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -745,10 +745,8 @@ theorem tendsto_addHaar_inter_smul_one_of_density_one_aux (s : Set E) (hs : Meas rw [← ENNReal.sub_mul]; swap · simp only [uzero, ENNReal.inv_eq_top, imp_true_iff, Ne, not_false_iff] congr 1 - apply - ENNReal.sub_eq_of_add_eq (ne_top_of_le_ne_top utop (measure_mono inter_subset_right)) - rw [inter_comm _ u, inter_comm _ u] - exact measure_inter_add_diff u vmeas + rw [inter_comm _ u, inter_comm _ u, eq_comm] + exact ENNReal.eq_sub_of_add_eq' utop (measure_inter_add_diff u vmeas) have L : Tendsto (fun r => μ (sᶜ ∩ closedBall x r) / μ (closedBall x r)) (𝓝[>] 0) (𝓝 0) := by have A : Tendsto (fun r => μ (closedBall x r) / μ (closedBall x r)) (𝓝[>] 0) (𝓝 1) := by apply tendsto_const_nhds.congr' _ diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 033d11b1cb4a5..b61c2e17ffc68 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -214,7 +214,7 @@ theorem measure_add_diff (hs : NullMeasurableSet s μ) (t : Set α) : theorem measure_diff' (s : Set α) (hm : NullMeasurableSet t μ) (h_fin : μ t ≠ ∞) : μ (s \ t) = μ (s ∪ t) - μ t := - Eq.symm <| ENNReal.sub_eq_of_add_eq h_fin <| by rw [add_comm, measure_add_diff hm, union_comm] + ENNReal.eq_sub_of_add_eq h_fin <| by rw [add_comm, measure_add_diff hm, union_comm] theorem measure_diff (h : s₂ ⊆ s₁) (h₂ : NullMeasurableSet s₂ μ) (h_fin : μ s₂ ≠ ∞) : μ (s₁ \ s₂) = μ s₁ - μ s₂ := by rw [measure_diff' _ h₂ h_fin, union_eq_self_of_subset_right h] diff --git a/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean b/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean index 291ecd8cf0105..ec687e378d87f 100644 --- a/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/SeparableMeasure.lean @@ -32,7 +32,7 @@ of separability in the metric space made by constant indicators equipped with th ## Main definitions -* `MeasureTheory.Measure.μ.MeasureDense 𝒜`: `𝒜` is a measure-dense family if it only contains +* `MeasureTheory.Measure.MeasureDense μ 𝒜`: `𝒜` is a measure-dense family if it only contains measurable sets and if the following condition is satisfied: if `s` is measurable with finite measure, then for any `ε > 0` there exists `t ∈ 𝒜` such that `μ (s ∆ t) < ε `. * `MeasureTheory.IsSeparable`: A measure is separable if there exists a countable and @@ -87,8 +87,7 @@ structure Measure.MeasureDense (μ : Measure X) (𝒜 : Set (Set X)) : Prop := /-- Each set has to be measurable. -/ measurable : ∀ s ∈ 𝒜, MeasurableSet s /-- Any measurable set can be approximated by sets in the family. -/ - approx : ∀ s, MeasurableSet s → μ s ≠ ∞ → ∀ (ε : ℝ), - 0 < ε → ∃ t ∈ 𝒜, μ (s ∆ t) < ENNReal.ofReal ε + approx : ∀ s, MeasurableSet s → μ s ≠ ∞ → ∀ ε : ℝ, 0 < ε → ∃ t ∈ 𝒜, μ (s ∆ t) < ENNReal.ofReal ε theorem Measure.MeasureDense.nonempty (h𝒜 : μ.MeasureDense 𝒜) : 𝒜.Nonempty := by rcases h𝒜.approx ∅ MeasurableSet.empty (by simp) 1 (by norm_num) with ⟨t, ht, -⟩ @@ -103,8 +102,8 @@ theorem Measure.MeasureDense.nonempty' (h𝒜 : μ.MeasureDense 𝒜) : /-- The set of measurable sets is measure-dense. -/ theorem measureDense_measurableSet : μ.MeasureDense {s | MeasurableSet s} where - measurable := fun _ h ↦ h - approx := fun s hs _ ε ε_pos ↦ ⟨s, hs, by simpa⟩ + measurable _ h := h + approx s hs _ ε ε_pos := ⟨s, hs, by simpa⟩ /-- If a family of sets `𝒜` is measure-dense in `X`, then any measurable set with finite measure can be approximated by sets in `𝒜` with finite measure. -/ @@ -155,9 +154,8 @@ theorem Measure.MeasureDense.indicatorConstLp_subset_closure (h𝒜 : μ.Measure with finite measure. -/ theorem Measure.MeasureDense.fin_meas (h𝒜 : μ.MeasureDense 𝒜) : μ.MeasureDense {s | s ∈ 𝒜 ∧ μ s ≠ ∞} where - measurable := fun s h ↦ h𝒜.measurable s h.1 - approx := by - intro s ms hμs ε ε_pos + measurable s h := h𝒜.measurable s h.1 + approx s ms hμs ε ε_pos := by rcases Measure.MeasureDense.fin_meas_approx h𝒜 ms hμs ε ε_pos with ⟨t, t_mem, hμt, hμst⟩ exact ⟨t, ⟨t_mem, hμt⟩, hμst⟩ @@ -165,12 +163,11 @@ theorem Measure.MeasureDense.fin_meas (h𝒜 : μ.MeasureDense 𝒜) : this algebra of sets is measure-dense. -/ theorem Measure.MeasureDense.of_generateFrom_isSetAlgebra_finite [IsFiniteMeasure μ] (h𝒜 : IsSetAlgebra 𝒜) (hgen : m = MeasurableSpace.generateFrom 𝒜) : μ.MeasureDense 𝒜 where - measurable := fun s hs ↦ hgen ▸ measurableSet_generateFrom hs - approx := by + measurable s hs := hgen ▸ measurableSet_generateFrom hs + approx s ms := by -- We want to show that any measurable set can be approximated by sets in `𝒜`. To do so, it is -- enough to show that such sets constitute a `σ`-algebra containing `𝒜`. This is contained in -- the theorem `generateFrom_induction`. - intro s ms have : MeasurableSet s ∧ ∀ (ε : ℝ), 0 < ε → ∃ t ∈ 𝒜, (μ (s ∆ t)).toReal < ε := by apply generateFrom_induction (p := fun s ↦ MeasurableSet s ∧ ∀ (ε : ℝ), 0 < ε → ∃ t ∈ 𝒜, (μ (s ∆ t)).toReal < ε) @@ -220,14 +217,14 @@ theorem Measure.MeasureDense.of_generateFrom_isSetAlgebra_finite [IsFiniteMeasur apply _root_.add_lt_add · rw [measure_diff (h_fin := measure_ne_top _ _), toReal_sub_of_le (ha := measure_ne_top _ _)] - apply lt_of_le_of_lt (sub_le_dist ..) - simp only [Finset.mem_range, Nat.lt_add_one_iff] - exact (dist_comm (α := ℝ) .. ▸ hN N (le_refl N)) - exact (measure_mono <| iUnion_subset <| - fun i ↦ iUnion_subset (fun _ ↦ subset_iUnion f i)) - exact iUnion_subset <| fun i ↦ iUnion_subset (fun _ ↦ subset_iUnion f i) - exact MeasurableSet.biUnion (countable_coe_iff.1 inferInstance) - (fun n _ ↦ (hf n).1.nullMeasurableSet) + · apply lt_of_le_of_lt (sub_le_dist ..) + simp only [Finset.mem_range, Nat.lt_add_one_iff] + exact (dist_comm (α := ℝ) .. ▸ hN N (le_refl N)) + · exact measure_mono <| iUnion_subset <| + fun i ↦ iUnion_subset fun _ ↦ subset_iUnion f i + · exact iUnion_subset <| fun i ↦ iUnion_subset (fun _ ↦ subset_iUnion f i) + · exact MeasurableSet.biUnion (countable_coe_iff.1 inferInstance) + (fun n _ ↦ (hf n).1.nullMeasurableSet) · calc (μ ((⋃ n ∈ (Finset.range (N + 1)), f n) ∆ (⋃ n ∈ (Finset.range (N + 1)), g ↑n))).toReal @@ -257,18 +254,17 @@ theorem Measure.MeasureDense.of_generateFrom_isSetAlgebra_sigmaFinite (h𝒜 : I (S : μ.FiniteSpanningSetsIn 𝒜) (hgen : m = MeasurableSpace.generateFrom 𝒜) : μ.MeasureDense 𝒜 where measurable s hs := hgen ▸ measurableSet_generateFrom hs - approx := by + approx s ms hμs ε ε_pos := by -- We use partial unions of (Sₙ) to get a monotone family spanning `X`. let T := Accumulate S.set - have T_mem : ∀ n, T n ∈ 𝒜 := fun n ↦ by + have T_mem (n) : T n ∈ 𝒜 := by simpa using h𝒜.biUnion_mem {k | k ≤ n}.toFinset (fun k _ ↦ S.set_mem k) - have T_finite : ∀ n, μ (T n) < ∞ := fun n ↦ by + have T_finite (n) : μ (T n) < ∞ := by simpa using measure_biUnion_lt_top {k | k ≤ n}.toFinset.finite_toSet (fun k _ ↦ S.finite k) have T_spanning : ⋃ n, T n = univ := S.spanning ▸ iUnion_accumulate -- We use the fact that we already know this is true for finite measures. As `⋃ n, T n = X`, -- we have that `μ ((T n) ∩ s) ⟶ μ s`. - intro s ms hμs ε ε_pos have mono : Monotone (fun n ↦ (T n) ∩ s) := fun m n hmn ↦ inter_subset_inter_left s (biUnion_subset_biUnion_left fun k hkm ↦ Nat.le_trans hkm hmn) have := tendsto_measure_iUnion_atTop (μ := μ) mono @@ -423,15 +419,14 @@ instance Lp.SecondCountableTopology [IsSeparable μ] [TopologicalSpace.Separable -- constant indicators with support in `𝒜₀`, and is denoted `D`. To make manipulations easier, -- we define the function `key` which given an integer `n` and two families of `n` elements -- in `u` and `𝒜₀` associates the corresponding sum. - let key : (n : ℕ) → (Fin n → u) → (Fin n → 𝒜₀) → (Lp E p μ) := - fun n d s ↦ ∑ i, indicatorConstLp p (h𝒜₀.measurable (s i) (Subtype.mem (s i))) - (s i).2.2 (d i : E) + let key (n : ℕ) (d : Fin n → u) (s : Fin n → 𝒜₀) : (Lp E p μ) := + ∑ i, indicatorConstLp p (h𝒜₀.measurable (s i) (Subtype.mem (s i))) (s i).2.2 (d i : E) let D := {s : Lp E p μ | ∃ n d t, s = key n d t} refine ⟨D, ?_, ?_⟩ · -- Countability directly follows from countability of `u` and `𝒜₀`. The function `f` below -- is the uncurryfied version of `key`, which is easier to manipulate as countability of the -- domain is automatically infered. - let f : (Σ n : ℕ, (Fin n → u) × (Fin n → 𝒜₀)) → Lp E p μ := fun nds ↦ key nds.1 nds.2.1 nds.2.2 + let f (nds : Σ n : ℕ, (Fin n → u) × (Fin n → 𝒜₀)) : Lp E p μ := key nds.1 nds.2.1 nds.2.2 have := count_𝒜₀.to_subtype have := countable_u.to_subtype have : D ⊆ range f := by diff --git a/Mathlib/NumberTheory/JacobiSum/Basic.lean b/Mathlib/NumberTheory/JacobiSum/Basic.lean index 5025b602f0542..1859c852cef47 100644 --- a/Mathlib/NumberTheory/JacobiSum/Basic.lean +++ b/Mathlib/NumberTheory/JacobiSum/Basic.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Stoll -/ import Mathlib.NumberTheory.GaussSum +import Mathlib.NumberTheory.MulChar.Lemmas +import Mathlib.RingTheory.RootsOfUnity.Lemmas /-! # Jacobi Sums @@ -226,3 +228,73 @@ lemma jacobiSum_mul_jacobiSum_inv (h : ringChar F' ≠ ringChar F) {χ φ : MulC ← mul_inv, gaussSum_mul_gaussSum_eq_card Hχφ ψ.prim] end field_field + +section image + +variable {F R : Type*} [Fintype F] [Field F] [CommRing R] [IsDomain R] + +/-- If `χ` and `φ` are multiplicative characters on a finite field `F` satisfying `χ^n = φ^n = 1` +and with values in an integral domain `R`, and `μ` is a primitive `n`th root of unity in `R`, +then the Jacobi sum `J(χ,φ)` is in `ℤ[μ] ⊆ R`. -/ +lemma jacobiSum_mem_algebraAdjoin_of_pow_eq_one {n : ℕ} (hn : n ≠ 0) {χ φ : MulChar F R} + (hχ : χ ^ n = 1) (hφ : φ ^ n = 1) {μ : R} (hμ : IsPrimitiveRoot μ n) : + jacobiSum χ φ ∈ Algebra.adjoin ℤ {μ} := + Subalgebra.sum_mem _ fun _ _ ↦ Subalgebra.mul_mem _ + (MulChar.apply_mem_algebraAdjoin_of_pow_eq_one hn hχ hμ _) + (MulChar.apply_mem_algebraAdjoin_of_pow_eq_one hn hφ hμ _) + +open Algebra in +private +lemma MulChar.exists_apply_sub_one_eq_mul_sub_one {n : ℕ} (hn : n ≠ 0) {χ : MulChar F R} {μ : R} + (hχ : χ ^ n = 1) (hμ : IsPrimitiveRoot μ n) {x : F} (hx : x ≠ 0) : + ∃ z ∈ Algebra.adjoin ℤ {μ}, χ x - 1 = z * (μ - 1) := by + obtain ⟨k, _, hk⟩ := exists_apply_eq_pow hn hχ hμ hx + refine hk ▸ ⟨(Finset.range k).sum (μ ^ ·), ?_, (geom_sum_mul μ k).symm⟩ + exact Subalgebra.sum_mem _ fun m _ ↦ Subalgebra.pow_mem _ (self_mem_adjoin_singleton _ μ) _ + +private +lemma MulChar.exists_apply_sub_one_mul_apply_sub_one {n : ℕ} (hn : n ≠ 0) {χ ψ : MulChar F R} + {μ : R} (hχ : χ ^ n = 1) (hψ : ψ ^ n = 1) (hμ : IsPrimitiveRoot μ n) (x : F) : + ∃ z ∈ Algebra.adjoin ℤ {μ}, (χ x - 1) * (ψ (1 - x) - 1) = z * (μ - 1) ^ 2 := by + rcases eq_or_ne x 0 with rfl | hx₀ + · exact ⟨0, Subalgebra.zero_mem _, by rw [sub_zero, ψ.map_one, sub_self, mul_zero, zero_mul]⟩ + rcases eq_or_ne x 1 with rfl | hx₁ + · exact ⟨0, Subalgebra.zero_mem _, by rw [χ.map_one, sub_self, zero_mul, zero_mul]⟩ + obtain ⟨z₁, hz₁, Hz₁⟩ := MulChar.exists_apply_sub_one_eq_mul_sub_one hn hχ hμ hx₀ + obtain ⟨z₂, hz₂, Hz₂⟩ := + MulChar.exists_apply_sub_one_eq_mul_sub_one hn hψ hμ (sub_ne_zero_of_ne hx₁.symm) + rewrite [Hz₁, Hz₂, sq] + exact ⟨z₁ * z₂, Subalgebra.mul_mem _ hz₁ hz₂, mul_mul_mul_comm ..⟩ + +/-- If `χ` and `ψ` are multiplicative characters of order dividing `n` on a finite field `F` +with values in an integral domain `R` and `μ` is a primitive `n`th root of unity in `R`, +then `J(χ,ψ) = -1 + z*(μ - 1)^2` for some `z ∈ ℤ[μ] ⊆ R`. (We assume that `#F ≡ 1 mod n`.) +Note that we do not state this as a divisbility in `R`, as this would give a weaker statement. -/ +lemma exists_jacobiSum_eq_neg_one_add [DecidableEq F] {n : ℕ} (hn : 2 < n) {χ ψ : MulChar F R} + {μ : R} (hχ : χ ^ n = 1) (hψ : ψ ^ n = 1) (hn' : n ∣ Fintype.card F - 1) + (hμ : IsPrimitiveRoot μ n) : + ∃ z ∈ Algebra.adjoin ℤ {μ}, jacobiSum χ ψ = -1 + z * (μ - 1) ^ 2 := by + obtain ⟨q, hq⟩ := hn' + rw [Nat.sub_eq_iff_eq_add NeZero.one_le] at hq + obtain ⟨z₁, hz₁, Hz₁⟩ := hμ.self_sub_one_pow_dvd_order hn + by_cases hχ₀ : χ = 1 <;> by_cases hψ₀ : ψ = 1 + · rw [hχ₀, hψ₀, jacobiSum_one_one] + refine ⟨q * z₁, Subalgebra.mul_mem _ (Subalgebra.natCast_mem _ q) hz₁, ?_⟩ + rw [hq, Nat.cast_add, Nat.cast_mul, Hz₁] + ring + · refine ⟨0, Subalgebra.zero_mem _, ?_⟩ + rw [hχ₀, jacobiSum_one_nontrivial hψ₀, zero_mul, add_zero] + · refine ⟨0, Subalgebra.zero_mem _, ?_⟩ + rw [jacobiSum_comm, hψ₀, jacobiSum_one_nontrivial hχ₀, zero_mul, add_zero] + · rw [jacobiSum_eq_aux, MulChar.sum_eq_zero_of_ne_one hχ₀, MulChar.sum_eq_zero_of_ne_one hψ₀, hq] + have H := MulChar.exists_apply_sub_one_mul_apply_sub_one (by omega) hχ hψ hμ + have Hcs x := (H x).choose_spec + refine ⟨-q * z₁ + ∑ x ∈ (univ \ {0, 1} : Finset F), (H x).choose, ?_, ?_⟩ + · refine Subalgebra.add_mem _ (Subalgebra.mul_mem _ (Subalgebra.neg_mem _ ?_) hz₁) ?_ + · exact Subalgebra.natCast_mem .. + · exact Subalgebra.sum_mem _ fun x _ ↦ (Hcs x).1 + · conv => enter [1, 2, 2, x]; rw [(Hcs x).2] + rw [← Finset.sum_mul, Nat.cast_add, Nat.cast_mul, Hz₁] + ring + +end image diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index 8e0ff1ca52eca..5970d297810cb 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -613,7 +613,7 @@ theorem mem_span_latticeBasis (x : (mixedSpace K)) : rfl theorem span_latticeBasis : - (Submodule.span ℤ (Set.range (latticeBasis K))) = (mixedEmbedding.integerLattice K) := + Submodule.span ℤ (Set.range (latticeBasis K)) = mixedEmbedding.integerLattice K := Submodule.ext_iff.mpr (mem_span_latticeBasis K) instance : DiscreteTopology (mixedEmbedding.integerLattice K) := by diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean index f9b153340e1c4..223002432594b 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean @@ -250,22 +250,22 @@ variable (K) in /-- The set of images by `mixedEmbedding` of algebraic integers of `K` contained in the fundamental cone. -/ def integralPoint : Set (mixedSpace K) := - fundamentalCone K ∩ (mixedEmbedding.integerLattice K) + fundamentalCone K ∩ mixedEmbedding.integerLattice K theorem mem_integralPoint {a : mixedSpace K} : - a ∈ integralPoint K ↔ a ∈ fundamentalCone K ∧ ∃ x : (𝓞 K), mixedEmbedding K x = a:= by + a ∈ integralPoint K ↔ a ∈ fundamentalCone K ∧ ∃ x : 𝓞 K, mixedEmbedding K x = a := by simp only [integralPoint, Set.mem_inter_iff, SetLike.mem_coe, LinearMap.mem_range, AlgHom.toLinearMap_apply, RingHom.toIntAlgHom_coe, RingHom.coe_comp, Function.comp_apply] /-- If `a` is an integral point, then there is a *unique* algebraic integer in `𝓞 K` such that `mixedEmbedding K x = a`. -/ theorem exists_unique_preimage_of_integralPoint {a : mixedSpace K} (ha : a ∈ integralPoint K) : - ∃! x : (𝓞 K), mixedEmbedding K x = a := by + ∃! x : 𝓞 K, mixedEmbedding K x = a := by obtain ⟨_, ⟨x, rfl⟩⟩ := mem_integralPoint.mp ha refine Function.Injective.existsUnique_of_mem_range ?_ (Set.mem_range_self x) exact (mixedEmbedding_injective K).comp RingOfIntegers.coe_injective -theorem integralPoint_ne_zero (a : integralPoint K) : (a : mixedSpace K) ≠ 0 := by +theorem integralPoint_ne_zero (a : integralPoint K) : (a : mixedSpace K) ≠ 0 := by by_contra! exact a.prop.1.2 (this.symm ▸ mixedEmbedding.norm.map_zero') diff --git a/Mathlib/Order/Filter/AtTopBot/ModEq.lean b/Mathlib/Order/Filter/AtTopBot/ModEq.lean index f7974d0c8d968..d45c17cc0d74b 100644 --- a/Mathlib/Order/Filter/AtTopBot/ModEq.lean +++ b/Mathlib/Order/Filter/AtTopBot/ModEq.lean @@ -3,7 +3,10 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Algebra.Order.Ring.Abs +import Mathlib.Algebra.Order.Ring.Basic +import Mathlib.Algebra.Order.Ring.Nat +import Mathlib.Algebra.Ring.Divisibility.Basic +import Mathlib.Algebra.Ring.Int import Mathlib.Data.Nat.ModEq import Mathlib.Order.Filter.AtTopBot.Monoid diff --git a/Mathlib/Order/Filter/Ultrafilter.lean b/Mathlib/Order/Filter/Ultrafilter.lean index 3d1ae9a87f5fc..6f238409f1d35 100644 --- a/Mathlib/Order/Filter/Ultrafilter.lean +++ b/Mathlib/Order/Filter/Ultrafilter.lean @@ -465,4 +465,8 @@ theorem ofComapInfPrincipal_eq_of_map (h : m '' s ∈ g) : (ofComapInfPrincipal _ ≤ ↑g ⊓ (𝓟 <| m '' s) := inf_le_inf_right _ map_comap_le _ = ↑g := inf_of_le_left (le_principal_iff.mpr h) +theorem eq_of_le_pure {X : Type _} {α : Filter X} (hα : α.NeBot) {x y : X} + (hx : α ≤ pure x) (hy : α ≤ pure y) : x = y := + Filter.pure_injective (hα.le_pure_iff.mp hx ▸ hα.le_pure_iff.mp hy) + end Ultrafilter diff --git a/Mathlib/Probability/Distributions/Gamma.lean b/Mathlib/Probability/Distributions/Gamma.lean index 0548d04bfd3c0..562c9b3c3b6f9 100644 --- a/Mathlib/Probability/Distributions/Gamma.lean +++ b/Mathlib/Probability/Distributions/Gamma.lean @@ -30,8 +30,8 @@ open MeasureTheory Real Set Filter Topology lemma lintegral_Iic_eq_lintegral_Iio_add_Icc {y z : ℝ} (f : ℝ → ℝ≥0∞) (hzy : z ≤ y) : ∫⁻ x in Iic y, f x = (∫⁻ x in Iio z, f x) + ∫⁻ x in Icc z y, f x := by rw [← Iio_union_Icc_eq_Iic hzy, lintegral_union measurableSet_Icc] - rw [Set.disjoint_iff] - rintro x ⟨h1 : x < _, h2, _⟩ + simp_rw [Set.disjoint_iff_forall_ne, mem_Iio, mem_Icc] + intros linarith namespace ProbabilityTheory @@ -49,8 +49,9 @@ def gammaPDF (a r x : ℝ) : ℝ≥0∞ := ENNReal.ofReal (gammaPDFReal a r x) lemma gammaPDF_eq (a r x : ℝ) : - gammaPDF a r x = ENNReal.ofReal (if 0 ≤ x then - r ^ a / (Gamma a) * x ^ (a-1) * exp (-(r * x)) else 0) := rfl + gammaPDF a r x = + ENNReal.ofReal (if 0 ≤ x then r ^ a / (Gamma a) * x ^ (a-1) * exp (-(r * x)) else 0) := + rfl lemma gammaPDF_of_neg {a r x : ℝ} (hx : x < 0) : gammaPDF a r x = 0 := by simp only [gammaPDF_eq, if_neg (not_le.mpr hx), ENNReal.ofReal_zero] diff --git a/Mathlib/RingTheory/Congruence/Basic.lean b/Mathlib/RingTheory/Congruence/Basic.lean index 1a9c8ef824451..c6c46c89d1b9b 100644 --- a/Mathlib/RingTheory/Congruence/Basic.lean +++ b/Mathlib/RingTheory/Congruence/Basic.lean @@ -126,6 +126,16 @@ theorem ext' {c d : RingCon R} (H : ⇑c = ⇑d) : c = d := DFunLike.coe_injecti theorem ext {c d : RingCon R} (H : ∀ x y, c x y ↔ d x y) : c = d := ext' <| by ext; apply H +/-- +Pulling back a `RingCon` across a ring homomorphism. +-/ +def comap {R R' F : Type*} [Add R] [Add R'] + [FunLike F R R'] [AddHomClass F R R'] [Mul R] [Mul R'] [MulHomClass F R R'] + (J : RingCon R') (f : F) : + RingCon R where + __ := J.toCon.comap f (map_mul f) + __ := J.toAddCon.comap f (map_add f) + end Basic section Quotient diff --git a/Mathlib/RingTheory/HahnSeries/Multiplication.lean b/Mathlib/RingTheory/HahnSeries/Multiplication.lean index 327200bad91d8..153ac24993c5d 100644 --- a/Mathlib/RingTheory/HahnSeries/Multiplication.lean +++ b/Mathlib/RingTheory/HahnSeries/Multiplication.lean @@ -231,7 +231,7 @@ theorem add_smul [AddCommMonoid R] [SMulWithZero R V] {x y : HahnSeries Γ R} rw [smul_coeff_left hwf, HahnSeries.add_coeff', of_symm_add] · simp_all only [Pi.add_apply, HahnSeries.add_coeff'] rw [smul_coeff_left hwf Set.subset_union_right, - smul_coeff_left hwf Set.subset_union_left] + smul_coeff_left hwf Set.subset_union_left] simp only [HahnSeries.add_coeff, h, sum_add_distrib] · intro b simp_all only [Set.isPWO_union, HahnSeries.isPWO_support, and_self, HahnSeries.mem_support, diff --git a/Mathlib/RingTheory/IsTensorProduct.lean b/Mathlib/RingTheory/IsTensorProduct.lean index d146835dbbfcd..b02662937d43f 100644 --- a/Mathlib/RingTheory/IsTensorProduct.lean +++ b/Mathlib/RingTheory/IsTensorProduct.lean @@ -318,11 +318,40 @@ theorem IsBaseChange.comp {f : M →ₗ[R] N} (hf : IsBaseChange S f) {g : N → ext rfl +/-- If `N` is the base change of `M` to `S` and `O` the base change of `M` to `T`, then +`O` is the base change of `N` to `T`. -/ +lemma IsBaseChange.of_comp {f : M →ₗ[R] N} (hf : IsBaseChange S f) {h : N →ₗ[S] O} + (hc : IsBaseChange T ((h : N →ₗ[R] O) ∘ₗ f)) : + IsBaseChange T h := by + apply IsBaseChange.of_lift_unique + intro Q _ _ _ _ r + letI : Module R Q := inferInstanceAs (Module R (RestrictScalars R S Q)) + haveI : IsScalarTower R S Q := IsScalarTower.of_algebraMap_smul fun r ↦ congrFun rfl + haveI : IsScalarTower R T Q := IsScalarTower.of_algebraMap_smul fun r x ↦ by + simp [IsScalarTower.algebraMap_apply R S T] + let r' : M →ₗ[R] Q := r ∘ₗ f + let q : O →ₗ[T] Q := hc.lift r' + refine ⟨q, ?_, ?_⟩ + · apply hf.algHom_ext' + simp [LinearMap.comp_assoc, hc.lift_comp] + · intro q' hq' + apply hc.algHom_ext' + apply_fun LinearMap.restrictScalars R at hq' + rw [← LinearMap.comp_assoc] + rw [show q'.restrictScalars R ∘ₗ h.restrictScalars R = _ from hq', hc.lift_comp] + +/-- If `N` is the base change `M` to `S`, then `O` is the base change of `M` to `T` if and +only if `O` is the base change of `N` to `T`. -/ +lemma IsBaseChange.comp_iff {f : M →ₗ[R] N} (hf : IsBaseChange S f) {h : N →ₗ[S] O} : + IsBaseChange T ((h : N →ₗ[R] O) ∘ₗ f) ↔ IsBaseChange T h := + ⟨fun hc ↦ IsBaseChange.of_comp hf hc, fun hh ↦ IsBaseChange.comp hf hh⟩ + variable {R' S' : Type*} [CommSemiring R'] [CommSemiring S'] variable [Algebra R R'] [Algebra S S'] [Algebra R' S'] [Algebra R S'] variable [IsScalarTower R R' S'] [IsScalarTower R S S'] open IsScalarTower (toAlgHom) +open IsScalarTower (algebraMap_apply) variable (R S R' S') @@ -471,4 +500,33 @@ theorem Algebra.IsPushout.algHom_ext [H : Algebra.IsPushout R S R' S'] {A : Type · intro s₁ s₂ e₁ e₂ rw [map_add, map_add, e₁, e₂] +/-- +Let the following be a commutative diagram of rings +``` + R → S → T + ↓ ↓ ↓ + R' → S' → T' +``` +where the left-hand square is a pushout. Then the following are equivalent: +- the big rectangle is a pushout. +- the right-hand square is a pushout. + +Note that this is essentially the isomorphism `T ⊗[S] (S ⊗[R] R') ≃ₐ[T] T ⊗[R] R'`. +-/ +lemma Algebra.IsPushout.comp_iff {T' : Type*} [CommRing T'] [Algebra R T'] + [Algebra S' T'] [Algebra S T'] [Algebra T T'] [Algebra R' T'] + [IsScalarTower R T T'] [IsScalarTower S T T'] [IsScalarTower S S' T'] + [IsScalarTower R R' T'] [IsScalarTower R S' T'] [IsScalarTower R' S' T'] + [Algebra.IsPushout R S R' S'] : + Algebra.IsPushout R T R' T' ↔ Algebra.IsPushout S T S' T' := by + let f : R' →ₗ[R] S' := (IsScalarTower.toAlgHom R R' S').toLinearMap + haveI : IsScalarTower R S T' := IsScalarTower.of_algebraMap_eq <| fun x ↦ by + rw [algebraMap_apply R S' T', algebraMap_apply R S S', ← algebraMap_apply S S' T'] + have heq : (toAlgHom S S' T').toLinearMap.restrictScalars R ∘ₗ f = + (toAlgHom R R' T').toLinearMap := by + ext x + simp [f, ← IsScalarTower.algebraMap_apply] + rw [isPushout_iff, isPushout_iff, ← heq, IsBaseChange.comp_iff] + exact Algebra.IsPushout.out + end IsBaseChange diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index e2ae574e5cae7..669be5c1d5cd1 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -11,6 +11,7 @@ import Mathlib.RingTheory.HahnSeries.Summable import Mathlib.RingTheory.PowerSeries.Inverse import Mathlib.FieldTheory.RatFunc.AsPolynomial import Mathlib.RingTheory.Localization.FractionRing +import Mathlib.Topology.UniformSpace.Cauchy /-! # Laurent Series @@ -26,16 +27,23 @@ import Mathlib.RingTheory.Localization.FractionRing * Embedding of rational functions into Laurent series, provided as a coercion, utilizing the underlying `RatFunc.coeAlgHom`. * Study of the `X`-Adic valuation on the ring of Laurent series over a field +* In `LaurentSeries.uniformContinuous_coeff` we show that sending a Laurent series to its `d`th +coefficient is uniformly continuous, ensuring that it sends a Cauchy filter `ℱ` in `LaurentSeries K` +to a Cauchy filter in `K`: since this latter is given the discrete topology, this provides an +element `LaurentSeries.Cauchy.coeff ℱ d` in `K` that serves as `d`th coefficient of the Laurent +series to which the filter `ℱ` converges. ## Main Results * Basic properties of Hasse derivatives ### About the `X`-Adic valuation: * The (integral) valuation of a power series is the order of the first non-zero coefficient, see -`intValuation_le_iff_coeff_lt_eq_zero`. +`LaurentSeries.intValuation_le_iff_coeff_lt_eq_zero`. * The valuation of a Laurent series is the order of the first non-zero coefficient, see -`valuation_le_iff_coeff_lt_eq_zero`. +`LaurentSeries.valuation_le_iff_coeff_lt_eq_zero`. * Every Laurent series of valuation less than `(1 : ℤₘ₀)` comes from a power series, see -`val_le_one_iff_eq_coe`. +`LaurentSeries.val_le_one_iff_eq_coe`. +* The uniform space of `LaurentSeries` over a field is complete, formalized in the instance +`instLaurentSeriesComplete`. ## Implementation details * Since `LaurentSeries` is just an abbreviation of `HahnSeries ℤ _`, the definition of the @@ -464,6 +472,7 @@ end RatFunc namespace LaurentSeries + open IsDedekindDomain.HeightOneSpectrum PowerSeries RatFunc instance : Valued (LaurentSeries K) ℤₘ₀ := Valued.mk' (PowerSeries.idealX K).valuation @@ -609,3 +618,169 @@ theorem val_le_one_iff_eq_coe (f : LaurentSeries K) : Valued.v f ≤ (1 : ℤₘ end LaurentSeries end AdicValuation +namespace LaurentSeries +section Complete + +open Filter + +open scoped Multiplicative + +variable {K : Type*} [Field K] + +/- Sending a Laurent series to its `d`-th coefficient is uniformly continuous (independently of the + uniformity with which `K` is endowed). -/ +theorem uniformContinuous_coeff {uK : UniformSpace K} (d : ℤ) : + UniformContinuous fun f : LaurentSeries K ↦ f.coeff d := by + refine uniformContinuous_iff_eventually.mpr fun S hS ↦ eventually_iff_exists_mem.mpr ?_ + let γ : ℤₘ₀ˣ := Units.mk0 (↑(Multiplicative.ofAdd (-(d + 1)))) WithZero.coe_ne_zero + use {P | Valued.v (P.snd - P.fst) < ↑γ} + refine ⟨(Valued.hasBasis_uniformity (LaurentSeries K) ℤₘ₀).mem_of_mem (by tauto), fun P hP ↦ ?_⟩ + rw [eq_coeff_of_valuation_sub_lt K (le_of_lt hP) (lt_add_one _)] + exact mem_uniformity_of_eq hS rfl + +/-- Since extracting coefficients is uniformly continuous, every Cauchy filter in +`laurentSeries K` gives rise to a Cauchy filter in `K` for every `d : ℤ`, and such Cauchy filter +in `K` converges to a principal filter -/ +def Cauchy.coeff {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) : ℤ → K := + let _ : UniformSpace K := ⊥ + fun d ↦ UniformSpace.DiscreteUnif.cauchyConst rfl <| hℱ.map (uniformContinuous_coeff d) + +theorem Cauchy.coeff_tendsto {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) (D : ℤ) : + Tendsto (fun f : LaurentSeries K ↦ f.coeff D) ℱ (𝓟 {coeff hℱ D}) := + let _ : UniformSpace K := ⊥ + le_of_eq <| UniformSpace.DiscreteUnif.eq_const_of_cauchy (by rfl) + (hℱ.map (uniformContinuous_coeff D)) ▸ (principal_singleton _).symm + +/- For every Cauchy filter of Laurent series, there is a `N` such that the `n`-th coefficient +vanishes for all `n ≤ N` and almost all series in the filter. This is an auxiliary lemma used +to construct the limit of the Cauchy filter as a Laurent series, ensuring that the support of the +limit is `PWO`. +The result is true also for more general Hahn Series indexed over a partially ordered group `Γ` +beyond the special case `Γ = ℤ`, that corresponds to Laurent Series: nevertheless the proof below +does not generalise, as it relies on the study of the `X`-adic valuation attached to the height-one +prime `X`, and this is peculiar to the one-variable setting. In the future we should prove this +result in full generality and deduce the case `Γ = ℤ` from that one.-/ +lemma Cauchy.exists_lb_eventual_support {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) : + ∃ N, ∀ᶠ f : LaurentSeries K in ℱ, ∀ n < N, f.coeff n = (0 : K) := by + let entourage : Set (LaurentSeries K × LaurentSeries K) := + {P : LaurentSeries K × LaurentSeries K | + Valued.v (P.snd - P.fst) < ((Multiplicative.ofAdd 0 : Multiplicative ℤ) : ℤₘ₀)} + let ζ := Units.mk0 (G₀ := ℤₘ₀) _ (WithZero.coe_ne_zero (a := (Multiplicative.ofAdd 0))) + obtain ⟨S, ⟨hS, ⟨T, ⟨hT, H⟩⟩⟩⟩ := mem_prod_iff.mp <| Filter.le_def.mp hℱ.2 entourage + <| (Valued.hasBasis_uniformity (LaurentSeries K) ℤₘ₀).mem_of_mem (i := ζ) (by tauto) + obtain ⟨f, hf⟩ := forall_mem_nonempty_iff_neBot.mpr hℱ.1 (S ∩ T) (inter_mem_iff.mpr ⟨hS, hT⟩) + obtain ⟨N, hN⟩ : ∃ N : ℤ, ∀ g : LaurentSeries K, + Valued.v (g - f) ≤ ↑(Multiplicative.ofAdd (0 : ℤ)) → ∀ n < N, g.coeff n = 0 := by + by_cases hf : f = 0 + · refine ⟨0, fun x hg ↦ ?_⟩ + rw [hf, sub_zero] at hg + exact (valuation_le_iff_coeff_lt_eq_zero K).mp hg + · refine ⟨min (f.2.isWF.min (HahnSeries.support_nonempty_iff.mpr hf)) 0 - 1, fun _ hg n hn ↦ ?_⟩ + rw [eq_coeff_of_valuation_sub_lt K hg (d := 0)] + · exact Function.nmem_support.mp fun h ↦ + f.2.isWF.not_lt_min (HahnSeries.support_nonempty_iff.mpr hf) h + <| lt_trans hn <| Int.sub_one_lt_iff.mpr <| min_le_left _ _ + exact lt_of_lt_of_le hn <| le_of_lt (Int.sub_one_lt_of_le <| min_le_right _ _) + use N + apply mem_of_superset (inter_mem hS hT) + intro g hg + have h_prod : (f, g) ∈ entourage := Set.prod_mono (Set.inter_subset_left (t := T)) + (Set.inter_subset_right (s := S)) |>.trans H <| Set.mem_prod.mpr ⟨hf, hg⟩ + exact hN g (le_of_lt h_prod) + +/- The support of `Cauchy.coeff` has a lower bound. -/ +theorem Cauchy.exists_lb_support {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) : + ∃ N, ∀ n, n < N → coeff hℱ n = 0 := by + let _ : UniformSpace K := ⊥ + obtain ⟨N, hN⟩ := exists_lb_eventual_support hℱ + refine ⟨N, fun n hn ↦ Ultrafilter.eq_of_le_pure (hℱ.map (uniformContinuous_coeff n)).1 + ((principal_singleton _).symm ▸ coeff_tendsto _ _) ?_⟩ + simp only [pure_zero, nonpos_iff] + apply Filter.mem_of_superset hN (fun _ ha ↦ ha _ hn) + +/- The support of `Cauchy.coeff` is bounded below -/ +theorem Cauchy.coeff_support_bddBelow {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) : + BddBelow (coeff hℱ).support := by + refine ⟨(exists_lb_support hℱ).choose, fun d hd ↦ ?_⟩ + by_contra hNd + exact hd ((exists_lb_support hℱ).choose_spec d (not_le.mp hNd)) + +/-- To any Cauchy filter ℱ of `LaurentSeries K`, we can attach a laurent series that is the limit +of the filter. Its `d`-th coefficient is defined as the limit of `Cauchy.coeff hℱ d`, which is +again Cauchy but valued in the discrete space `K`. That sufficiently negative coefficients vanish +follows from `Cauchy.coeff_support_bddBelow` -/ +def Cauchy.limit {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) : LaurentSeries K := + HahnSeries.mk (coeff hℱ) <| Set.IsWF.isPWO (coeff_support_bddBelow _).wellFoundedOn_lt + +/- The following lemma shows that for every `d` smaller than the minimum between the integers +produced in `Cauchy.exists_lb_eventual_support` and `Cauchy.exists_lb_support`, for almost all +series in `ℱ` the `d`th coefficient coincides with the `d`th coefficient of `Cauchy.coeff hℱ`. -/ +theorem Cauchy.exists_lb_coeff_ne {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) : + ∃ N, ∀ᶠ f : LaurentSeries K in ℱ, ∀ d < N, coeff hℱ d = f.coeff d := by + obtain ⟨⟨N₁, hN₁⟩, ⟨N₂, hN₂⟩⟩ := exists_lb_eventual_support hℱ, exists_lb_support hℱ + refine ⟨min N₁ N₂, ℱ.3 hN₁ fun _ hf d hd ↦ ?_⟩ + rw [hf d (lt_of_lt_of_le hd (min_le_left _ _)), hN₂ d (lt_of_lt_of_le hd (min_le_right _ _))] + +/- Given a Cauchy filter `ℱ` in the Laurent Series and a bound `D`, for almost all series in the +filter the coefficients below `D` coincide with `Caucy.coeff hℱ`-/ +theorem Cauchy.coeff_eventually_equal {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) {D : ℤ} : + ∀ᶠ f : LaurentSeries K in ℱ, ∀ d, d < D → coeff hℱ d = f.coeff d := by + -- `φ` sends `d` to the set of Laurent Series having `d`th coefficient equal to `ℱ.coeff`. + let φ : ℤ → Set (LaurentSeries K) := fun d ↦ {f | coeff hℱ d = f.coeff d} + have intersec₁ : + (⋂ n ∈ Set.Iio D, φ n) ⊆ {x : LaurentSeries K | ∀ d : ℤ, d < D → coeff hℱ d = x.coeff d} := by + intro _ hf + simpa only [Set.mem_iInter] using hf + -- The goal is now to show that the intersection of all `φ d` (for `d < D`) is in `ℱ`. + let ℓ := (exists_lb_coeff_ne hℱ).choose + let N := max ℓ D + have intersec₂ : ⋂ n ∈ Set.Iio D, φ n ⊇ (⋂ n ∈ Set.Iio ℓ, φ n) ∩ (⋂ n ∈ Set.Icc ℓ N, φ n) := by + simp only [Set.mem_Iio, Set.mem_Icc, Set.subset_iInter_iff] + intro i hi x hx + simp only [Set.mem_inter_iff, Set.mem_iInter, and_imp] at hx + by_cases H : i < ℓ + exacts [hx.1 _ H, hx.2 _ (le_of_not_lt H) <| le_of_lt <| lt_max_of_lt_right hi] + suffices (⋂ n ∈ Set.Iio ℓ, φ n) ∩ (⋂ n ∈ Set.Icc ℓ N, φ n) ∈ ℱ by + exact ℱ.sets_of_superset this <| intersec₂.trans intersec₁ + /- To show that the intersection we have in sight is in `ℱ`, we use that it contains a double + intersection (an infinite and a finite one): by general properties of filters, we are reduced + to show that both terms are in `ℱ`, which is easy in light of their definition. -/ + · simp only [Set.mem_Iio, Set.mem_Ico, inter_mem_iff] + constructor + · have := (exists_lb_coeff_ne hℱ).choose_spec + rw [Filter.eventually_iff] at this + convert this + ext + simp only [Set.mem_iInter, Set.mem_setOf_eq]; rfl + · rw [biInter_mem (Set.finite_Icc ℓ N)] + intro _ _ + apply coeff_tendsto hℱ + simp only [principal_singleton, mem_pure]; rfl + + +open scoped Topology + +/- The main result showing that the Cauchy filter tends to the `Cauchy.limit`-/ +theorem Cauchy.eventually_mem_nhds {ℱ : Filter (LaurentSeries K)} (hℱ : Cauchy ℱ) + {U : Set (LaurentSeries K)} (hU : U ∈ 𝓝 (Cauchy.limit hℱ)) : ∀ᶠ f in ℱ, f ∈ U := by + obtain ⟨γ, hU₁⟩ := Valued.mem_nhds.mp hU + suffices ∀ᶠ f in ℱ, f ∈ {y : LaurentSeries K | Valued.v (y - limit hℱ) < ↑γ} by + apply this.mono fun _ hf ↦ hU₁ hf + set D := -(Multiplicative.toAdd (WithZero.unzero γ.ne_zero) - 1) with hD₀ + have hD : ((Multiplicative.ofAdd (-D) : Multiplicative ℤ) : ℤₘ₀) < γ := by + rw [← WithZero.coe_unzero γ.ne_zero, WithZero.coe_lt_coe, hD₀, neg_neg, ofAdd_sub, + ofAdd_toAdd, div_lt_comm, div_self', ← ofAdd_zero, Multiplicative.ofAdd_lt] + exact zero_lt_one + apply coeff_eventually_equal hℱ |>.mono + intro _ hf + apply lt_of_le_of_lt (valuation_le_iff_coeff_lt_eq_zero K |>.mpr _) hD + intro n hn + rw [HahnSeries.sub_coeff, sub_eq_zero, hf n hn |>.symm]; rfl + +/- Laurent Series with coefficients in a field are complete w.r.t. the `X`-adic valuation -/ +instance instLaurentSeriesComplete : CompleteSpace (LaurentSeries K) := + ⟨fun hℱ ↦ ⟨Cauchy.limit hℱ, fun _ hS ↦ Cauchy.eventually_mem_nhds hℱ hS⟩⟩ + +end Complete + +end LaurentSeries diff --git a/Mathlib/RingTheory/MvPowerSeries/LexOrder.lean b/Mathlib/RingTheory/MvPowerSeries/LexOrder.lean index a714c9a16fc57..13ecad6b47dce 100644 --- a/Mathlib/RingTheory/MvPowerSeries/LexOrder.lean +++ b/Mathlib/RingTheory/MvPowerSeries/LexOrder.lean @@ -104,13 +104,10 @@ theorem le_lexOrder_iff {φ : MvPowerSeries σ R} {w : WithTop (Lex (σ →₀ intro h' have hφ : φ ≠ 0 := by rw [ne_eq, ← lexOrder_eq_top_iff_eq_zero] - intro h'' - rw [h'', ← not_le] at h' - apply h' - exact le_top + exact ne_top_of_lt h' obtain ⟨d, hd⟩ := exists_finsupp_eq_lexOrder_of_ne_zero hφ refine coeff_ne_zero_of_lexOrder hd.symm (h d ?_) - exact (lt_of_eq_of_lt hd.symm h') + rwa [← hd] theorem min_lexOrder_le {φ ψ : MvPowerSeries σ R} : min (lexOrder φ) (lexOrder ψ) ≤ lexOrder (φ + ψ) := by diff --git a/Mathlib/RingTheory/Norm/Basic.lean b/Mathlib/RingTheory/Norm/Basic.lean index 98235e2bf157f..142c31df3e1f0 100644 --- a/Mathlib/RingTheory/Norm/Basic.lean +++ b/Mathlib/RingTheory/Norm/Basic.lean @@ -74,7 +74,7 @@ theorem PowerBasis.norm_gen_eq_prod_roots [Algebra R F] (pb : PowerBasis R S) have := minpoly.monic pb.isIntegral_gen rw [PowerBasis.norm_gen_eq_coeff_zero_minpoly, ← pb.natDegree_minpoly, RingHom.map_mul, ← coeff_map, - prod_roots_eq_coeff_zero_of_monic_of_split (this.map _) ((splits_id_iff_splits _).2 hf), + prod_roots_eq_coeff_zero_of_monic_of_splits (this.map _) ((splits_id_iff_splits _).2 hf), this.natDegree_map, map_pow, ← mul_assoc, ← mul_pow] simp only [map_neg, _root_.map_one, neg_mul, neg_neg, one_pow, one_mul] diff --git a/Mathlib/RingTheory/Presentation.lean b/Mathlib/RingTheory/Presentation.lean index 1fbea429322dc..2cdda618659b5 100644 --- a/Mathlib/RingTheory/Presentation.lean +++ b/Mathlib/RingTheory/Presentation.lean @@ -179,6 +179,7 @@ private lemma span_range_relation_eq_ker_localizationAway : show Ideal.span {C r * X () - 1} = Ideal.comap _ (RingHom.ker (mvPolynomialQuotientEquiv S r)) simp [RingHom.ker_equiv, ← RingHom.ker_eq_comap_bot] +variable (S) in /-- If `S` is the localization of `R` away from `r`, we can construct a natural presentation of `S` as `R`-algebra with a single generator `X` and the relation `r * X - 1 = 0`. -/ @[simps relation, simps (config := .lemmasOnly) rels] @@ -190,12 +191,15 @@ noncomputable def localizationAway : Presentation R S where simp only [Generators.localizationAway_vars, Set.range_const] apply span_range_relation_eq_ker_localizationAway r -instance localizationAway_isFinite : (localizationAway r (S := S)).IsFinite where +instance localizationAway_isFinite : (localizationAway S r).IsFinite where finite_vars := inferInstanceAs <| Finite Unit finite_rels := inferInstanceAs <| Finite Unit +instance : Fintype (localizationAway S r).rels := + inferInstanceAs (Fintype Unit) + @[simp] -lemma localizationAway_dimension_zero : (localizationAway r (S := S)).dimension = 0 := by +lemma localizationAway_dimension_zero : (localizationAway S r).dimension = 0 := by simp [Presentation.dimension, localizationAway, Generators.localizationAway_vars] end Localization @@ -408,7 +412,12 @@ noncomputable def comp : Presentation R T where (fun rp ↦ MvPolynomial.rename Sum.inr <| P.relation rp) span_range_relation_eq_ker := Q.span_range_relation_eq_ker_comp P -lemma comp_relation_map (r : Q.rels) : +@[simp] +lemma comp_relation_inr (r : P.rels) : + (Q.comp P).relation (Sum.inr r) = rename Sum.inr (P.relation r) := + rfl + +lemma comp_aeval_relation_inl (r : Q.rels) : aeval (Sum.elim X (MvPolynomial.C ∘ P.val)) ((Q.comp P).relation (Sum.inl r)) = Q.relation r := by show (Q.aux P) _ = _ diff --git a/Mathlib/RingTheory/Smooth/StandardSmooth.lean b/Mathlib/RingTheory/Smooth/StandardSmooth.lean index d8064eb6a7e8c..f9ed2c1090626 100644 --- a/Mathlib/RingTheory/Smooth/StandardSmooth.lean +++ b/Mathlib/RingTheory/Smooth/StandardSmooth.lean @@ -62,10 +62,6 @@ Finally, for ring homomorphisms we define: ## TODO -- Show that the canonical presentation for localization away from an element is standard smooth - of relative dimension 0. -- Show that the composition of submersive presentations of relative dimensions `n` and `m` is - submersive of relative dimension `n + m`. - Show that the module of Kaehler differentials of a standard smooth `R`-algebra `S` of relative dimension `n` is `S`-free of rank `n`. In particular this shows that the relative dimension is independent of the choice of the standard smooth presentation. @@ -87,9 +83,9 @@ in June 2024. universe t t' w w' u v -open TensorProduct Classical +open TensorProduct MvPolynomial Classical -variable (n : ℕ) +variable (n m : ℕ) namespace Algebra @@ -161,7 +157,7 @@ lemma jacobian_eq_jacobiMatrix_det : P.jacobian = algebraMap P.Ring S P.jacobiMa lemma jacobiMatrix_apply (i j : P.rels) : P.jacobiMatrix i j = MvPolynomial.pderiv (P.map i) (P.relation j) := by - simp [jacobiMatrix, LinearMap.toMatrix, differential] + simp [jacobiMatrix, LinearMap.toMatrix, differential, basis] end Matrix @@ -182,6 +178,7 @@ instance (h : Function.Bijective (algebraMap R S)) : Fintype (ofBijectiveAlgebra instance (h : Function.Bijective (algebraMap R S)) : Fintype (ofBijectiveAlgebraMap h).rels := inferInstanceAs (Fintype PEmpty) +@[simp] lemma ofBijectiveAlgebraMap_jacobian (h : Function.Bijective (algebraMap R S)) : (ofBijectiveAlgebraMap h).jacobian = 1 := by have : (algebraMap (ofBijectiveAlgebraMap h).Ring S).mapMatrix @@ -190,7 +187,167 @@ lemma ofBijectiveAlgebraMap_jacobian (h : Function.Bijective (algebraMap R S)) : contradiction rw [jacobian_eq_jacobiMatrix_det, RingHom.map_det, this, Matrix.det_one] -open MvPolynomial +section Localization + +variable (r : R) [IsLocalization.Away r S] + +variable (S) in +/-- If `S` is the localization of `R` at `r`, this is the canonical submersive presentation +of `S` as `R`-algebra. -/ +@[simps map] +noncomputable def localizationAway : PreSubmersivePresentation R S where + __ := Presentation.localizationAway S r + map _ := () + map_inj _ _ h := h + relations_finite := inferInstanceAs <| Finite Unit + +instance : Fintype (localizationAway S r).rels := + inferInstanceAs (Fintype Unit) + +instance : DecidableEq (localizationAway S r).rels := + inferInstanceAs (DecidableEq Unit) + +@[simp] +lemma localizationAway_jacobiMatrix : + (localizationAway S r).jacobiMatrix = Matrix.diagonal (fun () ↦ MvPolynomial.C r) := by + have h : (pderiv ()) (C r * X () - 1) = C r := by simp + ext (i : Unit) (j : Unit) : 1 + rwa [jacobiMatrix_apply] + +@[simp] +lemma localizationAway_jacobian : (localizationAway S r).jacobian = algebraMap R S r := by + rw [jacobian_eq_jacobiMatrix_det, localizationAway_jacobiMatrix] + simp [show Fintype.card (localizationAway r (S := S)).rels = 1 from rfl] + +end Localization + +section Composition + +variable {T} [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] +variable (Q : PreSubmersivePresentation S T) (P : PreSubmersivePresentation R S) + +/-- Given an `R`-algebra `S` and an `S`-algebra `T` with pre-submersive presentations, +this is the canonical pre-submersive presentation of `T` as an `R`-algebra. -/ +@[simps map] +noncomputable def comp : PreSubmersivePresentation R T where + __ := Q.toPresentation.comp P.toPresentation + map := Sum.elim (fun rq ↦ Sum.inl <| Q.map rq) (fun rp ↦ Sum.inr <| P.map rp) + map_inj := Function.Injective.sum_elim ((Sum.inl_injective).comp (Q.map_inj)) + ((Sum.inr_injective).comp (P.map_inj)) <| by simp + relations_finite := inferInstanceAs <| Finite (Q.rels ⊕ P.rels) + +/-- The dimension of the composition of two finite submersive presentations is +the sum of the dimensions. -/ +lemma dimension_comp_eq_dimension_add_dimension [Q.IsFinite] [P.IsFinite] : + (Q.comp P).dimension = Q.dimension + P.dimension := by + simp only [Presentation.dimension] + erw [Presentation.comp_rels, Generators.comp_vars] + have : Nat.card P.rels ≤ Nat.card P.vars := + card_relations_le_card_vars_of_isFinite P + have : Nat.card Q.rels ≤ Nat.card Q.vars := + card_relations_le_card_vars_of_isFinite Q + simp only [Nat.card_sum] + omega + +section + +/-! +### Jacobian of composition + +Let `S` be an `R`-algebra and `T` be an `S`-algebra with presentations `P` and `Q` respectively. +In this section we compute the jacobian of the composition of `Q` and `P` to be +the product of the jacobians. For this we use a block decomposition of the jacobi matrix and show +that the upper-right block vanishes, the upper-left block has determinant jacobian of `Q` and +the lower-right block has determinant jacobian of `P`. + +-/ + +variable [Fintype (Q.comp P).rels] + +private lemma jacobiMatrix_comp_inl_inr (i : Q.rels) (j : P.rels) : + (Q.comp P).jacobiMatrix (Sum.inl i) (Sum.inr j) = 0 := by + rw [jacobiMatrix_apply] + refine MvPolynomial.pderiv_eq_zero_of_not_mem_vars (fun hmem ↦ ?_) + apply MvPolynomial.vars_rename at hmem + simp at hmem + +private lemma jacobiMatrix_comp_₁₂ : (Q.comp P).jacobiMatrix.toBlocks₁₂ = 0 := by + ext i j : 1 + simp [Matrix.toBlocks₁₂, jacobiMatrix_comp_inl_inr] + +section Q + +variable [Fintype Q.rels] + +private lemma jacobiMatrix_comp_inl_inl (i j : Q.rels) : + aeval (Sum.elim X (MvPolynomial.C ∘ P.val)) + ((Q.comp P).jacobiMatrix (Sum.inl j) (Sum.inl i)) = Q.jacobiMatrix j i := by + rw [jacobiMatrix_apply, jacobiMatrix_apply, comp_map, Sum.elim_inl, + ← Q.comp_aeval_relation_inl P.toPresentation] + apply aeval_sum_elim_pderiv_inl + +private lemma jacobiMatrix_comp_₁₁_det : + (aeval (Q.comp P).val) (Q.comp P).jacobiMatrix.toBlocks₁₁.det = Q.jacobian := by + rw [jacobian_eq_jacobiMatrix_det, AlgHom.map_det (aeval (Q.comp P).val), RingHom.map_det] + congr + ext i j : 1 + simp only [Matrix.map_apply, RingHom.mapMatrix_apply, ← Q.jacobiMatrix_comp_inl_inl P] + apply aeval_sum_elim + +end Q + +section P + +variable [Fintype P.rels] + +private lemma jacobiMatrix_comp_inr_inr (i j : P.rels) : + (Q.comp P).jacobiMatrix (Sum.inr i) (Sum.inr j) = + MvPolynomial.rename Sum.inr (P.jacobiMatrix i j) := by + rw [jacobiMatrix_apply, jacobiMatrix_apply] + simp only [comp_map, Sum.elim_inr] + apply pderiv_rename Sum.inr_injective + +private lemma jacobiMatrix_comp_₂₂_det : + (aeval (Q.comp P).val) (Q.comp P).jacobiMatrix.toBlocks₂₂.det = algebraMap S T P.jacobian := by + rw [jacobian_eq_jacobiMatrix_det] + rw [AlgHom.map_det (aeval (Q.comp P).val), RingHom.map_det, RingHom.map_det] + congr + ext i j : 1 + simp only [Matrix.toBlocks₂₂, AlgHom.mapMatrix_apply, Matrix.map_apply, Matrix.of_apply, + RingHom.mapMatrix_apply, Generators.algebraMap_apply, map_aeval, coe_eval₂Hom] + rw [jacobiMatrix_comp_inr_inr, ← IsScalarTower.algebraMap_eq] + simp only [aeval, AlgHom.coe_mk, coe_eval₂Hom] + generalize P.jacobiMatrix i j = p + induction' p using MvPolynomial.induction_on with a p q hp hq p i hp + · simp only [algHom_C, algebraMap_eq, eval₂_C] + erw [MvPolynomial.eval₂_C] + · simp [hp, hq] + · simp only [map_mul, rename_X, eval₂_mul, hp, eval₂_X] + erw [Generators.comp_val] + simp + +end P + +end + +/-- The jacobian of the composition of presentations is the product of the jacobians. -/ +@[simp] +lemma comp_jacobian_eq_jacobian_smul_jacobian : (Q.comp P).jacobian = P.jacobian • Q.jacobian := by + cases nonempty_fintype Q.rels + cases nonempty_fintype P.rels + letI : Fintype (Q.comp P).rels := inferInstanceAs <| Fintype (Q.rels ⊕ P.rels) + rw [jacobian_eq_jacobiMatrix_det, ← Matrix.fromBlocks_toBlocks ((Q.comp P).jacobiMatrix), + jacobiMatrix_comp_₁₂] + convert_to + (aeval (Q.comp P).val) (Q.comp P).jacobiMatrix.toBlocks₁₁.det * + (aeval (Q.comp P).val) (Q.comp P).jacobiMatrix.toBlocks₂₂.det = P.jacobian • Q.jacobian + · simp only [Generators.algebraMap_apply, ← map_mul] + congr + convert Matrix.det_fromBlocks_zero₁₂ (Q.comp P).jacobiMatrix.toBlocks₁₁ + (Q.comp P).jacobiMatrix.toBlocks₂₁ (Q.comp P).jacobiMatrix.toBlocks₂₂ + · rw [jacobiMatrix_comp_₁₁_det, jacobiMatrix_comp_₂₂_det, mul_comm, Algebra.smul_def] + +end Composition section BaseChange @@ -204,6 +361,7 @@ noncomputable def baseChange : PreSubmersivePresentation T (T ⊗[R] S) where map_inj := P.map_inj relations_finite := P.relations_finite +@[simp] lemma baseChange_jacobian : (P.baseChange T).jacobian = 1 ⊗ₜ P.jacobian := by classical cases nonempty_fintype P.rels @@ -258,6 +416,37 @@ noncomputable def ofBijectiveAlgebraMap (h : Function.Bijective (algebraMap R S) noncomputable def id : SubmersivePresentation.{t, w} R R := ofBijectiveAlgebraMap Function.bijective_id +section Composition + +variable {R S T} [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] +variable (Q : SubmersivePresentation S T) (P : SubmersivePresentation R S) + +/-- Given an `R`-algebra `S` and an `S`-algebra `T` with submersive presentations, +this is the canonical submersive presentation of `T` as an `R`-algebra. -/ +noncomputable def comp : SubmersivePresentation R T where + __ := Q.toPreSubmersivePresentation.comp P.toPreSubmersivePresentation + jacobian_isUnit := by + rw [comp_jacobian_eq_jacobian_smul_jacobian, Algebra.smul_def, IsUnit.mul_iff] + exact ⟨RingHom.isUnit_map _ <| P.jacobian_isUnit, Q.jacobian_isUnit⟩ + isFinite := Presentation.comp_isFinite Q.toPresentation P.toPresentation + +end Composition + +section Localization + +variable {R} (r : R) [IsLocalization.Away r S] + +/-- If `S` is the localization of `R` at `r`, this is the canonical submersive presentation +of `S` as `R`-algebra. -/ +noncomputable def localizationAway : SubmersivePresentation R S where + __ := PreSubmersivePresentation.localizationAway S r + jacobian_isUnit := by + rw [localizationAway_jacobian] + apply IsLocalization.map_units' (⟨r, 1, by simp⟩ : Submonoid.powers r) + isFinite := Presentation.localizationAway_isFinite r + +end Localization + section BaseChange variable (T) [CommRing T] [Algebra R T] (P : SubmersivePresentation R S) @@ -316,6 +505,37 @@ instance IsStandardSmoothOfRelativeDimension.id : IsStandardSmoothOfRelativeDimension.{t, w} 0 R R := IsStandardSmoothOfRelativeDimension.of_algebraMap_bijective Function.bijective_id +section Composition + +variable (R S T) [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] + +lemma IsStandardSmooth.trans [IsStandardSmooth.{t, w} R S] [IsStandardSmooth.{t', w'} S T] : + IsStandardSmooth.{max t t', max w w'} R T where + out := by + obtain ⟨⟨P⟩⟩ := ‹IsStandardSmooth R S› + obtain ⟨⟨Q⟩⟩ := ‹IsStandardSmooth S T› + exact ⟨Q.comp P⟩ + +lemma IsStandardSmoothOfRelativeDimension.trans [IsStandardSmoothOfRelativeDimension.{t, w} n R S] + [IsStandardSmoothOfRelativeDimension.{t', w'} m S T] : + IsStandardSmoothOfRelativeDimension.{max t t', max w w'} (m + n) R T where + out := by + obtain ⟨P, hP⟩ := ‹IsStandardSmoothOfRelativeDimension n R S› + obtain ⟨Q, hQ⟩ := ‹IsStandardSmoothOfRelativeDimension m S T› + refine ⟨Q.comp P, hP ▸ hQ ▸ ?_⟩ + apply PreSubmersivePresentation.dimension_comp_eq_dimension_add_dimension + +end Composition + +lemma IsStandardSmooth.localization_away (r : R) [IsLocalization.Away r S] : + IsStandardSmooth.{0, 0} R S where + out := ⟨SubmersivePresentation.localizationAway S r⟩ + +lemma IsStandardSmoothOfRelativeDimension.localization_away (r : R) [IsLocalization.Away r S] : + IsStandardSmoothOfRelativeDimension.{0, 0} 0 R S where + out := ⟨SubmersivePresentation.localizationAway S r, + Presentation.localizationAway_dimension_zero r⟩ + section BaseChange variable (T) [CommRing T] [Algebra R T] diff --git a/Mathlib/RingTheory/TwoSidedIdeal/BigOperators.lean b/Mathlib/RingTheory/TwoSidedIdeal/BigOperators.lean new file mode 100644 index 0000000000000..8ed5a222e9ed2 --- /dev/null +++ b/Mathlib/RingTheory/TwoSidedIdeal/BigOperators.lean @@ -0,0 +1,74 @@ +/- +Copyright (c) 2024 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ + +import Mathlib.RingTheory.Congruence.BigOperators +import Mathlib.RingTheory.TwoSidedIdeal.Basic + +/-! +# Interactions between `∑, ∏` and two sided ideals + +-/ + +namespace TwoSidedIdeal + +section sum + +variable {R : Type*} [NonUnitalNonAssocRing R] (I : TwoSidedIdeal R) + +lemma listSum_mem {ι : Type*} (l : List ι) (f : ι → R) (hl : ∀ x ∈ l, f x ∈ I) : + (l.map f).sum ∈ I := by + rw [mem_iff, ← List.sum_map_zero] + exact I.ringCon.listSum l hl + +lemma multisetSum_mem {ι : Type*} (s : Multiset ι) (f : ι → R) (hs : ∀ x ∈ s, f x ∈ I) : + (s.map f).sum ∈ I := by + rw [mem_iff, ← Multiset.sum_map_zero] + exact I.ringCon.multisetSum s hs + +lemma finsetSum_mem {ι : Type*} (s : Finset ι) (f : ι → R) (hs : ∀ x ∈ s, f x ∈ I) : + s.sum f ∈ I := by + rw [mem_iff, ← Finset.sum_const_zero] + exact I.ringCon.finsetSum s hs + +end sum + +section prod + +section ring + +variable {R : Type*} [Ring R] (I : TwoSidedIdeal R) + +lemma listProd_mem {ι : Type*} (l : List ι) (f : ι → R) (hl : ∃ x ∈ l, f x ∈ I) : + (l.map f).prod ∈ I := by + induction l with + | nil => simp only [List.not_mem_nil, false_and, exists_false] at hl + | cons x l ih => + simp only [List.mem_cons, exists_eq_or_imp] at hl + rcases hl with h | hal + · simpa only [List.map_cons, List.prod_cons] using I.mul_mem_right _ _ h + · simpa only [List.map_cons, List.prod_cons] using I.mul_mem_left _ _ <| ih hal + +end ring + +section commRing + +variable {R : Type*} [CommRing R] (I : TwoSidedIdeal R) + +lemma multiSetProd_mem {ι : Type*} (s : Multiset ι) (f : ι → R) (hs : ∃ x ∈ s, f x ∈ I) : + (s.map f).prod ∈ I := by + rcases s + simpa using listProd_mem (hl := hs) + +lemma finsetProd_mem {ι : Type*} (s : Finset ι) (f : ι → R) (hs : ∃ x ∈ s, f x ∈ I) : + s.prod f ∈ I := by + rcases s + simpa using multiSetProd_mem (hs := hs) + +end commRing + +end prod + +end TwoSidedIdeal diff --git a/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean b/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean new file mode 100644 index 0000000000000..ea08a0c3e8866 --- /dev/null +++ b/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean @@ -0,0 +1,272 @@ +/- +Copyright (c) 2024 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang, Jireh Loreaux +-/ + +import Mathlib.RingTheory.TwoSidedIdeal.Lattice +import Mathlib.RingTheory.Congruence.Opposite +import Mathlib.Algebra.BigOperators.Ring +import Mathlib.Data.Fintype.BigOperators +import Mathlib.RingTheory.Ideal.Basic +import Mathlib.Order.GaloisConnection + +/-! +# Operations on two-sided ideals + +This file defines operations on two-sided ideals of a ring `R`. + +## Main definitions and results + +- `TwoSidedIdeal.span`: the span of `s ⊆ R` is the smallest two-sided ideal containing the set. +- `TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure_nonunital`: in an associative but non-unital + ring, an element `x` is in the two-sided ideal spanned by `s` if and only if `x` is in the closure + of `s ∪ {y * a | y ∈ s, a ∈ R} ∪ {a * y | y ∈ s, a ∈ R} ∪ {a * y * b | y ∈ s, a, b ∈ R}` as an + additive subgroup. +- `TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure`: in a unital and associative ring, an + element `x` is in the two-sided ideal spanned by `s` if and only if `x` is in the closure of + `{a*y*b | a, b ∈ R, y ∈ s}` as an additive subgroup. + + +- `TwoSidedIdeal.comap`: pullback of a two-sided ideal; defined as the preimage of a + two-sided ideal. +- `TwoSidedIdeal.map`: pushforward of a two-sided ideal; defined as the span of the image of a + two-sided ideal. +- `TwoSidedIdeal.ker`: the kernel of a ring homomorphism as a two-sided ideal. + +- `TwoSidedIdeal.gc`: `fromIdeal` and `asIdeal` form a Galois connection where + `fromIdeal : Ideal R → TwoSidedIdeal R` is defined as the smallest two-sided ideal containing an + ideal and `asIdeal : TwoSidedIdeal R → Ideal R` the inclusion map. +-/ + +namespace TwoSidedIdeal + +section NonUnitalNonAssocRing + +variable {R S : Type*} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] +variable {F : Type*} [FunLike F R S] +variable (f : F) + +/-- +The smallest two-sided ideal containing a set. +-/ +abbrev span (s : Set R) : TwoSidedIdeal R := + { ringCon := ringConGen (fun a b ↦ a - b ∈ s) } + +lemma subset_span {s : Set R} : s ⊆ (span s : Set R) := by + intro x hx + rw [SetLike.mem_coe, mem_iff] + exact RingConGen.Rel.of _ _ (by simpa using hx) + +lemma mem_span_iff {s : Set R} {x} : + x ∈ span s ↔ ∀ (I : TwoSidedIdeal R), s ⊆ I → x ∈ I := by + refine ⟨?_, fun h => h _ subset_span⟩ + delta span + rw [RingCon.ringConGen_eq] + intro h I hI + refine sInf_le (α := RingCon R) ?_ h + intro x y hxy + specialize hI hxy + rwa [SetLike.mem_coe, ← rel_iff] at hI + +lemma span_mono {s t : Set R} (h : s ⊆ t) : span s ≤ span t := by + intro x hx + rw [mem_span_iff] at hx ⊢ + exact fun I hI => hx I <| h.trans hI + +/-- +Pushout of a two-sided ideal. Defined as the span of the image of a two-sided ideal under a ring +homomorphism. +-/ +def map (I : TwoSidedIdeal R) : TwoSidedIdeal S := + span (f '' I) + +lemma map_mono {I J : TwoSidedIdeal R} (h : I ≤ J) : + map f I ≤ map f J := + span_mono <| Set.image_mono h + +variable [NonUnitalRingHomClass F R S] + +/-- +Preimage of a two-sided ideal, as a two-sided ideal. -/ +def comap (I : TwoSidedIdeal S) : TwoSidedIdeal R := +{ ringCon := I.ringCon.comap f } + +lemma mem_comap {I : TwoSidedIdeal S} {x : R} : + x ∈ I.comap f ↔ f x ∈ I := by + simp [comap, RingCon.comap, mem_iff] + +/-- +The kernel of a ring homomorphism, as a two-sided ideal. +-/ +def ker : TwoSidedIdeal R := + .mk' + {r | f r = 0} (map_zero _) (by rintro _ _ (h1 : f _ = 0) (h2 : f _ = 0); simp [h1, h2]) + (by rintro _ (h : f _ = 0); simp [h]) (by rintro _ _ (h : f _ = 0); simp [h]) + (by rintro _ _ (h : f _ = 0); simp [h]) + +lemma mem_ker {x : R} : x ∈ ker f ↔ f x = 0 := by + delta ker; rw [mem_mk']; rfl + · rintro _ _ (h1 : f _ = 0) (h2 : f _ = 0); simp [h1, h2] + · rintro _ (h : f _ = 0); simp [h] + · rintro _ _ (h : f _ = 0); simp [h] + · rintro _ _ (h : f _ = 0); simp [h] + +end NonUnitalNonAssocRing + +section NonUnitalRing + +variable {R : Type*} [NonUnitalRing R] + +open AddSubgroup in +/-- If `s : Set R` is absorbing under multiplication, then its `TwoSidedIdeal.span` coincides with +its `AddSubgroup.closure`, as sets. -/ +lemma mem_span_iff_mem_addSubgroup_closure_absorbing {s : Set R} + (h_left : ∀ x y, y ∈ s → x * y ∈ s) (h_right : ∀ y x, y ∈ s → y * x ∈ s) {z : R} : + z ∈ span s ↔ z ∈ closure s := by + have h_left' {x y} (hy : y ∈ closure s) : x * y ∈ closure s := by + have := (AddMonoidHom.mulLeft x).map_closure s ▸ mem_map_of_mem _ hy + refine closure_mono ?_ this + rintro - ⟨y, hy, rfl⟩ + exact h_left x y hy + have h_right' {y x} (hy : y ∈ closure s) : y * x ∈ closure s := by + have := (AddMonoidHom.mulRight x).map_closure s ▸ mem_map_of_mem _ hy + refine closure_mono ?_ this + rintro - ⟨y, hy, rfl⟩ + exact h_right y x hy + let I : TwoSidedIdeal R := .mk' (closure s) (AddSubgroup.zero_mem _) + (AddSubgroup.add_mem _) (AddSubgroup.neg_mem _) h_left' h_right' + suffices z ∈ span s ↔ z ∈ I by simpa only [I, mem_mk', SetLike.mem_coe] + rw [mem_span_iff] + -- Suppose that for every ideal `J` with `s ⊆ J`, then `z ∈ J`. Apply this to `I` to get `z ∈ I`. + refine ⟨fun h ↦ h I fun x hx ↦ ?mem_closure_of_forall, fun hz J hJ ↦ ?mem_ideal_of_subset⟩ + case mem_closure_of_forall => simpa only [I, SetLike.mem_coe, mem_mk'] using subset_closure hx + /- Conversely, suppose that `z ∈ I` and that `J` is any ideal containing `s`. Then by the + induction principle for `AddSubgroup`, we must also have `z ∈ J`. -/ + case mem_ideal_of_subset => + simp only [I, SetLike.mem_coe, mem_mk'] at hz + induction hz using closure_induction' with + | mem x hx => exact hJ hx + | one => exact zero_mem _ + | mul x _ y _ hx hy => exact J.add_mem hx hy + | inv x _ hx => exact J.neg_mem hx + +open Pointwise Set + +lemma set_mul_subset {s : Set R} {I : TwoSidedIdeal R} (h : s ⊆ I) (t : Set R): + t * s ⊆ I := by + rintro - ⟨r, -, x, hx, rfl⟩ + exact mul_mem_left _ _ _ (h hx) + +lemma subset_mul_set {s : Set R} {I : TwoSidedIdeal R} (h : s ⊆ I) (t : Set R): + s * t ⊆ I := by + rintro - ⟨x, hx, r, -, rfl⟩ + exact mul_mem_right _ _ _ (h hx) + +lemma mem_span_iff_mem_addSubgroup_closure_nonunital {s : Set R} {z : R} : + z ∈ span s ↔ z ∈ AddSubgroup.closure (s ∪ s * univ ∪ univ * s ∪ univ * s * univ) := by + trans z ∈ span (s ∪ s * univ ∪ univ * s ∪ univ * s * univ) + · refine ⟨(span_mono (by simp only [Set.union_assoc, Set.subset_union_left]) ·), fun h ↦ ?_⟩ + refine mem_span_iff.mp h (span s) ?_ + simp only [union_subset_iff, union_assoc] + exact ⟨subset_span, subset_mul_set subset_span _, set_mul_subset subset_span _, + subset_mul_set (set_mul_subset subset_span _) _⟩ + · refine mem_span_iff_mem_addSubgroup_closure_absorbing ?_ ?_ + · rintro x y (((hy | ⟨y, hy, r, -, rfl⟩) | ⟨r, -, y, hy, rfl⟩) | + ⟨-, ⟨r', -, y, hy, rfl⟩, r, -, rfl⟩) + · exact .inl <| .inr <| ⟨x, mem_univ _, y, hy, rfl⟩ + · exact .inr <| ⟨x * y, ⟨x, mem_univ _, y, hy, rfl⟩, r, mem_univ _, mul_assoc ..⟩ + · exact .inl <| .inr <| ⟨x * r, mem_univ _, y, hy, mul_assoc ..⟩ + · refine .inr <| ⟨x * r' * y, ⟨x * r', mem_univ _, y, hy, ?_⟩, ⟨r, mem_univ _, ?_⟩⟩ + all_goals simp [mul_assoc] + · rintro y x (((hy | ⟨y, hy, r, -, rfl⟩) | ⟨r, -, y, hy, rfl⟩) | + ⟨-, ⟨r', -, y, hy, rfl⟩, r, -, rfl⟩) + · exact .inl <| .inl <| .inr ⟨y, hy, x, mem_univ _, rfl⟩ + · exact .inl <| .inl <| .inr ⟨y, hy, r * x, mem_univ _, (mul_assoc ..).symm⟩ + · exact .inr <| ⟨r * y, ⟨r, mem_univ _, y, hy, rfl⟩, x, mem_univ _, rfl⟩ + · refine .inr <| ⟨r' * y, ⟨r', mem_univ _, y, hy, rfl⟩, r * x, mem_univ _, ?_⟩ + simp [mul_assoc] + +end NonUnitalRing + +section Ring + +variable {R : Type*} [Ring R] + +open Pointwise Set in +lemma mem_span_iff_mem_addSubgroup_closure {s : Set R} {z : R} : + z ∈ span s ↔ z ∈ AddSubgroup.closure (univ * s * univ) := by + trans z ∈ span (univ * s * univ) + · refine ⟨(span_mono (fun x hx ↦ ?_) ·), fun hz ↦ ?_⟩ + · exact ⟨1 * x, ⟨1, mem_univ _, x, hx, rfl⟩, 1, mem_univ _, by simp⟩ + · exact mem_span_iff.mp hz (span s) <| subset_mul_set (set_mul_subset subset_span _) _ + · refine mem_span_iff_mem_addSubgroup_closure_absorbing ?_ ?_ + · intro x y hy + rw [mul_assoc] at hy ⊢ + obtain ⟨r, -, y, hy, rfl⟩ := hy + exact ⟨x * r, mem_univ _, y, hy, mul_assoc ..⟩ + · rintro - x ⟨y, hy, r, -, rfl⟩ + exact ⟨y, hy, r * x, mem_univ _, (mul_assoc ..).symm⟩ + +/-- Given an ideal `I`, `span I` is the smallest two-sided ideal containing `I`. -/ +def fromIdeal : Ideal R →o TwoSidedIdeal R where + toFun I := span I + monotone' _ _ := span_mono + +lemma mem_fromIdeal {I : Ideal R} {x : R} : + x ∈ fromIdeal I ↔ x ∈ span I := by simp [fromIdeal] + +/-- Every two-sided ideal is also a left ideal. -/ +def asIdeal : TwoSidedIdeal R →o Ideal R where + toFun I := + { carrier := I + add_mem' := I.add_mem + zero_mem' := I.zero_mem + smul_mem' := fun r x hx => I.mul_mem_left r x hx } + monotone' _ _ h _ h' := h h' + +lemma mem_asIdeal {I : TwoSidedIdeal R} {x : R} : + x ∈ asIdeal I ↔ x ∈ I := by simp [asIdeal] + +lemma gc : GaloisConnection fromIdeal (asIdeal (R := R)) := + fun I J => ⟨fun h x hx ↦ h <| mem_span_iff.2 fun _ H ↦ H hx, fun h x hx ↦ by + simp only [fromIdeal, OrderHom.coe_mk, mem_span_iff] at hx + exact hx _ h⟩ + +@[simp] +lemma coe_asIdeal {I : TwoSidedIdeal R} : (asIdeal I : Set R) = I := rfl + +/-- Every two-sided ideal is also a right ideal. -/ +def asIdealOpposite : TwoSidedIdeal R →o Ideal Rᵐᵒᵖ where + toFun I := asIdeal ⟨I.ringCon.op⟩ + monotone' I J h x h' := by + simp only [mem_asIdeal, mem_iff, RingCon.op_iff, MulOpposite.unop_zero] at h' ⊢ + exact J.rel_iff _ _ |>.2 <| h <| I.rel_iff 0 x.unop |>.1 h' + +lemma mem_asIdealOpposite {I : TwoSidedIdeal R} {x : Rᵐᵒᵖ} : + x ∈ asIdealOpposite I ↔ x.unop ∈ I := by + simpa [asIdealOpposite, asIdeal, TwoSidedIdeal.mem_iff, RingCon.op_iff] using + ⟨I.ringCon.symm, I.ringCon.symm⟩ + +end Ring + +section CommRing + +variable {R : Type*} [CommRing R] + +/-- +When the ring is commutative, two-sided ideals are exactly the same as left ideals. +-/ +def orderIsoIdeal : TwoSidedIdeal R ≃o Ideal R where + toFun := asIdeal + invFun := fromIdeal + map_rel_iff' := ⟨fun h _ hx ↦ h hx, fun h ↦ asIdeal.monotone' h⟩ + left_inv _ := SetLike.ext fun _ ↦ mem_span_iff.trans <| by aesop + right_inv J := SetLike.ext fun x ↦ mem_span_iff.trans + ⟨fun h ↦ mem_mk' _ _ _ _ _ _ _ |>.1 <| h (mk' + J J.zero_mem J.add_mem J.neg_mem (J.mul_mem_left _) (J.mul_mem_right _)) + (fun x => by simp [mem_mk']), by aesop⟩ + +end CommRing + +end TwoSidedIdeal diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index edce0ce9c8677..f1887c8c8f025 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -323,8 +323,7 @@ theorem WfDvdMonoid.of_exists_prime_factors : WfDvdMonoid α := rw [dif_neg ane0] by_cases h : b = 0 · simp [h, lt_top_iff_ne_top] - · rw [dif_neg h] - erw [WithTop.coe_lt_coe] + · rw [dif_neg h, Nat.cast_lt] have cne0 : c ≠ 0 := by refine mt (fun con => ?_) h rw [b_eq, con, mul_zero] @@ -390,8 +389,8 @@ theorem MulEquiv.uniqueFactorizationMonoid (e : α ≃* β) (hα : UniqueFactori he ▸ e.prime_iff.1 (hp c hc), Units.map e.toMonoidHom u, by - erw [Multiset.prod_hom, ← map_mul e, h] - simp⟩ + rw [Multiset.prod_hom, toMonoidHom_eq_coe, Units.coe_map, MonoidHom.coe_coe, ← map_mul e, h, + apply_symm_apply]⟩ theorem MulEquiv.uniqueFactorizationMonoid_iff (e : α ≃* β) : UniqueFactorizationMonoid α ↔ UniqueFactorizationMonoid β := diff --git a/Mathlib/Tactic/Linter/Lint.lean b/Mathlib/Tactic/Linter/Lint.lean index e1345febe9b36..6d00f05354dde 100644 --- a/Mathlib/Tactic/Linter/Lint.lean +++ b/Mathlib/Tactic/Linter/Lint.lean @@ -320,6 +320,12 @@ register_option linter.style.longFile : Nat := { descr := "enable the longFile linter" } +/-- The number of lines that the `longFile` linter considers the default. -/ +register_option linter.style.longFileDefValue : Nat := { + defValue := 1500 + descr := "a soft upper bound on the number of lines of each file" +} + namespace Style.longFile @[inherit_doc Mathlib.Linter.linter.style.longFile] @@ -327,7 +333,7 @@ def longFileLinter : Linter where run := withSetOptionIn fun stx ↦ do let linterBound := linter.style.longFile.get (← getOptions) if linterBound == 0 then return - let defValue := 1500 + let defValue := linter.style.longFileDefValue.get (← getOptions) let smallOption := match stx with | `(set_option linter.style.longFile $x) => TSyntax.getNat ⟨x.raw⟩ ≤ defValue | _ => false diff --git a/Mathlib/Tactic/ModCases.lean b/Mathlib/Tactic/ModCases.lean index 9846cca299457..fd1a5afdd358f 100644 --- a/Mathlib/Tactic/ModCases.lean +++ b/Mathlib/Tactic/ModCases.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Heather Macbeth -/ import Mathlib.Data.Int.ModEq +import Mathlib.Tactic.HaveI /-! # `mod_cases` tactic diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index 69e02cd2a9319..b526794778cbb 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -609,7 +609,7 @@ theorem continuous_of_cover_nhds {ι : Sort*} {f : α → β} {s : ι → Set α rw [ContinuousAt, ← nhdsWithin_eq_nhds.2 hi] exact hf _ _ (mem_of_mem_nhds hi) -theorem continuousOn_empty (f : α → β) : ContinuousOn f ∅ := fun _ => False.elim +@[simp] theorem continuousOn_empty (f : α → β) : ContinuousOn f ∅ := fun _ => False.elim @[simp] theorem continuousOn_singleton (f : α → β) (a : α) : ContinuousOn f {a} := diff --git a/Mathlib/Topology/FiberPartition.lean b/Mathlib/Topology/FiberPartition.lean index c01b90f3fe77b..8e21ee229b9e1 100644 --- a/Mathlib/Topology/FiberPartition.lean +++ b/Mathlib/Topology/FiberPartition.lean @@ -29,7 +29,7 @@ variable [TopologicalSpace S] /-- The canonical map from the disjoint union induced by `f` to `S`. -/ @[simps apply] def sigmaIsoHom : C((x : Fiber f) × x.val, S) where - toFun := fun ⟨a, x⟩ ↦ x.val + toFun | ⟨a, x⟩ => x.val lemma sigmaIsoHom_inj : Function.Injective (sigmaIsoHom f) := by rintro ⟨⟨_, _, rfl⟩, ⟨_, hx⟩⟩ ⟨⟨_, _, rfl⟩, ⟨_, hy⟩⟩ h @@ -43,7 +43,7 @@ lemma sigmaIsoHom_surj : Function.Surjective (sigmaIsoHom f) := /-- The inclusion map from a component of the disjoint union induced by `f` into `S`. -/ def sigmaIncl (a : Fiber f) : C(a.val, S) where - toFun := fun x ↦ x.val + toFun x := x.val /-- The inclusion map from a fiber of a composition into the intermediate fiber. -/ def sigmaInclIncl {X : Type*} (g : Y → X) (a : Fiber (g ∘ f)) @@ -53,7 +53,7 @@ def sigmaInclIncl {X : Type*} (g : Y → X) (a : Fiber (g ∘ f)) have := x.prop simp only [sigmaIncl, ContinuousMap.coe_mk, Fiber.mem_iff_eq_image, comp_apply] at this rw [Fiber.mem_iff_eq_image, Fiber.mk_image, this, ← Fiber.map_preimage_eq_image] - rfl⟩ + simp [sigmaIncl]⟩ variable (l : LocallyConstant S Y) [CompactSpace S] diff --git a/Mathlib/Topology/Instances/Real.lean b/Mathlib/Topology/Instances/Real.lean index 5b895c2cdeb63..16325c2f3e0ab 100644 --- a/Mathlib/Topology/Instances/Real.lean +++ b/Mathlib/Topology/Instances/Real.lean @@ -3,16 +3,17 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Data.Real.Star -import Mathlib.Algebra.Algebra.Basic +import Mathlib.Algebra.Module.Rat +import Mathlib.Algebra.Module.Submodule.Lattice import Mathlib.Algebra.Periodic +import Mathlib.Data.Real.Star import Mathlib.Topology.Algebra.Order.Archimedean import Mathlib.Topology.Algebra.Order.Field -import Mathlib.Topology.Algebra.UniformMulAction import Mathlib.Topology.Algebra.Star +import Mathlib.Topology.Algebra.UniformMulAction import Mathlib.Topology.Instances.Int -import Mathlib.Topology.Order.Bornology import Mathlib.Topology.Metrizable.Basic +import Mathlib.Topology.Order.Bornology /-! # Topological properties of ℝ diff --git a/Mathlib/Topology/KrullDimension.lean b/Mathlib/Topology/KrullDimension.lean index ea4e77fb0194c..b5eee2e982ae4 100644 --- a/Mathlib/Topology/KrullDimension.lean +++ b/Mathlib/Topology/KrullDimension.lean @@ -22,3 +22,42 @@ closed irreducible sets. -/ noncomputable def topologicalKrullDim (T : Type*) [TopologicalSpace T] : WithBot ℕ∞ := krullDim (IrreducibleCloseds T) + +variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] + +/-- +Map induced on irreducible closed subsets by a closed continuous map `f`. +This is just a wrapper around the image of `f` together with proofs that it +preserves irreducibility (by continuity) and closedness (since `f` is closed). +-/ +def IrreducibleCloseds.map {f : X → Y} (hf1 : Continuous f) (hf2 : IsClosedMap f) + (c : IrreducibleCloseds X) : + IrreducibleCloseds Y where + carrier := f '' c + is_irreducible' := c.is_irreducible'.image f hf1.continuousOn + is_closed' := hf2 c c.is_closed' + +/-- +Taking images under a closed embedding is strictly monotone on the preorder of irreducible closeds. +-/ +lemma IrreducibleCloseds.map_strictMono {f : X → Y} (hf : ClosedEmbedding f) : + StrictMono (IrreducibleCloseds.map hf.continuous hf.isClosedMap) := + fun ⦃_ _⦄ UltV ↦ hf.inj.image_strictMono UltV + +/-- +If `f : X → Y` is a closed embedding, then the Krull dimension of `X` is less than or equal +to the Krull dimension of `Y`. +-/ +theorem ClosedEmbedding.topologicalKrullDim_le (f : X → Y) (hf : ClosedEmbedding f) : + topologicalKrullDim X ≤ topologicalKrullDim Y := + krullDim_le_of_strictMono _ (IrreducibleCloseds.map_strictMono hf) + +/-- The topological Krull dimension is invariant under homeomorphisms -/ +theorem IsHomeomorph.topologicalKrullDim_eq (f : X → Y) (h : IsHomeomorph f) : + topologicalKrullDim X = topologicalKrullDim Y := + have fwd : topologicalKrullDim X ≤ topologicalKrullDim Y := + ClosedEmbedding.topologicalKrullDim_le f h.closedEmbedding + have bwd : topologicalKrullDim Y ≤ topologicalKrullDim X := + ClosedEmbedding.topologicalKrullDim_le (h.homeomorph f).symm + (h.homeomorph f).symm.closedEmbedding + le_antisymm fwd bwd diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index b57b69003857a..3e05609edfe68 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -136,6 +136,31 @@ theorem mem_idRel {a b : α} : (a, b) ∈ @idRel α ↔ a = b := theorem idRel_subset {s : Set (α × α)} : idRel ⊆ s ↔ ∀ a, (a, a) ∈ s := by simp [subset_def] +theorem eq_singleton_left_of_prod_subset_idRel {X : Type _} {S T : Set X} (hS : S.Nonempty) + (hT : T.Nonempty) (h_diag : S ×ˢ T ⊆ idRel) : ∃ x, S = {x} := by + rcases hS, hT with ⟨⟨s, hs⟩, ⟨t, ht⟩⟩ + refine ⟨s, eq_singleton_iff_nonempty_unique_mem.mpr ⟨⟨s, hs⟩, fun x hx ↦ ?_⟩⟩ + rw [prod_subset_iff] at h_diag + replace hs := h_diag s hs t ht + replace hx := h_diag x hx t ht + simp only [idRel, mem_setOf_eq] at hx hs + rwa [← hs] at hx + +theorem eq_singleton_right_prod_subset_idRel {X : Type _} {S T : Set X} (hS : S.Nonempty) + (hT : T.Nonempty) (h_diag : S ×ˢ T ⊆ idRel) : ∃ x, T = {x} := by + rw [Set.prod_subset_iff] at h_diag + replace h_diag := fun x hx y hy => (h_diag y hy x hx).symm + exact eq_singleton_left_of_prod_subset_idRel hT hS (prod_subset_iff.mpr h_diag) + +theorem eq_singleton_prod_subset_idRel {X : Type _} {S T : Set X} (hS : S.Nonempty) + (hT : T.Nonempty) (h_diag : S ×ˢ T ⊆ idRel) : ∃ x, S = {x} ∧ T = {x} := by + obtain ⟨⟨x, hx⟩, ⟨y, hy⟩⟩ := eq_singleton_left_of_prod_subset_idRel hS hT h_diag, + eq_singleton_right_prod_subset_idRel hS hT h_diag + refine ⟨x, ⟨hx, ?_⟩⟩ + rw [hy, Set.singleton_eq_singleton_iff] + exact (Set.prod_subset_iff.mp h_diag x (by simp only [hx, Set.mem_singleton]) y + (by simp only [hy, Set.mem_singleton])).symm + /-- The composition of relations -/ def compRel (r₁ r₂ : Set (α × α)) := { p : α × α | ∃ z : α, (p.1, z) ∈ r₁ ∧ (z, p.2) ∈ r₂ } diff --git a/Mathlib/Topology/UniformSpace/Cauchy.lean b/Mathlib/Topology/UniformSpace/Cauchy.lean index 7907f4159bea3..a1a7590a0efa4 100644 --- a/Mathlib/Topology/UniformSpace/Cauchy.lean +++ b/Mathlib/Topology/UniformSpace/Cauchy.lean @@ -787,4 +787,29 @@ theorem secondCountable_of_separable [SeparableSpace α] : SecondCountableTopolo refine ⟨_, ⟨y, hys, k, rfl⟩, (hts k).subset hxy, fun z hz => ?_⟩ exact hUV (ball_subset_of_comp_subset (hk hxy) hUU' (hk hz)) +section DiscreteUniformity + +open Filter + +/-- A Cauchy filter in a discrete uniform space is contained in a principal filter-/ +theorem DiscreteUnif.cauchy_le_pure {X : Type _} {uX : UniformSpace X} + (hX : uX = ⊥) {α : Filter X} (hα : Cauchy α) : ∃ x : X, α = pure x := by + rcases hα with ⟨α_ne_bot, α_le⟩ + rw [hX, bot_uniformity, le_principal_iff, mem_prod_iff] at α_le + obtain ⟨S, ⟨hS, ⟨T, ⟨hT, H⟩⟩⟩⟩ := α_le + obtain ⟨x, rfl⟩ := eq_singleton_left_of_prod_subset_idRel (α_ne_bot.nonempty_of_mem hS) + (Filter.nonempty_of_mem hT) H + exact ⟨x, α_ne_bot.le_pure_iff.mp <| le_pure_iff.mpr hS⟩ + +/-- A constant to which a Cauchy filter in a discrete uniform space converges. -/ +noncomputable def DiscreteUnif.cauchyConst {X : Type _} {uX : UniformSpace X} + (hX : uX = ⊥) {α : Filter X} (hα : Cauchy α) : X := + (DiscreteUnif.cauchy_le_pure hX hα).choose + +theorem DiscreteUnif.eq_const_of_cauchy {X : Type _} {uX : UniformSpace X} (hX : uX = ⊥) + {α : Filter X} (hα : Cauchy α) : α = pure (DiscreteUnif.cauchyConst hX hα) := + (DiscreteUnif.cauchy_le_pure hX hα).choose_spec + +end DiscreteUniformity + end UniformSpace diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index f1113456868a1..4a3b5c75b8fd9 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -356,7 +356,9 @@ Single Variable Real Analysis: Weierstrass trigonometric approximation theorem: 'span_fourier_closure_eq_top' Convexity: convex functions of a real variable: 'ConvexOn' - continuity and differentiability of convex functions: 'https://en.wikipedia.org/wiki/Convex_function#Functions_of_one_variable' + continuity and differentiability of convex functions: + continuity: 'ConvexOn.continuousOn' + differentiability: 'https://en.wikipedia.org/wiki/Convex_function#Functions_of_one_variable' characterizations of convexity: 'convexOn_of_deriv2_nonneg' convexity inequalities: 'analysis/mean_inequalities.html' diff --git a/lake-manifest.json b/lake-manifest.json index 3e190c4de8899..9161d8bacca4b 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "bf12ff6041cbab6eba6b54d9467baed807bb2bfd", + "rev": "4756e0fc48acce0cc808df0ad149de5973240df6", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "50aaaf78b7db5bd635c19c660d59ed31b9bc9b5a", + "rev": "ff420521a0c098891f4f44ecda9dd7ff57b50bad", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -55,7 +55,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "fb7841a6f4fb389ec0e47dd4677844d49906af3c", + "rev": "e285a7ade149c551c17a4b24f127e1ef782e4bb1", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "2ba60fa2c384a94735454db11a2d523612eaabff", + "rev": "c9e106b0541a3b1c4bf82017aca78a50cc3e5ca2", "name": "LeanSearchClient", "manifestFile": "lake-manifest.json", "inputRev": "main", diff --git a/lean-toolchain b/lean-toolchain index 98556ba065e2a..89985206aca4e 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.12.0-rc1 +leanprover/lean4:v4.12.0 diff --git a/test/Lint.lean b/test/Lint.lean index c4f49cc99687e..25bf9c9147d13 100644 --- a/test/Lint.lean +++ b/test/Lint.lean @@ -294,70 +294,3 @@ note: this linter can be disabled with `set_option linter.style.longLine false` #guard_msgs in set_option linter.style.longLine true in #check " \" " - -/- -# Testing the `longFile` linter - -Things to note: -* `set_option linter.style.longFile 0` disables the linter, allowing us to set a value smaller than - `1500` without triggering the warning for setting a small value for the option; -* `guard_msgs ... in #exit` and `set_option ... in #exit` allow processing of the file *beyond* - `#exit`, since they wrap `#exit` inside an anonymous section, - making Lean active again *after* that anonymous section. - --/ - -section longFile - -/-- -warning: The default value of the `longFile` linter is 1500. -The current value of 1500 does not exceed the allowed bound. -Please, remove the `set_option linter.style.longFile 1500`. --/ -#guard_msgs in --- Do not allow setting a "small" `longFile` linter option -set_option linter.style.longFile 1500 - -/-- -warning: using 'exit' to interrupt Lean ---- -warning: The default value of the `longFile` linter is 1500. -This file is 331 lines long which does not exceed the allowed bound. -Please, remove the `set_option linter.style.longFile 1600`. --/ -#guard_msgs in --- Do not allow unnecessarily increasing the `longFile` linter option -set_option linter.style.longFile 1600 in -#exit - -/-- -warning: using 'exit' to interrupt Lean ---- -warning: This file is 346 lines long, but the limit is 10. - -You can extend the allowed length of the file using `set_option linter.style.longFile 1500`. -You can completely disable this linter by setting the length limit to `0`. --/ -#guard_msgs in --- First, we silence the linter, so that we can set a default value smaller than 1500. -set_option linter.style.longFile 0 in --- Next, we test that the `longFile` linter warns when a file exceeds the allowed value. -set_option linter.style.longFile 10 in -#exit - -/-- -warning: using 'exit' to interrupt Lean ---- -warning: The default value of the `longFile` linter is 1500. -This file is 361 lines long which does not exceed the allowed bound. -Please, remove the `set_option linter.style.longFile 1700`. --/ -#guard_msgs in --- First, we silence the linter, so that we can set a default value smaller than 1500. -set_option linter.style.longFile 0 in --- If we set the allowed bound for the `longFile` linter that is too large, --- the linter tells us to use a smaller bound. -set_option linter.style.longFile 1700 in -#exit - -end longFile diff --git a/test/LongFile.lean b/test/LongFile.lean new file mode 100644 index 0000000000000..faec16c404806 --- /dev/null +++ b/test/LongFile.lean @@ -0,0 +1,68 @@ +import Mathlib.Tactic.Linter.Lint + +/- +# Testing the `longFile` linter + +Things to note: +* `set_option linter.style.longFile 0` disables the linter, allowing us to set a value smaller than + `1500` without triggering the warning for setting a small value for the option; +* `guard_msgs ... in #exit` and `set_option ... in #exit` allow processing of the file *beyond* + `#exit`, since they wrap `#exit` inside an anonymous section, + making Lean active again *after* that anonymous section. + +-/ + +section longFile + +/-- +warning: The default value of the `longFile` linter is 1500. +The current value of 1500 does not exceed the allowed bound. +Please, remove the `set_option linter.style.longFile 1500`. +-/ +#guard_msgs in +-- Do not allow setting a "small" `longFile` linter option +set_option linter.style.longFile 1500 + +/-- +warning: using 'exit' to interrupt Lean +--- +warning: The default value of the `longFile` linter is 1500. +This file is 36 lines long which does not exceed the allowed bound. +Please, remove the `set_option linter.style.longFile 1600`. +-/ +#guard_msgs in +-- Do not allow unnecessarily increasing the `longFile` linter option +set_option linter.style.longFile 1600 in +#exit + +/-- +warning: using 'exit' to interrupt Lean +--- +warning: This file is 51 lines long, but the limit is 10. + +You can extend the allowed length of the file using `set_option linter.style.longFile 1500`. +You can completely disable this linter by setting the length limit to `0`. +-/ +#guard_msgs in +-- First, we silence the linter, so that we can set a default value smaller than 1500. +set_option linter.style.longFile 0 in +-- Next, we test that the `longFile` linter warns when a file exceeds the allowed value. +set_option linter.style.longFile 10 in +#exit + +/-- +warning: using 'exit' to interrupt Lean +--- +warning: The default value of the `longFile` linter is 1500. +This file is 66 lines long which does not exceed the allowed bound. +Please, remove the `set_option linter.style.longFile 1700`. +-/ +#guard_msgs in +-- First, we silence the linter, so that we can set a default value smaller than 1500. +set_option linter.style.longFile 0 in +-- If we set the allowed bound for the `longFile` linter that is too large, +-- the linter tells us to use a smaller bound. +set_option linter.style.longFile 1700 in +#exit + +end longFile