diff --git a/LeanAPAP.lean b/LeanAPAP.lean index 093a197c4a..57aa9081c6 100644 --- a/LeanAPAP.lean +++ b/LeanAPAP.lean @@ -1,5 +1,6 @@ import LeanAPAP.Extras.BSG import LeanAPAP.FiniteField.Basic +import LeanAPAP.Mathlib.Algebra.Algebra.Rat import LeanAPAP.Mathlib.Algebra.Group.Action.Defs import LeanAPAP.Mathlib.Algebra.Group.AddChar import LeanAPAP.Mathlib.Algebra.Order.Group.Unbundled.Basic @@ -44,9 +45,12 @@ import LeanAPAP.Prereqs.Function.Indicator.Basic import LeanAPAP.Prereqs.Function.Indicator.Defs import LeanAPAP.Prereqs.Function.Translate import LeanAPAP.Prereqs.LargeSpec -import LeanAPAP.Prereqs.LpNorm.Compact +import LeanAPAP.Prereqs.LpNorm.Compact.Defs +import LeanAPAP.Prereqs.LpNorm.Compact.Inner import LeanAPAP.Prereqs.LpNorm.Discrete.Basic import LeanAPAP.Prereqs.LpNorm.Discrete.Defs +import LeanAPAP.Prereqs.LpNorm.Discrete.Inner import LeanAPAP.Prereqs.LpNorm.Weighted import LeanAPAP.Prereqs.MarcinkiewiczZygmund +import LeanAPAP.Prereqs.NNLpNorm import LeanAPAP.Prereqs.Rudin diff --git a/LeanAPAP/FiniteField/Basic.lean b/LeanAPAP/FiniteField/Basic.lean index 3aae2650f0..fb70d4d9c7 100644 --- a/LeanAPAP/FiniteField/Basic.lean +++ b/LeanAPAP/FiniteField/Basic.lean @@ -31,22 +31,22 @@ lemma global_dichotomy (hA : A.Nonempty) (hγC : γ ≤ C.dens) (hγ : 0 < γ) _ ≤ _ := div_le_div_of_nonneg_right hAC (card G).cast_nonneg _ = |⟪balance (μ A) ∗ balance (μ A), μ C⟫_[ℝ]| := ?_ _ ≤ ‖balance (μ_[ℝ] A) ∗ balance (μ A)‖_[p] * ‖μ_[ℝ] C‖_[NNReal.conjExponent p] := - abs_l2Inner_le_lpNorm_mul_lpNorm hp'' _ _ + abs_dL2Inner_le_dLpNorm_mul_dLpNorm hp'' _ _ _ ≤ ‖balance (μ_[ℝ] A) ○ balance (μ A)‖_[p] * (card G ^ (-(p : ℝ)⁻¹) * γ ^ (-(p : ℝ)⁻¹)) := - mul_le_mul (lpNorm_conv_le_lpNorm_dconv' (by positivity) (even_two_mul _) _) ?_ + mul_le_mul (dLpNorm_conv_le_dLpNorm_dconv' (by positivity) (even_two_mul _) _) ?_ (by positivity) (by positivity) _ = ‖balance (μ_[ℝ] A) ○ balance (μ A)‖_[↑(2 * ⌈γ.curlog⌉₊), const _ (card G)⁻¹] * γ ^ (-(p : ℝ)⁻¹) := ?_ _ ≤ _ := mul_le_mul_of_nonneg_left ?_ $ by positivity - · rw [← balance_conv, balance, l2Inner_sub_left, l2Inner_const_left, expect_conv, sum_mu ℝ hA, + · rw [← balance_conv, balance, dL2Inner_sub_left, dL2Inner_const_left, expect_conv, sum_mu ℝ hA, expect_mu ℝ hA, sum_mu ℝ hC, conj_trivial, one_mul, mul_one, ← mul_inv_cancel₀, ← mul_sub, abs_mul, abs_of_nonneg, mul_div_cancel_left₀] <;> positivity - · rw [lpNorm_mu hp''.symm.one_le hC, hp''.symm.coe.inv_sub_one, NNReal.coe_natCast, ← mul_rpow] - rw [nnratCast_dens, le_div_iff₀, mul_comm] at hγC + · rw [dLpNorm_mu hp''.symm.one_le hC, hp''.symm.coe.inv_sub_one, NNReal.coe_natCast, ← mul_rpow] + rw [nnratCast_dens, le_div_iff, mul_comm] at hγC refine rpow_le_rpow_of_nonpos ?_ hγC (neg_nonpos.2 ?_) all_goals positivity · simp_rw [Nat.cast_mul, Nat.cast_two, p] - rw [wlpNorm_const_right, mul_assoc, mul_left_comm, NNReal.coe_inv, inv_rpow, rpow_neg] + rw [wLpNorm_const_right, mul_assoc, mul_left_comm, NNReal.coe_inv, inv_rpow, rpow_neg] push_cast any_goals norm_cast; rw [Nat.succ_le_iff] rfl @@ -86,7 +86,7 @@ lemma ap_in_ff (hS : S.Nonempty) (hα₀ : 0 < α) (hε₀ : 0 < ε) (hε₁ : calc _ = ⟪μ V', μ A₁ ∗ μ A₂ ○ 𝟭 S⟫_[ℝ] := by sorry - -- rw [conv_assoc, conv_l2Inner, ← conj_l2Inner] + -- rw [conv_assoc, conv_dL2Inner, ← conj_dL2Inner] -- simp _ = _ := sorry @@ -103,7 +103,7 @@ lemma di_in_ff (hq₃ : 3 ≤ q) (hq : q.Prime) (hε₀ : 0 < ε) (hε₁ : ε < refine ⟨⊤, univ, _⟩ rw [AffineSubspace.direction_top] simp only [AffineSubspace.top_coe, coe_univ, eq_self_iff_true, finrank_top, tsub_self, - Nat.cast_zero, indicate_empty, zero_mul, lpNorm_zero, true_and_iff, + Nat.cast_zero, indicate_empty, zero_mul, nnLpNorm_zero, true_and_iff, Finset.card_empty, zero_div] at hαA ⊢ exact ⟨by positivity, mul_nonpos_of_nonneg_of_nonpos (by positivity) hαA⟩ have hγ₁ : γ ≤ 1 := hγC.trans (by norm_cast; exact dens_le_one) @@ -118,13 +118,13 @@ lemma di_in_ff (hq₃ : 3 ≤ q) (hq : q.Prime) (hε₀ : 0 < ε) (hε₁ : ε < simp [smul_dconv, dconv_smul, smul_smul] · simp [card_univ, show (card G : ℂ) ≠ 0 by sorry] · simp only [comp_const, Nonneg.coe_inv, NNReal.coe_natCast] - rw [← ENNReal.coe_one, lpNorm_const one_ne_zero] + rw [← ENNReal.coe_one, dLpNorm_const one_ne_zero] sorry -- simp only [Nonneg.coe_one, inv_one, rpow_one, norm_inv, norm_coe_nat, -- mul_inv_cancel₀ (show (card G : ℝ) ≠ 0 by positivity)] · have hγ' : (1 : ℝ≥0) ≤ 2 * ⌈γ.curlog⌉₊ := sorry sorry - -- simpa [wlpNorm_nsmul hγ', ← nsmul_eq_mul, div_le_iff₀' (show (0 : ℝ) < card G by positivity), + -- simpa [wLpNorm_nsmul hγ', ← nsmul_eq_mul, div_le_iff' (show (0 : ℝ) < card G by positivity), -- ← div_div, *] using global_dichotomy hA hγC hγ hAC sorry @@ -207,7 +207,7 @@ theorem ff (hq₃ : 3 ≤ q) (hq : q.Prime) {A : Finset G} (hA₀ : A.Nonempty) _ ≤ 2 := by norm_num · positivity all_goals positivity - rw [hB.l2Inner_mu_conv_mu_mu_two_smul_mu] at hBV + rw [hB.dL2Inner_mu_conv_mu_mu_two_smul_mu] at hBV suffices h : (q ^ (n - 65 * curlog α ^ 9) : ℝ) ≤ q ^ (n / 2) by rw [rpow_le_rpow_left_iff ‹_›, sub_le_comm, sub_half, div_le_iff₀' zero_lt_two, ← mul_assoc] at h norm_num at h diff --git a/LeanAPAP/Mathlib/Algebra/Algebra/Rat.lean b/LeanAPAP/Mathlib/Algebra/Algebra/Rat.lean new file mode 100644 index 0000000000..83976519bd --- /dev/null +++ b/LeanAPAP/Mathlib/Algebra/Algebra/Rat.lean @@ -0,0 +1,30 @@ +import Mathlib.Algebra.Algebra.Rat +import Mathlib.Algebra.Module.Basic + +variable {α β : Type*} + +instance [Monoid α] [AddCommMonoid β] [Module ℚ≥0 β] [DistribMulAction α β] : + SMulCommClass ℚ≥0 α β where + smul_comm q a b := by + rw [← q.num_div_den, div_eq_mul_inv] + simp_rw [mul_smul, inv_natCast_smul_comm, Nat.cast_smul_eq_nsmul] + rw [smul_comm a q.num] + +instance [Monoid α] [AddCommMonoid β] [Module ℚ≥0 β] [DistribMulAction α β] : + SMulCommClass α ℚ≥0 β := .symm .. + +instance [Semiring α] [Module ℚ≥0 α] : IsScalarTower ℚ≥0 α α where + smul_assoc q a b := sorry + +instance [Monoid α] [AddCommGroup β] [Module ℚ β] [DistribMulAction α β] : + SMulCommClass ℚ α β where + smul_comm q a b := by + rw [← q.num_div_den, div_eq_mul_inv] + simp_rw [mul_smul, inv_natCast_smul_comm, Int.cast_smul_eq_zsmul] + rw [smul_comm a q.num] + +instance [Monoid α] [AddCommGroup β] [Module ℚ β] [DistribMulAction α β] : SMulCommClass α ℚ β := + .symm .. + +instance [Ring α] [Module ℚ α] : IsScalarTower ℚ α α where + smul_assoc q a b := sorry diff --git a/LeanAPAP/Physics/AlmostPeriodicity.lean b/LeanAPAP/Physics/AlmostPeriodicity.lean index ec36585ebb..780ef86de1 100644 --- a/LeanAPAP/Physics/AlmostPeriodicity.lean +++ b/LeanAPAP/Physics/AlmostPeriodicity.lean @@ -10,6 +10,7 @@ import LeanAPAP.Prereqs.MarcinkiewiczZygmund # Almost-periodicity -/ +open MeasureTheory open scoped Pointwise Combinatorics.Additive namespace Finset @@ -85,7 +86,8 @@ end open Finset Real open scoped BigOperators Pointwise NNReal ENNReal -variable {G : Type*} [Fintype G] {A S : Finset G} {f : G → ℂ} {ε K : ℝ} {k m : ℕ} +variable {G : Type*} [Fintype G] [MeasurableSpace G] [DiscreteMeasurableSpace G] {A S : Finset G} + {f : G → ℂ} {ε K : ℝ} {k m : ℕ} lemma lemma28_end (hε : 0 < ε) (hm : 1 ≤ m) (hk : 64 * m / ε ^ 2 ≤ k) : (8 * m) ^ m * k ^ (m - 1) * A.card ^ k * k * (2 * ‖f‖_[2 * m]) ^ (2 * m) ≤ @@ -95,7 +97,7 @@ lemma lemma28_end (hε : 0 < ε) (hm : 1 ≤ m) (hk : 64 * m / ε ^ 2 ≤ k) : refine mul_ne_zero two_pos.ne' ?_ rw [← pos_iff_ne_zero, ← Nat.succ_le_iff] exact hm - rw [mul_pow (2 : ℝ), ← hmeq, ← lpNorm_pow_eq_sum hm' f, ← mul_assoc, ← mul_assoc, + rw [mul_pow (2 : ℝ), ← hmeq, ← dLpNorm_pow_eq_sum_norm hm' f, ← mul_assoc, ← mul_assoc, mul_right_comm _ (A.card ^ k : ℝ), mul_right_comm _ (A.card ^ k : ℝ), mul_right_comm _ (A.card ^ k : ℝ)] gcongr ?_ * _ * _ @@ -119,6 +121,22 @@ variable [DecidableEq G] [AddCommGroup G] local notation:70 s:70 " ^^ " n:71 => Fintype.piFinset fun _ : Fin n ↦ s +lemma lemma28_part_one (hm : 1 ≤ m) (x : G) : + ∑ a in A ^^ k, ‖∑ i, f (x - a i) - (k • (mu A ∗ f)) x‖ ^ (2 * m) ≤ + (8 * m) ^ m * k ^ (m - 1) * + ∑ a in A ^^ k, ∑ i, ‖f (x - a i) - (mu A ∗ f) x‖ ^ (2 * m) := by + let f' : G → ℂ := fun a ↦ f (x - a) - (mu A ∗ f) x + refine (RCLike.marcinkiewicz_zygmund (by linarith only [hm]) f' ?_).trans_eq' ?_ + · intro i + rw [Fintype.sum_piFinset_apply, sum_sub_distrib] + simp only [sub_eq_zero, sum_const, indicate_apply] + rw [← Pi.smul_apply (card A), ← smul_conv, card_smul_mu, conv_eq_sum_sub'] + simp only [boole_mul, indicate_apply] + rw [← sum_filter, filter_mem_eq_inter, univ_inter, sub_self, smul_zero] + congr with a : 1 + simp only [sum_sub_distrib, Pi.smul_apply, sum_const, card_fin] +variable [DecidableEq G] [AddCommGroup G] + namespace AlmostPeriodicity def LProp (k m : ℕ) (ε : ℝ) (f : G → ℂ) (A : Finset G) (a : Fin k → G) : Prop := @@ -143,21 +161,6 @@ lemma lemma28_markov (hε : 0 < ε) (hm : 1 ≤ m) congr with a : 3 refine pow_le_pow_iff_left ?_ ?_ ?_ <;> positivity -lemma lemma28_part_one (hm : 1 ≤ m) (x : G) : - ∑ a in A ^^ k, ‖∑ i, f (x - a i) - (k • (mu A ∗ f)) x‖ ^ (2 * m) ≤ - (8 * m) ^ m * k ^ (m - 1) * - ∑ a in A ^^ k, ∑ i, ‖f (x - a i) - (mu A ∗ f) x‖ ^ (2 * m) := by - let f' : G → ℂ := fun a ↦ f (x - a) - (mu A ∗ f) x - refine (RCLike.marcinkiewicz_zygmund (by linarith only [hm]) f' ?_).trans_eq' ?_ - · intro i - rw [Fintype.sum_piFinset_apply, sum_sub_distrib] - simp only [sub_eq_zero, sum_const, indicate_apply] - rw [← Pi.smul_apply (card A), ← smul_conv, card_smul_mu, conv_eq_sum_sub'] - simp only [boole_mul, indicate_apply] - rw [← sum_filter, filter_mem_eq_inter, univ_inter, sub_self, smul_zero] - congr with a : 1 - simp only [sum_sub_distrib, Pi.smul_apply, sum_const, card_fin] - lemma lemma28_part_two (hm : 1 ≤ m) (hA : A.Nonempty) : (8 * m) ^ m * k ^ (m - 1) * ∑ a in A ^^ k, ∑ i, ‖τ (a i) f - mu A ∗ f‖_[2 * m] ^ (2 * m) ≤ (8 * m) ^ m * k ^ (m - 1) * ∑ a in A ^^ k, ∑ i : Fin k, (2 * ‖f‖_[2 * m]) ^ (2 * m) := by @@ -166,16 +169,16 @@ lemma lemma28_part_two (hm : 1 ≤ m) (hA : A.Nonempty) : have hm' : 1 < 2 * m := (Nat.mul_le_mul_left 2 hm).trans_lt' $ by norm_num1 have hm'' : (1 : ℝ≥0∞) ≤ 2 * m := by rw [← hmeq, Nat.one_le_cast]; exact hm'.le gcongr - refine (lpNorm_sub_le hm'' _ _).trans ?_ - rw [lpNorm_translate, two_mul ‖f‖_[2 * m], add_le_add_iff_left] + refine (dLpNorm_sub_le hm'' _ _).trans ?_ + rw [dLpNorm_translate, two_mul ‖f‖_[2 * m], add_le_add_iff_left] have hmeq' : ((2 * m : ℝ≥0) : ℝ≥0∞) = 2 * m := by rw [ENNReal.coe_mul, ENNReal.coe_two, ENNReal.coe_natCast] have : (1 : ℝ≥0) < 2 * m := by rw [← Nat.cast_two, ← Nat.cast_mul, Nat.one_lt_cast] exact hm' rw [← hmeq', conv_comm] - refine (lpNorm_conv_le this.le _ _).trans ?_ - rw [l1Norm_mu hA, mul_one] + refine (dLpNorm_conv_le this.le _ _).trans ?_ + rw [dL1Norm_mu hA, mul_one] lemma lemma28 (hε : 0 < ε) (hm : 1 ≤ m) (hk : (64 : ℝ) * m / ε ^ 2 ≤ k) : (A.card ^ k : ℝ) / 2 ≤ (l k m ε f A).card := by @@ -191,7 +194,7 @@ lemma lemma28 (hε : 0 < ε) (hm : 1 ≤ m) (hk : (64 : ℝ) * m / ε ^ 2 ≤ k) have hm' : 2 * m ≠ 0 := by linarith have hmeq : ((2 * m : ℕ) : ℝ≥0∞) = 2 * m := by rw [Nat.cast_mul, Nat.cast_two] rw [← hmeq, mul_pow] - simp only [lpNorm_pow_eq_sum hm'] + simp only [dLpNorm_pow_eq_sum_nnnorm hm'] rw [sum_comm] have : ∀ x : G, ∑ a in A ^^ k, ‖∑ i, f (x - a i) - (k • (mu A ∗ f)) x‖ ^ (2 * m) ≤ @@ -203,7 +206,7 @@ lemma lemma28 (hε : 0 < ε) (hm : 1 ≤ m) (hk : (64 : ℝ) * m / ε ^ 2 ≤ k) simp only [@sum_comm _ _ G] have (a : Fin k → G) (i : Fin k) : ∑ x, ‖f (x - a i) - (mu A ∗ f) x‖ ^ (2 * m) = ‖τ (a i) f - mu A ∗ f‖_[2 * m] ^ (2 * m) := by - rw [← hmeq, lpNorm_pow_eq_sum hm'] + rw [← hmeq, dLpNorm_pow_eq_sum_nnnorm hm'] simp only [Pi.sub_apply, translate_apply] simp only [this] have : @@ -234,18 +237,18 @@ lemma just_the_triangle_inequality {t : G} {a : Fin k → G} (ha : a ∈ l k m refine Finset.sum_congr rfl fun j _ ↦ ?_ rw [sub_right_comm] have h₄₁ : ‖τ t f₁ - k • (mu A ∗ f)‖_[2 * m] = ‖τ (-t) (τ t f₁ - k • (mu A ∗ f))‖_[2 * m] := by - rw [lpNorm_translate] + rw [dLpNorm_translate] have h₄ : ‖τ t f₁ - k • (mu A ∗ f)‖_[2 * m] = ‖f₁ - τ (-t) (k • (mu A ∗ f))‖_[2 * m] := by rw [h₄₁, translate_sub_right, translate_translate] simp have h₅₁ : ‖τ (-t) (k • (mu A ∗ f)) - f₁‖_[2 * m] ≤ k * ε * ‖f‖_[2 * m] := by - rwa [lpNorm_sub_comm, ← h₄, ← h₃] + rwa [dLpNorm_sub_comm, ← h₄, ← h₃] have : (0 : ℝ) < k := by positivity refine le_of_mul_le_mul_left ?_ this - rw [← nsmul_eq_mul, ← lpNorm_nsmul hp _ (_ - mu A ∗ f), nsmul_sub, ← + rw [← nsmul_eq_mul, ← dLpNorm_nsmul hp _ (_ - mu A ∗ f), nsmul_sub, ← translate_smul_right (-t) (mu A ∗ f) k, mul_assoc, mul_left_comm, two_mul ((k : ℝ) * _), ← mul_assoc] - exact (lpNorm_sub_le_lpNorm_sub_add_lpNorm_sub hp _ _).trans (add_le_add h₅₁ h₁) + exact (dLpNorm_sub_le_dLpNorm_sub_add_dLpNorm_sub hp _ _).trans (add_le_add h₅₁ h₁) lemma big_shifts_step2 (L : Finset (Fin k → G)) (hk : k ≠ 0) : (∑ x in L + S.piDiag (Fin k), ∑ l in L, ∑ s in S.piDiag (Fin k), ite (l + s = x) (1 : ℝ) 0) ^ 2 @@ -381,7 +384,7 @@ lemma almost_periodicity (ε : ℝ) (hε : 0 < ε) (hε' : ε ≤ 1) (m : ℕ) ( refine (mul_le_mul_of_nonneg_right this (Nat.cast_nonneg _)).trans ?_ rw [one_mul, Nat.cast_le] exact card_le_univ _ - simp only [mu_empty, zero_conv, translate_zero_right, sub_self, lpNorm_zero] + simp only [mu_empty, zero_conv, translate_zero_right, sub_self, nnLpNorm_zero] positivity let k := ⌈(64 : ℝ) * m / (ε / 2) ^ 2⌉₊ have hk : k ≠ 0 := by positivity @@ -438,19 +441,19 @@ theorem linfty_almost_periodicity (ε : ℝ) (hε₀ : 0 < ε) (hε₁ : ε ≤ = (F ∗ μ C) x := by simp [sub_conv, F] _ = ∑ y, F y * μ C (x - y) := conv_eq_sum_sub' .. _ = ∑ y, F y * μ (x +ᵥ -C) y := by simp [neg_add_eq_sub] - rw [linftyNorm_eq_ciSup] + rw [dLinftyNorm_eq_iSup] refine ciSup_le fun x ↦ ?_ calc ‖(τ t (μ A ∗ 𝟭 B ∗ μ C) - μ A ∗ 𝟭 B ∗ μ C : G → ℂ) x‖ = ‖∑ y, F y * μ (x +ᵥ -C) y‖ := by rw [this] _ ≤ ∑ y, ‖F y * μ (x +ᵥ -C) y‖ := norm_sum_le _ _ - _ = ‖F * μ (x +ᵥ -C)‖_[1] := by rw [l1Norm_eq_sum]; rfl - _ ≤ ‖F‖_[M] * ‖μ_[ℂ] (x +ᵥ -C)‖_[NNReal.conjExponent M] := l1Norm_mul_le hM _ _ + _ = ‖F * μ (x +ᵥ -C)‖_[1] := by rw [dL1Norm_eq_sum_nnnorm]; rfl + _ ≤ ‖F‖_[M] * ‖μ_[ℂ] (x +ᵥ -C)‖_[NNReal.conjExponent M] := dL1Norm_mul_le hM _ _ _ ≤ ε / exp 1 * B.card ^ (M : ℝ)⁻¹ * ‖μ_[ℂ] (x +ᵥ -C)‖_[NNReal.conjExponent M] := by gcongr - simpa only [← ENNReal.coe_natCast, lpNorm_indicate hM₀] using hT _ ht + simpa only [← ENNReal.coe_natCast, dLpNorm_indicate hM₀] using hT _ ht _ = ε * ((C.card / B.card) ^ (-(M : ℝ)⁻¹) / exp 1) := by - rw [← mul_comm_div, lpNorm_mu hM.symm.one_le hC.neg.vadd_finset, card_vadd_finset, + rw [← mul_comm_div, dLpNorm_mu hM.symm.one_le hC.neg.vadd_finset, card_vadd_finset, card_neg, hM.symm.coe.inv_sub_one, div_rpow, mul_assoc, NNReal.coe_natCast, rpow_neg, rpow_neg, ← div_eq_mul_inv, inv_div_inv] <;> positivity _ ≤ ε := mul_le_of_le_one_right (by positivity) $ (div_le_one $ by positivity).2 ?_ @@ -482,12 +485,12 @@ theorem linfty_almost_periodicity_boosted (ε : ℝ) (hε₀ : 0 < ε) (hε₁ : ‖μ T ∗^ k ∗ F - F‖_[∞] = ‖𝔼 a ∈ T ^^ k, (τ (∑ i, a i) F - F)‖_[∞] := by rw [mu_iterConv_conv, expect_sub_distrib, expect_const hT'.piFinset_const] - _ ≤ 𝔼 a ∈ T ^^ k, ‖τ (∑ i, a i) F - F‖_[∞] := lpNorm_expect_le le_top _ _ + _ ≤ 𝔼 a ∈ T ^^ k, ‖τ (∑ i, a i) F - F‖_[∞] := dLpNorm_expect_le le_top _ _ _ ≤ 𝔼 _a ∈ T ^^ k, ε := expect_le_expect fun x hx ↦ ?_ _ = ε := by rw [expect_const hT'.piFinset_const] calc ‖τ (∑ i, x i) F - F‖_[⊤] - _ ≤ ∑ i, ‖τ (x i) F - F‖_[⊤] := lpNorm_translate_sum_sub_le le_top _ _ _ + _ ≤ ∑ i, ‖τ (x i) F - F‖_[⊤] := dLpNorm_translate_sum_sub_le le_top _ _ _ _ ≤ ∑ _i, ε / k := sum_le_sum fun i _ ↦ hT _ $ Fintype.mem_piFinset.1 hx _ _ = ε := by simp only [sum_const, card_fin, nsmul_eq_mul]; rw [mul_div_cancel₀]; positivity diff --git a/LeanAPAP/Physics/DRC.lean b/LeanAPAP/Physics/DRC.lean index 4677ab14f3..86735fe95a 100644 --- a/LeanAPAP/Physics/DRC.lean +++ b/LeanAPAP/Physics/DRC.lean @@ -8,7 +8,7 @@ import LeanAPAP.Prereqs.LpNorm.Weighted # Dependent Random Choice -/ -open Real Finset Fintype Function +open Real Finset Fintype Function MeasureTheory open scoped NNReal Pointwise variable {G : Type*} [DecidableEq G] [Fintype G] [AddCommGroup G] {p : ℕ} {B₁ B₂ A : Finset G} @@ -22,7 +22,7 @@ private lemma lemma_0 (p : ℕ) (B₁ B₂ A : Finset G) (f : G → ℝ) : ∑ s, ⟪𝟭_[ℝ] (B₁ ∩ c p A s) ○ 𝟭 (B₂ ∩ c p A s), f⟫_[ℝ] = (B₁.card * B₂.card) • ∑ x, (μ_[ℝ] B₁ ○ μ B₂) x * (𝟭 A ○ 𝟭 A) x ^ p * f x := by simp_rw [mul_assoc] - simp only [l2Inner_eq_sum, RCLike.conj_to_real, mul_sum, sum_mul, smul_sum, + simp only [dL2Inner_eq_sum, RCLike.conj_to_real, mul_sum, sum_mul, smul_sum, @sum_comm _ _ (Fin p → G), sum_dconv_mul, dconv_apply_sub, Fintype.sum_pow, map_indicate] congr with b₁ congr with b₂ @@ -41,15 +41,17 @@ private lemma sum_cast_c (p : ℕ) (B A : Finset G) : ∑ s, ((B ∩ c p A s).card : ℝ) = A.card ^ p * B.card := by rw [← Nat.cast_sum, sum_c]; norm_cast +variable [MeasurableSpace G] [DiscreteMeasurableSpace G] + /-- If `A` is nonempty, and `B₁` and `B₂` intersect, then the `μ B₁ ○ μ B₂`-weighted Lp norm of `𝟭 A ○ 𝟭 A` is positive. -/ -private lemma lpNorm_conv_pos (hp : p ≠ 0) (hB : (B₁ ∩ B₂).Nonempty) (hA : A.Nonempty) : +private lemma dLpNorm_conv_pos (hp : p ≠ 0) (hB : (B₁ ∩ B₂).Nonempty) (hA : A.Nonempty) : 0 < ‖𝟭_[ℝ] A ○ 𝟭 A‖_[p, μ B₁ ○ μ B₂] ^ p := by - rw [wlpNorm_pow_eq_sum] + rw [wLpNorm_pow_eq_sum] refine sum_pos' (fun x _ ↦ by positivity) ⟨0, mem_univ _, smul_pos ?_ $ pow_pos ?_ _⟩ · rwa [pos_iff_ne_zero, ← Function.mem_support, support_dconv, support_mu, support_mu, ← coe_sub, mem_coe, zero_mem_sub_iff, not_disjoint_iff_nonempty_inter] <;> exact mu_nonneg - · rw [norm_pos_iff, ← Function.mem_support, support_dconv, support_indicate] + · rw [nnnorm_pos, ← Function.mem_support, support_dconv, support_indicate] exact hA.to_set.zero_mem_sub all_goals exact indicate_nonneg -- positivity · positivity @@ -66,7 +68,7 @@ lemma drc (hp₂ : 2 ≤ p) (f : G → ℝ≥0) (hf : ∃ x, x ∈ B₁ - B₂ have := hB.mono inter_subset_left have := hB.mono inter_subset_right have hp₀ : p ≠ 0 := by positivity - have := lpNorm_conv_pos hp₀ hB hA + have := dLpNorm_conv_pos hp₀ hB hA set M : ℝ := 2 ⁻¹ * ‖𝟭_[ℝ] A ○ 𝟭 A‖_[p, μ B₁ ○ μ B₂] ^ p * (sqrt B₁.card * sqrt B₂.card) / A.card ^ p with hM_def @@ -85,7 +87,7 @@ lemma drc (hp₂ : 2 ≤ p) (f : G → ℝ≥0) (hf : ∃ x, x ∈ B₁ - B₂ have hg : ∀ s, 0 ≤ g s := fun s ↦ by rw [hg_def]; dsimp; positivity have hgB : ∑ s, g s = B₁.card * B₂.card * ‖𝟭_[ℝ] A ○ 𝟭 A‖_[p, μ B₁ ○ μ B₂] ^ p := by have hAdconv : 0 ≤ 𝟭_[ℝ] A ○ 𝟭 A := dconv_nonneg indicate_nonneg indicate_nonneg - simpa only [wlpNorm_pow_eq_sum hp₀, l2Inner_eq_sum, sum_dconv, sum_indicate, Pi.one_apply, + simpa only [wLpNorm_pow_eq_sum hp₀, dL2Inner_eq_sum, sum_dconv, sum_indicate, Pi.one_apply, RCLike.inner_apply, RCLike.conj_to_real, norm_of_nonneg (hAdconv _), mul_one, nsmul_eq_mul, Nat.cast_mul, ← hg_def, NNReal.smul_def, NNReal.coe_dconv, NNReal.coe_comp_mu] using lemma_0 p B₁ B₂ A 1 @@ -96,7 +98,7 @@ lemma drc (hp₂ : 2 ≤ p) (f : G → ℝ≥0) (hf : ∃ x, x ∈ B₁ - B₂ refine ⟨_, inter_subset_left (s₂ := c p A s), _, inter_subset_left (s₂ := c p A s), ?_⟩ simp only [indicate_apply, mem_filter, mem_univ, true_and_iff, boole_mul] at hs split_ifs at hs with h; swap - · simp only [zero_mul, l2Inner_eq_sum, Function.comp_apply, RCLike.inner_apply, + · simp only [zero_mul, dL2Inner_eq_sum, Function.comp_apply, RCLike.inner_apply, RCLike.conj_to_real] at hs have : 0 ≤ 𝟭_[ℝ] (A₁ s) ○ 𝟭 (A₂ s) := dconv_nonneg indicate_nonneg indicate_nonneg -- positivity @@ -110,7 +112,7 @@ lemma drc (hp₂ : 2 ≤ p) (f : G → ℝ≥0) (hf : ∃ x, x ∈ B₁ - B₂ positivity refine ⟨(lt_of_mul_lt_mul_left (hs.trans_eq' ?_) $ hg s).le, this.trans $ mul_le_of_le_one_right ?_ $ div_le_one_of_le ?_ ?_, this.trans $ mul_le_of_le_one_left ?_ $ div_le_one_of_le ?_ ?_⟩ - · simp_rw [A₁, A₂, g, ← card_smul_mu, smul_dconv, dconv_smul, l2Inner_smul_left, star_trivial, + · simp_rw [A₁, A₂, g, ← card_smul_mu, smul_dconv, dconv_smul, dL2Inner_smul_left, star_trivial, nsmul_eq_mul, mul_assoc] any_goals positivity all_goals exact Nat.cast_le.2 $ card_mono inter_subset_left @@ -180,7 +182,7 @@ lemma sifting (B₁ B₂ : Finset G) (hε : 0 < ε) (hε₁ : ε ≤ 1) (hδ : 0 simp_rw [nonempty_iff_ne_empty] rintro c r h rfl simp [pow_mul', (zero_lt_four' ℝ).not_le, inv_mul_le_iff (zero_lt_four' ℝ), mul_assoc, - div_nonpos_iff, mul_nonpos_iff, (pow_pos (lpNorm_conv_pos hp₀.ne' hB hA) 2).not_le] at h + div_nonpos_iff, mul_nonpos_iff, (pow_pos (dLpNorm_conv_pos hp₀.ne' hB hA) 2).not_le] at h norm_cast at h simp [hp₀, hA.ne_empty] at h have hA₁ : A₁.Nonempty := aux _ _ hcard₁ @@ -190,12 +192,12 @@ lemma sifting (B₁ B₂ : Finset G) (hε : 0 < ε) (hε₁ : ε ≤ 1) (hδ : 0 calc _ = ∑ x in (s p ε B₁ B₂ A)ᶜ, (μ A₁ ○ μ A₂) x := ?_ _ = ⟪μ_[ℝ] A₁ ○ μ A₂, (↑) ∘ 𝟭_[ℝ≥0] ((s (↑p) ε B₁ B₂ A)ᶜ)⟫_[ℝ] := by - simp [l2Inner_eq_sum, -mem_compl, -mem_s, apply_ite NNReal.toReal, indicate_apply] - _ ≤ _ := (le_div_iff₀ $ lpNorm_conv_pos hp₀.ne' hB hA).2 h + simp [dL2Inner_eq_sum, -mem_compl, -mem_s, apply_ite NNReal.toReal, indicate_apply] + _ ≤ _ := (le_div_iff₀ $ dLpNorm_conv_pos hp₀.ne' hB hA).2 h _ ≤ _ := ?_ · simp_rw [sub_eq_iff_eq_add', sum_add_sum_compl, sum_dconv, map_mu] rw [sum_mu _ hA₁, sum_mu _ hA₂, one_mul] - rw [div_le_iff₀ (lpNorm_conv_pos hp₀.ne' hB hA), ← le_div_iff₀' (zero_lt_two' ℝ)] + rw [div_le_iff₀ (dLpNorm_conv_pos hp₀.ne' hB hA), ← le_div_iff₀' (zero_lt_two' ℝ)] simp only [apply_ite NNReal.toReal, indicate_apply, NNReal.coe_one, NNReal.coe_zero, mul_boole, sum_ite_mem, univ_inter, mul_div_right_comm] calc @@ -213,10 +215,10 @@ lemma sifting (B₁ B₂ : Finset G) (hε : 0 < ε) (hε₁ : ε ≤ 1) (hδ : 0 := ?_ _ ≤ _ := (mul_le_of_le_one_left (mul_nonneg (hp.pow_nonneg _) $ hp.pow_nonneg _) $ - mul_le_one l1Norm_mu_le_one lpNorm_nonneg l1Norm_mu_le_one) + mul_le_one dL1Norm_mu_le_one nnLpNorm_nonneg dL1Norm_mu_le_one) _ ≤ _ := mul_le_mul_of_nonneg_right ?_ $ hp.pow_nonneg _ · have : 0 ≤ μ_[ℝ] B₁ ○ μ B₂ := dconv_nonneg mu_nonneg mu_nonneg - simp_rw [← l1Norm_dconv mu_nonneg mu_nonneg, l1Norm_eq_sum, norm_of_nonneg (this _), sum_mul, + simp_rw [← dL1Norm_dconv mu_nonneg mu_nonneg, dL1Norm_eq_sum_nnnorm, norm_of_nonneg (this _), sum_mul, mul_pow] calc (1 - ε) ^ p ≤ exp (-ε) ^ p := pow_le_pow_left (sub_nonneg.2 hε₁) (one_sub_le_exp_neg _) _ @@ -242,8 +244,8 @@ lemma sifting_cor (hε : 0 < ε) (hε₁ : ε ≤ 1) (hδ : 0 < δ) (hp : Even p rw [nnratCast_dens, le_div_iff₀, ← mul_div_right_comm] calc _ = ‖𝟭_[ℝ] A ○ 𝟭 A‖_[1, μ univ] := by - simp [mu, wlpNorm_smul_right, hp₀, l1Norm_dconv, card_univ, inv_mul_eq_div] - _ ≤ _ := wlpNorm_mono_right (one_le_two.trans $ by norm_cast) _ _ + simp [mu, wLpNorm_smul_right, hp₀, dL1Norm_dconv, card_univ, inv_mul_eq_div] + _ ≤ _ := wLpNorm_mono_right (one_le_two.trans $ by norm_cast) _ _ · exact Nat.cast_pos.2 hA.card_pos obtain ⟨A₁, -, A₂, -, h, hcard₁, hcard₂⟩ := sifting univ univ hε hε₁ hδ hp hp₂ hpε (by simp [univ_nonempty]) hA (by simpa) diff --git a/LeanAPAP/Physics/Unbalancing.lean b/LeanAPAP/Physics/Unbalancing.lean index 3936a82760..09f88d7e09 100644 --- a/LeanAPAP/Physics/Unbalancing.lean +++ b/LeanAPAP/Physics/Unbalancing.lean @@ -1,17 +1,18 @@ import Mathlib.Algebra.Order.Group.PosPart import Mathlib.Data.Complex.ExponentialBounds -import LeanAPAP.Prereqs.Convolution.Discrete.Defs +import LeanAPAP.Prereqs.Convolution.Discrete.Basic +import LeanAPAP.Prereqs.LpNorm.Discrete.Inner import LeanAPAP.Prereqs.LpNorm.Weighted /-! # Unbalancing -/ -open Finset Real -open scoped ComplexConjugate ComplexOrder NNReal +open Finset Real MeasureTheory +open scoped ComplexConjugate ComplexOrder NNReal ENNReal -variable {G : Type*} [Fintype G] [DecidableEq G] [AddCommGroup G] {ν : G → ℝ≥0} {f : G → ℝ} - {g h : G → ℂ} {ε : ℝ} {p : ℕ} +variable {G : Type*} [Fintype G] [MeasurableSpace G] [DiscreteMeasurableSpace G] [DecidableEq G] + [AddCommGroup G] {ν : G → ℝ≥0} {f : G → ℝ} {g h : G → ℂ} {ε : ℝ} {p : ℕ} /-- Note that we do the physical proof in order to avoid the Fourier transform. -/ lemma pow_inner_nonneg' {f : G → ℂ} (hf : f = g ○ g) (hν : (↑) ∘ ν = h ○ h) (k : ℕ) : @@ -20,7 +21,7 @@ lemma pow_inner_nonneg' {f : G → ℂ} (hf : f = g ○ g) (hν : (↑) ∘ ν = ⟪f ^ k, (↑) ∘ ν⟫_[ℂ] = ∑ z : Fin k → G, (‖∑ x, (∏ i, conj (g (x + z i))) * h x‖ : ℂ) ^ 2 by rw [this] positivity - rw [hf, hν, l2Inner_eq_sum] + rw [hf, hν, dL2Inner_eq_sum] simp only [WithLp.equiv_symm_pi_apply, Pi.pow_apply, RCLike.inner_apply, map_pow] simp_rw [dconv_apply h, mul_sum] --TODO: Please make `conv` work here :( @@ -43,7 +44,7 @@ lemma pow_inner_nonneg' {f : G → ℂ} (hf : f = g ○ g) (hν : (↑) ∘ ν = lemma pow_inner_nonneg {f : G → ℝ} (hf : (↑) ∘ f = g ○ g) (hν : (↑) ∘ ν = h ○ h) (k : ℕ) : (0 : ℝ) ≤ ⟪(↑) ∘ ν, f ^ k⟫_[ℝ] := by sorry - -- simpa [← Complex.zero_le_real, starRingEnd_apply, l2Inner_eq_sum] + -- simpa [← Complex.zero_le_real, starRingEnd_apply, dL2Inner_eq_sum] -- using pow_inner_nonneg' hf hν k private lemma log_ε_pos (hε₀ : 0 < ε) (hε₁ : ε ≤ 1) : 0 < log (3 / ε) := @@ -56,16 +57,18 @@ private lemma p'_pos (hp : 5 ≤ p) (hε₀ : 0 < ε) (hε₁ : ε ≤ 1) : 0 < private lemma unbalancing' (p : ℕ) (hp : 5 ≤ p) (hp₁ : Odd p) (hε₀ : 0 < ε) (hε₁ : ε ≤ 1) (hf : (↑) ∘ f = g ○ g) (hν : (↑) ∘ ν = h ○ h) (hν₁ : ‖((↑) ∘ ν : G → ℝ)‖_[1] = 1) (hε : ε ≤ ‖f‖_[p, ν]) : - 1 + ε / 2 ≤ ‖f + 1‖_[(⟨24 / ε * log (3 / ε) * p, (p'_pos hp hε₀ hε₁).le⟩ : ℝ≥0), ν] := by - simp only [l1Norm_eq_sum, NNReal.norm_eq, Function.comp_apply] at hν₁ + 1 + ε / 2 ≤ ‖f + 1‖_[.ofReal (24 / ε * log (3 / ε) * p), ν] := by + simp only [dL1Norm_eq_sum_nnnorm, NNReal.nnnorm_eq, Function.comp_apply] at hν₁ obtain hf₁ | hf₁ := le_total 2 ‖f + 1‖_[2 * p, ν] · calc 1 + ε / 2 ≤ 1 + 1 / 2 := add_le_add_left (div_le_div_of_nonneg_right hε₁ zero_le_two) _ _ ≤ 2 := by norm_num _ ≤ ‖f + 1‖_[2 * p, ν] := hf₁ - _ ≤ _ := wlpNorm_mono_right (NNReal.coe_le_coe.1 ?_) _ _ + _ ≤ _ := wLpNorm_mono_right ?_ _ _ + norm_cast + rw [ENNReal.natCast_le_ofReal (by positivity)] push_cast - refine mul_le_mul_of_nonneg_right ?_ ?_ + gcongr calc 2 ≤ 24 / 1 * 0.6931471803 := by norm_num _ ≤ 24 / ε * log (3 / ε) := @@ -82,13 +85,15 @@ private lemma unbalancing' (p : ℕ) (hp : 5 ≤ p) (hp₁ : Odd p) (hε₀ : 0 _ ≤ ⟪((↑) ∘ ν : G → ℝ), f ^ p⟫_[ℝ] + ∑ i, ↑(ν i) * ((f ^ (p - 1)) i * |f| i) := (le_add_of_nonneg_left $ pow_inner_nonneg hf hν _) _ = _ := ?_ - rw [wlpNorm_pow_eq_sum hp₁.pos.ne'] - dsimp - refine sum_congr rfl fun i _ ↦ ?_ - rw [← abs_of_nonneg ((Nat.Odd.sub_odd hp₁ odd_one).pow_nonneg $ f _), abs_pow, - pow_sub_one_mul hp₁.pos.ne'] - simp [l2Inner_eq_sum, ← sum_add_distrib, ← mul_add, ← pow_sub_one_mul hp₁.pos.ne' (f _), - mul_sum, mul_left_comm (2 : ℝ), add_abs_eq_two_nsmul_posPart] + · norm_cast + rw [wLpNorm_pow_eq_sum hp₁.pos.ne'] + push_cast + dsimp + refine sum_congr rfl fun i _ ↦ ?_ + rw [← abs_of_nonneg ((Nat.Odd.sub_odd hp₁ odd_one).pow_nonneg $ f _), abs_pow, + pow_sub_one_mul hp₁.pos.ne', NNReal.smul_def, smul_eq_mul] + · simp [dL2Inner_eq_sum, ← sum_add_distrib, ← mul_add, ← pow_sub_one_mul hp₁.pos.ne' (f _), + mul_sum, mul_left_comm (2 : ℝ), add_abs_eq_two_nsmul_posPart] set P := univ.filter fun i ↦ 0 ≤ f i set T := univ.filter fun i ↦ 3 / 4 * ε ≤ f i have hTP : T ⊆ P := monotone_filter_right _ fun i ↦ le_trans $ by positivity @@ -106,7 +111,7 @@ private lemma unbalancing' (p : ℕ) (hp : 5 ≤ p) (hp₁ : Odd p) (hε₀ : 0 _ ≤ ∑ i in P \ T, ↑(ν i) * (3 / 4 * ε) ^ p := sum_le_sum fun i hi ↦ ?_ _ = (3 / 4) ^ p * ε ^ p * ∑ i in P \ T, (ν i : ℝ) := by rw [← sum_mul, mul_comm, mul_pow] _ ≤ 4⁻¹ * ε ^ p * ∑ i, (ν i : ℝ) := ?_ - _ = 4⁻¹ * ε ^ p := by rw [hν₁, mul_one] + _ = 4⁻¹ * ε ^ p := by norm_cast; simp [hν₁] · simp only [mem_sdiff, mem_filter, mem_univ, true_and_iff, not_le, P, T] at hi exact mul_le_mul_of_nonneg_left (pow_le_pow_left hi.1 hi.2.le _) (by positivity) · refine mul_le_mul (mul_le_mul_of_nonneg_right (le_trans (pow_le_pow_of_le_one ?_ ?_ hp) ?_) @@ -116,8 +121,12 @@ private lemma unbalancing' (p : ℕ) (hp : 5 ≤ p) (hp₁ : Odd p) (hε₀ : 0 | norm_num replace hf₁ : ‖f‖_[2 * p, ν] ≤ 3 := by calc - _ ≤ ‖f + 1‖_[2 * p, ν] + ‖(1 : G → ℝ)‖_[2 * p, ν] := wlpNorm_le_add_wlpNorm_add hp' _ _ _ - _ ≤ (2 + 1 : ℝ) := (add_le_add hf₁ (by rw [wlpNorm_one, hν₁, one_rpow]; positivity)) + _ ≤ ‖f + 1‖_[2 * p, ν] + ‖(1 : G → ℝ)‖_[2 * p, ν] := + wLpNorm_le_add_wLpNorm_add (mod_cast hp') _ _ _ + _ ≤ 2 + 1 := by + gcongr + have : 2 * (p : ℝ≥0∞) ≠ 0 := mul_ne_zero two_ne_zero (by positivity) + simp [wLpNorm_one, ENNReal.mul_ne_top, *] _ = 3 := by norm_num replace hp' := zero_lt_one.trans_le hp' have : 4⁻¹ * ε ^ p ≤ sqrt (∑ i in T, ν i) * 3 ^ p := by @@ -130,17 +139,16 @@ private lemma unbalancing' (p : ℕ) (hp : 5 ≤ p) (hp₁ : Odd p) (hε₀ : 0 _ = ∑ i in T, sqrt (ν i) * sqrt (ν i * |(f ^ (2 * p)) i|) := by simp [← mul_assoc, pow_mul'] _ ≤ sqrt (∑ i in T, ν i) * sqrt (∑ i in T, ν i * |(f ^ (2 * p)) i|) := (sum_sqrt_mul_sqrt_le _ (fun i ↦ ?_) fun i ↦ ?_) - _ ≤ sqrt (∑ i in T, ν i) * sqrt (∑ i, ν i * |(f ^ (2 * p)) i|) := - (mul_le_mul_of_nonneg_left (sqrt_le_sqrt $ sum_le_univ_sum_of_nonneg fun i ↦ ?_) ?_) + _ ≤ sqrt (∑ i in T, ν i) * sqrt (∑ i, ν i * |(f ^ (2 * p)) i|) := by + gcongr; exact sum_le_univ_sum_of_nonneg fun i ↦ by positivity _ = sqrt (∑ i in T, ν i) * ‖f‖_[2 * ↑p, ν] ^ p := ?_ - _ ≤ _ := mul_le_mul_of_nonneg_left (pow_le_pow_left wlpNorm_nonneg hf₁ _) ?_ + _ ≤ _ := by gcongr; exact mod_cast hf₁ any_goals positivity - rw [wlpNorm_eq_sum hp'.ne', NNReal.coe_mul, mul_inv, rpow_mul, NNReal.coe_natCast, - rpow_inv_natCast_pow] - simp only [wlpNorm_eq_sum hp'.ne', sqrt_eq_rpow, Nonneg.coe_one, rpow_two, - abs_nonneg, NNReal.smul_def, rpow_mul, Pi.pow_apply, abs_pow, norm_eq_abs, mul_pow, - rpow_natCast, smul_eq_mul, pow_mul, one_div, NNReal.coe_two] - all_goals positivity + rw [wLpNorm_eq_sum (mod_cast hp'.ne') (by simp [ENNReal.mul_ne_top])] + norm_cast + rw [← NNReal.rpow_mul_natCast] + have : (p : ℝ) ≠ 0 := by positivity + simp [mul_comm, this, Real.sqrt_eq_rpow] set p' := 24 / ε * log (3 / ε) * p have hp' : 0 < p' := p'_pos hp hε₀ hε₁ have : 1 - 8⁻¹ * ε ≤ (∑ i in T, ↑(ν i)) ^ p'⁻¹ := by @@ -177,10 +185,10 @@ private lemma unbalancing' (p : ℕ) (hp : 5 ≤ p) (hp₁ : Odd p) (hε₀ : 0 _ ≤ (∑ i, ↑(ν i) * |f i + 1| ^ p') ^ p'⁻¹ := rpow_le_rpow ?_ (sum_le_sum_of_subset_of_nonneg (subset_univ _) fun i _ _ ↦ ?_) ?_ _ = _ := by - rw [wlpNorm_eq_sum (NNReal.coe_ne_zero.1 _)] - simp [NNReal.smul_def, hp'.ne', p'] - dsimp + rw [wLpNorm_eq_sum] + simp [NNReal.smul_def, hp'.ne', p', (p'_pos hp hε₀ hε₁).le] positivity + simp all_goals positivity /-- The unbalancing step. Note that we do the physical proof in order to avoid the Fourier @@ -188,8 +196,7 @@ transform. -/ lemma unbalancing (p : ℕ) (hp : p ≠ 0) (ε : ℝ) (hε₀ : 0 < ε) (hε₁ : ε ≤ 1) (ν : G → ℝ≥0) (f : G → ℝ) (g h : G → ℂ) (hf : (↑) ∘ f = g ○ g) (hν : (↑) ∘ ν = h ○ h) (hν₁ : ‖((↑) ∘ ν : G → ℝ)‖_[1] = 1) (hε : ε ≤ ‖f‖_[p, ν]) : - 1 + ε / 2 ≤‖f + 1‖_[(⟨120 / ε * log (3 / ε) * p, by - have := log_ε_pos hε₀ hε₁; positivity⟩ : ℝ≥0), ν] := by + 1 + ε / 2 ≤‖f + 1‖_[.ofReal (120 / ε * log (3 / ε) * p), ν] := by have := log_ε_pos hε₀ hε₁ have := calc @@ -197,15 +204,15 @@ lemma unbalancing (p : ℕ) (hp : p ≠ 0) (ε : ℝ) (hε₀ : 0 < ε) (hε₁ _ ≤ 2 * p + 3 := add_le_add_right (mul_le_mul_left' (Nat.one_le_iff_ne_zero.2 hp) _) _ calc - _ ≤ ‖f + 1‖_[(⟨_, (p'_pos this hε₀ hε₁).le⟩ : ℝ≥0), ν] := + _ ≤ ↑‖f + 1‖_[.ofReal (24 / ε * log (3 / ε) * ↑(2 * p + 3)), ν] := unbalancing' (2 * p + 3) this ((even_two_mul _).add_odd $ by decide) hε₀ hε₁ hf hν hν₁ $ hε.trans $ - wlpNorm_mono_right + wLpNorm_mono_right (Nat.cast_le.2 $ le_add_of_le_left $ le_mul_of_one_le_left' one_le_two) _ _ - _ ≤ _ := wlpNorm_mono_right ?_ _ _ + _ ≤ _ := wLpNorm_mono_right ?_ _ _ + rw [← Nat.one_le_iff_ne_zero] at hp + gcongr ENNReal.ofReal ?_ calc - _ ≤ 24 / ε * log (3 / ε) * ↑(2 * p + 3 * p) := - mul_le_mul_of_nonneg_left (Nat.cast_le.2 $ add_le_add_left - (le_mul_of_one_le_right ?_ $ Nat.one_le_iff_ne_zero.2 hp) _) ?_ + _ = 24 / ε * log (3 / ε) * ↑(2 * p + 3 * 1) := by simp + _ ≤ 24 / ε * log (3 / ε) * ↑(2 * p + 3 * p) := by gcongr _ = _ := by push_cast; ring - all_goals positivity diff --git a/LeanAPAP/Prereqs/AddChar/Basic.lean b/LeanAPAP/Prereqs/AddChar/Basic.lean index bc3a2e599c..b55a3f0d00 100644 --- a/LeanAPAP/Prereqs/AddChar/Basic.lean +++ b/LeanAPAP/Prereqs/AddChar/Basic.lean @@ -1,10 +1,10 @@ import LeanAPAP.Mathlib.Algebra.Group.AddChar import LeanAPAP.Prereqs.Expect.Basic -import LeanAPAP.Prereqs.LpNorm.Discrete.Defs +import LeanAPAP.Prereqs.LpNorm.Discrete.Inner open Finset hiding card open Fintype (card) -open Function +open Function MeasureTheory open scoped BigOperators ComplexConjugate DirectSum variable {G H R : Type*} @@ -15,8 +15,8 @@ variable [AddGroup G] section RCLike variable [RCLike R] -protected lemma l2Inner_self [Fintype G] (ψ : AddChar G R) : - ⟪(ψ : G → R), ψ⟫_[R] = Fintype.card G := l2Inner_self_of_norm_eq_one ψ.norm_apply +protected lemma dL2Inner_self [Fintype G] (ψ : AddChar G R) : + ⟪(ψ : G → R), ψ⟫_[R] = Fintype.card G := dL2Inner_self_of_norm_eq_one ψ.norm_apply end RCLike @@ -45,26 +45,26 @@ variable [AddCommGroup G] section RCLike variable [RCLike R] {ψ₁ ψ₂ : AddChar G R} -lemma l2Inner_eq [Fintype G] (ψ₁ ψ₂ : AddChar G R) : +lemma dL2Inner_eq [Fintype G] (ψ₁ ψ₂ : AddChar G R) : ⟪(ψ₁ : G → R), ψ₂⟫_[R] = if ψ₁ = ψ₂ then ↑(card G) else 0 := by split_ifs with h - · rw [h, AddChar.l2Inner_self] + · rw [h, AddChar.dL2Inner_self] have : ψ₁⁻¹ * ψ₂ ≠ 1 := by rwa [Ne, inv_mul_eq_one] - simp_rw [l2Inner_eq_sum, ← inv_apply_eq_conj] + simp_rw [dL2Inner_eq_sum, ← inv_apply_eq_conj] simpa [map_neg_eq_inv] using sum_eq_zero_iff_ne_zero.2 this -lemma l2Inner_eq_zero_iff_ne [Fintype G] : ⟪(ψ₁ : G → R), ψ₂⟫_[R] = 0 ↔ ψ₁ ≠ ψ₂ := by - rw [l2Inner_eq, Ne.ite_eq_right_iff (Nat.cast_ne_zero.2 Fintype.card_ne_zero)] +lemma dL2Inner_eq_zero_iff_ne [Fintype G] : ⟪(ψ₁ : G → R), ψ₂⟫_[R] = 0 ↔ ψ₁ ≠ ψ₂ := by + rw [dL2Inner_eq, Ne.ite_eq_right_iff (Nat.cast_ne_zero.2 Fintype.card_ne_zero)] -lemma l2Inner_eq_card_iff_eq [Fintype G] : ⟪(ψ₁ : G → R), ψ₂⟫_[R] = card G ↔ ψ₁ = ψ₂ := by - rw [l2Inner_eq, Ne.ite_eq_left_iff (Nat.cast_ne_zero.2 Fintype.card_ne_zero)] +lemma dL2Inner_eq_card_iff_eq [Fintype G] : ⟪(ψ₁ : G → R), ψ₂⟫_[R] = card G ↔ ψ₁ = ψ₂ := by + rw [dL2Inner_eq, Ne.ite_eq_left_iff (Nat.cast_ne_zero.2 Fintype.card_ne_zero)] variable (G R) protected lemma linearIndependent [Finite G] : LinearIndependent R ((⇑) : AddChar G R → G → R) := by cases nonempty_fintype G - exact linearIndependent_of_ne_zero_of_l2Inner_eq_zero AddChar.coe_ne_zero fun ψ₁ ψ₂ ↦ - l2Inner_eq_zero_iff_ne.2 + exact linearIndependent_of_ne_zero_of_dL2Inner_eq_zero AddChar.coe_ne_zero fun ψ₁ ψ₂ ↦ + dL2Inner_eq_zero_iff_ne.2 noncomputable instance instFintype [Finite G] : Fintype (AddChar G R) := @Fintype.ofFinite _ (AddChar.linearIndependent G R).finite diff --git a/LeanAPAP/Prereqs/Chang.lean b/LeanAPAP/Prereqs/Chang.lean index 36b6644865..0f232e2555 100644 --- a/LeanAPAP/Prereqs/Chang.lean +++ b/LeanAPAP/Prereqs/Chang.lean @@ -25,7 +25,7 @@ lemma α_pos (hf : f ≠ 0) : 0 < α f := by unfold α; positivity lemma α_le_one (f : G → ℂ) : α f ≤ 1 := by refine div_le_one_of_le (div_le_of_nonneg_of_le_mul ?_ ?_ ?_) ?_ any_goals positivity - rw [l1Norm_eq_sum, l2Norm_sq_eq_sum] + rw [dL1Norm_eq_sum_nnnorm, dL2Norm_sq_eq_sum_nnnorm] exact sq_sum_le_card_mul_sum_sq lemma general_hoelder (hη : 0 ≤ η) (ν : G → ℝ≥0) (hfν : ∀ x, f x ≠ 0 → 1 ≤ ν x) @@ -44,11 +44,11 @@ lemma general_hoelder (hη : 0 ≤ η) (ν : G → ℝ≥0) (hfν : ∀ x, f x _ ≤ _ := inner_le_weight_mul_Lp_of_nonneg _ (p := m) ?_ _ _ (fun _ ↦ norm_nonneg _) fun _ ↦ norm_nonneg _ _ = ‖f‖_[1] ^ (1 - (m : ℝ)⁻¹) * (∑ x, ‖f x‖ * ‖∑ γ in Δ, c γ * conj (γ x)‖ ^ m) ^ (m⁻¹ : ℝ) := - by simp_rw [l1Norm_eq_sum, rpow_natCast] + by simp_rw [dL1Norm_eq_sum_nnnorm, rpow_natCast] rotate_left · rw [← nsmul_eq_mul'] exact card_nsmul_le_sum _ _ _ fun x hx ↦ mem_largeSpec.1 $ hΔ hx - · simp_rw [mul_sum, mul_comm (f _), mul_assoc (c _), @sum_comm _ _ G, ← mul_sum, ← l2Inner_eq_sum, + · simp_rw [mul_sum, mul_comm (f _), mul_assoc (c _), @sum_comm _ _ G, ← mul_sum, ← dL2Inner_eq_sum, ← dft_apply, ← hc, ← RCLike.ofReal_sum, RCLike.norm_ofReal] exact le_abs_self _ · norm_cast @@ -75,7 +75,7 @@ lemma general_hoelder (hη : 0 ≤ η) (ν : G → ℝ≥0) (hfν : ∀ x, f x _ ≤ (∑ x, ‖f x‖ ^ 2) * ∑ x, (sqrt (ν x) * ‖∑ γ in Δ, c γ * conj (γ x)‖ ^ m) ^ 2 := sum_mul_sq_le_sq_mul_sq _ _ _ _ ≤ ‖f‖_[2] ^ 2 * ∑ x, ν x * (‖∑ γ in Δ, c γ * conj (γ x)‖ ^ 2) ^ m := by - simp_rw [l2Norm_sq_eq_sum, mul_pow, sq_sqrt (NNReal.coe_nonneg _), pow_right_comm]; rfl + simp_rw [dL2Norm_sq_eq_sum_nnnorm, mul_pow, sq_sqrt (NNReal.coe_nonneg _), pow_right_comm]; rfl rw [mul_rotate', mul_left_comm, mul_pow, mul_pow, ← pow_mul', ← pow_mul', ← div_le_iff₀', mul_div_assoc, mul_div_assoc] at this calc @@ -88,7 +88,7 @@ lemma general_hoelder (hη : 0 ≤ η) (ν : G → ℝ≥0) (hfν : ∀ x, f x (norm_sum_le _ _).trans $ sum_le_sum fun _ _ ↦ norm_sum_le _ _ _ = _ := by simp [energy, norm_c, -Complex.norm_eq_abs, norm_prod] · push_cast - simp_rw [← RCLike.conj_mul, dft_apply, l2Inner_eq_sum, map_sum, map_mul, RCLike.conj_conj, + simp_rw [← RCLike.conj_mul, dft_apply, dL2Inner_eq_sum, map_sum, map_mul, RCLike.conj_conj, mul_pow, sum_pow', sum_mul, mul_sum, @sum_comm _ _ G, ← AddChar.inv_apply_eq_conj, ← AddChar.neg_apply', prod_mul_prod_comm, ← AddChar.add_apply, ← AddChar.sum_apply, mul_left_comm (Algebra.cast (ν _ : ℝ) : ℂ), ← mul_sum, ← sub_eq_add_neg, sum_sub_distrib, @@ -126,7 +126,7 @@ lemma AddDissociated.boringEnergy_le [DecidableEq G] {s : Finset G} obtain rfl | hn := eq_or_ne n 0 · simp calc - _ = ‖dft (𝟭 s)‖ₙ_[↑(2 * n)] ^ (2 * n) := by rw [nlpNorm_dft_indicate_pow] + _ = ‖dft (𝟭 s)‖ₙ_[↑(2 * n)] ^ (2 * n) := by rw [cLpNorm_dft_indicate_pow] _ ≤ (4 * rexp 2⁻¹ * sqrt ↑(2 * n) * ‖dft (𝟭 s)‖ₙ_[2]) ^ (2 * n) := by gcongr refine rudin_ineq (le_mul_of_one_le_right zero_le_two $ Nat.one_le_iff_ne_zero.2 hn) @@ -134,7 +134,7 @@ lemma AddDissociated.boringEnergy_le [DecidableEq G] {s : Finset G} rwa [cft_dft, support_comp_eq_preimage, support_indicate, Set.preimage_comp, Set.neg_preimage, addDissociated_neg, AddEquiv.addDissociated_preimage] _ = _ := by - simp_rw [mul_pow, pow_mul, nl2Norm_dft_indicate] + simp_rw [mul_pow, pow_mul, cL2Norm_dft_indicate] rw [← exp_nsmul, sq_sqrt, sq_sqrt] simp_rw [← mul_pow] simp [changConst] diff --git a/LeanAPAP/Prereqs/Convolution/Norm.lean b/LeanAPAP/Prereqs/Convolution/Norm.lean index cb1d8f979d..eaf44b2a0d 100644 --- a/LeanAPAP/Prereqs/Convolution/Norm.lean +++ b/LeanAPAP/Prereqs/Convolution/Norm.lean @@ -1,6 +1,8 @@ import Mathlib.Data.Real.StarOrdered import LeanAPAP.Prereqs.Convolution.Discrete.Defs import LeanAPAP.Prereqs.LpNorm.Discrete.Basic +import LeanAPAP.Prereqs.Convolution.Order +import LeanAPAP.Prereqs.LpNorm.Compact.Defs /-! # Norm of a convolution @@ -9,120 +11,121 @@ This file characterises the L1-norm of the convolution of two functions and prov convolution inequality. -/ -open Finset Function Real +open Finset Function Real MeasureTheory open scoped ComplexConjugate ENNReal NNReal Pointwise variable {α β : Type*} [Fintype α] [DecidableEq α] [AddCommGroup α] section RCLike -variable [RCLike β] +variable [RCLike β] {p : ℝ≥0∞} -@[simp] lemma lpNorm_trivChar (p : ℝ≥0∞) : ‖(trivChar : α → β)‖_[p] = 1 := by - obtain _ | p := p - · simp only [ENNReal.none_eq_top, linftyNorm_eq_ciSup, trivChar_apply, apply_ite, norm_one, - norm_zero] - exact IsLUB.ciSup_eq ⟨by aesop (add simp mem_upperBounds), fun x hx ↦ hx ⟨0, if_pos rfl⟩⟩ - obtain rfl | hp := eq_or_ne p 0 - · simp [l0Norm_eq_card] - · simp [lpNorm_eq_sum hp, apply_ite, hp] +lemma conv_eq_dL2Inner (f g : α → β) (a : α) : (f ∗ g) a = ⟪conj f, τ a fun x ↦ g (-x)⟫_[β] := by + simp [dL2Inner_eq_sum, conv_eq_sum_sub', map_sum] -lemma conv_eq_l2Inner (f g : α → β) (a : α) : (f ∗ g) a = ⟪conj f, τ a fun x ↦ g (-x)⟫_[β] := by - simp [l2Inner_eq_sum, conv_eq_sum_sub', map_sum] +lemma dconv_eq_dL2Inner (f g : α → β) (a : α) : (f ○ g) a = conj ⟪f, τ a g⟫_[β] := by + simp [dL2Inner_eq_sum, dconv_eq_sum_sub', map_sum] -lemma dconv_eq_l2Inner (f g : α → β) (a : α) : (f ○ g) a = conj ⟪f, τ a g⟫_[β] := by - simp [l2Inner_eq_sum, dconv_eq_sum_sub', map_sum] - -lemma l2Inner_dconv (f g h : α → β) : ⟪f, g ○ h⟫_[β] = ⟪conj g, conj f ∗ conj h⟫_[β] := by +lemma dL2Inner_dconv (f g h : α → β) : ⟪f, g ○ h⟫_[β] = ⟪conj g, conj f ∗ conj h⟫_[β] := by calc _ = ∑ b, ∑ a, g a * conj (h b) * conj (f (a - b)) := by - simp_rw [l2Inner, mul_comm _ ((_ ○ _) _), sum_dconv_mul]; exact sum_comm + simp_rw [dL2Inner, mul_comm _ ((_ ○ _) _), sum_dconv_mul]; exact sum_comm _ = ∑ b, ∑ a, conj (f a) * conj (h b) * g (a + b) := by simp_rw [← Fintype.sum_prod_type'] exact Fintype.sum_equiv ((Equiv.refl _).prodShear Equiv.subRight) _ _ (by simp [mul_rotate, mul_right_comm]) _ = _ := by - simp_rw [l2Inner, mul_comm _ ((_ ∗ _) _), sum_conv_mul, Pi.conj_apply, RCLike.conj_conj] + simp_rw [dL2Inner, mul_comm _ ((_ ∗ _) _), sum_conv_mul, Pi.conj_apply, RCLike.conj_conj] exact sum_comm -lemma l2Inner_conv (f g h : α → β) : ⟪f, g ∗ h⟫_[β] = ⟪conj g, conj f ○ conj h⟫_[β] := by - simp_rw [l2Inner_dconv, RCLike.conj_conj] +lemma dL2Inner_conv (f g h : α → β) : ⟪f, g ∗ h⟫_[β] = ⟪conj g, conj f ○ conj h⟫_[β] := by + simp_rw [dL2Inner_dconv, RCLike.conj_conj] + +lemma dconv_dL2Inner (f g h : α → β) : ⟪f ○ g, h⟫_[β] = ⟪conj h ∗ conj g, conj f⟫_[β] := by + rw [dL2Inner_anticomm, dL2Inner_anticomm (_ ∗ _), conj_dconv, dL2Inner_dconv]; simp -lemma dconv_l2Inner (f g h : α → β) : ⟪f ○ g, h⟫_[β] = ⟪conj h ∗ conj g, conj f⟫_[β] := by - rw [l2Inner_anticomm, l2Inner_anticomm (_ ∗ _), conj_dconv, l2Inner_dconv]; simp +lemma conv_dL2Inner (f g h : α → β) : ⟪f ∗ g, h⟫_[β] = ⟪conj h ○ conj g, conj f⟫_[β] := by + rw [dL2Inner_anticomm, dL2Inner_anticomm (_ ○ _), conj_conv, dL2Inner_conv]; simp -lemma conv_l2Inner (f g h : α → β) : ⟪f ∗ g, h⟫_[β] = ⟪conj h ○ conj g, conj f⟫_[β] := by - rw [l2Inner_anticomm, l2Inner_anticomm (_ ○ _), conj_conv, l2Inner_conv]; simp +variable [MeasurableSpace α] [DiscreteMeasurableSpace α] + +@[simp] lemma dLpNorm_trivChar (hp : p ≠ 0) : ‖(trivChar : α → β)‖_[p] = 1 := by + obtain _ | p := p + · simp only [ENNReal.none_eq_top, dLinftyNorm_eq_iSup, trivChar_apply, apply_ite, nnnorm_one, + nnnorm_zero] + exact IsLUB.ciSup_eq ⟨by aesop (add simp mem_upperBounds), fun x hx ↦ hx ⟨0, if_pos rfl⟩⟩ + · simp at hp + simp [dLpNorm_eq_sum_nnnorm hp, apply_ite, hp] --- TODO: This proof would feel much less painful if we had `ℝ≥0`-valued Lp-norms. /-- A special case of **Young's convolution inequality**. -/ -lemma lpNorm_conv_le {p : ℝ≥0} (hp : 1 ≤ p) (f g : α → β) : ‖f ∗ g‖_[p] ≤ ‖f‖_[p] * ‖g‖_[1] := by +lemma dLpNorm_conv_le {p : ℝ≥0} (hp : 1 ≤ p) (f g : α → β) : ‖f ∗ g‖_[p] ≤ ‖f‖_[p] * ‖g‖_[1] := by obtain rfl | hp := hp.eq_or_lt - · simp_rw [ENNReal.coe_one, l1Norm_eq_sum, sum_mul_sum, conv_eq_sum_sub'] + · simp_rw [ENNReal.coe_one, dL1Norm_eq_sum_nnnorm, sum_mul_sum, conv_eq_sum_sub'] calc - ∑ x, ‖∑ y, f y * g (x - y)‖ ≤ ∑ x, ∑ y, ‖f y * g (x - y)‖ := - sum_le_sum fun x _ ↦ norm_sum_le _ _ + ∑ x, ‖∑ y, f y * g (x - y)‖₊ ≤ ∑ x, ∑ y, ‖f y * g (x - y)‖₊ := + sum_le_sum fun x _ ↦ nnnorm_sum_le _ _ _ = _ := ?_ rw [sum_comm] - simp_rw [norm_mul] + simp_rw [nnnorm_mul] exact sum_congr rfl fun x _ ↦ Fintype.sum_equiv (Equiv.subRight x) _ _ fun _ ↦ rfl have hp₀ := zero_lt_one.trans hp - rw [← rpow_le_rpow_iff _ (mul_nonneg _ _) hp₀, mul_rpow] - any_goals exact lpNorm_nonneg + rw [← NNReal.rpow_le_rpow_iff hp₀, NNReal.mul_rpow] dsimp - simp_rw [lpNorm_rpow_eq_sum hp₀.ne', conv_eq_sum_sub'] + simp_rw [dLpNorm_rpow_eq_sum_nnnorm hp₀.ne', conv_eq_sum_sub'] have hpconj : (p : ℝ).IsConjExponent (1 - (p : ℝ)⁻¹)⁻¹ := ⟨hp, by simp_rw [inv_inv, add_sub_cancel]⟩ - have (x) : ‖∑ y, f y * g (x - y)‖ ^ (p : ℝ) ≤ - (∑ y, ‖f y‖ ^ (p : ℝ) * ‖g (x - y)‖) * (∑ y, ‖g (x - y)‖) ^ (p - 1 : ℝ) := by - rw [← le_rpow_inv_iff_of_pos (norm_nonneg _), mul_rpow, ← rpow_mul, sub_one_mul, + have (x) : ‖∑ y, f y * g (x - y)‖₊ ^ (p : ℝ) ≤ + (∑ y, ‖f y‖₊ ^ (p : ℝ) * ‖g (x - y)‖₊) * (∑ y, ‖g (x - y)‖₊) ^ (p - 1 : ℝ) := by + rw [← NNReal.le_rpow_inv_iff_of_pos, NNReal.mul_rpow, ← NNReal.rpow_mul, sub_one_mul, mul_inv_cancel₀] any_goals positivity calc - _ ≤ ∑ y, ‖f y * g (x - y)‖ := norm_sum_le _ _ - _ = ∑ y, ‖f y‖ * ‖g (x - y)‖ ^ (p : ℝ)⁻¹ * ‖g (x - y)‖ ^ (1 - (p : ℝ)⁻¹) := ?_ - _ ≤ _ := inner_le_Lp_mul_Lq _ _ _ hpconj + _ ≤ ∑ y, ‖f y * g (x - y)‖₊ := nnnorm_sum_le _ _ + _ = ∑ y, ‖f y‖₊ * ‖g (x - y)‖₊ ^ (p : ℝ)⁻¹ * ‖g (x - y)‖₊ ^ (1 - (p : ℝ)⁻¹) := ?_ + _ ≤ _ := NNReal.inner_le_Lp_mul_Lq _ _ _ hpconj _ = _ := ?_ · congr with t - rw [norm_mul, mul_assoc, ← rpow_add' (norm_nonneg _), add_sub_cancel, rpow_one] + rw [nnnorm_mul, mul_assoc, ← NNReal.rpow_add', add_sub_cancel, NNReal.rpow_one] simp · have : 1 - (p : ℝ)⁻¹ ≠ 0 := sub_ne_zero.2 (inv_ne_one.2 $ NNReal.coe_ne_one.2 hp.ne').symm - simp only [abs_mul, abs_rpow_of_nonneg, mul_rpow, rpow_nonneg, hp₀.ne', this, - abs_norm, norm_nonneg, rpow_inv_rpow, Ne, NNReal.coe_eq_zero, not_false_iff, one_div, - rpow_rpow_inv, div_inv_eq_mul, one_mul] + simp only [abs_mul, abs_rpow_of_nonneg, NNReal.mul_rpow, rpow_nonneg, hp₀.ne', this, + abs_norm, norm_nonneg, NNReal.rpow_inv_rpow, Ne, NNReal.coe_eq_zero, not_false_iff, one_div, + NNReal.rpow_rpow_inv, div_inv_eq_mul, one_mul] + calc - ∑ x, ‖∑ y, f y * g (x - y)‖ ^ (p : ℝ) ≤ - ∑ x, (∑ y, ‖f y‖ ^ (p : ℝ) * ‖g (x - y)‖) * (∑ y, ‖g (x - y)‖) ^ (p - 1 : ℝ) := + ∑ x, ‖∑ y, f y * g (x - y)‖₊ ^ (p : ℝ) ≤ + ∑ x, (∑ y, ‖f y‖₊ ^ (p : ℝ) * ‖g (x - y)‖₊) * (∑ y, ‖g (x - y)‖₊) ^ (p - 1 : ℝ) := sum_le_sum fun i _ ↦ this _ _ = _ := ?_ - have hg : ∀ x, ∑ y, ‖g (x - y)‖ = ‖g‖_[1] := by - simp_rw [l1Norm_eq_sum] + have hg : ∀ x, ∑ y, ‖g (x - y)‖₊ = ‖g‖_[1] := by + simp_rw [dL1Norm_eq_sum_nnnorm] exact fun x ↦ Fintype.sum_equiv (Equiv.subLeft _) _ _ fun _ ↦ rfl - have hg' : ∀ y, ∑ x, ‖g (x - y)‖ = ‖g‖_[1] := by - simp_rw [l1Norm_eq_sum] + have hg' : ∀ y, ∑ x, ‖g (x - y)‖₊ = ‖g‖_[1] := by + simp_rw [dL1Norm_eq_sum_nnnorm] exact fun x ↦ Fintype.sum_equiv (Equiv.subRight _) _ _ fun _ ↦ rfl simp_rw [hg] rw [← sum_mul, sum_comm] simp_rw [← mul_sum, hg'] - rw [← sum_mul, mul_assoc, ← rpow_one_add' lpNorm_nonneg, add_sub_cancel] + rw [← sum_mul, mul_assoc, ← NNReal.rpow_one_add', add_sub_cancel] · rw [add_sub_cancel] positivity /-- A special case of **Young's convolution inequality**. -/ -lemma lpNorm_dconv_le {p : ℝ≥0} (hp : 1 ≤ p) (f g : α → β) : ‖f ○ g‖_[p] ≤ ‖f‖_[p] * ‖g‖_[1] := by - simpa only [conv_conjneg, lpNorm_conjneg] using lpNorm_conv_le hp f (conjneg g) +lemma dLpNorm_dconv_le {p : ℝ≥0} (hp : 1 ≤ p) (f g : α → β) : ‖f ○ g‖_[p] ≤ ‖f‖_[p] * ‖g‖_[1] := by + simpa only [conv_conjneg, dLpNorm_conjneg] using dLpNorm_conv_le hp f (conjneg g) end RCLike section Real -variable {f g : α → ℝ} {n : ℕ} +variable [MeasurableSpace α] [DiscreteMeasurableSpace α] {f g : α → ℝ} {n : ℕ} --TODO: Include `f : α → ℂ` -lemma l1Norm_conv (hf : 0 ≤ f) (hg : 0 ≤ g) : ‖f ∗ g‖_[1] = ‖f‖_[1] * ‖g‖_[1] := by +lemma dL1Norm_conv (hf : 0 ≤ f) (hg : 0 ≤ g) : ‖f ∗ g‖_[1] = ‖f‖_[1] * ‖g‖_[1] := by + ext have : ∀ x, 0 ≤ ∑ y, f y * g (x - y) := fun x ↦ sum_nonneg fun y _ ↦ mul_nonneg (hf _) (hg _) - simp only [l1Norm_eq_sum, ← sum_conv, conv_eq_sum_sub', norm_of_nonneg (this _), - norm_of_nonneg (hf _), norm_of_nonneg (hg _)] + simp [dL1Norm_eq_sum_nnnorm, ← sum_conv, conv_eq_sum_sub', nnnorm_of_nonneg (this _), + nnnorm_of_nonneg (hf _), nnnorm_of_nonneg (hg _)] -lemma l1Norm_dconv (hf : 0 ≤ f) (hg : 0 ≤ g) : ‖f ○ g‖_[1] = ‖f‖_[1] * ‖g‖_[1] := by - simpa using l1Norm_conv hf (conjneg_nonneg.2 hg) +lemma dL1Norm_dconv (hf : 0 ≤ f) (hg : 0 ≤ g) : ‖f ○ g‖_[1] = ‖f‖_[1] * ‖g‖_[1] := by + simpa using dL1Norm_conv hf (conjneg_nonneg.2 hg) end Real diff --git a/LeanAPAP/Prereqs/Convolution/ThreeAP.lean b/LeanAPAP/Prereqs/Convolution/ThreeAP.lean index a6badb143d..b0b69a2599 100644 --- a/LeanAPAP/Prereqs/Convolution/ThreeAP.lean +++ b/LeanAPAP/Prereqs/Convolution/ThreeAP.lean @@ -2,23 +2,23 @@ import Mathlib.Combinatorics.Additive.AP.Three.Defs import Mathlib.Data.Real.StarOrdered import LeanAPAP.Prereqs.Convolution.Discrete.Defs import LeanAPAP.Prereqs.Function.Indicator.Defs -import LeanAPAP.Prereqs.LpNorm.Discrete.Defs +import LeanAPAP.Prereqs.LpNorm.Discrete.Inner /-! # The convolution characterisation of 3AP-free sets -/ -open Finset Fintype Function +open Finset Fintype Function MeasureTheory open scoped Pointwise variable {G : Type*} [AddCommGroup G] [DecidableEq G] [Fintype G] {s : Finset G} -lemma ThreeAPFree.l2Inner_mu_conv_mu_mu_two_smul_mu (hG : Odd (card G)) +lemma ThreeAPFree.dL2Inner_mu_conv_mu_mu_two_smul_mu (hG : Odd (card G)) (hs : ThreeAPFree (s : Set G)) : ⟪μ s ∗ μ s, μ (s.image (2 • ·))⟫_[ℝ] = (s.card ^ 2 : ℝ)⁻¹ := by obtain rfl | hs' := s.eq_empty_or_nonempty · simp - simp only [l2Inner_eq_sum, sum_conv_mul, ← sum_product', RCLike.conj_to_real] + simp only [dL2Inner_eq_sum, sum_conv_mul, ← sum_product', RCLike.conj_to_real] rw [← diag_union_offDiag univ, sum_union (disjoint_diag_offDiag _), sum_diag, ← sum_add_sum_compl s, @sum_eq_card_nsmul _ _ _ _ _ (s.card ^ 3 : ℝ)⁻¹, nsmul_eq_mul, Finset.sum_eq_zero, Finset.sum_eq_zero, add_zero, add_zero, pow_succ', mul_inv, diff --git a/LeanAPAP/Prereqs/Energy.lean b/LeanAPAP/Prereqs/Energy.lean index eab58479f7..4b61ddc385 100644 --- a/LeanAPAP/Prereqs/Energy.lean +++ b/LeanAPAP/Prereqs/Energy.lean @@ -45,25 +45,25 @@ lemma boringEnergy_eq (n : ℕ) (s : Finset G) : boringEnergy n s = ∑ x, (𝟭 @[simp] lemma boringEnergy_one (s : Finset G) : boringEnergy 1 s = s.card := by simp [boringEnergy_eq, indicate_apply] -lemma nlpNorm_dft_indicate_pow (n : ℕ) (s : Finset G) : +lemma cLpNorm_dft_indicate_pow (n : ℕ) (s : Finset G) : ‖dft (𝟭 s)‖ₙ_[↑(2 * n)] ^ (2 * n) = boringEnergy n s := by obtain rfl | hn := n.eq_zero_or_pos · simp refine Complex.ofReal_injective ?_ calc _ = ⟪dft (𝟭 s ∗^ n), dft (𝟭 s ∗^ n)⟫ₙ_[ℂ] := ?_ - _ = ⟪𝟭 s ∗^ n, 𝟭 s ∗^ n⟫_[ℂ] := nl2Inner_dft _ _ + _ = ⟪𝟭 s ∗^ n, 𝟭 s ∗^ n⟫_[ℂ] := ndL2Inner_dft _ _ _ = _ := ?_ - · rw [nlpNorm_pow_eq_expect] + · rw [cLpNorm_pow_eq_expect_nnnorm] simp_rw [pow_mul', ← norm_pow _ n, Complex.ofReal_expect, Complex.ofReal_pow, - ← Complex.conj_mul', nl2Inner_eq_expect, dft_iterConv_apply] + ← Complex.conj_mul', ndL2Inner_eq_expect, dft_iterConv_apply] positivity - · simp only [l2Inner_eq_sum, boringEnergy_eq, Complex.ofReal_mul, Complex.ofReal_natCast, + · simp only [dL2Inner_eq_sum, boringEnergy_eq, Complex.ofReal_mul, Complex.ofReal_natCast, Complex.ofReal_sum, Complex.ofReal_pow, mul_eq_mul_left_iff, Nat.cast_eq_zero, Fintype.card_ne_zero, or_false, sq, (((indicate_isSelfAdjoint _).iterConv _).apply _).conj_eq, Complex.coe_iterConv, Complex.ofReal_comp_indicate] -lemma nl2Norm_dft_indicate (s : Finset G) : ‖dft (𝟭 s)‖ₙ_[2] = sqrt s.card := by - rw [eq_comm, sqrt_eq_iff_eq_sq] - simpa [eq_comm] using nlpNorm_dft_indicate_pow 1 s +lemma cL2Norm_dft_indicate (s : Finset G) : ‖dft (𝟭 s)‖ₙ_[2] = sqrt s.card := by + rw [eq_comm, sqrt_eq_iff_sq_eq] + simpa using cLpNorm_dft_indicate_pow 1 s all_goals positivity diff --git a/LeanAPAP/Prereqs/FourierTransform/Compact.lean b/LeanAPAP/Prereqs/FourierTransform/Compact.lean index 1ba57ac97f..126559f89e 100644 --- a/LeanAPAP/Prereqs/FourierTransform/Compact.lean +++ b/LeanAPAP/Prereqs/FourierTransform/Compact.lean @@ -26,50 +26,50 @@ lemma cft_apply (f : α → ℂ) (ψ : AddChar α ℂ) : cft f ψ = ⟪ψ, f⟫ @[simp] lemma cft_zero : cft (0 : α → ℂ) = 0 := by ext; simp [cft_apply] @[simp] lemma cft_add (f g : α → ℂ) : cft (f + g) = cft f + cft g := by - ext; simp [nl2Inner_add_right, cft_apply] + ext; simp [ndL2Inner_add_right, cft_apply] @[simp] lemma cft_neg (f : α → ℂ) : cft (-f) = - cft f := by ext; simp [cft_apply] @[simp] lemma cft_sub (f g : α → ℂ) : cft (f - g) = cft f - cft g := by - ext; simp [nl2Inner_sub_right, cft_apply] + ext; simp [ndL2Inner_sub_right, cft_apply] @[simp] lemma cft_const (a : ℂ) (hψ : ψ ≠ 0) : cft (const α a) ψ = 0 := by - simp only [cft_apply, nl2Inner_eq_expect, const_apply, ← expect_mul, ← map_expect, + simp only [cft_apply, ndL2Inner_eq_expect, const_apply, ← expect_mul, ← map_expect, expect_eq_zero_iff_ne_zero.2 hψ, map_zero, zero_mul] @[simp] lemma cft_smul [DistribSMul γ ℂ] [Star γ] [StarModule γ ℂ] [IsScalarTower γ ℂ ℂ] [SMulCommClass γ ℂ ℂ] (c : γ) (f : α → ℂ) : cft (c • f) = c • cft f := by have := SMulCommClass.symm γ ℂ ℂ ext - simp [nl2Inner_smul_right, cft_apply] + simp [ndL2Inner_smul_right, cft_apply] /-- **Parseval-Plancherel identity** for the discrete Fourier transform. -/ -@[simp] lemma l2Inner_cft (f g : α → ℂ) : ⟪cft f, cft g⟫_[ℂ] = ⟪f, g⟫ₙ_[ℂ] := by +@[simp] lemma dL2Inner_cft (f g : α → ℂ) : ⟪cft f, cft g⟫_[ℂ] = ⟪f, g⟫ₙ_[ℂ] := by classical unfold cft - simp_rw [l2Inner_eq_sum, nl2Inner_eq_expect, map_expect, map_mul, starRingEnd_self_apply, + simp_rw [dL2Inner_eq_sum, ndL2Inner_eq_expect, map_expect, map_mul, starRingEnd_self_apply, expect_mul, mul_expect, ← expect_sum_comm, mul_mul_mul_comm _ (conj $ f _), ← sum_mul, ← AddChar.inv_apply_eq_conj, ← map_neg_eq_inv, ← map_add_eq_mul, AddChar.sum_apply_eq_ite] simp [add_neg_eq_zero, card_univ, Fintype.card_ne_zero, NNRat.smul_def (α := ℂ)] /-- **Parseval-Plancherel identity** for the discrete Fourier transform. -/ -@[simp] lemma l2Norm_cft (f : α → ℂ) : ‖cft f‖_[2] = ‖f‖ₙ_[2] := - (sq_eq_sq lpNorm_nonneg nlpNorm_nonneg).1 $ Complex.ofReal_injective $ by - push_cast; simpa only [nl2Inner_self, l2Inner_self] using l2Inner_cft f f +@[simp] lemma dL2Norm_cft (f : α → ℂ) : ‖cft f‖_[2] = ‖f‖ₙ_[2] := + (sq_eq_sq nnLpNorm_nonneg cLpNorm_nonneg).1 $ Complex.ofReal_injective $ by + push_cast; simpa only [ndL2Inner_self, dL2Inner_self] using dL2Inner_cft f f /-- **Fourier inversion** for the discrete Fourier transform. -/ lemma cft_inversion (f : α → ℂ) (a : α) : ∑ ψ, cft f ψ * ψ a = f a := by - classical simp_rw [cft, nl2Inner_eq_expect, expect_mul, ← expect_sum_comm, mul_right_comm _ (f _), + classical simp_rw [cft, ndL2Inner_eq_expect, expect_mul, ← expect_sum_comm, mul_right_comm _ (f _), ← sum_mul, ← AddChar.inv_apply_eq_conj, inv_mul_eq_div, ← map_sub_eq_div, AddChar.sum_apply_eq_ite, sub_eq_zero, ite_mul, zero_mul, Fintype.expect_ite_eq] simp [add_neg_eq_zero, card_univ, NNRat.smul_def (α := ℂ), Fintype.card_ne_zero] lemma dft_cft_doubleDualEmb (f : α → ℂ) (a : α) : dft (cft f) (doubleDualEmb a) = f (-a) := by - simp only [← cft_inversion f (-a), mul_comm (conj _), dft_apply, l2Inner_eq_sum, map_neg_eq_inv, + simp only [← cft_inversion f (-a), mul_comm (conj _), dft_apply, dL2Inner_eq_sum, map_neg_eq_inv, AddChar.inv_apply_eq_conj, doubleDualEmb_apply] lemma cft_dft_doubleDualEmb (f : α → ℂ) (a : α) : cft (dft f) (doubleDualEmb a) = f (-a) := by - simp only [← dft_inversion f (-a), mul_comm (conj _), cft_apply, nl2Inner_eq_expect, + simp only [← dft_inversion f (-a), mul_comm (conj _), cft_apply, ndL2Inner_eq_expect, map_neg_eq_inv, AddChar.inv_apply_eq_conj, doubleDualEmb_apply] lemma dft_cft (f : α → ℂ) : dft (cft f) = f ∘ doubleDualEquiv.symm ∘ Neg.neg := @@ -84,16 +84,16 @@ lemma cft_injective : Injective (cft : (α → ℂ) → AddChar α ℂ → ℂ) funext fun a ↦ (cft_inversion _ _).symm.trans $ by rw [h, cft_inversion] lemma cft_inv (ψ : AddChar α ℂ) (hf : IsSelfAdjoint f) : cft f ψ⁻¹ = conj (cft f ψ) := by - simp_rw [cft_apply, nl2Inner_eq_expect, map_expect, AddChar.inv_apply', map_mul, + simp_rw [cft_apply, ndL2Inner_eq_expect, map_expect, AddChar.inv_apply', map_mul, AddChar.inv_apply_eq_conj, Complex.conj_conj, (hf.apply _).conj_eq] @[simp] lemma cft_conj (f : α → ℂ) (ψ : AddChar α ℂ) : cft (conj f) ψ = conj (cft f ψ⁻¹) := by - simp only [cft_apply, nl2Inner_eq_expect, map_expect, map_mul, ← inv_apply', ← inv_apply_eq_conj, + simp only [cft_apply, ndL2Inner_eq_expect, map_expect, map_mul, ← inv_apply', ← inv_apply_eq_conj, inv_inv, Pi.conj_apply] lemma cft_conjneg_apply (f : α → ℂ) (ψ : AddChar α ℂ) : cft (conjneg f) ψ = conj (cft f ψ) := by - simp only [cft_apply, nl2Inner_eq_expect, conjneg_apply, map_expect, map_mul, RCLike.conj_conj] + simp only [cft_apply, ndL2Inner_eq_expect, conjneg_apply, map_expect, map_mul, RCLike.conj_conj] refine Fintype.expect_equiv (Equiv.neg _) _ _ fun i ↦ ?_ simp only [Equiv.neg_apply, ← inv_apply_eq_conj, ← inv_apply', inv_apply] @@ -105,14 +105,14 @@ lemma cft_conjneg (f : α → ℂ) : cft (conjneg f) = conj (cft f) := funext $ lemma cft_dilate (f : α → ℂ) (ψ : AddChar α ℂ) (hn : (card α).Coprime n) : cft (dilate f n) ψ = cft f (ψ ^ n) := by - simp_rw [cft_apply, nl2Inner_eq_expect, dilate] + simp_rw [cft_apply, ndL2Inner_eq_expect, dilate] rw [← Nat.card_eq_fintype_card] at hn refine (Fintype.expect_bijective _ hn.nsmul_right_bijective _ _ ?_).symm simp only [pow_apply, ← map_nsmul_eq_pow, zmod_val_inv_nsmul_nsmul hn, forall_const] @[simp] lemma cft_trivNChar [DecidableEq α] : cft (trivNChar : α → ℂ) = 1 := by ext - simp [trivChar_apply, cft_apply, nl2Inner_eq_expect, ← map_expect, card_univ, + simp [trivChar_apply, cft_apply, ndL2Inner_eq_expect, ← map_expect, card_univ, NNRat.smul_def (α := ℂ)] @[simp] lemma cft_one : cft (1 : α → ℂ) = trivChar := @@ -121,12 +121,12 @@ lemma cft_dilate (f : α → ℂ) (ψ : AddChar α ℂ) (hn : (card α).Coprime variable [DecidableEq α] @[simp] lemma cft_indicate_zero (s : Finset α) : cft (𝟭 s) 0 = s.dens := by - simp only [cft_apply, nl2Inner_eq_expect, expect_indicate, AddChar.zero_apply, map_one, one_mul, + simp only [cft_apply, ndL2Inner_eq_expect, expect_indicate, AddChar.zero_apply, map_one, one_mul, dens, NNRat.smul_def (α := ℂ), div_eq_inv_mul] simp lemma cft_nconv_apply (f g : α → ℂ) (ψ : AddChar α ℂ) : cft (f ∗ₙ g) ψ = cft f ψ * cft g ψ := by - simp_rw [cft, nl2Inner_eq_expect, nconv_eq_expect_sub', mul_expect, expect_mul, ← expect_product', + simp_rw [cft, ndL2Inner_eq_expect, nconv_eq_expect_sub', mul_expect, expect_mul, ← expect_product', univ_product_univ] refine Fintype.expect_equiv ((Equiv.prodComm _ _).trans $ ((Equiv.refl _).prodShear Equiv.subRight).trans $ Equiv.prodComm _ _) _ _ fun (a, b) ↦ ?_ @@ -152,21 +152,21 @@ lemma cft_ndconv (f g : α → ℂ) : cft (f ○ₙ g) = cft f * conj (cft g) := @[simp] lemma cft_iterNConv_apply (f : α → ℂ) (n : ℕ) (ψ : AddChar α ℂ) : cft (f ∗^ₙ n) ψ = cft f ψ ^ n := congr_fun (cft_iterNConv _ _) _ --- lemma l2Norm_iterNConv (f : α → ℂ) (n : ℕ) : ‖f ∗^ₙ n‖ₙ_[2] = ‖f ^ n‖_[2] := by --- rw [← l2Norm_cft, cft_iterNConv, ← ENNReal.coe_two, lpNorm_pow] +-- lemma dL2Norm_iterNConv (f : α → ℂ) (n : ℕ) : ‖f ∗^ₙ n‖ₙ_[2] = ‖f ^ n‖_[2] := by +-- rw [← dL2Norm_cft, cft_iterNConv, ← ENNReal.coe_two, dLpNorm_pow] -- norm_cast -- refine (sq_eq_sq (by positivity) $ by positivity).1 ?_ --- rw [← ENNReal.coe_two, lpNorm_pow, ← pow_mul', ← Complex.ofReal_inj] +-- rw [← ENNReal.coe_two, dLpNorm_pow, ← pow_mul', ← Complex.ofReal_inj] -- push_cast -- simp_rw [pow_mul, ← Complex.mul_conj', conj_iterConv_apply, mul_pow] -lemma nlpNorm_nconv_le_nlpNorm_ndconv (hn₀ : n ≠ 0) (hn : Even n) (f : α → ℂ) : +lemma cLpNorm_nconv_le_cLpNorm_ndconv (hn₀ : n ≠ 0) (hn : Even n) (f : α → ℂ) : ‖f ∗ₙ f‖ₙ_[n] ≤ ‖f ○ₙ f‖ₙ_[n] := by cases isEmpty_or_nonempty α · rw [Subsingleton.elim (f ∗ₙ f) (f ○ₙ f)] refine le_of_pow_le_pow_left hn₀ (by positivity) ?_ obtain ⟨n, rfl⟩ := hn.two_dvd - simp_rw [nlpNorm_pow_eq_expect hn₀, ← cft_inversion (f ∗ₙ f), ← cft_inversion (f ○ₙ f), + simp_rw [cLpNorm_pow_eq_expect_nnnorm hn₀, ← cft_inversion (f ∗ₙ f), ← cft_inversion (f ○ₙ f), cft_nconv, cft_ndconv, Pi.mul_apply] rw [← Real.norm_of_nonneg (expect_nonneg fun i _ ↦ ?_), ← Complex.norm_real] rw [Complex.ofReal_expect (univ : Finset α)] @@ -195,8 +195,8 @@ lemma nlpNorm_nconv_le_nlpNorm_ndconv (hn₀ : n ≠ 0) (hn : Even n) (f : α -- refine expect_congr rfl fun x _ ↦ expect_congr rfl fun a _ ↦ prod_congr rfl fun i _ ↦ _ -- ring ---TODO: Can we unify with `nlpNorm_nconv_le_nlpNorm_ndconv`? -lemma nlpNorm_nconv_le_nlpNorm_ndconv' (hn₀ : n ≠ 0) (hn : Even n) (f : α → ℝ) : +--TODO: Can we unify with `cLpNorm_nconv_le_cLpNorm_ndconv`? +lemma cLpNorm_nconv_le_cLpNorm_ndconv' (hn₀ : n ≠ 0) (hn : Even n) (f : α → ℝ) : ‖f ∗ₙ f‖ₙ_[n] ≤ ‖f ○ₙ f‖ₙ_[n] := by - simpa only [← Complex.coe_comp_nconv, ← Complex.coe_comp_ndconv, Complex.nlpNorm_coe_comp] using - nlpNorm_nconv_le_nlpNorm_ndconv hn₀ hn ((↑) ∘ f) + simpa only [← Complex.coe_comp_nconv, ← Complex.coe_comp_ndconv, Complex.cLpNorm_coe_comp] using + cLpNorm_nconv_le_cLpNorm_ndconv hn₀ hn ((↑) ∘ f) diff --git a/LeanAPAP/Prereqs/FourierTransform/Discrete.lean b/LeanAPAP/Prereqs/FourierTransform/Discrete.lean index 8cb380ab89..182b05d6c5 100644 --- a/LeanAPAP/Prereqs/FourierTransform/Discrete.lean +++ b/LeanAPAP/Prereqs/FourierTransform/Discrete.lean @@ -1,5 +1,6 @@ import LeanAPAP.Prereqs.AddChar.PontryaginDuality -import LeanAPAP.Prereqs.LpNorm.Compact +import LeanAPAP.Prereqs.LpNorm.Compact.Inner +import LeanAPAP.Prereqs.LpNorm.Discrete.Inner import LeanAPAP.Prereqs.Convolution.Discrete.Defs /-! @@ -9,7 +10,7 @@ This file defines the discrete Fourier transform and shows the Parseval-Plancher Fourier inversion formula for it. -/ -open AddChar Finset Function +open AddChar Finset Function MeasureTheory open Fintype (card) open scoped BigOperators ComplexConjugate ComplexOrder @@ -23,37 +24,37 @@ lemma dft_apply (f : α → ℂ) (ψ : AddChar α ℂ) : dft f ψ = ⟪ψ, f⟫_ @[simp] lemma dft_zero : dft (0 : α → ℂ) = 0 := by ext; simp [dft_apply] @[simp] lemma dft_add (f g : α → ℂ) : dft (f + g) = dft f + dft g := by - ext; simp [l2Inner_add_right, dft_apply] + ext; simp [dL2Inner_add_right, dft_apply] @[simp] lemma dft_neg (f : α → ℂ) : dft (-f) = - dft f := by ext; simp [dft_apply] @[simp] lemma dft_sub (f g : α → ℂ) : dft (f - g) = dft f - dft g := by - ext; simp [l2Inner_sub_right, dft_apply] + ext; simp [dL2Inner_sub_right, dft_apply] @[simp] lemma dft_const (a : ℂ) (hψ : ψ ≠ 0) : dft (const α a) ψ = 0 := by - simp only [dft_apply, l2Inner_eq_sum, const_apply, ← sum_mul, ← map_sum, + simp only [dft_apply, dL2Inner_eq_sum, const_apply, ← sum_mul, ← map_sum, sum_eq_zero_iff_ne_zero.2 hψ, map_zero, zero_mul] @[simp] lemma dft_smul [DistribSMul γ ℂ] [Star γ] [StarModule γ ℂ] [SMulCommClass γ ℂ ℂ] (c : γ) - (f : α → ℂ) : dft (c • f) = c • dft f := by ext; simp [l2Inner_smul_right, dft_apply] + (f : α → ℂ) : dft (c • f) = c • dft f := by ext; simp [dL2Inner_smul_right, dft_apply] /-- **Parseval-Plancherel identity** for the discrete Fourier transform. -/ -@[simp] lemma nl2Inner_dft (f g : α → ℂ) : ⟪dft f, dft g⟫ₙ_[ℂ] = ⟪f, g⟫_[ℂ] := by +@[simp] lemma ndL2Inner_dft (f g : α → ℂ) : ⟪dft f, dft g⟫ₙ_[ℂ] = ⟪f, g⟫_[ℂ] := by classical unfold dft - simp_rw [l2Inner_eq_sum, nl2Inner_eq_expect, map_sum, map_mul, starRingEnd_self_apply, sum_mul, + simp_rw [dL2Inner_eq_sum, ndL2Inner_eq_expect, map_sum, map_mul, starRingEnd_self_apply, sum_mul, mul_sum, expect_sum_comm, mul_mul_mul_comm _ (conj $ f _), ← expect_mul, ← AddChar.inv_apply_eq_conj, ← map_neg_eq_inv, ← map_add_eq_mul, AddChar.expect_apply_eq_ite, add_neg_eq_zero, boole_mul, Fintype.sum_ite_eq] /-- **Parseval-Plancherel identity** for the discrete Fourier transform. -/ -@[simp] lemma nl2Norm_dft (f : α → ℂ) : ‖dft f‖ₙ_[2] = ‖f‖_[2] := - (sq_eq_sq nlpNorm_nonneg lpNorm_nonneg).1 $ Complex.ofReal_injective $ by - push_cast; simpa only [nl2Inner_self, l2Inner_self] using nl2Inner_dft f f +@[simp] lemma cL2Norm_dft (f : α → ℂ) : ‖dft f‖ₙ_[2] = ‖f‖_[2] := + (sq_eq_sq cLpNorm_nonneg nnLpNorm_nonneg).1 $ Complex.ofReal_injective $ by + push_cast; simpa only [ndL2Inner_self, dL2Inner_self] using ndL2Inner_dft f f /-- **Fourier inversion** for the discrete Fourier transform. -/ lemma dft_inversion (f : α → ℂ) (a : α) : 𝔼 ψ, dft f ψ * ψ a = f a := by - classical simp_rw [dft, l2Inner_eq_sum, sum_mul, expect_sum_comm, mul_right_comm _ (f _), + classical simp_rw [dft, dL2Inner_eq_sum, sum_mul, expect_sum_comm, mul_right_comm _ (f _), ← expect_mul, ← AddChar.inv_apply_eq_conj, inv_mul_eq_div, ← map_sub_eq_div, AddChar.expect_apply_eq_ite, sub_eq_zero, boole_mul, Fintype.sum_ite_eq] @@ -65,7 +66,7 @@ lemma dft_inversion' (f : α → ℂ) (a : α) : ∑ ψ : AddChar α ℂ, dft f lemma dft_dft_doubleDualEmb (f : α → ℂ) (a : α) : dft (dft f) (doubleDualEmb a) = card α * f (-a) := by - simp only [← dft_inversion f (-a), mul_comm (conj _), dft_apply, l2Inner_eq_sum, map_neg_eq_inv, + simp only [← dft_inversion f (-a), mul_comm (conj _), dft_apply, dL2Inner_eq_sum, map_neg_eq_inv, AddChar.inv_apply_eq_conj, doubleDualEmb_apply, ← Fintype.card_mul_expect, AddChar.card_eq] lemma dft_dft (f : α → ℂ) : dft (dft f) = card α * f ∘ doubleDualEquiv.symm ∘ Neg.neg := @@ -77,16 +78,16 @@ lemma dft_injective : Injective (dft : (α → ℂ) → AddChar α ℂ → ℂ) funext fun a ↦ (dft_inversion _ _).symm.trans $ by rw [h, dft_inversion] lemma dft_inv (ψ : AddChar α ℂ) (hf : IsSelfAdjoint f) : dft f ψ⁻¹ = conj (dft f ψ) := by - simp_rw [dft_apply, l2Inner_eq_sum, map_sum, AddChar.inv_apply', map_mul, + simp_rw [dft_apply, dL2Inner_eq_sum, map_sum, AddChar.inv_apply', map_mul, AddChar.inv_apply_eq_conj, Complex.conj_conj, (hf.apply _).conj_eq] @[simp] lemma dft_conj (f : α → ℂ) (ψ : AddChar α ℂ) : dft (conj f) ψ = conj (dft f ψ⁻¹) := by - simp only [dft_apply, l2Inner_eq_sum, map_sum, map_mul, ← inv_apply', ← inv_apply_eq_conj, + simp only [dft_apply, dL2Inner_eq_sum, map_sum, map_mul, ← inv_apply', ← inv_apply_eq_conj, inv_inv, Pi.conj_apply] lemma dft_conjneg_apply (f : α → ℂ) (ψ : AddChar α ℂ) : dft (conjneg f) ψ = conj (dft f ψ) := by - simp only [dft_apply, l2Inner_eq_sum, conjneg_apply, map_sum, map_mul, RCLike.conj_conj] + simp only [dft_apply, dL2Inner_eq_sum, conjneg_apply, map_sum, map_mul, RCLike.conj_conj] refine Fintype.sum_equiv (Equiv.neg α) _ _ fun i ↦ ?_ simp only [Equiv.neg_apply, ← inv_apply_eq_conj, ← inv_apply', inv_apply] @@ -101,13 +102,13 @@ lemma dft_comp_neg_apply (f : α → ℂ) (ψ : AddChar α ℂ) : lemma dft_dilate (f : α → ℂ) (ψ : AddChar α ℂ) (hn : (card α).Coprime n) : dft (dilate f n) ψ = dft f (ψ ^ n) := by - simp_rw [dft_apply, l2Inner_eq_sum, dilate] + simp_rw [dft_apply, dL2Inner_eq_sum, dilate] rw [← Nat.card_eq_fintype_card] at hn refine (Fintype.sum_bijective _ hn.nsmul_right_bijective _ _ ?_).symm simp only [pow_apply, ← map_nsmul_eq_pow, zmod_val_inv_nsmul_nsmul hn, forall_const] @[simp] lemma dft_trivChar [DecidableEq α] : dft (trivChar : α → ℂ) = 1 := by - ext; simp [trivChar_apply, dft_apply, l2Inner_eq_sum, ← map_sum] + ext; simp [trivChar_apply, dft_apply, dL2Inner_eq_sum, ← map_sum] @[simp] lemma dft_one : dft (1 : α → ℂ) = card α • trivChar := dft_injective $ by classical rw [dft_smul, dft_trivChar, dft_dft, Pi.one_comp, nsmul_eq_mul] @@ -115,10 +116,10 @@ lemma dft_dilate (f : α → ℂ) (ψ : AddChar α ℂ) (hn : (card α).Coprime variable [DecidableEq α] @[simp] lemma dft_indicate_zero (A : Finset α) : dft (𝟭 A) 0 = A.card := by - simp only [dft_apply, l2Inner_eq_sum, sum_indicate, AddChar.zero_apply, map_one, one_mul] + simp only [dft_apply, dL2Inner_eq_sum, sum_indicate, AddChar.zero_apply, map_one, one_mul] lemma dft_conv_apply (f g : α → ℂ) (ψ : AddChar α ℂ) : dft (f ∗ g) ψ = dft f ψ * dft g ψ := by - simp_rw [dft, l2Inner_eq_sum, conv_eq_sum_sub', mul_sum, sum_mul, ← sum_product', + simp_rw [dft, dL2Inner_eq_sum, conv_eq_sum_sub', mul_sum, sum_mul, ← sum_product', univ_product_univ] refine Fintype.sum_equiv ((Equiv.prodComm _ _).trans $ ((Equiv.refl _).prodShear Equiv.subRight).trans $ Equiv.prodComm _ _) _ _ fun (a, b) ↦ ?_ @@ -142,12 +143,12 @@ lemma dft_dconv (f g : α → ℂ) : dft (f ○ g) = dft f * conj (dft g) := fun @[simp] lemma dft_iterConv_apply (f : α → ℂ) (n : ℕ) (ψ : AddChar α ℂ) : dft (f ∗^ n) ψ = dft f ψ ^ n := congr_fun (dft_iterConv _ _) _ -lemma lpNorm_conv_le_lpNorm_dconv (hn₀ : n ≠ 0) (hn : Even n) (f : α → ℂ) : +lemma dLpNorm_conv_le_dLpNorm_dconv (hn₀ : n ≠ 0) (hn : Even n) (f : α → ℂ) : ‖f ∗ f‖_[n] ≤ ‖f ○ f‖_[n] := by refine le_of_pow_le_pow_left hn₀ (by positivity) ?_ obtain ⟨k, hnk⟩ := hn.two_dvd calc ‖f ∗ f‖_[n] ^ n - = ∑ x, ‖(f ∗ f) x‖ ^ n := lpNorm_pow_eq_sum hn₀ _ + = ∑ x, ‖(f ∗ f) x‖ ^ n := dLpNorm_pow_eq_sum_nnnorm hn₀ _ _ = ∑ x, ‖(𝔼 ψ, dft f ψ ^ 2 * ψ x)‖ ^ n := by simp_rw [← norm_pow, ← dft_inversion (f ∗ f), dft_conv_apply, sq] _ ≤ ∑ x, ‖𝔼 ψ, ‖dft f ψ‖ ^ 2 * ψ x‖ ^ n := Complex.le_of_eq_sum_of_eq_sum_norm @@ -155,12 +156,12 @@ lemma lpNorm_conv_le_lpNorm_dconv (hn₀ : n ≠ 0) (hn : Even n) (f : α → (∏ i, dft f (ψ.2 i) ^ 2) * ∑ x, (∑ i, ψ.1 i - ∑ i, ψ.2 i) x) univ (by positivity) ?_ ?_ _ = ∑ x, ‖(f ○ f) x‖ ^ n := by simp_rw [← norm_pow, ← dft_inversion (f ○ f), dft_dconv_apply, Complex.mul_conj'] - _ = ‖f ○ f‖_[n] ^ n := (lpNorm_pow_eq_sum hn₀ _).symm + _ = ‖f ○ f‖_[n] ^ n := (dLpNorm_pow_eq_sum_nnnorm hn₀ _).symm · push_cast simp_rw [hnk, pow_mul, ← Complex.conj_mul', map_expect, mul_pow, expect_pow, expect_mul_expect] sorry sorry --- simp_rw [lpNorm_pow_eq_sum hn₀, mul_sum, ← mul_pow, ← nsmul_eq_mul, ← norm_nsmul, nsmul_eq_mul, +-- simp_rw [dLpNorm_pow_eq_sum_nnnorm hn₀, mul_sum, ← mul_pow, ← nsmul_eq_mul, ← norm_nsmul, nsmul_eq_mul, -- ← dft_inversion', dft_conv, dft_dconv, Pi.mul_apply] -- rw [← Real.norm_of_nonneg (sum_nonneg fun i _ ↦ ?_), ← Complex.norm_real] -- rw [Complex.ofReal_sum (univ : Finset α)] @@ -188,8 +189,8 @@ lemma lpNorm_conv_le_lpNorm_dconv (hn₀ : n ≠ 0) (hn : Even n) (f : α → -- refine sum_congr rfl fun x _ ↦ sum_congr rfl fun a _ ↦ prod_congr rfl fun i _ ↦ _ -- ring ---TODO: Can we unify with `lpNorm_conv_le_lpNorm_dconv`? -lemma lpNorm_conv_le_lpNorm_dconv' (hn₀ : n ≠ 0) (hn : Even n) (f : α → ℝ) : +--TODO: Can we unify with `dLpNorm_conv_le_dLpNorm_dconv`? +lemma dLpNorm_conv_le_dLpNorm_dconv' (hn₀ : n ≠ 0) (hn : Even n) (f : α → ℝ) : ‖f ∗ f‖_[n] ≤ ‖f ○ f‖_[n] := by - simpa only [← Complex.coe_comp_conv, ← Complex.coe_comp_dconv, Complex.lpNorm_coe_comp] using - lpNorm_conv_le_lpNorm_dconv hn₀ hn ((↑) ∘ f) + simpa only [← Complex.coe_comp_conv, ← Complex.coe_comp_dconv, Complex.dLpNorm_coe_comp] using + dLpNorm_conv_le_dLpNorm_dconv hn₀ hn ((↑) ∘ f) diff --git a/LeanAPAP/Prereqs/LargeSpec.lean b/LeanAPAP/Prereqs/LargeSpec.lean index 1247a09fd5..120e6d04db 100644 --- a/LeanAPAP/Prereqs/LargeSpec.lean +++ b/LeanAPAP/Prereqs/LargeSpec.lean @@ -17,7 +17,7 @@ noncomputable def largeSpec (f : G → ℂ) (η : ℝ) : Finset (AddChar G ℂ) @[simp] lemma mem_largeSpec : ψ ∈ largeSpec f η ↔ η * ‖f‖_[1] ≤ ‖dft f ψ‖ := by simp [largeSpec] lemma largeSpec_anti (f : G → ℂ) : Antitone (largeSpec f) := fun η ν h ψ ↦ by - simp_rw [mem_largeSpec]; exact (mul_le_mul_of_nonneg_right h lpNorm_nonneg).trans + simp_rw [mem_largeSpec]; exact (mul_le_mul_of_nonneg_right h nnLpNorm_nonneg).trans @[simp] lemma largeSpec_zero_left (η : ℝ) : largeSpec (0 : G → ℂ) η = univ := by simp [largeSpec] @[simp] lemma largeSpec_zero_right (f : G → ℂ) : largeSpec f 0 = univ := by simp [largeSpec] diff --git a/LeanAPAP/Prereqs/LpNorm/Compact.lean b/LeanAPAP/Prereqs/LpNorm/Compact.lean deleted file mode 100644 index 408274ca37..0000000000 --- a/LeanAPAP/Prereqs/LpNorm/Compact.lean +++ /dev/null @@ -1,541 +0,0 @@ -import Mathlib.Data.Finset.Density -import LeanAPAP.Prereqs.LpNorm.Discrete.Basic - -/-! -# Normalised Lp norms --/ - -open Finset hiding card -open Function Real -open Fintype (card) -open scoped BigOperators ComplexConjugate ENNReal NNReal - -variable {ι 𝕜 : Type*} [Fintype ι] - -/-! ### Lp norm -/ - -section NormedAddCommGroup -variable {α : ι → Type*} [∀ i, NormedAddCommGroup (α i)] {p q : ℝ≥0∞} {f g h : ∀ i, α i} - -/-- The Lp norm of a function with the compact normalisation. -/ -noncomputable def nlpNorm (p : ℝ≥0∞) (f : ∀ i, α i) : ℝ := ‖f‖_[p] / card ι ^ p.toReal⁻¹ - -notation "‖" f "‖ₙ_[" p "]" => nlpNorm p f - -lemma nlpNorm_eq_expect' (hp : p.toReal ≠ 0) (f : ∀ i, α i) : - ‖f‖ₙ_[p] = (𝔼 i, ‖f i‖ ^ p.toReal) ^ p.toReal⁻¹ := by - rw [nlpNorm, lpNorm_eq_sum', ← div_rpow, Fintype.expect_eq_sum_div_card (α := ℝ)] <;> positivity - -lemma nlpNorm_eq_expect'' {p : ℝ} (hp : 0 < p) (f : ∀ i, α i) : - ‖f‖ₙ_[p.toNNReal] = (𝔼 i, ‖f i‖ ^ p) ^ p⁻¹ := by - rw [nlpNorm_eq_expect'] <;> simp [hp.ne', hp.le] - -lemma nlpNorm_eq_expect {p : ℝ≥0} (hp : p ≠ 0) (f : ∀ i, α i) : - ‖f‖ₙ_[p] = (𝔼 i, ‖f i‖ ^ (p : ℝ)) ^ (p⁻¹ : ℝ) := nlpNorm_eq_expect' (by simpa using hp) _ - -lemma nlpNorm_rpow_eq_expect {p : ℝ≥0} (hp : p ≠ 0) (f : ∀ i, α i) : - ‖f‖ₙ_[p] ^ (p : ℝ) = 𝔼 i, ‖f i‖ ^ (p : ℝ) := by - rw [nlpNorm_eq_expect hp, rpow_inv_rpow] <;> positivity - -lemma nlpNorm_pow_eq_expect {p : ℕ} (hp : p ≠ 0) (f : ∀ i, α i) : - ‖f‖ₙ_[p] ^ p = 𝔼 i, ‖f i‖ ^ p := by - simpa using nlpNorm_rpow_eq_expect (Nat.cast_ne_zero.2 hp) f - -lemma nl2Norm_sq_eq_expect (f : ∀ i, α i) : ‖f‖ₙ_[2] ^ 2 = 𝔼 i, ‖f i‖ ^ 2 := by - simpa using nlpNorm_pow_eq_expect two_ne_zero _ - -lemma nl2Norm_eq_expect (f : ∀ i, α i) : ‖f‖ₙ_[2] = sqrt (𝔼 i, ‖f i‖ ^ 2) := by - simpa [sqrt_eq_rpow] using nlpNorm_eq_expect two_ne_zero _ - -lemma nl1Norm_eq_expect (f : ∀ i, α i) : ‖f‖ₙ_[1] = 𝔼 i, ‖f i‖ := by simp [nlpNorm_eq_expect'] - -lemma nl0Norm_eq_card (f : ∀ i, α i) : ‖f‖ₙ_[0] = {i | f i ≠ 0}.toFinite.toFinset.card := by - simp [l0Norm_eq_card, nlpNorm] - -lemma nlinftyNorm_eq_ciSup (f : ∀ i, α i) : ‖f‖ₙ_[∞] = ⨆ i, ‖f i‖ := by - simp [nlpNorm, linftyNorm_eq_ciSup] - -@[simp] lemma nlpNorm_zero : ‖(0 : ∀ i, α i)‖ₙ_[p] = 0 := by simp [nlpNorm] - -@[simp] lemma nlpNorm_of_isEmpty [IsEmpty ι] (p : ℝ≥0∞) (f : ∀ i, α i) : ‖f‖ₙ_[p] = 0 := by - simp [nlpNorm] - -@[simp] lemma nlpNorm_norm (p : ℝ≥0∞) (f : ∀ i, α i) : ‖fun i ↦ ‖f i‖‖ₙ_[p] = ‖f‖ₙ_[p] := by - simp [nlpNorm] - -@[simp] lemma nlpNorm_neg (f : ∀ i, α i) : ‖-f‖ₙ_[p] = ‖f‖ₙ_[p] := by simp [← nlpNorm_norm _ (-f)] - -lemma nlpNorm_sub_comm (f g : ∀ i, α i) : ‖f - g‖ₙ_[p] = ‖g - f‖ₙ_[p] := by - simp [← nlpNorm_neg (f - g)] - -@[simp] lemma nlpNorm_nonneg : 0 ≤ ‖f‖ₙ_[p] := by unfold nlpNorm; positivity - -@[simp] lemma nlpNorm_eq_zero [Nonempty ι] : ‖f‖ₙ_[p] = 0 ↔ f = 0 := by - obtain p | p := p - · simp [nlinftyNorm_eq_ciSup, ENNReal.none_eq_top, ← sup'_univ_eq_ciSup, le_antisymm_iff, - Function.funext_iff] - obtain rfl | hp := eq_or_ne p 0 - · simp [nl0Norm_eq_card, eq_empty_iff_forall_not_mem, Function.funext_iff] - · rw [← rpow_eq_zero nlpNorm_nonneg (NNReal.coe_ne_zero.2 hp)] - simp [nlpNorm_rpow_eq_expect hp, Fintype.expect_eq_zero_iff_of_nonneg, rpow_nonneg, - Function.funext_iff, rpow_eq_zero _ (NNReal.coe_ne_zero.2 hp), Pi.le_def] - -@[simp] lemma nlpNorm_pos [Nonempty ι] : 0 < ‖f‖ₙ_[p] ↔ f ≠ 0 := - nlpNorm_nonneg.gt_iff_ne.trans nlpNorm_eq_zero.not - -lemma nlpNorm_mono_right (hpq : p ≤ q) (f : ∀ i, α i) : ‖f‖ₙ_[p] ≤ ‖f‖ₙ_[q] := sorry - -section one_le - -lemma nlpNorm_add_le (hp : 1 ≤ p) (f g : ∀ i, α i) : ‖f + g‖ₙ_[p] ≤ ‖f‖ₙ_[p] + ‖g‖ₙ_[p] := by - simp only [nlpNorm, ← add_div] - exact div_le_div_of_nonneg_right (lpNorm_add_le hp _ _) (by positivity) - -lemma nlpNorm_sub_le (hp : 1 ≤ p) (f g : ∀ i, α i) : ‖f - g‖ₙ_[p] ≤ ‖f‖ₙ_[p] + ‖g‖ₙ_[p] := by - simp only [nlpNorm, ← add_div] - exact div_le_div_of_nonneg_right (lpNorm_sub_le hp _ _) (by positivity) - -lemma nlpNorm_le_nlpNorm_add_nlpNorm_sub' (hp : 1 ≤ p) (f g : ∀ i, α i) : - ‖f‖ₙ_[p] ≤ ‖g‖ₙ_[p] + ‖f - g‖ₙ_[p] := by - simp only [nlpNorm, ← add_div] - exact div_le_div_of_nonneg_right (lpNorm_le_lpNorm_add_lpNorm_sub' hp _ _) (by positivity) - -lemma nlpNorm_le_nlpNorm_add_nlpNorm_sub (hp : 1 ≤ p) (f g : ∀ i, α i) : - ‖f‖ₙ_[p] ≤ ‖g‖ₙ_[p] + ‖g - f‖ₙ_[p] := by - simp only [nlpNorm, ← add_div] - exact div_le_div_of_nonneg_right (lpNorm_le_lpNorm_add_lpNorm_sub hp _ _) (by positivity) - -lemma nlpNorm_le_add_nlpNorm_add (hp : 1 ≤ p) (f g : ∀ i, α i) : - ‖f‖ₙ_[p] ≤ ‖f + g‖ₙ_[p] + ‖g‖ₙ_[p] := by - simp only [nlpNorm, ← add_div] - exact div_le_div_of_nonneg_right (lpNorm_le_add_lpNorm_add hp _ _) (by positivity) - -lemma nlpNorm_sub_le_nlpNorm_sub_add_nlpNorm_sub (hp : 1 ≤ p) (f g : ∀ i, α i) : - ‖f - h‖ₙ_[p] ≤ ‖f - g‖ₙ_[p] + ‖g - h‖ₙ_[p] := by - simp only [nlpNorm, ← add_div] - exact div_le_div_of_nonneg_right (lpNorm_sub_le_lpNorm_sub_add_lpNorm_sub hp _ _) (by positivity) - -variable [NormedField 𝕜] [∀ i, NormedSpace 𝕜 (α i)] - --- TODO: `p ≠ 0` is enough -lemma nlpNorm_smul (hp : 1 ≤ p) (c : 𝕜) (f : ∀ i, α i) : ‖c • f‖ₙ_[p] = ‖c‖ * ‖f‖ₙ_[p] := by - simp only [nlpNorm, mul_div_assoc, lpNorm_smul hp] - -variable [∀ i, NormedSpace ℝ (α i)] - -lemma nlpNorm_nsmul (hp : 1 ≤ p) (n : ℕ) (f : ∀ i, α i) : ‖n • f‖ₙ_[p] = n • ‖f‖ₙ_[p] := by - simp only [nlpNorm, nsmul_eq_mul, mul_div_assoc, lpNorm_nsmul hp] - -end one_le -end NormedAddCommGroup - -section NormedAddCommGroup -variable {α : Type*} [NormedAddCommGroup α] {p : ℝ≥0∞} - -@[simp] -lemma nlpNorm_const [Nonempty ι] (hp : p ≠ 0) (a : α) : ‖const ι a‖ₙ_[p] = ‖a‖ := by - obtain _ | p := p - · simp [nlinftyNorm_eq_ciSup] - have : (card ι : ℝ) ^ (p : ℝ)⁻¹ ≠ 0 := by positivity - simp [nlpNorm, ENNReal.coe_ne_coe.1 hp, mul_div_cancel_left₀ _ this] - -end NormedAddCommGroup - -section RCLike -variable [RCLike 𝕜] {p : ℝ≥0∞} {f g : ι → 𝕜} - -@[simp] lemma nlpNorm_one [Nonempty ι] (hp : p ≠ 0) : ‖(1 : ι → 𝕜)‖ₙ_[p] = 1 := - (nlpNorm_const hp 1).trans $ by simp - -lemma nlpNorm_natCast_mul (hp : 1 ≤ p) (n : ℕ) (f : ι → 𝕜) : - ‖(n : ι → 𝕜) * f‖ₙ_[p] = n * ‖f‖ₙ_[p] := by simpa only [nsmul_eq_mul] using nlpNorm_nsmul hp n f - -lemma nlpNorm_natCast_mul' (hp : 1 ≤ p) (n : ℕ) (f : ι → 𝕜) : - ‖(n * f ·)‖ₙ_[p] = n * ‖f‖ₙ_[p] := nlpNorm_natCast_mul hp _ _ - -lemma nlpNorm_mul_natCast (hp : 1 ≤ p) (f : ι → 𝕜) (n : ℕ) : - ‖f * (n : ι → 𝕜)‖ₙ_[p] = ‖f‖ₙ_[p] * n := by - simpa only [mul_comm] using nlpNorm_natCast_mul hp n f - -lemma nlpNorm_mul_natCast' (hp : 1 ≤ p) (f : ι → 𝕜) (n : ℕ) : - ‖(f · * n)‖ₙ_[p] = ‖f‖ₙ_[p] * n := nlpNorm_mul_natCast hp _ _ - -lemma nlpNorm_div_natCast' (hp : 1 ≤ p) (f : ι → 𝕜) (n : ℕ) : - ‖(f · / n)‖ₙ_[p] = ‖f‖ₙ_[p] / n := by simp [nlpNorm, lpNorm_div_natCast' hp, div_right_comm] - -lemma nlpNorm_div_natCast (hp : 1 ≤ p) (f : ι → 𝕜) (n : ℕ) : - ‖f / (n : ι → 𝕜)‖ₙ_[p] = ‖f‖ₙ_[p] / n := nlpNorm_div_natCast' hp _ _ - -end RCLike - -section Real -variable {p : ℝ≥0} {f g : ι → ℝ} - -lemma nlpNorm_mono (hf : 0 ≤ f) (hfg : f ≤ g) : ‖f‖ₙ_[p] ≤ ‖g‖ₙ_[p] := - div_le_div_of_nonneg_right (lpNorm_mono hf hfg) $ by positivity - -end Real - -/-! #### Inner product -/ - -section Semifield -variable [Semifield 𝕜] [CharZero 𝕜] [StarRing 𝕜] {γ : Type*} [DistribSMul γ 𝕜] - -/-- Inner product giving rise to the L2 norm with the compact normalisation. -/ -def nl2Inner (f g : ι → 𝕜) : 𝕜 := 𝔼 i, conj (f i) * g i - -notation "⟪" f ", " g "⟫ₙ_[" 𝕜 "]" => @nl2Inner _ 𝕜 _ _ _ _ f g - -lemma nl2Inner_eq_expect (f g : ι → 𝕜) : ⟪f, g⟫ₙ_[𝕜] = 𝔼 i, conj (f i) * g i := rfl -lemma nl2Inner_eq_l2Inner_div_card (f g : ι → 𝕜) : ⟪f, g⟫ₙ_[𝕜] = ⟪f, g⟫_[𝕜] / card ι := - Fintype.expect_eq_sum_div_card _ - -@[simp] lemma conj_nl2Inner (f g : ι → 𝕜) : conj ⟪f, g⟫ₙ_[𝕜] = ⟪g, f⟫ₙ_[𝕜] := by - simp [nl2Inner_eq_expect, map_expect, mul_comm] - -@[simp] lemma nl2Inner_zero_left (g : ι → 𝕜) : ⟪0, g⟫ₙ_[𝕜] = 0 := by simp [nl2Inner] -@[simp] lemma nl2Inner_zero_right (f : ι → 𝕜) : ⟪f, 0⟫ₙ_[𝕜] = 0 := by simp [nl2Inner] - -@[simp] lemma nl2Inner_of_isEmpty [IsEmpty ι] (f g : ι → 𝕜) : ⟪f, g⟫ₙ_[𝕜] = 0 := by - simp [nl2Inner_eq_l2Inner_div_card] - -@[simp] lemma nl2Inner_const_left (a : 𝕜) (f : ι → 𝕜) : - ⟪const _ a, f⟫ₙ_[𝕜] = conj a * 𝔼 x, f x := by - simp only [nl2Inner, const_apply, mul_expect] - -@[simp] -lemma nl2Inner_const_right (f : ι → 𝕜) (a : 𝕜) : ⟪f, const _ a⟫ₙ_[𝕜] = (𝔼 x, conj (f x)) * a := by - simp only [nl2Inner, const_apply, expect_mul] - -lemma nl2Inner_add_left (f₁ f₂ g : ι → 𝕜) : ⟪f₁ + f₂, g⟫ₙ_[𝕜] = ⟪f₁, g⟫ₙ_[𝕜] + ⟪f₂, g⟫ₙ_[𝕜] := by - simp_rw [nl2Inner, Pi.add_apply, map_add, add_mul, expect_add_distrib] - -lemma nl2Inner_add_right (f g₁ g₂ : ι → 𝕜) : ⟪f, g₁ + g₂⟫ₙ_[𝕜] = ⟪f, g₁⟫ₙ_[𝕜] + ⟪f, g₂⟫ₙ_[𝕜] := by - simp_rw [nl2Inner, Pi.add_apply, mul_add, expect_add_distrib] - -lemma nl2Inner_smul_left [Star γ] [StarModule γ 𝕜] [SMulCommClass γ ℚ≥0 𝕜] - [IsScalarTower γ 𝕜 𝕜] (c : γ) (f g : ι → 𝕜) : - ⟪c • f, g⟫ₙ_[𝕜] = star c • ⟪f, g⟫ₙ_[𝕜] := by - simp only [nl2Inner, Pi.smul_apply, smul_mul_assoc, smul_expect, starRingEnd_apply, - star_smul] - -lemma nl2Inner_smul_right [Star γ] [StarModule γ 𝕜] [SMulCommClass γ ℚ≥0 𝕜] [IsScalarTower γ 𝕜 𝕜] - [SMulCommClass γ 𝕜 𝕜] (c : γ) (f g : ι → 𝕜) : ⟪f, c • g⟫ₙ_[𝕜] = c • ⟪f, g⟫ₙ_[𝕜] := by - simp only [nl2Inner, Pi.smul_apply, mul_smul_comm, smul_expect, starRingEnd_apply, - star_smul] - -lemma smul_nl2Inner_left [InvolutiveStar γ] [StarModule γ 𝕜] [SMulCommClass γ ℚ≥0 𝕜] - [IsScalarTower γ 𝕜 𝕜] (c : γ) (f g : ι → 𝕜) : c • ⟪f, g⟫ₙ_[𝕜] = ⟪star c • f, g⟫ₙ_[𝕜] := by - rw [nl2Inner_smul_left, star_star] - -end Semifield - -section Field -variable [Field 𝕜] [CharZero 𝕜] [StarRing 𝕜] - -@[simp] lemma nl2Inner_neg_left (f g : ι → 𝕜) : ⟪-f, g⟫ₙ_[𝕜] = -⟪f, g⟫ₙ_[𝕜] := by simp [nl2Inner] -@[simp] lemma nl2Inner_neg_right (f g : ι → 𝕜) : ⟪f, -g⟫ₙ_[𝕜] = -⟪f, g⟫ₙ_[𝕜] := by simp [nl2Inner] - -lemma nl2Inner_sub_left (f₁ f₂ g : ι → 𝕜) : ⟪f₁ - f₂, g⟫ₙ_[𝕜] = ⟪f₁, g⟫ₙ_[𝕜] - ⟪f₂, g⟫ₙ_[𝕜] := by - simp_rw [sub_eq_add_neg, nl2Inner_add_left, nl2Inner_neg_left] - -lemma nl2Inner_sub_right (f g₁ g₂ : ι → 𝕜) : ⟪f, g₁ - g₂⟫ₙ_[𝕜] = ⟪f, g₁⟫ₙ_[𝕜] - ⟪f, g₂⟫ₙ_[𝕜] := by - simp_rw [sub_eq_add_neg, nl2Inner_add_right, nl2Inner_neg_right] - -end Field - -section LinearOrderedSemifield -variable [LinearOrderedSemifield 𝕜] [PosSMulMono ℚ≥0 𝕜] [CharZero 𝕜] [StarRing 𝕜] - [StarOrderedRing 𝕜] {f g : ι → 𝕜} - -lemma nl2Inner_nonneg (hf : 0 ≤ f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫ₙ_[𝕜] := - expect_nonneg fun _ _ ↦ mul_nonneg (star_nonneg_iff.2 $ hf _) $ hg _ - -end LinearOrderedSemifield - -section LinearOrderedField -variable [LinearOrderedField 𝕜] [StarRing 𝕜] [TrivialStar 𝕜] {f g : ι → 𝕜} - ---TODO: Can we remove the `TrivialStar` assumption? -lemma abs_nl2Inner_le_nl2Inner_abs : |⟪f, g⟫ₙ_[𝕜]| ≤ ⟪|f|, |g|⟫ₙ_[𝕜] := - (abs_expect_le_expect_abs _ _).trans_eq $ - expect_congr rfl fun i _ ↦ by simp only [abs_mul, conj_trivial, Pi.abs_apply] - -end LinearOrderedField - -section RCLike -variable {κ : Type*} [RCLike 𝕜] {f : ι → 𝕜} - -@[simp] lemma nl2Inner_self (f : ι → 𝕜) : ⟪f, f⟫ₙ_[𝕜] = (‖f‖ₙ_[2] : 𝕜) ^ 2 := by - simp_rw [← algebraMap.coe_pow, nl2Norm_sq_eq_expect, nl2Inner, - algebraMap.coe_expect _ (α := ℝ) (β := 𝕜), RCLike.ofReal_pow, RCLike.conj_mul] - -lemma nl2Inner_self_of_norm_eq_one [Nonempty ι] (hf : ∀ x, ‖f x‖ = 1) : ⟪f, f⟫ₙ_[𝕜] = 1 := by - simp [-nl2Inner_self, nl2Inner, RCLike.conj_mul, hf] - -lemma linearIndependent_of_ne_zero_of_nl2Inner_eq_zero {v : κ → ι → 𝕜} (hz : ∀ k, v k ≠ 0) - (ho : Pairwise fun k l ↦ ⟪v k, v l⟫ₙ_[𝕜] = 0) : LinearIndependent 𝕜 v := by - cases isEmpty_or_nonempty ι - · have : IsEmpty κ := ⟨fun k ↦ hz k $ Subsingleton.elim _ _⟩ - exact linearIndependent_empty_type - · exact linearIndependent_of_ne_zero_of_l2Inner_eq_zero hz $ by - simpa [nl2Inner_eq_l2Inner_div_card] using ho - -end RCLike - -section nlpNorm -variable {α β : Type*} [Fintype α] {p : ℝ≥0∞} - -@[simp] lemma nlpNorm_conj [RCLike β] (f : α → β) : ‖conj f‖ₙ_[p] = ‖f‖ₙ_[p] := by simp [nlpNorm] - -variable [AddCommGroup α] - -@[simp] -lemma nlpNorm_translate [NormedAddCommGroup β] (a : α) (f : α → β) : ‖τ a f‖ₙ_[p] = ‖f‖ₙ_[p] := by - simp [nlpNorm] - -@[simp] lemma nlpNorm_conjneg [RCLike β] (f : α → β) : ‖conjneg f‖ₙ_[p] = ‖f‖ₙ_[p] := by - simp [nlpNorm] - -end nlpNorm - -section RCLike -variable {α β : Type*} [Fintype α] - -lemma nl1Norm_mul [RCLike β] (f g : α → β) : - ‖f * g‖ₙ_[1] = ⟪fun i ↦ ‖f i‖, fun i ↦ ‖g i‖⟫ₙ_[ℝ] := by simp [nl2Inner, nl1Norm_eq_expect] - -end RCLike - -/-- **Cauchy-Schwarz inequality** -/ -lemma nl2Inner_le_l2Norm_mul_l2Norm (f g : ι → ℝ) : ⟪f, g⟫ₙ_[ℝ] ≤ ‖f‖ₙ_[2] * ‖g‖ₙ_[2] := by - simp only [nlpNorm, div_mul_div_comm, ← sq, ENNReal.toReal_ofNat, ← one_div, ← sqrt_eq_rpow] - rw [sq_sqrt, nl2Inner_eq_l2Inner_div_card (𝕜 := ℝ)] - refine div_le_div_of_nonneg_right (l2Inner_le_l2Norm_mul_l2Norm _ _) ?_ - all_goals positivity - -namespace Mathlib.Meta.Positivity -open Lean Meta Qq Function - -private alias ⟨_, nlpNorm_pos_of_ne_zero⟩ := nlpNorm_pos - -private lemma nlpNorm_pos_of_pos {α : ι → Type*} [Nonempty ι] [∀ i, NormedAddCommGroup (α i)] - [∀ i, Preorder (α i)] {p : ℝ≥0∞} {f : ∀ i, α i} (hf : 0 < f) : 0 < ‖f‖ₙ_[p] := - nlpNorm_pos_of_ne_zero hf.ne' - -section LinearOrderedSemifield -variable [LinearOrderedSemifield 𝕜] [Module ℚ≥0 𝕜] [StarRing 𝕜] [StarOrderedRing 𝕜] {f g : ι → 𝕜} - -private lemma nl2Inner_nonneg_of_nonneg_of_nonneg (hf : 0 ≤ f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫ₙ_[𝕜] := - sorry - -private lemma nl2Inner_nonneg_of_pos_of_nonneg (hf : 0 < f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫ₙ_[𝕜] := - nl2Inner_nonneg_of_nonneg_of_nonneg hf.le hg - -private lemma nl2Inner_nonneg_of_nonneg_of_pos (hf : 0 ≤ f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫ₙ_[𝕜] := - nl2Inner_nonneg_of_nonneg_of_nonneg hf hg.le - -private lemma nl2Inner_nonneg_of_pos_of_pos (hf : 0 < f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫ₙ_[𝕜] := - nl2Inner_nonneg_of_nonneg_of_nonneg hf.le hg.le - -end LinearOrderedSemifield - -/-- The `positivity` extension which identifies expressions of the form `‖f‖ₙ_[p]`. -/ -@[positivity ‖_‖ₙ_[_]] def evalNLpNorm : PositivityExt where eval {u} α _z _p e := do - if let 0 := u then -- lean4#3060 means we can't combine this with the match below - match α, e with - | ~q(ℝ), ~q(@nlpNorm $ι $instι $α $instnorm $p $f) => - try - let _pα ← synthInstanceQ (q(∀ i, PartialOrder ($α i)) : Q(Type (max u_1 u_2))) - let _instιno ← synthInstanceQ (q(Nonempty $ι) : Q(Prop)) - assumeInstancesCommute - match ← core (q(inferInstance) : Q(Zero (∀ i, $α i))) q(inferInstance) f with - | .positive pf => return .positive q(nlpNorm_pos_of_pos $pf) - | .nonzero pf => return .positive q(nlpNorm_pos_of_ne_zero $pf) - | _ => return .nonnegative q(nlpNorm_nonneg) - catch _ => - assumeInstancesCommute - if let some pf ← findLocalDeclWithType? q($f ≠ 0) then - let pf : Q($f ≠ 0) := .fvar pf - let _instιno ← synthInstanceQ q(Nonempty $ι) - return .positive q(nlpNorm_pos_of_ne_zero $pf) - else - return .nonnegative q(nlpNorm_nonneg) - | _ => throwError "not nlpNorm" - else - throwError "not nlpNorm" - -/-- The `positivity` extension which identifies expressions of the form `⟪f, g⟫_[𝕜]`. -/ -@[positivity ⟪_, _⟫ₙ_[_]] def evalNL2Inner : PositivityExt where eval {u 𝕜} _ _ e := do - match e with - | ~q(@nl2Inner $ι _ $instι $instfield $instmod $inststar $f $g) => - let _p𝕜 ← synthInstanceQ q(LinearOrderedSemifield $𝕜) - let _p𝕜 ← synthInstanceQ q(StarRing $𝕜) - let _p𝕜 ← synthInstanceQ q(StarOrderedRing $𝕜) - assumeInstancesCommute - match ← core q(inferInstance) q(inferInstance) f, - ← core q(inferInstance) q(inferInstance) g with - | .positive pf, .positive pg => return .nonnegative q(nl2Inner_nonneg_of_pos_of_pos $pf $pg) - | .positive pf, .nonnegative pg => - return .nonnegative q(nl2Inner_nonneg_of_pos_of_nonneg $pf $pg) - | .nonnegative pf, .positive pg => - return .nonnegative q(nl2Inner_nonneg_of_nonneg_of_pos $pf $pg) - | .nonnegative pf, .nonnegative pg => - return .nonnegative q(nl2Inner_nonneg_of_nonneg_of_nonneg $pf $pg) - | _, _ => return .none - | _ => throwError "not nl2Inner" - -section Examples -section NormedAddCommGroup -variable {α : ι → Type*} [∀ i, NormedAddCommGroup (α i)] {w : ι → ℝ≥0} {f : ∀ i, α i} - -example {p : ℝ≥0∞} : 0 ≤ ‖f‖ₙ_[p] := by positivity -example {p : ℝ≥0∞} [Nonempty ι] (hf : f ≠ 0) : 0 < ‖f‖ₙ_[p] := by positivity -example {p : ℝ≥0∞} [Nonempty ι] {f : ι → ℝ} (hf : 0 < f) : 0 < ‖f‖ₙ_[p] := by positivity - -end NormedAddCommGroup - -section Complex -variable {w : ι → ℝ≥0} {f : ι → ℂ} - -example {p : ℝ≥0∞} : 0 ≤ ‖f‖ₙ_[p] := by positivity -example {p : ℝ≥0∞} [Nonempty ι] (hf : f ≠ 0) : 0 < ‖f‖ₙ_[p] := by positivity -example {p : ℝ≥0∞} [Nonempty ι] {f : ι → ℝ} (hf : 0 < f) : 0 < ‖f‖ₙ_[p] := by positivity - -end Complex - -section LinearOrderedSemifield -variable [LinearOrderedSemifield 𝕜] [StarRing 𝕜] [StarOrderedRing 𝕜] {f g : ι → 𝕜} - -example (hf : 0 < f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫ₙ_[𝕜] := by positivity -example (hf : 0 < f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫ₙ_[𝕜] := by positivity -example (hf : 0 ≤ f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫ₙ_[𝕜] := by positivity -example (hf : 0 ≤ f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫ₙ_[𝕜] := by positivity - -end LinearOrderedSemifield -end Examples -end Mathlib.Meta.Positivity - -/-! ### Hölder inequality -/ - -section nlpNorm -variable {α : Type*} [Fintype α] {p q : ℝ≥0} {f g : α → ℝ} - -@[simp] -lemma nlpNorm_abs (p : ℝ≥0∞) (f : α → ℝ) : ‖|f|‖ₙ_[p] = ‖f‖ₙ_[p] := by simpa using nlpNorm_norm p f - -lemma nl1Norm_mul_of_nonneg (hf : 0 ≤ f) (hg : 0 ≤ g) : ‖f * g‖ₙ_[1] = ⟪f, g⟫ₙ_[ℝ] := by - convert nl1Norm_mul f g using 2 <;> ext a <;> refine (norm_of_nonneg ?_).symm; exacts [hf _, hg _] - -lemma nlpNorm_rpow (hp : p ≠ 0) (hq : q ≠ 0) (hf : 0 ≤ f) : - ‖f ^ (q : ℝ)‖ₙ_[p] = ‖f‖ₙ_[p * q] ^ (q : ℝ) := by - refine rpow_left_injOn (NNReal.coe_ne_zero.2 hp) nlpNorm_nonneg (by dsimp; positivity) ?_ - dsimp - rw [← rpow_mul nlpNorm_nonneg, ← mul_comm, ← ENNReal.coe_mul, ← NNReal.coe_mul, - nlpNorm_rpow_eq_expect hp, nlpNorm_rpow_eq_expect (mul_ne_zero hq hp)] - simp [abs_rpow_of_nonneg (hf _), ← rpow_mul] - -lemma nl1Norm_rpow (hq : q ≠ 0) (hf : 0 ≤ f) : ‖f ^ (q : ℝ)‖ₙ_[1] = ‖f‖ₙ_[q] ^ (q : ℝ) := by - simpa only [ENNReal.coe_one, one_mul] using nlpNorm_rpow one_ne_zero hq hf - -lemma nlpNorm_eq_l1Norm_rpow (hp : p ≠ 0) (f : α → ℝ) : - ‖f‖ₙ_[p] = ‖|f| ^ (p : ℝ)‖ₙ_[1] ^ (p⁻¹ : ℝ) := by - simp [nlpNorm_eq_expect hp, nl1Norm_eq_expect, abs_rpow_of_nonneg] - -lemma nlpNorm_rpow' (hp : p ≠ 0) (hq : q ≠ 0) (f : α → ℝ) : - ‖f‖ₙ_[p] ^ (q : ℝ) = ‖|f| ^ (q : ℝ)‖ₙ_[p / q] := by - rw [← ENNReal.coe_div hq, nlpNorm_rpow (div_ne_zero hp hq) hq (abs_nonneg f), nlpNorm_abs, - ← ENNReal.coe_mul, div_mul_cancel₀ _ hq] - ---TODO: Generalise the following four to include `f g : α → ℂ` -/-- **Hölder's inequality**, binary case. -/ -lemma nl2Inner_le_nlpNorm_mul_nlpNorm (hpq : p.IsConjExponent q) (f g : α → ℝ) : - ⟪f, g⟫ₙ_[ℝ] ≤ ‖f‖ₙ_[p] * ‖g‖ₙ_[q] := by - cases isEmpty_or_nonempty α - · simp - have : 0 < (card α : ℝ) := by positivity - simpa [nl2Inner_eq_l2Inner_div_card, nlpNorm, div_mul_div_comm, ← rpow_add this, - hpq.coe.inv_add_inv_conj, div_le_div_right this] using l2Inner_le_lpNorm_mul_lpNorm hpq _ _ - -/-- **Hölder's inequality**, binary case. -/ -lemma abs_nl2Inner_le_nlpNorm_mul_nlpNorm (hpq : p.IsConjExponent q) (f g : α → ℝ) : - |⟪f, g⟫ₙ_[ℝ]| ≤ ‖f‖ₙ_[p] * ‖g‖ₙ_[q] := - abs_nl2Inner_le_nl2Inner_abs.trans $ - (nl2Inner_le_nlpNorm_mul_nlpNorm hpq _ _).trans_eq $ by simp_rw [nlpNorm_abs] - -/-- **Hölder's inequality**, binary case. -/ -lemma nlpNorm_mul_le (hp : p ≠ 0) (hq : q ≠ 0) (r : ℝ≥0) (hpqr : p⁻¹ + q⁻¹ = r⁻¹) (f g : α → ℝ) : - ‖f * g‖ₙ_[r] ≤ ‖f‖ₙ_[p] * ‖g‖ₙ_[q] := by - cases isEmpty_or_nonempty α - · simp - have : 0 < (card α : ℝ) := by positivity - simp only [nl2Inner_eq_l2Inner_div_card, nlpNorm, div_mul_div_comm, ← rpow_add this, - hpqr, div_le_div_right this, ← NNReal.coe_add, ← NNReal.coe_inv, ENNReal.coe_toReal] - exact div_le_div_of_nonneg_right (lpNorm_mul_le hp hq _ hpqr _ _) $ by positivity - -/-- **Hölder's inequality**, finitary case. -/ -lemma nlpNorm_prod_le {ι : Type*} {s : Finset ι} (hs : s.Nonempty) {p : ι → ℝ≥0} (hp : ∀ i, p i ≠ 0) - (q : ℝ≥0) (hpq : ∑ i ∈ s, (p i)⁻¹ = q⁻¹) (f : ι → α → ℝ) : - ‖∏ i ∈ s, f i‖ₙ_[q] ≤ ∏ i ∈ s, ‖f i‖ₙ_[p i] := by - cases isEmpty_or_nonempty α - · simp - have : 0 < (card α : ℝ) := by positivity - simp only [nl2Inner_eq_l2Inner_div_card, nlpNorm, prod_div_distrib, ← rpow_sum_of_pos this, - hpq, div_le_div_right this, ← NNReal.coe_sum, ← NNReal.coe_inv, ENNReal.coe_toReal] - exact div_le_div_of_nonneg_right (lpNorm_prod_le hs hp _ hpq _) $ by positivity - -end nlpNorm - -/-! ### Indicator -/ - -section indicate -variable {α β : Type*} [RCLike β] [Fintype α] [DecidableEq α] {s : Finset α} {p : ℝ≥0} - -lemma nlpNorm_rpow_indicate (hp : p ≠ 0) (s : Finset α) : ‖𝟭_[β] s‖ₙ_[p] ^ (p : ℝ) = s.dens := by - rw [nlpNorm, div_rpow] - simp [lpNorm_rpow_indicate hp, lpNorm_indicate, hp, dens] - all_goals positivity - -lemma nlpNorm_indicate (hp : p ≠ 0) (s : Finset α) : ‖𝟭_[β] s‖ₙ_[p] = s.dens ^ (p⁻¹ : ℝ) := by - refine (eq_rpow_inv ?_ ?_ ?_).2 (nlpNorm_rpow_indicate ?_ _) <;> positivity - -lemma nlpNorm_pow_indicate {p : ℕ} (hp : p ≠ 0) (s : Finset α) : - ‖𝟭_[β] s‖ₙ_[p] ^ (p : ℝ) = s.dens := by - simpa using nlpNorm_rpow_indicate (Nat.cast_ne_zero.2 hp) s - -lemma nl2Norm_sq_indicate (s : Finset α) : ‖𝟭_[β] s‖ₙ_[2] ^ 2 = s.dens := by - simpa using nlpNorm_pow_indicate two_ne_zero s - -lemma nl2Norm_indicate (s : Finset α) : ‖𝟭_[β] s‖ₙ_[2] = Real.sqrt s.dens := by - rw [eq_comm, sqrt_eq_iff_eq_sq, nl2Norm_sq_indicate] <;> positivity - -@[simp] lemma nl1Norm_indicate (s : Finset α) : ‖𝟭_[β] s‖ₙ_[1] = s.dens := by - simpa using nlpNorm_pow_indicate one_ne_zero s - -end indicate - -section mu -variable {α β : Type*} [RCLike β] [Fintype α] [DecidableEq α] {s : Finset α} {p : ℝ≥0} - -lemma nlpNorm_mu (hp : 1 ≤ p) (s : Finset α) : ‖μ_[β] s‖ₙ_[p] = s.dens ^ (p⁻¹ : ℝ) / s.card := by - rw [mu, nlpNorm_smul (ENNReal.one_le_coe_iff.2 hp) (s.card⁻¹ : β) (𝟭_[β] s), nlpNorm_indicate, - norm_inv, RCLike.norm_natCast, inv_mul_eq_div]; positivity - -lemma nl1Norm_mu (s : Finset α) : ‖μ_[β] s‖ₙ_[1] = s.dens / s.card := by - simpa using nlpNorm_mu le_rfl s - -end mu - -section -variable {α : Type*} [Fintype α] - -@[simp] -lemma RCLike.nlpNorm_coe_comp {𝕜 : Type*} [RCLike 𝕜] (p) (f : α → ℝ) : - ‖((↑) : ℝ → 𝕜) ∘ f‖ₙ_[p] = ‖f‖ₙ_[p] := by - simp only [← nlpNorm_norm _ (((↑) : ℝ → 𝕜) ∘ f), ← nlpNorm_norm _ f, Function.comp_apply, - RCLike.norm_ofReal, Real.norm_eq_abs] - -@[simp] lemma Complex.nlpNorm_coe_comp (p) (f : α → ℝ) : ‖((↑) : ℝ → ℂ) ∘ f‖ₙ_[p] = ‖f‖ₙ_[p] := - RCLike.nlpNorm_coe_comp _ _ - -end diff --git a/LeanAPAP/Prereqs/LpNorm/Compact/Defs.lean b/LeanAPAP/Prereqs/LpNorm/Compact/Defs.lean new file mode 100644 index 0000000000..76aee26345 --- /dev/null +++ b/LeanAPAP/Prereqs/LpNorm/Compact/Defs.lean @@ -0,0 +1,212 @@ +import Mathlib.Data.Finset.Density +import LeanAPAP.Mathlib.Probability.ConditionalProbability +import LeanAPAP.Prereqs.LpNorm.Discrete.Basic +import LeanAPAP.Prereqs.LpNorm.Discrete.Inner + +/-! +# Normalised Lp norms +-/ + +open Finset hiding card +open Function ProbabilityTheory Real +open Fintype (card) +open scoped BigOperators ComplexConjugate ENNReal NNReal + +variable {α 𝕜 R E : Type*} [MeasurableSpace α] + +/-! ### Lp norm -/ + +namespace MeasureTheory +section NormedAddCommGroup +variable [NormedAddCommGroup E] {p q : ℝ≥0∞} {f g h : α → E} + +/-- The Lp norm of a function with the compact normalisation. -/ +noncomputable def cLpNorm (p : ℝ≥0∞) (f : α → E) : ℝ≥0 := nnLpNorm f p .count[|Set.univ] + +notation "‖" f "‖ₙ_[" p "]" => cLpNorm p f + +@[simp] lemma cLpNorm_exponent_zero (f : α → E) : ‖f‖ₙ_[0] = 0 := by simp [cLpNorm] + +@[simp] lemma cLpNorm_zero (p : ℝ≥0∞) : ‖(0 : α → E)‖ₙ_[p] = 0 := by simp [cLpNorm] +@[simp] lemma cLpNorm_zero' (p : ℝ≥0∞) : ‖(fun _ ↦ 0 : α → E)‖ₙ_[p] = 0 := by simp [cLpNorm] + +@[simp] lemma cLpNorm_of_isEmpty [IsEmpty α] (f : α → E) (p : ℝ≥0∞) : ‖f‖ₙ_[p] = 0 := by + simp [cLpNorm] + +@[simp] lemma cLpNorm_neg (f : α → E) (p : ℝ≥0∞) : ‖-f‖ₙ_[p] = ‖f‖ₙ_[p] := by simp [cLpNorm] +@[simp] lemma cLpNorm_neg' (f : α → E) (p : ℝ≥0∞) : ‖fun x ↦ -f x‖ₙ_[p] = ‖f‖ₙ_[p] := by + simp [cLpNorm] + +lemma cLpNorm_sub_comm (f g : α → E) (p : ℝ≥0∞) : ‖f - g‖ₙ_[p] = ‖g - f‖ₙ_[p] := by + simp [cLpNorm, nnLpNorm_sub_comm] + +@[simp] lemma cLpNorm_norm (p : ℝ≥0∞) (f : α → E) : ‖fun i ↦ ‖f i‖‖ₙ_[p] = ‖f‖ₙ_[p] := + nnLpNorm_norm .. + +@[simp] lemma cLpNorm_abs (f : α → ℝ) (p : ℝ≥0∞) : ‖|f|‖ₙ_[p] = ‖f‖ₙ_[p] := nnLpNorm_abs .. +@[simp] lemma cLpNorm_abs' (f : α → ℝ) (p : ℝ≥0∞) : ‖fun i ↦ |f i|‖ₙ_[p] = ‖f‖ₙ_[p] := + nnLpNorm_abs' .. + +section NormedField +variable [NormedField 𝕜] {p : ℝ≥0∞} {f g : α → 𝕜} + +lemma cLpNorm_const_smul [Module 𝕜 E] [BoundedSMul 𝕜 E] (c : 𝕜) (f : α → E) : + ‖c • f‖ₙ_[p] = ‖c‖₊ * ‖f‖ₙ_[p] := by simp [cLpNorm, nnLpNorm_const_smul] + +lemma cLpNorm_nsmul [NormedSpace ℝ E] (n : ℕ) (f : α → E) : ‖n • f‖ₙ_[p] = n • ‖f‖ₙ_[p] := by + simp [cLpNorm, nnLpNorm_nsmul] + +variable [NormedSpace ℝ 𝕜] + +lemma cLpNorm_natCast_mul (n : ℕ) (f : α → 𝕜) (p : ℝ≥0∞) : ‖(n : α → 𝕜) * f‖ₙ_[p] = n * ‖f‖ₙ_[p] := + nnLpNorm_natCast_mul .. + +lemma cLpNorm_natCast_mul' (n : ℕ) (f : α → 𝕜) (p : ℝ≥0∞) : ‖(n * f ·)‖ₙ_[p] = n * ‖f‖ₙ_[p] := + nnLpNorm_natCast_mul' .. + +lemma cLpNorm_mul_natCast (f : α → 𝕜) (n : ℕ) (p : ℝ≥0∞) : ‖f * (n : α → 𝕜)‖ₙ_[p] = ‖f‖ₙ_[p] * n := + nnLpNorm_mul_natCast .. + +lemma cLpNorm_mul_natCast' (f : α → 𝕜) (n : ℕ) (p : ℝ≥0∞) : ‖(f · * n)‖ₙ_[p] = ‖f‖ₙ_[p] * n := + nnLpNorm_mul_natCast' .. + +lemma cLpNorm_div_natCast [CharZero 𝕜] {n : ℕ} (hn : n ≠ 0) (f : α → 𝕜) (p : ℝ≥0∞) : + ‖f / (n : α → 𝕜)‖ₙ_[p] = ‖f‖ₙ_[p] / n := nnLpNorm_div_natCast hn .. + +lemma cLpNorm_div_natCast' [CharZero 𝕜] {n : ℕ} (hn : n ≠ 0) (f : α → 𝕜) (p : ℝ≥0∞) : + ‖(f · / n)‖ₙ_[p] = ‖f‖ₙ_[p] / n := nnLpNorm_div_natCast' hn .. + +end NormedField + +section RCLike +variable {p : ℝ≥0∞} + +@[simp] lemma cLpNorm_conj [RCLike R] (f : α → R) : ‖conj f‖ₙ_[p] = ‖f‖ₙ_[p] := nnLpNorm_conj .. + +end RCLike + +section DiscreteMeasurableSpace +variable [DiscreteMeasurableSpace α] [Finite α] + +lemma cLpNorm_add_le (hp : 1 ≤ p) : ‖f + g‖ₙ_[p] ≤ ‖f‖ₙ_[p] + ‖g‖ₙ_[p] := + nnLpNorm_add_le .of_discrete .of_discrete hp + +lemma cLpNorm_sub_le (hp : 1 ≤ p) : ‖f - g‖ₙ_[p] ≤ ‖f‖ₙ_[p] + ‖g‖ₙ_[p] := + nnLpNorm_sub_le .of_discrete .of_discrete hp + +lemma cLpNorm_sum_le {ι : Type*} {s : Finset ι} {f : ι → α → E} (hp : 1 ≤ p) : + ‖∑ i ∈ s, f i‖ₙ_[p] ≤ ∑ i ∈ s, ‖f i‖ₙ_[p] := nnLpNorm_sum_le (fun _ _ ↦ .of_discrete) hp + +lemma cLpNorm_expect_le [Module ℚ≥0 E] [NormedSpace ℝ E] {ι : Type*} {s : Finset ι} {f : ι → α → E} + (hp : 1 ≤ p) : ‖𝔼 i ∈ s, f i‖ₙ_[p] ≤ 𝔼 i ∈ s, ‖f i‖ₙ_[p] := + nnLpNorm_expect_le (fun _ _ ↦ .of_discrete) hp + +lemma cLpNorm_le_cLpNorm_add_cLpNorm_sub' (hp : 1 ≤ p) : ‖f‖ₙ_[p] ≤ ‖g‖ₙ_[p] + ‖f - g‖ₙ_[p] := + nnLpNorm_le_nnLpNorm_add_nnLpNorm_sub' .of_discrete .of_discrete hp + +lemma cLpNorm_le_cLpNorm_add_cLpNorm_sub (hp : 1 ≤ p) : ‖f‖ₙ_[p] ≤ ‖g‖ₙ_[p] + ‖g - f‖ₙ_[p] := + nnLpNorm_le_nnLpNorm_add_nnLpNorm_sub .of_discrete .of_discrete hp + +lemma cLpNorm_le_add_cLpNorm_add (hp : 1 ≤ p) : ‖f‖ₙ_[p] ≤ ‖f + g‖ₙ_[p] + ‖g‖ₙ_[p] := + nnLpNorm_le_add_nnLpNorm_add .of_discrete .of_discrete hp + +lemma cLpNorm_sub_le_cLpNorm_sub_add_cLpNorm_sub (hp : 1 ≤ p) : + ‖f - h‖ₙ_[p] ≤ ‖f - g‖ₙ_[p] + ‖g - h‖ₙ_[p] := + nnLpNorm_sub_le_nnLpNorm_sub_add_nnLpNorm_sub .of_discrete .of_discrete .of_discrete hp + +end DiscreteMeasurableSpace + +variable [Fintype α] + +@[simp] lemma cLpNorm_const [Nonempty α] {p : ℝ≥0∞} (hp : p ≠ 0) (a : E) : + ‖fun _i : α ↦ a‖ₙ_[p] = ‖a‖₊ := by simp [cLpNorm, *] + +section NormedField +variable [NormedField 𝕜] {p : ℝ≥0∞} {f g : α → 𝕜} + +@[simp] lemma cLpNorm_one [Nonempty α] (hp : p ≠ 0) : + ‖(1 : α → 𝕜)‖ₙ_[p] = 1 := by simp [cLpNorm, *] + +end NormedField + +variable [DiscreteMeasurableSpace α] + +lemma cLpNorm_eq_expect_norm' (hp₀ : p ≠ 0) (hp : p ≠ ∞) (f : α → E) : + ‖f‖ₙ_[p] = (𝔼 i, ‖f i‖ ^ p.toReal) ^ p.toReal⁻¹ := by + simp [cLpNorm, coe_nnLpNorm_eq_integral_norm_rpow_toReal hp₀ hp .of_discrete, one_div, ← mul_sum, + integral_fintype, tsum_eq_sum' (s := univ) (by simp), ENNReal.coe_rpow_of_nonneg, cond_apply, + expect_eq_sum_div_card, div_eq_inv_mul] + +lemma cLpNorm_eq_expect_nnnorm' (hp₀ : p ≠ 0) (hp : p ≠ ∞) (f : α → E) : + ‖f‖ₙ_[p] = (𝔼 i, ‖f i‖₊ ^ p.toReal) ^ p.toReal⁻¹ := by + ext + simp [cLpNorm, coe_nnLpNorm_eq_integral_norm_rpow_toReal hp₀ hp .of_discrete, one_div, ← mul_sum, + integral_fintype, tsum_eq_sum' (s := univ) (by simp), ENNReal.coe_rpow_of_nonneg, cond_apply, + expect_eq_sum_div_card, div_eq_inv_mul] + +lemma cLpNorm_toNNReal_eq_expect_norm {p : ℝ} (hp : 0 < p) (f : α → E) : + ‖f‖ₙ_[p.toNNReal] = (𝔼 i, ‖f i‖ ^ p) ^ p⁻¹ := by + rw [cLpNorm_eq_expect_norm'] <;> simp [hp.le, hp] + +lemma cLpNorm_toNNReal_eq_expect_nnnorm {p : ℝ} (hp : 0 < p) (f : α → E) : + ‖f‖ₙ_[p.toNNReal] = (𝔼 i, ‖f i‖₊ ^ p) ^ p⁻¹ := by + rw [cLpNorm_eq_expect_nnnorm'] <;> simp [hp.le, hp] + +lemma cLpNorm_eq_expect_norm {p : ℝ≥0} (hp : p ≠ 0) (f : α → E) : + ‖f‖ₙ_[p] = (𝔼 i, ‖f i‖ ^ (p : ℝ)) ^ (p⁻¹ : ℝ) := + cLpNorm_eq_expect_norm' (by simpa using hp) (by simp) _ + +lemma cLpNorm_eq_expect_nnnorm {p : ℝ≥0} (hp : p ≠ 0) (f : α → E) : + ‖f‖ₙ_[p] = (𝔼 i, ‖f i‖₊ ^ (p : ℝ)) ^ (p⁻¹ : ℝ) := + cLpNorm_eq_expect_nnnorm' (by simpa using hp) (by simp) _ + +lemma cLpNorm_rpow_eq_expect_norm {p : ℝ≥0} (hp : p ≠ 0) (f : α → E) : + ‖f‖ₙ_[p] ^ (p : ℝ) = 𝔼 i, ‖f i‖ ^ (p : ℝ) := by + rw [cLpNorm_eq_expect_norm hp, Real.rpow_inv_rpow] <;> positivity + +lemma cLpNorm_rpow_eq_expect_nnnorm {p : ℝ≥0} (hp : p ≠ 0) (f : α → E) : + ‖f‖ₙ_[p] ^ (p : ℝ) = 𝔼 i, ‖f i‖₊ ^ (p : ℝ) := by + rw [cLpNorm_eq_expect_nnnorm hp, NNReal.rpow_inv_rpow]; positivity + +lemma cLpNorm_pow_eq_expect_norm {p : ℕ} (hp : p ≠ 0) (f : α → E) : + ‖f‖ₙ_[p] ^ p = 𝔼 i, ‖f i‖ ^ p := by + simpa using cLpNorm_rpow_eq_expect_norm (Nat.cast_ne_zero.2 hp) f + +lemma cLpNorm_pow_eq_expect_nnnorm {p : ℕ} (hp : p ≠ 0) (f : α → E) : + ‖f‖ₙ_[p] ^ p = 𝔼 i, ‖f i‖₊ ^ p := by + simpa using cLpNorm_rpow_eq_expect_nnnorm (Nat.cast_ne_zero.2 hp) f + +lemma cL2Norm_sq_eq_expect_norm (f : α → E) : ‖f‖ₙ_[2] ^ 2 = 𝔼 i, ‖f i‖ ^ 2 := by + simpa using cLpNorm_pow_eq_expect_norm two_ne_zero _ + +lemma cL2Norm_sq_eq_expect_nnnorm (f : α → E) : ‖f‖ₙ_[2] ^ 2 = 𝔼 i, ‖f i‖₊ ^ 2 := by + simpa using cLpNorm_pow_eq_expect_nnnorm two_ne_zero _ + +lemma cL2Norm_eq_expect_norm (f : α → E) : ‖f‖ₙ_[2] = (𝔼 i, ‖f i‖ ^ 2) ^ (2⁻¹ : ℝ) := by + simpa [sqrt_eq_rpow] using cLpNorm_eq_expect_norm two_ne_zero _ + +lemma cL2Norm_eq_expect_nnnorm (f : α → E) : ‖f‖ₙ_[2] = (𝔼 i, ‖f i‖₊ ^ 2) ^ (2⁻¹ : ℝ) := by + simpa [sqrt_eq_rpow] using cLpNorm_eq_expect_nnnorm two_ne_zero _ + +lemma cL1Norm_eq_expect_norm (f : α → E) : ‖f‖ₙ_[1] = 𝔼 i, ‖f i‖ := by + simp [cLpNorm_eq_expect_norm'] + +lemma cL1Norm_eq_expect_nnnorm (f : α → E) : ‖f‖ₙ_[1] = 𝔼 i, ‖f i‖₊ := by + simp [cLpNorm_eq_expect_nnnorm'] + +lemma cLinftyNorm_eq_iSup (f : α → E) : ‖f‖ₙ_[∞] = ⨆ i, ‖f i‖₊ := by + cases isEmpty_or_nonempty α + · simp + · simp [cLpNorm, nnLinftyNorm_eq_essSup] + sorry + +@[simp] lemma cLpNorm_eq_zero (hp : p ≠ 0) : ‖f‖ₙ_[p] = 0 ↔ f = 0 := by + simp [cLpNorm, nnLpNorm_eq_zero .of_discrete hp, ae_eq_top.2]; sorry + +@[simp] lemma cLpNorm_pos (hp : p ≠ 0) : 0 < ‖f‖ₙ_[p] ↔ f ≠ 0 := + pos_iff_ne_zero.trans (cLpNorm_eq_zero hp).not + +@[gcongr] lemma cLpNorm_mono_right (hpq : p ≤ q) : ‖f‖ₙ_[p] ≤ ‖f‖ₙ_[q] := sorry + +lemma cLpNorm_mono_real {g : α → ℝ} (h : ∀ x, ‖f x‖ ≤ g x) : ‖f‖ₙ_[p] ≤ ‖g‖ₙ_[p] := + nnLpNorm_mono_real .of_discrete h diff --git a/LeanAPAP/Prereqs/LpNorm/Compact/Inner.lean b/LeanAPAP/Prereqs/LpNorm/Compact/Inner.lean new file mode 100644 index 0000000000..be6f40012d --- /dev/null +++ b/LeanAPAP/Prereqs/LpNorm/Compact/Inner.lean @@ -0,0 +1,295 @@ +import Mathlib.Analysis.InnerProductSpace.PiL2 +import LeanAPAP.Mathlib.Algebra.Algebra.Rat +import LeanAPAP.Mathlib.Algebra.Group.Action.Defs +import LeanAPAP.Mathlib.Algebra.Order.Module.Defs +import LeanAPAP.Mathlib.Algebra.Star.Rat +import LeanAPAP.Mathlib.Data.NNReal.Basic +import LeanAPAP.Prereqs.Expect.Complex +import LeanAPAP.Prereqs.LpNorm.Compact.Defs + +/-! # Inner product -/ + +open Finset hiding card +open Function Real +open Fintype (card) +open scoped BigOperators ComplexConjugate ENNReal NNReal NNRat + +variable {ι R S : Type*} [Fintype ι] + +namespace MeasureTheory +section CommSemiring +variable [CommSemiring R] [StarRing R] [Module ℚ≥0 R] [DistribSMul S R] + +/-- Inner product giving rise to the L2 norm. -/ +def cL2Inner (f g : ι → R) : R := 𝔼 i, conj (f i) * g i + +notation "⟪" f ", " g "⟫ₙ_[" S "]" => cL2Inner (R := S) f g + +lemma cL2Inner_eq_smul_dL2Inner (f g : ι → R) : ⟪f, g⟫ₙ_[R] = (card ι : ℚ≥0)⁻¹ • ⟪f, g⟫_[R] := rfl + +lemma cL2Inner_eq_expect (f g : ι → R) : ⟪f, g⟫ₙ_[R] = 𝔼 i, conj (f i) * g i := rfl + +@[simp] lemma conj_cL2Inner (f g : ι → R) : conj ⟪f, g⟫ₙ_[R] = ⟪conj f, conj g⟫ₙ_[R] := by + simp [cL2Inner_eq_expect, map_expect] + exact Eq.trans (map_expect (starLinearEquiv ℚ≥0) _ _) (by simp [starRingEnd, starLinearEquiv]) + +lemma cL2Inner_anticomm (f g : ι → R) : ⟪f, g⟫ₙ_[R] = ⟪conj g, conj f⟫ₙ_[R] := by + simp [cL2Inner_eq_expect, map_sum, mul_comm] + +@[simp] lemma cL2Inner_zero_left (g : ι → R) : ⟪0, g⟫ₙ_[R] = 0 := by simp [cL2Inner_eq_expect] +@[simp] lemma cL2Inner_zero_right (f : ι → R) : ⟪f, 0⟫ₙ_[R] = 0 := by simp [cL2Inner_eq_expect] + +@[simp] lemma cL2Inner_of_isEmpty [IsEmpty ι] (f g : ι → R) : ⟪f, g⟫ₙ_[R] = 0 := by + simp [Subsingleton.elim f 0] + +@[simp] lemma cL2Inner_const_left (a : R) (f : ι → R) : + ⟪const _ a, f⟫ₙ_[R] = conj a * 𝔼 x, f x := by + simp only [cL2Inner_eq_expect, const_apply, mul_expect] + +@[simp] +lemma cL2Inner_const_right (f : ι → R) (a : R) : ⟪f, const _ a⟫ₙ_[R] = (𝔼 x, conj (f x)) * a := by + simp only [cL2Inner_eq_expect, const_apply, expect_mul] + +lemma cL2Inner_add_left (f₁ f₂ g : ι → R) : ⟪f₁ + f₂, g⟫ₙ_[R] = ⟪f₁, g⟫ₙ_[R] + ⟪f₂, g⟫ₙ_[R] := by + simp_rw [cL2Inner_eq_expect, Pi.add_apply, map_add, add_mul, expect_add_distrib] + +lemma cL2Inner_add_right (f g₁ g₂ : ι → R) : ⟪f, g₁ + g₂⟫ₙ_[R] = ⟪f, g₁⟫ₙ_[R] + ⟪f, g₂⟫ₙ_[R] := by + simp_rw [cL2Inner_eq_expect, Pi.add_apply, mul_add, expect_add_distrib] + +lemma cL2Inner_smul_left [Star S] [StarModule S R] [IsScalarTower S R R] (c : S) (f g : ι → R) : + ⟪c • f, g⟫ₙ_[R] = star c • ⟪f, g⟫ₙ_[R] := by + simp only [cL2Inner_eq_expect, Pi.smul_apply, smul_mul_assoc, smul_expect, starRingEnd_apply, + star_smul]; sorry + +lemma cL2Inner_smul_right [Star S] [StarModule S R] [SMulCommClass S R R] (c : S) (f g : ι → R) : + ⟪f, c • g⟫ₙ_[R] = c • ⟪f, g⟫ₙ_[R] := by + simp only [cL2Inner_eq_expect, Pi.smul_apply, mul_smul_comm, smul_expect, starRingEnd_apply, + star_smul]; sorry + +lemma smul_cL2Inner_left [InvolutiveStar S] [StarModule S R] [IsScalarTower S R R] (c : S) + (f g : ι → R) : c • ⟪f, g⟫ₙ_[R] = ⟪star c • f, g⟫ₙ_[R] := by rw [cL2Inner_smul_left, star_star] + +end CommSemiring + +section CommRing +variable [CommRing R] [StarRing R] [Module ℚ≥0 R] + +@[simp] +lemma cL2Inner_neg_left (f g : ι → R) : ⟪-f, g⟫ₙ_[R] = -⟪f, g⟫ₙ_[R] := by simp [cL2Inner_eq_expect] + +@[simp] +lemma cL2Inner_neg_right (f g : ι → R) : ⟪f, -g⟫ₙ_[R] = -⟪f, g⟫ₙ_[R] := by simp [cL2Inner_eq_expect] + +lemma cL2Inner_sub_left (f₁ f₂ g : ι → R) : ⟪f₁ - f₂, g⟫ₙ_[R] = ⟪f₁, g⟫ₙ_[R] - ⟪f₂, g⟫ₙ_[R] := by + simp_rw [sub_eq_add_neg, cL2Inner_add_left, cL2Inner_neg_left] + +lemma cL2Inner_sub_right (f g₁ g₂ : ι → R) : ⟪f, g₁ - g₂⟫ₙ_[R] = ⟪f, g₁⟫ₙ_[R] - ⟪f, g₂⟫ₙ_[R] := by + simp_rw [sub_eq_add_neg, cL2Inner_add_right, cL2Inner_neg_right] + +end CommRing + +section OrderedCommSemiring +variable [OrderedCommSemiring R] [StarRing R] [StarOrderedRing R] [Module ℚ≥0 R] + [PosSMulMono ℚ≥0 R] {f g : ι → R} + +lemma cL2Inner_nonneg (hf : 0 ≤ f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫ₙ_[R] := + expect_nonneg fun _ _ ↦ mul_nonneg (star_nonneg_iff.2 $ hf _) $ hg _ + +end OrderedCommSemiring + +section LinearOrderedCommRing +variable [LinearOrderedCommRing R] [StarRing R] [TrivialStar R] [Module ℚ≥0 R] + [PosSMulMono ℚ≥0 R] {f g : ι → R} + +--TODO: Can we remove the `TrivialStar` assumption? +lemma abs_cL2Inner_le_cL2Inner_abs : |⟪f, g⟫ₙ_[R]| ≤ ⟪|f|, |g|⟫ₙ_[R] := + (abs_expect_le_expect_abs _ _).trans_eq $ + expect_congr rfl fun i _ ↦ by simp only [abs_mul, conj_trivial, Pi.abs_apply] + +end LinearOrderedCommRing + +section RCLike +variable {κ : Type*} [RCLike R] {f : ι → R} + +@[simp] lemma cL2Inner_self {_ : MeasurableSpace ι} [DiscreteMeasurableSpace ι] (f : ι → R) : + ⟪f, f⟫ₙ_[R] = ((‖f‖ₙ_[2] : ℝ) : R) ^ 2 := by + simp_rw [← algebraMap.coe_pow, ← NNReal.coe_pow] + simp [cL2Norm_sq_eq_expect_nnnorm, cL2Inner_eq_expect, RCLike.conj_mul] + +lemma cL2Inner_self_of_norm_eq_one [Nonempty ι] (hf : ∀ x, ‖f x‖ = 1) : ⟪f, f⟫ₙ_[R] = 1 := by + simp [-cL2Inner_self, cL2Inner_eq_expect, RCLike.conj_mul, hf, card_univ] + +lemma linearIndependent_of_ne_zero_of_cL2Inner_eq_zero {v : κ → ι → R} (hz : ∀ k, v k ≠ 0) + (ho : Pairwise fun k l ↦ ⟪v k, v l⟫ₙ_[R] = 0) : LinearIndependent R v := by + cases isEmpty_or_nonempty ι + · have : IsEmpty κ := ⟨fun k ↦ hz k $ Subsingleton.elim _ _⟩ + exact linearIndependent_empty_type + · exact linearIndependent_of_ne_zero_of_dL2Inner_eq_zero hz $ by + simpa [cL2Inner_eq_smul_dL2Inner, ← NNRat.cast_smul_eq_nnqsmul R] using ho + +variable {mι : MeasurableSpace ι} [DiscreteMeasurableSpace ι] + +lemma cL1Norm_mul (f g : ι → R) : ‖f * g‖ₙ_[1] = ⟪fun i ↦ ‖f i‖, fun i ↦ ‖g i‖⟫ₙ_[ℝ] := by + simp [cL2Inner_eq_expect, cL1Norm_eq_expect_nnnorm] + +end RCLike + +variable {mι : MeasurableSpace ι} [DiscreteMeasurableSpace ι] + +/-- **Cauchy-Schwarz inequality** -/ +lemma cL2Inner_le_cL2Norm_mul_cL2Norm (f g : ι → ℝ) : ⟪f, g⟫ₙ_[ℝ] ≤ ‖f‖ₙ_[2] * ‖g‖ₙ_[2] := by + simp only [cL2Inner_eq_smul_dL2Inner, cL2Norm_eq_expect_nnnorm, ← NNReal.coe_mul, expect, + NNReal.coe_nnqsmul, ← NNRat.cast_smul_eq_nnqsmul ℝ≥0, smul_eq_mul, ← NNReal.mul_rpow, + mul_mul_mul_comm, ← sq] + simp only [NNReal.mul_rpow, ← dL2Norm_eq_sum_nnnorm, card_univ] + rw [← NNReal.rpow_two, NNReal.rpow_rpow_inv two_ne_zero, NNReal.smul_def, smul_eq_mul] + push_cast + gcongr + exact dL2Inner_le_dL2Norm_mul_dL2Norm _ _ + +end MeasureTheory + +namespace Mathlib.Meta.Positivity +open Lean Meta Qq Function MeasureTheory + +section OrderedCommSemiring +variable [OrderedCommSemiring R] [StarRing R] [StarOrderedRing R] [Module ℚ≥0 R] [PosSMulMono ℚ≥0 R] + {f g : ι → R} + +private lemma cL2Inner_nonneg_of_pos_of_nonneg (hf : 0 < f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫ₙ_[R] := + cL2Inner_nonneg hf.le hg + +private lemma cL2Inner_nonneg_of_nonneg_of_pos (hf : 0 ≤ f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫ₙ_[R] := + cL2Inner_nonneg hf hg.le + +private lemma cL2Inner_nonneg_of_pos_of_pos (hf : 0 < f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫ₙ_[R] := + cL2Inner_nonneg hf.le hg.le + +end OrderedCommSemiring + +/-- The `positivity` extension which identifies expressions of the form `⟪f, g⟫ₙ_[R]`. -/ +@[positivity ⟪_, _⟫ₙ_[_]] def evalCL2Inner : PositivityExt where eval {u R} _ _ e := do + match e with + | ~q(@cL2Inner $ι _ $instι $instring $inster $inststar $f $g) => + let _p𝕜 ← synthInstanceQ q(OrderedCommSemiring $R) + let _p𝕜 ← synthInstanceQ q(StarOrderedRing $R) + let _p𝕜 ← synthInstanceQ q(Module ℚ≥0 $R) + let _p𝕜 ← synthInstanceQ q(PosSMulMono ℚ≥0 $R) + assumeInstancesCommute + match ← core q(inferInstance) q(inferInstance) f, + ← core q(inferInstance) q(inferInstance) g with + | .positive pf, .positive pg => return .nonnegative q(cL2Inner_nonneg_of_pos_of_pos $pf $pg) + | .positive pf, .nonnegative pg => + return .nonnegative q(cL2Inner_nonneg_of_pos_of_nonneg $pf $pg) + | .nonnegative pf, .positive pg => + return .nonnegative q(cL2Inner_nonneg_of_nonneg_of_pos $pf $pg) + | .nonnegative pf, .nonnegative pg => return .nonnegative q(cL2Inner_nonneg $pf $pg) + | _, _ => return .none + | _ => throwError "not cL2Inner" + +section OrderedCommSemiring +variable [OrderedCommSemiring R] [StarRing R] [StarOrderedRing R] [Module ℚ≥0 R] [PosSMulMono ℚ≥0 R] + {f g : ι → R} + +example (hf : 0 < f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫ₙ_[R] := by positivity +example (hf : 0 < f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫ₙ_[R] := by positivity +example (hf : 0 ≤ f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫ₙ_[R] := by positivity +example (hf : 0 ≤ f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫ₙ_[R] := by positivity + +end OrderedCommSemiring +end Mathlib.Meta.Positivity + +/-! ### Hölder inequality -/ + +namespace MeasureTheory +section Real +variable {α : Type*} {mα : MeasurableSpace α} [DiscreteMeasurableSpace α] [Fintype α] {p q : ℝ≥0} + {f g : α → ℝ} + +lemma cL1Norm_mul_of_nonneg (hf : 0 ≤ f) (hg : 0 ≤ g) : ‖f * g‖ₙ_[1] = ⟪f, g⟫ₙ_[ℝ] := by + convert cL1Norm_mul f g using 2 <;> ext a <;> refine (norm_of_nonneg ?_).symm; exacts [hf _, hg _] + +/-- **Hölder's inequality**, binary case. -/ +lemma cL2Inner_le_cLpNorm_mul_cLpNorm (hpq : p.IsConjExponent q) (f g : α → ℝ) : + ⟪f, g⟫ₙ_[ℝ] ≤ ‖f‖ₙ_[p] * ‖g‖ₙ_[q] := by + have hp := hpq.ne_zero + have hq := hpq.symm.ne_zero + norm_cast at hp hq + rw [cL2Inner_eq_expect, expect_eq_sum_div_card, cLpNorm_eq_expect_nnnorm hp, + cLpNorm_eq_expect_nnnorm hq, expect_eq_sum_div_card, expect_eq_sum_div_card, + NNReal.div_rpow, NNReal.div_rpow, ← NNReal.coe_mul, div_mul_div_comm, ← NNReal.rpow_add', + hpq.coe.inv_add_inv_conj, NNReal.rpow_one] + push_cast + gcongr + rw [← dLpNorm_eq_sum_norm hp, ← dLpNorm_eq_sum_norm hq] + exact dL2Inner_le_dLpNorm_mul_dLpNorm hpq _ _ + · simp [hpq.coe.inv_add_inv_conj] + +/-- **Hölder's inequality**, binary case. -/ +lemma abs_cL2Inner_le_dLpNorm_mul_dLpNorm (hpq : p.IsConjExponent q) (f g : α → ℝ) : + |⟪f, g⟫ₙ_[ℝ]| ≤ ‖f‖ₙ_[p] * ‖g‖ₙ_[q] := + abs_cL2Inner_le_cL2Inner_abs.trans $ + (cL2Inner_le_cLpNorm_mul_cLpNorm hpq _ _).trans_eq $ by simp_rw [cLpNorm_abs] + +end Real + +section Hoelder +variable {α : Type*} {mα : MeasurableSpace α} [DiscreteMeasurableSpace α] [Fintype α] [RCLike R] + {p q : ℝ≥0} {f g : α → R} + +lemma norm_cL2Inner_le (f g : α → R) : ‖⟪f, g⟫ₙ_[R]‖₊ ≤ ⟪fun a ↦ ‖f a‖, fun a ↦ ‖g a‖⟫ₙ_[ℝ] := + (norm_sum_le _ _).trans $ by simp [cL2Inner] + +/-- **Hölder's inequality**, binary case. -/ +lemma norm_cL2Inner_le_dLpNorm_mul_dLpNorm (hpq : p.IsConjExponent q) (f g : α → R) : + ‖⟪f, g⟫ₙ_[R]‖₊ ≤ ‖f‖_[p] * ‖g‖_[q] := + calc + _ ≤ ⟪fun a ↦ ‖f a‖, fun a ↦ ‖g a‖⟫ₙ_[ℝ] := norm_cL2Inner_le _ _ + _ ≤ ‖fun a ↦ ‖f a‖‖_[p] * ‖fun a ↦ ‖g a‖‖_[q] := cL2Inner_le_cLpNorm_mul_cLpNorm hpq _ _ + _ = ‖f‖_[p] * ‖g‖_[q] := by simp_rw [dLpNorm_norm] + +/-- **Hölder's inequality**, binary case. -/ +lemma dLpNorm_mul_le (hp : p ≠ 0) (hq : q ≠ 0) (r : ℝ≥0) (hpqr : p⁻¹ + q⁻¹ = r⁻¹) (f g : α → R) : + ‖f * g‖_[r] ≤ ‖f‖_[p] * ‖g‖_[q] := by + have hr : r ≠ 0 := by + rintro rfl + simp [hp] at hpqr + have : (‖(f * g) ·‖ ^ (r : ℝ)) = (‖f ·‖ ^ (r : ℝ)) * (‖g ·‖ ^ (r : ℝ)) := by + ext; simp [mul_rpow, abs_mul] + rw [dLpNorm_eq_dL1Norm_rpow, NNReal.rpow_inv_le_iff_of_pos, this, ← NNReal.coe_le_coe] + push_cast + rw [dL1Norm_mul_of_nonneg, mul_rpow, ← NNReal.coe_rpow, ← NNReal.coe_rpow, dLpNorm_rpow', + dLpNorm_rpow', ← ENNReal.coe_div, ← ENNReal.coe_div] + refine cL2Inner_le_cLpNorm_mul_cLpNorm ⟨?_, ?_⟩ _ _ + · norm_cast + rw [div_eq_mul_inv, ← hpqr, mul_add, mul_inv_cancel₀ hp] + exact lt_add_of_pos_right _ (by positivity) + · norm_cast + simp [div_eq_mul_inv, hpqr, ← mul_add, hr] + any_goals intro a; dsimp + all_goals positivity + +/-- **Hölder's inequality**, binary case. -/ +lemma dL1Norm_mul_le (hpq : p.IsConjExponent q) (f g : α → R) : + ‖f * g‖_[1] ≤ ‖f‖_[p] * ‖g‖_[q] := + dLpNorm_mul_le (mod_cast hpq.ne_zero) (mod_cast hpq.symm.ne_zero) _ + (by simpa using hpq.inv_add_inv_conj) _ _ + +/-- **Hölder's inequality**, finitary case. -/ +lemma dLpNorm_prod_le {ι : Type*} {s : Finset ι} (hs : s.Nonempty) {p : ι → ℝ≥0} + (hp : ∀ i, p i ≠ 0) (q : ℝ≥0) + (hpq : 𝔼 i in s, (p i)⁻¹ = q⁻¹) (f : ι → α → R) : + ‖∏ i in s, f i‖_[q] ≤ ∏ i in s, ‖f i‖_[p i] := by + induction' s using Finset.cons_induction with i s hi ih generalizing q + · cases not_nonempty_empty hs + obtain rfl | hs := s.eq_empty_or_nonempty + · simp only [sum_cons, sum_empty, add_zero, inv_inj] at hpq + simp [← hpq] + simp_rw [prod_cons] + rw [sum_cons, ← inv_inv (𝔼 _ ∈ _, _ : ℝ≥0)] at hpq + refine (dLpNorm_mul_le (hp _) (inv_ne_zero (sum_pos (fun _ _ ↦ ?_) hs).ne') _ hpq _ _).trans + (mul_le_mul_left' (ih hs _ (inv_inv _).symm) _) + exact pos_iff_ne_zero.2 (inv_ne_zero $ hp _) + +end Hoelder +end MeasureTheory diff --git a/LeanAPAP/Prereqs/LpNorm/Discrete/Basic.lean b/LeanAPAP/Prereqs/LpNorm/Discrete/Basic.lean index fbb8b4b24d..3f80d1e16e 100644 --- a/LeanAPAP/Prereqs/LpNorm/Discrete/Basic.lean +++ b/LeanAPAP/Prereqs/LpNorm/Discrete/Basic.lean @@ -10,116 +10,90 @@ import LeanAPAP.Prereqs.LpNorm.Discrete.Defs open Finset Function Real open scoped BigOperators ComplexConjugate ENNReal NNReal -variable {ι 𝕜 : Type*} [Fintype ι] +namespace MeasureTheory +variable {ι G 𝕜 E R : Type*} [Fintype ι] {mι : MeasurableSpace ι} [DiscreteMeasurableSpace ι] -/-! ### Lp norm -/ +/-! ### Indicator -/ + +section Indicator +variable [RCLike R] [DecidableEq ι] {s : Finset ι} {p : ℝ≥0} + +lemma dLpNorm_rpow_indicate (hp : p ≠ 0) (s : Finset ι) : ‖𝟭_[R] s‖_[p] ^ (p : ℝ) = s.card := by + have : ∀ x, (ite (x ∈ s) 1 0 : ℝ) ^ (p : ℝ) = + ite (x ∈ s) (1 ^ (p : ℝ)) (0 ^ (p : ℝ)) := fun x ↦ by split_ifs <;> simp + simp [dLpNorm_rpow_eq_sum_nnnorm, hp, indicate_apply, apply_ite nnnorm, -sum_const, card_eq_sum_ones, + sum_boole, this, zero_rpow, filter_mem_eq_inter] + +lemma dLpNorm_indicate (hp : p ≠ 0) (s : Finset ι) : ‖𝟭_[R] s‖_[p] = s.card ^ (p⁻¹ : ℝ) := by + refine (NNReal.eq_rpow_inv_iff ?_).2 (dLpNorm_rpow_indicate ?_ _) <;> positivity + +lemma dLpNorm_pow_indicate {p : ℕ} (hp : p ≠ 0) (s : Finset ι) : + ‖𝟭_[R] s‖_[p] ^ (p : ℝ) = s.card := by + simpa using dLpNorm_rpow_indicate (Nat.cast_ne_zero.2 hp) s + +lemma dL2Norm_sq_indicate (s : Finset ι) : ‖𝟭_[R] s‖_[2] ^ 2 = s.card := by + simpa using dLpNorm_pow_indicate two_ne_zero s -section NormedAddCommGroup -variable {α : ι → Type*} [∀ i, NormedAddCommGroup (α i)] {p q : ℝ≥0∞} {f g h : ∀ i, α i} +lemma dL2Norm_indicate (s : Finset ι) : ‖𝟭_[R] s‖_[2] = NNReal.sqrt s.card := by + rw [eq_comm, NNReal.sqrt_eq_iff_eq_sq, dL2Norm_sq_indicate] -section one_le -variable [NormedField 𝕜] [∀ i, NormedSpace ℝ (α i)] +@[simp] lemma dL1Norm_indicate (s : Finset ι) : ‖𝟭_[R] s‖_[1] = s.card := by + simpa using dLpNorm_pow_indicate one_ne_zero s -lemma lpNorm_expect_le [∀ i, Module ℚ≥0 (α i)] (hp : 1 ≤ p) {κ : Type*} (s : Finset κ) - (f : κ → ∀ i, α i) : ‖𝔼 i ∈ s, f i‖_[p] ≤ 𝔼 i ∈ s, ‖f i‖_[p] := by +lemma dLpNorm_mu (hp : 1 ≤ p) (hs : s.Nonempty) : ‖μ_[R] s‖_[p] = s.card ^ ((p : ℝ)⁻¹ - 1) := by + rw [mu, dLpNorm_const_smul (s.card⁻¹ : R) (𝟭_[R] s), dLpNorm_indicate, nnnorm_inv, + RCLike.nnnorm_natCast, inv_mul_eq_div, ← NNReal.rpow_sub_one] <;> positivity + +lemma dLpNorm_mu_le (hp : 1 ≤ p) : ‖μ_[R] s‖_[p] ≤ s.card ^ (p⁻¹ - 1 : ℝ) := by obtain rfl | hs := s.eq_empty_or_nonempty · simp - refine (le_inv_smul_iff_of_pos $ by positivity).2 ?_ - rw [Nat.cast_smul_eq_nsmul, ← lpNorm_nsmul hp, card_smul_expect] - exact lpNorm_sum_le hp _ _ + · exact (dLpNorm_mu hp hs).le + +lemma dL1Norm_mu (hs : s.Nonempty) : ‖μ_[R] s‖_[1] = 1 := by simpa using dLpNorm_mu le_rfl hs -end one_le -end NormedAddCommGroup +lemma dL1Norm_mu_le_one : ‖μ_[R] s‖_[1] ≤ 1 := by simpa using dLpNorm_mu_le le_rfl -/-! #### Inner product -/ +end Indicator -section lpNorm -variable {α β : Type*} [AddCommGroup α] [Fintype α] {p : ℝ≥0∞} +/-! ### Translation -/ + +section dLpNorm +variable {mG : MeasurableSpace G} [DiscreteMeasurableSpace G] [AddCommGroup G] [Fintype G] + {p : ℝ≥0∞} @[simp] -lemma lpNorm_translate [NormedAddCommGroup β] (a : α) (f : α → β) : ‖τ a f‖_[p] = ‖f‖_[p] := by +lemma dLpNorm_translate [NormedAddCommGroup E] (a : G) (f : G → E) : ‖τ a f‖_[p] = ‖f‖_[p] := by obtain p | p := p - · simp only [linftyNorm_eq_ciSup, ENNReal.none_eq_top, translate_apply] + · simp only [dLinftyNorm_eq_iSup, ENNReal.none_eq_top, translate_apply] exact (Equiv.subRight _).iSup_congr fun _ ↦ rfl obtain rfl | hp := eq_or_ne p 0 - · simp only [l0Norm_eq_card, translate_apply, Ne, ENNReal.some_eq_coe, ENNReal.coe_zero, + · simp only [dLpNorm_exponent_zero, translate_apply, Ne, ENNReal.some_eq_coe, ENNReal.coe_zero, Nat.cast_inj] - exact card_equiv (Equiv.subRight a) (by simp) - · simp only [lpNorm_eq_sum hp, ENNReal.some_eq_coe, translate_apply] + · simp only [dLpNorm_eq_sum_nnnorm hp, ENNReal.some_eq_coe, translate_apply] congr 1 exact Fintype.sum_equiv (Equiv.subRight _) _ _ fun _ ↦ rfl -@[simp] lemma lpNorm_conjneg [RCLike β] (f : α → β) : ‖conjneg f‖_[p] = ‖f‖_[p] := by - simp only [conjneg, lpNorm_conj] +@[simp] lemma dLpNorm_conjneg [RCLike E] (f : G → E) : ‖conjneg f‖_[p] = ‖f‖_[p] := by + simp only [conjneg, dLpNorm_conj] obtain p | p := p - · simp only [linftyNorm_eq_ciSup, ENNReal.none_eq_top, conjneg, RCLike.norm_conj] + · simp only [dLinftyNorm_eq_iSup, ENNReal.none_eq_top, conjneg, RCLike.norm_conj] exact (Equiv.neg _).iSup_congr fun _ ↦ rfl obtain rfl | hp := eq_or_ne p 0 - · simp only [l0Norm_eq_card, Ne, ENNReal.some_eq_coe, ENNReal.coe_zero, Nat.cast_inj] - exact card_equiv (Equiv.neg _) (by simp) - · simp only [lpNorm_eq_sum hp, ENNReal.some_eq_coe] + · simp only [dLpNorm_exponent_zero, Ne, ENNReal.some_eq_coe, ENNReal.coe_zero, Nat.cast_inj] + · simp only [dLpNorm_eq_sum_nnnorm hp, ENNReal.some_eq_coe] congr 1 exact Fintype.sum_equiv (Equiv.neg _) _ _ fun _ ↦ rfl -lemma lpNorm_translate_sum_sub_le [NormedAddCommGroup β] (hp : 1 ≤ p) {ι : Type*} (s : Finset ι) - (a : ι → α) (f : α → β) : ‖τ (∑ i ∈ s, a i) f - f‖_[p] ≤ ∑ i ∈ s, ‖τ (a i) f - f‖_[p] := by +lemma dLpNorm_translate_sum_sub_le [NormedAddCommGroup E] (hp : 1 ≤ p) {ι : Type*} (s : Finset ι) + (a : ι → G) (f : G → E) : ‖τ (∑ i ∈ s, a i) f - f‖_[p] ≤ ∑ i ∈ s, ‖τ (a i) f - f‖_[p] := by induction' s using Finset.cons_induction with i s ih hs · simp calc _ = ‖τ (∑ j ∈ s, a j) (τ (a i) f - f) + (τ (∑ j ∈ s, a j) f - f)‖_[p] := by rw [sum_cons, translate_add', translate_sub_right, sub_add_sub_cancel] - _ ≤ ‖τ (∑ j ∈ s, a j) (τ (a i) f - f)‖_[p] + ‖(τ (∑ j ∈ s, a j) f - f)‖_[p] := - lpNorm_add_le hp _ _ + _ ≤ ‖τ (∑ j ∈ s, a j) (τ (a i) f - f)‖_[p] + ‖(τ (∑ j ∈ s, a j) f - f)‖_[p] := dLpNorm_add_le hp _ ≤ ‖τ (∑ j ∈ s, a j) (τ (a i) f - f)‖_[p] + ∑ j ∈ s, ‖(τ (a j) f - f)‖_[p] := add_le_add_left hs _ - _ = _ := by rw [lpNorm_translate, sum_cons] - -end lpNorm - -/-! ### Indicator -/ - -section indicate -variable {α β : Type*} [RCLike β] [Fintype α] [DecidableEq α] {s : Finset α} {p : ℝ≥0} - -lemma lpNorm_rpow_indicate (hp : p ≠ 0) (s : Finset α) : ‖𝟭_[β] s‖_[p] ^ (p : ℝ) = s.card := by - have : ∀ x, (ite (x ∈ s) 1 0 : ℝ) ^ (p : ℝ) = - ite (x ∈ s) (1 ^ (p : ℝ)) (0 ^ (p : ℝ)) := fun x ↦ by split_ifs <;> simp - simp [lpNorm_rpow_eq_sum, hp, indicate_apply, apply_ite Norm.norm, -sum_const, card_eq_sum_ones, - sum_boole, this, zero_rpow, filter_mem_eq_inter] - -lemma lpNorm_indicate (hp : p ≠ 0) (s : Finset α) : ‖𝟭_[β] s‖_[p] = s.card ^ (p⁻¹ : ℝ) := by - refine (eq_rpow_inv ?_ ?_ ?_).2 (lpNorm_rpow_indicate ?_ _) <;> positivity - -lemma lpNorm_pow_indicate {p : ℕ} (hp : p ≠ 0) (s : Finset α) : - ‖𝟭_[β] s‖_[p] ^ (p : ℝ) = s.card := by - simpa using lpNorm_rpow_indicate (Nat.cast_ne_zero.2 hp) s - -lemma l2Norm_sq_indicate (s : Finset α) : ‖𝟭_[β] s‖_[2] ^ 2 = s.card := by - simpa using lpNorm_pow_indicate two_ne_zero s - -lemma l2Norm_indicate (s : Finset α) : ‖𝟭_[β] s‖_[2] = Real.sqrt s.card := by - rw [eq_comm, sqrt_eq_iff_eq_sq, l2Norm_sq_indicate] <;> positivity - -@[simp] lemma l1Norm_indicate (s : Finset α) : ‖𝟭_[β] s‖_[1] = s.card := by - simpa using lpNorm_pow_indicate one_ne_zero s - -end indicate - -section mu -variable {α β : Type*} [RCLike β] [Fintype α] [DecidableEq α] {s : Finset α} {p : ℝ≥0} - -lemma lpNorm_mu (hp : 1 ≤ p) (hs : s.Nonempty) : ‖μ_[β] s‖_[p] = s.card ^ ((p : ℝ)⁻¹ - 1) := by - rw [mu, lpNorm_smul (ENNReal.one_le_coe_iff.2 hp) (s.card⁻¹ : β) (𝟭_[β] s), lpNorm_indicate, - norm_inv, RCLike.norm_natCast, inv_mul_eq_div, ← rpow_sub_one] <;> positivity - -lemma lpNorm_mu_le (hp : 1 ≤ p) : ‖μ_[β] s‖_[p] ≤ s.card ^ (p⁻¹ - 1 : ℝ) := by - obtain rfl | hs := s.eq_empty_or_nonempty - · simp - positivity - · exact (lpNorm_mu hp hs).le - -lemma l1Norm_mu (hs : s.Nonempty) : ‖μ_[β] s‖_[1] = 1 := by simpa using lpNorm_mu le_rfl hs - -lemma l1Norm_mu_le_one : ‖μ_[β] s‖_[1] ≤ 1 := by simpa using lpNorm_mu_le le_rfl + _ = _ := by rw [dLpNorm_translate, sum_cons] -end mu +end dLpNorm diff --git a/LeanAPAP/Prereqs/LpNorm/Discrete/Defs.lean b/LeanAPAP/Prereqs/LpNorm/Discrete/Defs.lean index 83407dec77..b70259a10a 100644 --- a/LeanAPAP/Prereqs/LpNorm/Discrete/Defs.lean +++ b/LeanAPAP/Prereqs/LpNorm/Discrete/Defs.lean @@ -1,561 +1,331 @@ -import Mathlib.Analysis.InnerProductSpace.PiL2 +import LeanAPAP.Mathlib.Analysis.RCLike.Basic +import LeanAPAP.Mathlib.Data.Fintype.Order +import LeanAPAP.Mathlib.MeasureTheory.Function.EssSup +import LeanAPAP.Mathlib.MeasureTheory.Measure.MeasureSpaceDef +import LeanAPAP.Mathlib.Order.Filter.Basic +import LeanAPAP.Prereqs.NNLpNorm /-! # Lp norms -/ open Finset Function Real -open scoped ComplexConjugate ENNReal NNReal NNRat +open scoped BigOperators ComplexConjugate ENNReal NNReal NNRat -variable {ι 𝕜 : Type*} [Fintype ι] +variable {α 𝕜 R E : Type*} [MeasurableSpace α] -/-! ### Lp norm -/ +namespace MeasureTheory +variable [NormedAddCommGroup E] {p q : ℝ≥0∞} {f g h : α → E} -section NormedAddCommGroup -variable {α : ι → Type*} [∀ i, NormedAddCommGroup (α i)] {p q : ℝ≥0∞} {f g h : ∀ i, α i} - -/-- The Lp norm of a function. -/ -noncomputable def lpNorm (p : ℝ≥0∞) (f : ∀ i, α i) : ℝ := ‖(WithLp.equiv p _).symm f‖ - -notation "‖" f "‖_[" p "]" => lpNorm p f - -lemma lpNorm_eq_sum' (hp : p.toReal ≠ 0) (f : ∀ i, α i) : - ‖f‖_[p] = (∑ i, ‖f i‖ ^ p.toReal) ^ p.toReal⁻¹ := by - rw [← one_div]; exact PiLp.norm_eq_sum (hp.lt_of_le' ENNReal.toReal_nonneg) _ - -lemma lpNorm_eq_sum'' {p : ℝ} (hp : 0 < p) (f : ∀ i, α i) : - ‖f‖_[p.toNNReal] = (∑ i, ‖f i‖ ^ p) ^ p⁻¹ := by rw [lpNorm_eq_sum'] <;> simp [hp.ne', hp.le] - -lemma lpNorm_eq_sum {p : ℝ≥0} (hp : p ≠ 0) (f : ∀ i, α i) : - ‖f‖_[p] = (∑ i, ‖f i‖ ^ (p : ℝ)) ^ (p⁻¹ : ℝ) := lpNorm_eq_sum' (by simpa using hp) _ - -lemma lpNorm_rpow_eq_sum {p : ℝ≥0} (hp : p ≠ 0) (f : ∀ i, α i) : - ‖f‖_[p] ^ (p : ℝ) = ∑ i, ‖f i‖ ^ (p : ℝ) := by - rw [lpNorm_eq_sum hp, rpow_inv_rpow (sum_nonneg fun i _ ↦ ?_)] <;> positivity - -lemma lpNorm_pow_eq_sum {p : ℕ} (hp : p ≠ 0) (f : ∀ i, α i) : ‖f‖_[p] ^ p = ∑ i, ‖f i‖ ^ p := by - simpa using lpNorm_rpow_eq_sum (Nat.cast_ne_zero.2 hp) f - -lemma l2Norm_sq_eq_sum (f : ∀ i, α i) : ‖f‖_[2] ^ 2 = ∑ i, ‖f i‖ ^ 2 := by - simpa using lpNorm_pow_eq_sum two_ne_zero _ - -lemma l2Norm_eq_sum (f : ∀ i, α i) : ‖f‖_[2] = sqrt (∑ i, ‖f i‖ ^ 2) := by - simpa [sqrt_eq_rpow] using lpNorm_eq_sum two_ne_zero _ - -lemma l1Norm_eq_sum (f : ∀ i, α i) : ‖f‖_[1] = ∑ i, ‖f i‖ := by simp [lpNorm_eq_sum'] - -lemma l0Norm_eq_card (f : ∀ i, α i) : ‖f‖_[0] = {i | f i ≠ 0}.toFinite.toFinset.card := - (PiLp.norm_eq_card _).trans $ by simp - -lemma linftyNorm_eq_ciSup (f : ∀ i, α i) : ‖f‖_[∞] = ⨆ i, ‖f i‖ := PiLp.norm_eq_ciSup _ - -@[simp] lemma lpNorm_zero : ‖(0 : ∀ i, α i)‖_[p] = 0 := by - obtain p | p := p; swap - obtain rfl | hp := @eq_zero_or_pos _ _ p - all_goals simp [linftyNorm_eq_ciSup, l0Norm_eq_card, lpNorm_eq_sum, *, ne_of_gt] +/-- The Lp norm of a function with the compact normalisation. -/ +noncomputable def dLpNorm (p : ℝ≥0∞) (f : α → E) : ℝ≥0 := nnLpNorm f p .count -@[simp] lemma lpNorm_of_isEmpty [IsEmpty ι] (p : ℝ≥0∞) (f : ∀ i, α i) : ‖f‖_[p] = 0 := by - simp [Subsingleton.elim f 0] +notation "‖" f "‖_[" p "]" => dLpNorm p f -@[simp] lemma lpNorm_norm (p : ℝ≥0∞) (f : ∀ i, α i) : ‖fun i ↦ ‖f i‖‖_[p] = ‖f‖_[p] := by - obtain p | p := p; swap - obtain rfl | hp := @eq_zero_or_pos _ _ p - all_goals simp [linftyNorm_eq_ciSup, l0Norm_eq_card, lpNorm_eq_sum, *, ne_of_gt] +@[simp] lemma dLpNorm_exponent_zero (f : α → E) : ‖f‖_[0] = 0 := by simp [dLpNorm] -@[simp] lemma lpNorm_neg (f : ∀ i, α i) : ‖-f‖_[p] = ‖f‖_[p] := by simp [← lpNorm_norm _ (-f)] +@[simp] lemma dLpNorm_zero (p : ℝ≥0∞) : ‖(0 : α → E)‖_[p] = 0 := by simp [dLpNorm] +@[simp] lemma dLpNorm_zero' (p : ℝ≥0∞) : ‖(fun _ ↦ 0 : α → E)‖_[p] = 0 := by simp [dLpNorm] -lemma lpNorm_sub_comm (f g : ∀ i, α i) : ‖f - g‖_[p] = ‖g - f‖_[p] := by - simp [← lpNorm_neg (f - g)] +@[simp] lemma dLpNorm_of_isEmpty [IsEmpty α] (f : α → E) (p : ℝ≥0∞) : ‖f‖_[p] = 0 := by + simp [dLpNorm] -@[simp] lemma lpNorm_nonneg : 0 ≤ ‖f‖_[p] := by - obtain p | p := p - · simp only [linftyNorm_eq_ciSup, ENNReal.none_eq_top] - exact Real.iSup_nonneg fun i ↦ norm_nonneg _ - obtain rfl | hp := eq_or_ne p 0 - · simp only [l0Norm_eq_card, lpNorm_eq_sum, ENNReal.some_eq_coe, ENNReal.coe_zero, *] - exact Nat.cast_nonneg _ - · simp only [lpNorm_eq_sum hp, ENNReal.some_eq_coe] - positivity +@[simp] lemma dLpNorm_neg (f : α → E) (p : ℝ≥0∞) : ‖-f‖_[p] = ‖f‖_[p] := by simp [dLpNorm] +@[simp] lemma dLpNorm_neg' (f : α → E) (p : ℝ≥0∞) : ‖fun x ↦ -f x‖_[p] = ‖f‖_[p] := by + simp [dLpNorm] -@[simp] lemma lpNorm_eq_zero : ‖f‖_[p] = 0 ↔ f = 0 := by - obtain p | p := p - · cases isEmpty_or_nonempty ι <;> - simp [linftyNorm_eq_ciSup, ENNReal.none_eq_top, ← sup'_univ_eq_ciSup, le_antisymm_iff, - Function.funext_iff] - obtain rfl | hp := eq_or_ne p 0 - · simp [l0Norm_eq_card, eq_empty_iff_forall_not_mem, Function.funext_iff] - · rw [← rpow_eq_zero lpNorm_nonneg (NNReal.coe_ne_zero.2 hp)] - simp [lpNorm_rpow_eq_sum hp, sum_eq_zero_iff_of_nonneg, rpow_nonneg, Function.funext_iff, - rpow_eq_zero _ (NNReal.coe_ne_zero.2 hp)] +lemma dLpNorm_sub_comm (f g : α → E) (p : ℝ≥0∞) : ‖f - g‖_[p] = ‖g - f‖_[p] := by + simp [dLpNorm, nnLpNorm_sub_comm] -@[simp] lemma lpNorm_pos : 0 < ‖f‖_[p] ↔ f ≠ 0 := lpNorm_nonneg.gt_iff_ne.trans lpNorm_eq_zero.not +@[simp] lemma dLpNorm_norm (f : α → E) (p : ℝ≥0∞) : ‖fun i ↦ ‖f i‖‖_[p] = ‖f‖_[p] := + nnLpNorm_norm .. -lemma lpNorm_mono_right (hpq : p ≤ q) (f : ∀ i, α i) : ‖f‖_[p] ≤ ‖f‖_[q] := sorry +@[simp] lemma dLpNorm_abs (f : α → ℝ) (p : ℝ≥0∞) : ‖|f|‖_[p] = ‖f‖_[p] := nnLpNorm_abs .. +@[simp] lemma dLpNorm_abs' (f : α → ℝ) (p : ℝ≥0∞) : ‖fun i ↦ |f i|‖_[p] = ‖f‖_[p] := + nnLpNorm_abs' .. -section one_le +section NormedField +variable [NormedField 𝕜] {p : ℝ≥0∞} {f g : α → 𝕜} -lemma lpNorm_add_le (hp : 1 ≤ p) (f g : ∀ i, α i) : ‖f + g‖_[p] ≤ ‖f‖_[p] + ‖g‖_[p] := - haveI := Fact.mk hp - norm_add_le _ _ +lemma dLpNorm_const_smul [Module 𝕜 E] [BoundedSMul 𝕜 E] (c : 𝕜) (f : α → E) : + ‖c • f‖_[p] = ‖c‖₊ * ‖f‖_[p] := by simp [dLpNorm, nnLpNorm_const_smul] -lemma lpNorm_sum_le (hp : 1 ≤ p) {κ : Type*} (s : Finset κ) (f : κ → ∀ i, α i) : - ‖∑ i ∈ s, f i‖_[p] ≤ ∑ i ∈ s, ‖f i‖_[p] := - haveI := Fact.mk hp - norm_sum_le _ _ +lemma dLpNorm_nsmul [NormedSpace ℝ E] (n : ℕ) (f : α → E) : ‖n • f‖_[p] = n • ‖f‖_[p] := by + simp [dLpNorm, nnLpNorm_nsmul] -lemma lpNorm_sub_le (hp : 1 ≤ p) (f g : ∀ i, α i) : ‖f - g‖_[p] ≤ ‖f‖_[p] + ‖g‖_[p] := - haveI := Fact.mk hp - norm_sub_le _ _ +variable [NormedSpace ℝ 𝕜] -lemma lpNorm_le_lpNorm_add_lpNorm_sub' (hp : 1 ≤ p) (f g : ∀ i, α i) : - ‖f‖_[p] ≤ ‖g‖_[p] + ‖f - g‖_[p] := - haveI := Fact.mk hp - norm_le_norm_add_norm_sub' _ _ +lemma dLpNorm_natCast_mul (n : ℕ) (f : α → 𝕜) (p : ℝ≥0∞) : ‖(n : α → 𝕜) * f‖_[p] = n * ‖f‖_[p] := + nnLpNorm_natCast_mul .. -lemma lpNorm_le_lpNorm_add_lpNorm_sub (hp : 1 ≤ p) (f g : ∀ i, α i) : - ‖f‖_[p] ≤ ‖g‖_[p] + ‖g - f‖_[p] := - haveI := Fact.mk hp - norm_le_norm_add_norm_sub _ _ +lemma dLpNorm_natCast_mul' (n : ℕ) (f : α → 𝕜) (p : ℝ≥0∞) : ‖(n * f ·)‖_[p] = n * ‖f‖_[p] := + nnLpNorm_natCast_mul' .. -lemma lpNorm_le_add_lpNorm_add (hp : 1 ≤ p) (f g : ∀ i, α i) : ‖f‖_[p] ≤ ‖f + g‖_[p] + ‖g‖_[p] := - haveI := Fact.mk hp - norm_le_add_norm_add _ _ +lemma dLpNorm_mul_natCast (f : α → 𝕜) (n : ℕ) (p : ℝ≥0∞) : ‖f * (n : α → 𝕜)‖_[p] = ‖f‖_[p] * n := + nnLpNorm_mul_natCast .. -lemma lpNorm_sub_le_lpNorm_sub_add_lpNorm_sub (hp : 1 ≤ p) (f g : ∀ i, α i) : - ‖f - h‖_[p] ≤ ‖f - g‖_[p] + ‖g - h‖_[p] := - haveI := Fact.mk hp - norm_sub_le_norm_sub_add_norm_sub _ _ _ - -variable [NormedField 𝕜] [∀ i, NormedSpace 𝕜 (α i)] - --- TODO: `p ≠ 0` is enough -lemma lpNorm_smul (hp : 1 ≤ p) (c : 𝕜) (f : ∀ i, α i) : ‖c • f‖_[p] = ‖c‖ * ‖f‖_[p] := - haveI := Fact.mk hp - norm_smul c _ - -variable [∀ i, NormedSpace ℝ (α i)] +lemma dLpNorm_mul_natCast' (f : α → 𝕜) (n : ℕ) (p : ℝ≥0∞) : ‖(f · * n)‖_[p] = ‖f‖_[p] * n := + nnLpNorm_mul_natCast' .. -lemma lpNorm_nsmul (hp : 1 ≤ p) (n : ℕ) (f : ∀ i, α i) : ‖n • f‖_[p] = n • ‖f‖_[p] := - haveI := Fact.mk hp - RCLike.norm_nsmul ℝ _ _ - -end one_le -end NormedAddCommGroup - -section NormedAddCommGroup -variable {α : Type*} [NormedAddCommGroup α] {p : ℝ≥0} +lemma dLpNorm_div_natCast [CharZero 𝕜] {n : ℕ} (hn : n ≠ 0) (f : α → 𝕜) (p : ℝ≥0∞) : + ‖f / (n : α → 𝕜)‖_[p] = ‖f‖_[p] / n := nnLpNorm_div_natCast hn .. -@[simp] -lemma lpNorm_const (hp : p ≠ 0) (a : α) : ‖const ι a‖_[p] = Fintype.card ι ^ (p⁻¹ : ℝ) * ‖a‖ := by - simp only [lpNorm_eq_sum hp, card_univ, mul_rpow, norm_nonneg, rpow_nonneg, - NNReal.coe_ne_zero.2 hp, rpow_rpow_inv, const_apply, sum_const, nsmul_eq_mul, Nat.cast_nonneg, - Ne, not_false_iff] +lemma dLpNorm_div_natCast' [CharZero 𝕜] {n : ℕ} (hn : n ≠ 0) (f : α → 𝕜) (p : ℝ≥0∞) : + ‖(f · / n)‖_[p] = ‖f‖_[p] / n := nnLpNorm_div_natCast' hn .. -end NormedAddCommGroup +end NormedField section RCLike -variable [RCLike 𝕜] {p : ℝ≥0} {f g : ι → 𝕜} - -@[simp] lemma lpNorm_one (hp : p ≠ 0) : ‖(1 : ι → 𝕜)‖_[p] = Fintype.card ι ^ (p⁻¹ : ℝ) := - (lpNorm_const hp 1).trans $ by simp - -lemma lpNorm_natCast_mul {p : ℝ≥0∞} (hp : 1 ≤ p) (n : ℕ) (f : ι → 𝕜) : - ‖(n : ι → 𝕜) * f‖_[p] = n * ‖f‖_[p] := by simpa only [nsmul_eq_mul] using lpNorm_nsmul hp n f - -lemma lpNorm_natCast_mul' {p : ℝ≥0∞} (hp : 1 ≤ p) (n : ℕ) (f : ι → 𝕜) : - ‖(n * f ·)‖_[p] = n * ‖f‖_[p] := lpNorm_natCast_mul hp _ _ - -lemma lpNorm_mul_natCast {p : ℝ≥0∞} (hp : 1 ≤ p) (f : ι → 𝕜) (n : ℕ) : - ‖f * (n : ι → 𝕜)‖_[p] = ‖f‖_[p] * n := by simpa only [mul_comm] using lpNorm_natCast_mul hp n f - -lemma lpNorm_mul_natCast' {p : ℝ≥0∞} (hp : 1 ≤ p) (f : ι → 𝕜) (n : ℕ) : - ‖(f · * n)‖_[p] = ‖f‖_[p] * n := lpNorm_mul_natCast hp _ _ - -lemma lpNorm_div_natCast {p : ℝ≥0∞} (hp : 1 ≤ p) (f : ι → 𝕜) (n : ℕ) : - ‖f / (n : ι → 𝕜)‖_[p] = ‖f‖_[p] / n := by - obtain rfl | hn := n.eq_zero_or_pos - · simp [Function.funext_iff] - · rw [eq_div_iff (by positivity), ← lpNorm_mul_natCast hp] - simp [Pi.mul_def, hn.ne'] +variable {p : ℝ≥0∞} -lemma lpNorm_div_natCast' {p : ℝ≥0∞} (hp : 1 ≤ p) (f : ι → 𝕜) (n : ℕ) : - ‖(f · / n)‖_[p] = ‖f‖_[p] / n := lpNorm_div_natCast hp _ _ +@[simp] lemma dLpNorm_conj [RCLike R] (f : α → R) : ‖conj f‖_[p] = ‖f‖_[p] := nnLpNorm_conj _ _ _ end RCLike -section Real -variable {p : ℝ≥0} {f g : ι → ℝ} - -lemma lpNorm_mono (hf : 0 ≤ f) (hfg : f ≤ g) : ‖f‖_[p] ≤ ‖g‖_[p] := by - obtain rfl | hp := eq_or_ne p 0 - · simp only [l0Norm_eq_card, ENNReal.some_eq_coe, ENNReal.coe_zero, Nat.cast_le] - exact - card_mono - (Set.Finite.toFinset_mono fun i ↦ mt fun hi ↦ ((hfg i).trans_eq hi).antisymm $ hf i) - have hp' := hp - rw [← pos_iff_ne_zero, ← NNReal.coe_pos] at hp - simp_rw [← rpow_le_rpow_iff lpNorm_nonneg lpNorm_nonneg hp, lpNorm_rpow_eq_sum hp', - norm_of_nonneg (hf _), norm_of_nonneg (hf.trans hfg _)] - exact sum_le_sum fun i _ ↦ rpow_le_rpow (hf _) (hfg _) hp.le - -end Real - -/-! #### Inner product -/ - -section CommSemiring -variable [CommSemiring 𝕜] [StarRing 𝕜] {γ : Type*} [DistribSMul γ 𝕜] - -/-- Inner product giving rise to the L2 norm. -/ -def l2Inner (f g : ι → 𝕜) : 𝕜 := ∑ i, conj (f i) * g i - -notation "⟪" f ", " g "⟫_[" 𝕜 "]" => @l2Inner _ 𝕜 _ _ _ f g - -lemma l2Inner_eq_sum (f g : ι → 𝕜) : ⟪f, g⟫_[𝕜] = ∑ i, conj (f i) * g i := rfl +section DiscreteMeasurableSpace +variable [DiscreteMeasurableSpace α] [Finite α] -@[simp] lemma conj_l2Inner (f g : ι → 𝕜) : conj ⟪f, g⟫_[𝕜] = ⟪conj f, conj g⟫_[𝕜] := by - simp [l2Inner_eq_sum, map_sum] +lemma dLpNorm_add_le (hp : 1 ≤ p) : ‖f + g‖_[p] ≤ ‖f‖_[p] + ‖g‖_[p] := + nnLpNorm_add_le .of_discrete .of_discrete hp -lemma l2Inner_anticomm (f g : ι → 𝕜) : ⟪f, g⟫_[𝕜] = ⟪conj g, conj f⟫_[𝕜] := by - simp [l2Inner_eq_sum, map_sum, mul_comm] +lemma dLpNorm_sub_le (hp : 1 ≤ p) : ‖f - g‖_[p] ≤ ‖f‖_[p] + ‖g‖_[p] := + nnLpNorm_sub_le .of_discrete .of_discrete hp -@[simp] lemma l2Inner_zero_left (g : ι → 𝕜) : ⟪0, g⟫_[𝕜] = 0 := by simp [l2Inner_eq_sum] -@[simp] lemma l2Inner_zero_right (f : ι → 𝕜) : ⟪f, 0⟫_[𝕜] = 0 := by simp [l2Inner_eq_sum] +lemma dLpNorm_sum_le {ι : Type*} {s : Finset ι} {f : ι → α → E} (hp : 1 ≤ p) : + ‖∑ i ∈ s, f i‖_[p] ≤ ∑ i ∈ s, ‖f i‖_[p] := nnLpNorm_sum_le (fun _ _ ↦ .of_discrete) hp -@[simp] lemma l2Inner_of_isEmpty [IsEmpty ι] (f g : ι → 𝕜) : ⟪f, g⟫_[𝕜] = 0 := by - simp [Subsingleton.elim f 0] +lemma dLpNorm_expect_le [Module ℚ≥0 E] [NormedSpace ℝ E] {ι : Type*} {s : Finset ι} {f : ι → α → E} + (hp : 1 ≤ p) : ‖𝔼 i ∈ s, f i‖_[p] ≤ 𝔼 i ∈ s, ‖f i‖_[p] := + nnLpNorm_expect_le (fun _ _ ↦ .of_discrete) hp -@[simp] lemma l2Inner_const_left (a : 𝕜) (f : ι → 𝕜) : ⟪const _ a, f⟫_[𝕜] = conj a * ∑ x, f x := by - simp only [l2Inner_eq_sum, const_apply, mul_sum] +lemma dLpNorm_le_dLpNorm_add_dLpNorm_sub' (hp : 1 ≤ p) : ‖f‖_[p] ≤ ‖g‖_[p] + ‖f - g‖_[p] := + nnLpNorm_le_nnLpNorm_add_nnLpNorm_sub' .of_discrete .of_discrete hp -@[simp] -lemma l2Inner_const_right (f : ι → 𝕜) (a : 𝕜) : ⟪f, const _ a⟫_[𝕜] = (∑ x, conj (f x)) * a := by - simp only [l2Inner_eq_sum, const_apply, sum_mul] - -lemma l2Inner_add_left (f₁ f₂ g : ι → 𝕜) : ⟪f₁ + f₂, g⟫_[𝕜] = ⟪f₁, g⟫_[𝕜] + ⟪f₂, g⟫_[𝕜] := by - simp_rw [l2Inner_eq_sum, Pi.add_apply, map_add, add_mul, sum_add_distrib] - -lemma l2Inner_add_right (f g₁ g₂ : ι → 𝕜) : ⟪f, g₁ + g₂⟫_[𝕜] = ⟪f, g₁⟫_[𝕜] + ⟪f, g₂⟫_[𝕜] := by - simp_rw [l2Inner_eq_sum, Pi.add_apply, mul_add, sum_add_distrib] - -lemma l2Inner_smul_left [Star γ] [StarModule γ 𝕜] [IsScalarTower γ 𝕜 𝕜] (c : γ) (f g : ι → 𝕜) : - ⟪c • f, g⟫_[𝕜] = star c • ⟪f, g⟫_[𝕜] := by - simp only [l2Inner_eq_sum, Pi.smul_apply, smul_mul_assoc, smul_sum, starRingEnd_apply, star_smul] +lemma dLpNorm_le_dLpNorm_add_dLpNorm_sub (hp : 1 ≤ p) : ‖f‖_[p] ≤ ‖g‖_[p] + ‖g - f‖_[p] := + nnLpNorm_le_nnLpNorm_add_nnLpNorm_sub .of_discrete .of_discrete hp -lemma l2Inner_smul_right [Star γ] [StarModule γ 𝕜] [SMulCommClass γ 𝕜 𝕜] (c : γ) (f g : ι → 𝕜) : - ⟪f, c • g⟫_[𝕜] = c • ⟪f, g⟫_[𝕜] := by - simp only [l2Inner_eq_sum, Pi.smul_apply, mul_smul_comm, smul_sum, starRingEnd_apply, star_smul] +lemma dLpNorm_le_add_dLpNorm_add (hp : 1 ≤ p) : ‖f‖_[p] ≤ ‖f + g‖_[p] + ‖g‖_[p] := + nnLpNorm_le_add_nnLpNorm_add .of_discrete .of_discrete hp -lemma smul_l2Inner_left [InvolutiveStar γ] [StarModule γ 𝕜] [IsScalarTower γ 𝕜 𝕜] (c : γ) - (f g : ι → 𝕜) : c • ⟪f, g⟫_[𝕜] = ⟪star c • f, g⟫_[𝕜] := by rw [l2Inner_smul_left, star_star] +lemma dLpNorm_sub_le_dLpNorm_sub_add_dLpNorm_sub (hp : 1 ≤ p) : + ‖f - h‖_[p] ≤ ‖f - g‖_[p] + ‖g - h‖_[p] := + nnLpNorm_sub_le_nnLpNorm_sub_add_nnLpNorm_sub .of_discrete .of_discrete .of_discrete hp -end CommSemiring +end DiscreteMeasurableSpace -section CommRing -variable [CommRing 𝕜] [StarRing 𝕜] +variable [Fintype α] @[simp] -lemma l2Inner_neg_left (f g : ι → 𝕜) : ⟪-f, g⟫_[𝕜] = -⟪f, g⟫_[𝕜] := by simp [l2Inner_eq_sum] +lemma dLpNorm_const [Nonempty α] {p : ℝ≥0∞} (hp : p ≠ 0) (a : E) : + ‖fun _i : α ↦ a‖_[p] = ‖a‖₊ * Fintype.card α ^ (p.toReal⁻¹ : ℝ) := by simp [dLpNorm, *] @[simp] -lemma l2Inner_neg_right (f g : ι → 𝕜) : ⟪f, -g⟫_[𝕜] = -⟪f, g⟫_[𝕜] := by simp [l2Inner_eq_sum] - -lemma l2Inner_sub_left (f₁ f₂ g : ι → 𝕜) : ⟪f₁ - f₂, g⟫_[𝕜] = ⟪f₁, g⟫_[𝕜] - ⟪f₂, g⟫_[𝕜] := by - simp_rw [sub_eq_add_neg, l2Inner_add_left, l2Inner_neg_left] +lemma dLpNorm_const' {p : ℝ≥0∞} (hp₀ : p ≠ 0) (hp : p ≠ ∞) (a : E) : + ‖fun _i : α ↦ a‖_[p] = ‖a‖₊ * Fintype.card α ^ (p.toReal⁻¹ : ℝ) := by simp [dLpNorm, *] -lemma l2Inner_sub_right (f g₁ g₂ : ι → 𝕜) : ⟪f, g₁ - g₂⟫_[𝕜] = ⟪f, g₁⟫_[𝕜] - ⟪f, g₂⟫_[𝕜] := by - simp_rw [sub_eq_add_neg, l2Inner_add_right, l2Inner_neg_right] +section NormedField +variable [NormedField 𝕜] {p : ℝ≥0∞} {f g : α → 𝕜} -end CommRing +@[simp] lemma dLpNorm_one [Nonempty α] (hp : p ≠ 0) : + ‖(1 : α → 𝕜)‖_[p] = Fintype.card α ^ (p.toReal⁻¹ : ℝ) := by simp [dLpNorm, *] -section OrderedCommSemiring -variable [OrderedCommSemiring 𝕜] [StarRing 𝕜] [StarOrderedRing 𝕜] {f g : ι → 𝕜} +@[simp] lemma dLpNorm_one' (hp₀ : p ≠ 0) (hp : p ≠ ∞) : + ‖(1 : α → 𝕜)‖_[p] = Fintype.card α ^ (p.toReal⁻¹ : ℝ) := by simp [dLpNorm, *] -lemma l2Inner_nonneg (hf : 0 ≤ f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫_[𝕜] := - sum_nonneg fun _ _ ↦ mul_nonneg (star_nonneg_iff.2 $ hf _) $ hg _ +end NormedField -end OrderedCommSemiring +variable [DiscreteMeasurableSpace α] -section LinearOrderedCommRing -variable [LinearOrderedCommRing 𝕜] [StarRing 𝕜] [TrivialStar 𝕜] {f g : ι → 𝕜} - ---TODO: Can we remove the `TrivialStar` assumption? -lemma abs_l2Inner_le_l2Inner_abs : |⟪f, g⟫_[𝕜]| ≤ ⟪|f|, |g|⟫_[𝕜] := - (abs_sum_le_sum_abs _ _).trans_eq $ - sum_congr rfl fun i _ ↦ by simp only [abs_mul, conj_trivial, Pi.abs_apply] +lemma dLpNorm_eq_sum_norm' (hp₀ : p ≠ 0) (hp : p ≠ ∞) (f : α → E) : + ‖f‖_[p] = (∑ i, ‖f i‖ ^ p.toReal) ^ p.toReal⁻¹ := by + simp [dLpNorm, coe_nnLpNorm_eq_integral_norm_rpow_toReal hp₀ hp .of_discrete, one_div, + integral_fintype, tsum_eq_sum' (s := univ) (by simp), ENNReal.coe_rpow_of_nonneg] -end LinearOrderedCommRing +lemma dLpNorm_eq_sum_nnnorm' (hp₀ : p ≠ 0) (hp : p ≠ ∞) (f : α → E) : + ‖f‖_[p] = (∑ i, ‖f i‖₊ ^ p.toReal) ^ p.toReal⁻¹ := by + ext; push_cast; exact dLpNorm_eq_sum_norm' hp₀ hp f -section RCLike -variable {κ : Type*} [RCLike 𝕜] {f : ι → 𝕜} +lemma dLpNorm_toNNReal_eq_sum_norm {p : ℝ} (hp : 0 < p) (f : α → E) : + ‖f‖_[p.toNNReal] = (∑ i, ‖f i‖ ^ p) ^ p⁻¹ := by + rw [dLpNorm_eq_sum_nnnorm'] <;> simp [hp.ne', hp.le, hp, ← ENNReal.coe_rpow_of_nonneg] -lemma l2Inner_eq_inner (f g : ι → 𝕜) : - ⟪f, g⟫_[𝕜] = inner ((WithLp.equiv 2 _).symm f) ((WithLp.equiv 2 _).symm g) := rfl +lemma dLpNorm_toNNReal_eq_sum_nnnorm {p : ℝ} (hp : 0 < p) (f : α → E) : + ‖f‖_[p.toNNReal] = (∑ i, ‖f i‖₊ ^ p) ^ p⁻¹ := by + rw [dLpNorm_eq_sum_nnnorm'] <;> simp [hp.ne', hp.le, hp, ← ENNReal.coe_rpow_of_nonneg] -lemma inner_eq_l2Inner (f g : PiLp 2 fun _i : ι ↦ 𝕜) : - inner f g = ⟪WithLp.equiv 2 _ f, WithLp.equiv 2 _ g⟫_[𝕜] := rfl +lemma dLpNorm_eq_sum_norm {p : ℝ≥0} (hp : p ≠ 0) (f : α → E) : + ‖f‖_[p] = (∑ i, ‖f i‖ ^ (p : ℝ)) ^ (p⁻¹ : ℝ) := + dLpNorm_eq_sum_norm' (by simpa using hp) (by simp) _ -@[simp] lemma l2Inner_self (f : ι → 𝕜) : ⟪f, f⟫_[𝕜] = (‖f‖_[2] : 𝕜) ^ 2 := by - simp_rw [← algebraMap.coe_pow, l2Norm_sq_eq_sum, l2Inner_eq_sum, algebraMap.coe_sum, - RCLike.ofReal_pow, RCLike.conj_mul] +lemma dLpNorm_eq_sum_nnnorm {p : ℝ≥0} (hp : p ≠ 0) (f : α → E) : + ‖f‖_[p] = (∑ i, ‖f i‖₊ ^ (p : ℝ)) ^ (p⁻¹ : ℝ) := + dLpNorm_eq_sum_nnnorm' (by simpa using hp) (by simp) _ -lemma l2Inner_self_of_norm_eq_one (hf : ∀ x, ‖f x‖ = 1) : ⟪f, f⟫_[𝕜] = Fintype.card ι := by - simp [-l2Inner_self, l2Inner_eq_sum, RCLike.conj_mul, hf, card_univ] +lemma dLpNorm_rpow_eq_sum_norm {p : ℝ≥0} (hp : p ≠ 0) (f : α → E) : + ‖f‖_[p] ^ (p : ℝ) = ∑ i, ‖f i‖ ^ (p : ℝ) := by + rw [dLpNorm_eq_sum_norm hp, Real.rpow_inv_rpow (by positivity) (mod_cast hp)] -lemma linearIndependent_of_ne_zero_of_l2Inner_eq_zero {v : κ → ι → 𝕜} (hz : ∀ k, v k ≠ 0) - (ho : Pairwise fun k l ↦ ⟪v k, v l⟫_[𝕜] = 0) : LinearIndependent 𝕜 v := by - simp_rw [l2Inner_eq_inner] at ho - have := linearIndependent_of_ne_zero_of_inner_eq_zero ?_ ho - exacts [this, hz] +lemma dLpNorm_rpow_eq_sum_nnnorm {p : ℝ≥0} (hp : p ≠ 0) (f : α → E) : + ‖f‖_[p] ^ (p : ℝ) = ∑ i, ‖f i‖₊ ^ (p : ℝ) := by + rw [dLpNorm_eq_sum_nnnorm hp, NNReal.rpow_inv_rpow (mod_cast hp)] -end RCLike +lemma dLpNorm_pow_eq_sum_norm {p : ℕ} (hp : p ≠ 0) (f : α → E) : ‖f‖_[p] ^ p = ∑ i, ‖f i‖ ^ p := by + simpa using dLpNorm_rpow_eq_sum_norm (Nat.cast_ne_zero.2 hp) f -section lpNorm -variable {α β : Type*} [Fintype α] {p : ℝ≥0∞} +lemma dLpNorm_pow_eq_sum_nnnorm {p : ℕ} (hp : p ≠ 0) (f : α → E) : + ‖f‖_[p] ^ p = ∑ i, ‖f i‖₊ ^ p := by + simpa using dLpNorm_rpow_eq_sum_nnnorm (Nat.cast_ne_zero.2 hp) f -@[simp] lemma lpNorm_conj [RCLike β] (f : α → β) : ‖conj f‖_[p] = ‖f‖_[p] := by - obtain p | p := p; swap; obtain rfl | hp := eq_or_ne p 0 - all_goals - simp only [linftyNorm_eq_ciSup, lpNorm_eq_sum, l0Norm_eq_card, ENNReal.some_eq_coe, - ENNReal.none_eq_top, ENNReal.coe_zero, Pi.conj_apply, RCLike.norm_conj, map_ne_zero, *] - · simp only [lpNorm_eq_sum hp, Pi.conj_apply, RCLike.norm_conj] +lemma dL2Norm_sq_eq_sum_norm (f : α → E) : ‖f‖_[2] ^ 2 = ∑ i, ‖f i‖ ^ 2 := by + simpa using dLpNorm_pow_eq_sum_norm two_ne_zero _ -end lpNorm +lemma dL2Norm_sq_eq_sum_nnnorm (f : α → E) : ‖f‖_[2] ^ 2 = ∑ i, ‖f i‖₊ ^ 2 := by + simpa using dLpNorm_pow_eq_sum_nnnorm two_ne_zero _ -section RCLike -variable {α β : Type*} [Fintype α] +lemma dL2Norm_eq_sum_norm (f : α → E) : ‖f‖_[2] = (∑ i, ‖f i‖ ^ 2) ^ (2⁻¹ : ℝ) := by + simpa [sqrt_eq_rpow] using dLpNorm_eq_sum_norm two_ne_zero _ -lemma l1Norm_mul [RCLike β] (f g : α → β) : - ‖f * g‖_[1] = ⟪fun i ↦ ‖f i‖, fun i ↦ ‖g i‖⟫_[ℝ] := by simp [l2Inner_eq_sum, l1Norm_eq_sum] +lemma dL2Norm_eq_sum_nnnorm (f : α → E) : ‖f‖_[2] = (∑ i, ‖f i‖₊ ^ 2) ^ (2⁻¹ : ℝ) := by + simpa [sqrt_eq_rpow] using dLpNorm_eq_sum_nnnorm two_ne_zero _ -end RCLike +lemma dL1Norm_eq_sum_norm (f : α → E) : ‖f‖_[1] = ∑ i, ‖f i‖ := by simp [dLpNorm_eq_sum_norm'] +lemma dL1Norm_eq_sum_nnnorm (f : α → E) : ‖f‖_[1] = ∑ i, ‖f i‖₊ := by simp [dLpNorm_eq_sum_nnnorm'] -/-- **Cauchy-Schwarz inequality** -/ -lemma l2Inner_le_l2Norm_mul_l2Norm (f g : ι → ℝ) : ⟪f, g⟫_[ℝ] ≤ ‖f‖_[2] * ‖g‖_[2] := - real_inner_le_norm ((WithLp.equiv 2 _).symm f) _ +@[simp] lemma dLinftyNorm_eq_iSup (f : α → E) : ‖f‖_[∞] = ⨆ i, ‖f i‖₊ := by + cases isEmpty_or_nonempty α + · simp + · simp [dLpNorm, nnLinftyNorm_eq_essSup] -namespace Mathlib.Meta.Positivity -open Lean Meta Qq Function +@[simp] lemma dLpNorm_eq_zero (hp : p ≠ 0) : ‖f‖_[p] = 0 ↔ f = 0 := by + simp [dLpNorm, nnLpNorm_eq_zero .of_discrete hp, ae_eq_top.2] -private alias ⟨_, lpNorm_pos_of_ne_zero⟩ := lpNorm_pos +@[simp] lemma dLpNorm_pos (hp : p ≠ 0) : 0 < ‖f‖_[p] ↔ f ≠ 0 := + pos_iff_ne_zero.trans (dLpNorm_eq_zero hp).not -private lemma lpNorm_pos_of_pos {α : ι → Type*} [∀ i, NormedAddCommGroup (α i)] - [∀ i, Preorder (α i)] {p : ℝ≥0∞} {f : ∀ i, α i} (hf : 0 < f) : 0 < ‖f‖_[p] := - lpNorm_pos_of_ne_zero hf.ne' +@[gcongr] lemma dLpNorm_mono_right (hpq : p ≤ q) : ‖f‖_[p] ≤ ‖f‖_[q] := sorry -section OrderedCommSemiring -variable [OrderedCommSemiring 𝕜] [StarRing 𝕜] [StarOrderedRing 𝕜] {f g : ι → 𝕜} +lemma dLpNorm_mono_real {g : α → ℝ} (h : ∀ x, ‖f x‖ ≤ g x) : ‖f‖_[p] ≤ ‖g‖_[p] := + nnLpNorm_mono_real .of_discrete h -private lemma l2Inner_nonneg_of_pos_of_nonneg (hf : 0 < f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫_[𝕜] := - l2Inner_nonneg hf.le hg +end MeasureTheory -private lemma l2Inner_nonneg_of_nonneg_of_pos (hf : 0 ≤ f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫_[𝕜] := - l2Inner_nonneg hf hg.le +namespace Mathlib.Meta.Positivity +open Lean Meta Qq Function MeasureTheory -private lemma l2Inner_nonneg_of_pos_of_pos (hf : 0 < f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫_[𝕜] := - l2Inner_nonneg hf.le hg.le +private alias ⟨_, dLpNorm_pos_of_ne_zero⟩ := dLpNorm_pos -end OrderedCommSemiring +private lemma dLpNorm_pos_of_pos {α E : Type*} {_ : MeasurableSpace α} [DiscreteMeasurableSpace α] + [Fintype α] [NormedAddCommGroup E] [Preorder E] {p : ℝ≥0∞} {f : α → E} + (hp : p ≠ 0) (hf : 0 < f) : 0 < ‖f‖_[p] := dLpNorm_pos_of_ne_zero hp hf.ne' /-- The `positivity` extension which identifies expressions of the form `‖f‖_[p]`. -/ @[positivity ‖_‖_[_]] def evalLpNorm : PositivityExt where eval {u} α _z _p e := do - if let 0 := u then -- lean4#3060 means we can't combine this with the match below - match α, e with - | ~q(ℝ), ~q(@lpNorm $ι $instι $α $instnorm $p $f) => + match u, α, e with + | 0, ~q(ℝ≥0), ~q(@nnLpNorm $ι $E $instιmeas $instEnorm $f $p $μ) => + match μ with + | ~q(Measure.count) => + let some pp := (← core q(inferInstance) q(inferInstance) p).toNonzero _ _ | failure try - let _pα ← synthInstanceQ (q(∀ i, PartialOrder ($α i)) : Q(Type (max u_1 u_2))) + let _pE ← synthInstanceQ q(PartialOrder $E) assumeInstancesCommute + let _ ← synthInstanceQ q(Fintype $ι) + let _ ← synthInstanceQ q(DiscreteMeasurableSpace $ι) match ← core q(inferInstance) q(inferInstance) f with - | .positive pf => return .positive q(lpNorm_pos_of_pos $pf) - | .nonzero pf => return .positive q(lpNorm_pos_of_ne_zero $pf) - | _ => return .nonnegative q(lpNorm_nonneg) + | .positive pf => + return .positive q(dLpNorm_pos_of_pos $pp $pf) + | .nonzero pf => return .positive q(dLpNorm_pos_of_ne_zero $pp $pf) + | _ => failure catch _ => assumeInstancesCommute - if let some pf ← findLocalDeclWithType? q($f ≠ 0) then - let pf : Q($f ≠ 0) := .fvar pf - return .positive q(lpNorm_pos_of_ne_zero $pf) - else - return .nonnegative q(lpNorm_nonneg) - | _ => throwError "not lpNorm" - else - throwError "not lpNorm" - -/-- The `positivity` extension which identifies expressions of the form `⟪f, g⟫_[𝕜]`. -/ -@[positivity ⟪_, _⟫_[_]] def evalL2Inner : PositivityExt where eval {u 𝕜} _ _ e := do - match e with - | ~q(@l2Inner $ι _ $instι $instring $inststar $f $g) => - let _p𝕜 ← synthInstanceQ q(OrderedCommSemiring $𝕜) - let _p𝕜 ← synthInstanceQ q(StarOrderedRing $𝕜) - assumeInstancesCommute - match ← core q(inferInstance) q(inferInstance) f, - ← core q(inferInstance) q(inferInstance) g with - | .positive pf, .positive pg => return .nonnegative q(l2Inner_nonneg_of_pos_of_pos $pf $pg) - | .positive pf, .nonnegative pg => - return .nonnegative q(l2Inner_nonneg_of_pos_of_nonneg $pf $pg) - | .nonnegative pf, .positive pg => - return .nonnegative q(l2Inner_nonneg_of_nonneg_of_pos $pf $pg) - | .nonnegative pf, .nonnegative pg => return .nonnegative q(l2Inner_nonneg $pf $pg) - | _, _ => return .none - | _ => throwError "not l2Inner" + let some pf ← findLocalDeclWithType? q($f ≠ 0) | failure + let pf : Q($f ≠ 0) := .fvar pf + let _ ← synthInstanceQ q(Fintype $ι) + let _ ← synthInstanceQ q(DiscreteMeasurableSpace $ι) + return .positive q(dLpNorm_pos_of_ne_zero $pp $pf) + | _ => throwError "not dLpNorm" + | _ => throwError "not dLpNorm" section Examples section NormedAddCommGroup -variable {α : ι → Type*} [∀ i, NormedAddCommGroup (α i)] {w : ι → ℝ≥0} {f : ∀ i, α i} +variable [Fintype α] [DiscreteMeasurableSpace α] [NormedAddCommGroup E] [PartialOrder E] + {w : α → ℝ≥0} {f : α → E} example {p : ℝ≥0∞} : 0 ≤ ‖f‖_[p] := by positivity -example {p : ℝ≥0∞} (hf : f ≠ 0) : 0 < ‖f‖_[p] := by positivity -example {p : ℝ≥0∞} {f : ι → ℝ} (hf : 0 < f) : 0 < ‖f‖_[p] := by positivity +-- example {p : ℝ≥0∞} (hp : p ≠ 0) (hf : f ≠ 0) : 0 < ‖f‖_[p] := by positivity +-- example {p : ℝ≥0∞} (hp : p ≠ 0) {f : α → ℝ} (hf : 0 < f) : 0 < ‖f‖_[p] := by positivity end NormedAddCommGroup section Complex -variable {w : ι → ℝ≥0} {f : ι → ℂ} +variable [DiscreteMeasurableSpace α] {w : α → ℝ≥0} {f : α → ℂ} example {p : ℝ≥0∞} : 0 ≤ ‖f‖_[p] := by positivity -example {p : ℝ≥0∞} (hf : f ≠ 0) : 0 < ‖f‖_[p] := by positivity -example {p : ℝ≥0∞} {f : ι → ℝ} (hf : 0 < f) : 0 < ‖f‖_[p] := by positivity +-- example {p : ℝ≥0∞} (hp : p ≠ 0) (hf : f ≠ 0) : 0 < ‖f‖_[p] := by positivity +-- example {p : ℝ≥0∞} (hp : p ≠ 0) {f : α → ℝ} (hf : 0 < f) : 0 < ‖f‖_[p] := by positivity end Complex - -section OrderedCommSemiring -variable [OrderedCommSemiring 𝕜] [StarRing 𝕜] [StarOrderedRing 𝕜] {f g : ι → 𝕜} - -example (hf : 0 < f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫_[𝕜] := by positivity -example (hf : 0 < f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫_[𝕜] := by positivity -example (hf : 0 ≤ f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫_[𝕜] := by positivity -example (hf : 0 ≤ f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫_[𝕜] := by positivity - -end OrderedCommSemiring end Examples end Mathlib.Meta.Positivity /-! ### Hölder inequality -/ +namespace MeasureTheory section Real -variable {α : Type*} [Fintype α] {p q : ℝ≥0} {f g : α → ℝ} - -@[simp] -lemma lpNorm_abs (p : ℝ≥0∞) (f : α → ℝ) : ‖|f|‖_[p] = ‖f‖_[p] := by simpa using lpNorm_norm p f +variable {α : Type*} {mα : MeasurableSpace α} [DiscreteMeasurableSpace α] [Fintype α] {p q : ℝ≥0} + {f g : α → ℝ} -lemma l1Norm_mul_of_nonneg (hf : 0 ≤ f) (hg : 0 ≤ g) : ‖f * g‖_[1] = ⟪f, g⟫_[ℝ] := by - convert l1Norm_mul f g using 2 <;> ext a <;> refine (norm_of_nonneg ?_).symm; exacts [hf _, hg _] - -lemma lpNorm_rpow (hp : p ≠ 0) (hq : q ≠ 0) (hf : 0 ≤ f) : +lemma dLpNorm_rpow (hp : p ≠ 0) (hq : q ≠ 0) (hf : 0 ≤ f) : ‖f ^ (q : ℝ)‖_[p] = ‖f‖_[p * q] ^ (q : ℝ) := by - refine rpow_left_injOn (NNReal.coe_ne_zero.2 hp) lpNorm_nonneg (by dsimp; positivity) ?_ + refine NNReal.rpow_left_injective (NNReal.coe_ne_zero.2 hp) ?_ dsimp - rw [← rpow_mul lpNorm_nonneg, ← mul_comm, ← ENNReal.coe_mul, ← NNReal.coe_mul, - lpNorm_rpow_eq_sum hp, lpNorm_rpow_eq_sum (mul_ne_zero hq hp)] + rw [← NNReal.rpow_mul, ← mul_comm, ← ENNReal.coe_mul, ← NNReal.coe_mul, + dLpNorm_rpow_eq_sum_nnnorm hp, dLpNorm_rpow_eq_sum_nnnorm (mul_ne_zero hq hp)] + ext simp [abs_rpow_of_nonneg (hf _), ← rpow_mul] -lemma lpNorm_pow (hp : p ≠ 0) {q : ℕ} (hq : q ≠ 0) (f : α → ℂ) : +lemma dLpNorm_pow (hp : p ≠ 0) {q : ℕ} (hq : q ≠ 0) (f : α → ℂ) : ‖f ^ q‖_[p] = ‖f‖_[p * q] ^ q := by - refine rpow_left_injOn (NNReal.coe_ne_zero.2 hp) lpNorm_nonneg (by dsimp; positivity) ?_ + refine NNReal.rpow_left_injective (NNReal.coe_ne_zero.2 hp) ?_ dsimp - rw [← rpow_natCast_mul lpNorm_nonneg, ← mul_comm, ← ENNReal.coe_natCast, ← ENNReal.coe_mul, - ← NNReal.coe_natCast, ← NNReal.coe_mul, lpNorm_rpow_eq_sum hp, lpNorm_rpow_eq_sum] - simp [← rpow_natCast_mul] + rw [← NNReal.rpow_natCast_mul, ← mul_comm, ← ENNReal.coe_natCast, ← ENNReal.coe_mul, + ← NNReal.coe_natCast, ← NNReal.coe_mul, dLpNorm_rpow_eq_sum_nnnorm hp, dLpNorm_rpow_eq_sum_nnnorm] + simp [← NNReal.rpow_natCast_mul] positivity -lemma l1Norm_rpow (hq : q ≠ 0) (hf : 0 ≤ f) : ‖f ^ (q : ℝ)‖_[1] = ‖f‖_[q] ^ (q : ℝ) := by - simpa only [ENNReal.coe_one, one_mul] using lpNorm_rpow one_ne_zero hq hf - -lemma l1Norm_pow {q : ℕ} (hq : q ≠ 0) (f : α → ℂ) : ‖f ^ q‖_[1] = ‖f‖_[q] ^ q := by - simpa only [ENNReal.coe_one, one_mul] using lpNorm_pow one_ne_zero hq f +lemma dL1Norm_rpow (hq : q ≠ 0) (hf : 0 ≤ f) : ‖f ^ (q : ℝ)‖_[1] = ‖f‖_[q] ^ (q : ℝ) := by + simpa only [ENNReal.coe_one, one_mul] using dLpNorm_rpow one_ne_zero hq hf -/-- **Hölder's inequality**, binary case. -/ -lemma l2Inner_le_lpNorm_mul_lpNorm (hpq : p.IsConjExponent q) (f g : α → ℝ) : - ⟪f, g⟫_[ℝ] ≤ ‖f‖_[p] * ‖g‖_[q] := by - have hp := hpq.ne_zero - have hq := hpq.symm.ne_zero - norm_cast at hp hq - simpa [l2Inner_eq_sum, lpNorm_eq_sum, *] using inner_le_Lp_mul_Lq _ f g hpq.coe - -/-- **Hölder's inequality**, binary case. -/ -lemma abs_l2Inner_le_lpNorm_mul_lpNorm (hpq : p.IsConjExponent q) (f g : α → ℝ) : - |⟪f, g⟫_[ℝ]| ≤ ‖f‖_[p] * ‖g‖_[q] := - abs_l2Inner_le_l2Inner_abs.trans $ - (l2Inner_le_lpNorm_mul_lpNorm hpq _ _).trans_eq $ by simp_rw [lpNorm_abs] +lemma dL1Norm_pow {q : ℕ} (hq : q ≠ 0) (f : α → ℂ) : ‖f ^ q‖_[1] = ‖f‖_[q] ^ q := by + simpa only [ENNReal.coe_one, one_mul] using dLpNorm_pow one_ne_zero hq f end Real section Hoelder -variable {α : Type*} [Fintype α] [RCLike 𝕜] {p q : ℝ≥0} {f g : α → 𝕜} +variable {α : Type*} {mα : MeasurableSpace α} [DiscreteMeasurableSpace α] [Fintype α] [RCLike 𝕜] + {p q : ℝ≥0} {f g : α → 𝕜} -lemma lpNorm_eq_l1Norm_rpow (hp : p ≠ 0) (f : α → 𝕜) : +lemma dLpNorm_eq_dL1Norm_rpow (hp : p ≠ 0) (f : α → 𝕜) : ‖f‖_[p] = ‖fun a ↦ ‖f a‖ ^ (p : ℝ)‖_[1] ^ (p⁻¹ : ℝ) := by - simp [lpNorm_eq_sum hp, l1Norm_eq_sum, abs_rpow_of_nonneg] + ext; simp [dLpNorm_eq_sum_nnnorm hp, dL1Norm_eq_sum_nnnorm, abs_rpow_of_nonneg] -lemma lpNorm_rpow' (hp : p ≠ 0) (hq : q ≠ 0) (f : α → 𝕜) : +lemma dLpNorm_rpow' (hp : p ≠ 0) (hq : q ≠ 0) (f : α → 𝕜) : ‖f‖_[p] ^ (q : ℝ) = ‖(fun a ↦ ‖f a‖) ^ (q : ℝ)‖_[p / q] := by - rw [← ENNReal.coe_div hq, lpNorm_rpow (div_ne_zero hp hq) hq (fun _ ↦ norm_nonneg _), lpNorm_norm, + rw [← ENNReal.coe_div hq, dLpNorm_rpow (div_ne_zero hp hq) hq (fun _ ↦ norm_nonneg _), dLpNorm_norm, ← ENNReal.coe_mul, div_mul_cancel₀ _ hq] -lemma norm_l2Inner_le (f g : α → 𝕜) : ‖⟪f, g⟫_[𝕜]‖ ≤ ⟪fun a ↦ ‖f a‖, fun a ↦ ‖g a‖⟫_[ℝ] := - (norm_sum_le _ _).trans $ by simp [l2Inner] - -/-- **Hölder's inequality**, binary case. -/ -lemma norm_l2Inner_le_lpNorm_mul_lpNorm (hpq : p.IsConjExponent q) (f g : α → 𝕜) : - ‖⟪f, g⟫_[𝕜]‖ ≤ ‖f‖_[p] * ‖g‖_[q] := - calc - _ ≤ ⟪fun a ↦ ‖f a‖, fun a ↦ ‖g a‖⟫_[ℝ] := norm_l2Inner_le _ _ - _ ≤ ‖fun a ↦ ‖f a‖‖_[p] * ‖fun a ↦ ‖g a‖‖_[q] := l2Inner_le_lpNorm_mul_lpNorm hpq _ _ - _ = ‖f‖_[p] * ‖g‖_[q] := by simp_rw [lpNorm_norm] - -/-- **Hölder's inequality**, binary case. -/ -lemma lpNorm_mul_le (hp : p ≠ 0) (hq : q ≠ 0) (r : ℝ≥0) (hpqr : p⁻¹ + q⁻¹ = r⁻¹) (f g : α → 𝕜) : - ‖f * g‖_[r] ≤ ‖f‖_[p] * ‖g‖_[q] := by - have hr : r ≠ 0 := by - rintro rfl - simp [hp] at hpqr - have : (‖(f * g) ·‖ ^ (r : ℝ)) = (‖f ·‖ ^ (r : ℝ)) * (‖g ·‖ ^ (r : ℝ)) := by - ext; simp [mul_rpow, abs_mul] - rw [lpNorm_eq_l1Norm_rpow, rpow_inv_le_iff_of_pos, this, l1Norm_mul_of_nonneg, - mul_rpow lpNorm_nonneg lpNorm_nonneg, lpNorm_rpow', lpNorm_rpow', ← ENNReal.coe_div, ← - ENNReal.coe_div] - refine l2Inner_le_lpNorm_mul_lpNorm ⟨?_, ?_⟩ _ _ - · norm_cast - rw [div_eq_mul_inv, ← hpqr, mul_add, mul_inv_cancel₀ hp] - exact lt_add_of_pos_right _ (by positivity) - · norm_cast - simp [div_eq_mul_inv, hpqr, ← mul_add, hr] - any_goals intro a; dsimp - all_goals positivity - -/-- **Hölder's inequality**, binary case. -/ -lemma l1Norm_mul_le (hpq : p.IsConjExponent q) (f g : α → 𝕜) : - ‖f * g‖_[1] ≤ ‖f‖_[p] * ‖g‖_[q] := - lpNorm_mul_le (mod_cast hpq.ne_zero) (mod_cast hpq.symm.ne_zero) _ - (by simpa using hpq.inv_add_inv_conj) _ _ - -/-- **Hölder's inequality**, finitary case. -/ -lemma lpNorm_prod_le {ι : Type*} {s : Finset ι} (hs : s.Nonempty) {p : ι → ℝ≥0} - (hp : ∀ i, p i ≠ 0) (q : ℝ≥0) - (hpq : ∑ i in s, (p i)⁻¹ = q⁻¹) (f : ι → α → 𝕜) : - ‖∏ i in s, f i‖_[q] ≤ ∏ i in s, ‖f i‖_[p i] := by - induction' s using Finset.cons_induction with i s hi ih generalizing q - · cases not_nonempty_empty hs - obtain rfl | hs := s.eq_empty_or_nonempty - · simp only [sum_cons, sum_empty, add_zero, inv_inj] at hpq - simp [← hpq] - simp_rw [prod_cons] - rw [sum_cons, ← inv_inv (∑ _ ∈ _, _ : ℝ≥0)] at hpq - refine - (lpNorm_mul_le (hp _) (inv_ne_zero (sum_pos (fun _ _ ↦ ?_) hs).ne') _ hpq _ _).trans - (mul_le_mul_of_nonneg_left (ih hs _ (inv_inv _).symm) lpNorm_nonneg) - exact pos_iff_ne_zero.2 (inv_ne_zero $ hp _) - end Hoelder section -variable {α : Type*} [Fintype α] +variable {α : Type*} {mα : MeasurableSpace α} @[simp] -lemma RCLike.lpNorm_coe_comp {𝕜 : Type*} [RCLike 𝕜] (p) (f : α → ℝ) : +lemma RCLike.dLpNorm_coe_comp [RCLike 𝕜] (p) (f : α → ℝ) : ‖((↑) : ℝ → 𝕜) ∘ f‖_[p] = ‖f‖_[p] := by - simp only [← lpNorm_norm _ (((↑) : ℝ → 𝕜) ∘ f), ← lpNorm_norm _ f, Function.comp_apply, + simp only [← dLpNorm_norm (((↑) : ℝ → 𝕜) ∘ f), ← dLpNorm_norm f, Function.comp_apply, RCLike.norm_ofReal, Real.norm_eq_abs] -@[simp] lemma Complex.lpNorm_coe_comp (p) (f : α → ℝ) : ‖((↑) : ℝ → ℂ) ∘ f‖_[p] = ‖f‖_[p] := - RCLike.lpNorm_coe_comp _ _ +@[simp] lemma Complex.dLpNorm_coe_comp (p) (f : α → ℝ) : ‖((↑) : ℝ → ℂ) ∘ f‖_[p] = ‖f‖_[p] := + RCLike.dLpNorm_coe_comp _ _ end +end MeasureTheory diff --git a/LeanAPAP/Prereqs/LpNorm/Discrete/Inner.lean b/LeanAPAP/Prereqs/LpNorm/Discrete/Inner.lean new file mode 100644 index 0000000000..5f2ad82336 --- /dev/null +++ b/LeanAPAP/Prereqs/LpNorm/Discrete/Inner.lean @@ -0,0 +1,265 @@ +import Mathlib.Analysis.InnerProductSpace.PiL2 +import LeanAPAP.Prereqs.LpNorm.Discrete.Defs + +/-! # Inner product -/ + +open Finset Function Real +open scoped ComplexConjugate ENNReal NNReal NNRat + +variable {ι R S : Type*} [Fintype ι] + +namespace MeasureTheory +section CommSemiring +variable [CommSemiring R] [StarRing R] [DistribSMul S R] + +/-- Inner product giving rise to the L2 norm. -/ +def dL2Inner (f g : ι → R) : R := ∑ i, conj (f i) * g i + +notation "⟪" f ", " g "⟫_[" S "]" => dL2Inner (R := S) f g + +lemma dL2Inner_eq_sum (f g : ι → R) : ⟪f, g⟫_[R] = ∑ i, conj (f i) * g i := rfl + +@[simp] lemma conj_dL2Inner (f g : ι → R) : conj ⟪f, g⟫_[R] = ⟪conj f, conj g⟫_[R] := by + simp [dL2Inner_eq_sum, map_sum] + +lemma dL2Inner_anticomm (f g : ι → R) : ⟪f, g⟫_[R] = ⟪conj g, conj f⟫_[R] := by + simp [dL2Inner_eq_sum, map_sum, mul_comm] + +@[simp] lemma dL2Inner_zero_left (g : ι → R) : ⟪0, g⟫_[R] = 0 := by simp [dL2Inner_eq_sum] +@[simp] lemma dL2Inner_zero_right (f : ι → R) : ⟪f, 0⟫_[R] = 0 := by simp [dL2Inner_eq_sum] + +@[simp] lemma dL2Inner_of_isEmpty [IsEmpty ι] (f g : ι → R) : ⟪f, g⟫_[R] = 0 := by + simp [Subsingleton.elim f 0] + +@[simp] lemma dL2Inner_const_left (a : R) (f : ι → R) : ⟪const _ a, f⟫_[R] = conj a * ∑ x, f x := by + simp only [dL2Inner_eq_sum, const_apply, mul_sum] + +@[simp] +lemma dL2Inner_const_right (f : ι → R) (a : R) : ⟪f, const _ a⟫_[R] = (∑ x, conj (f x)) * a := by + simp only [dL2Inner_eq_sum, const_apply, sum_mul] + +lemma dL2Inner_add_left (f₁ f₂ g : ι → R) : ⟪f₁ + f₂, g⟫_[R] = ⟪f₁, g⟫_[R] + ⟪f₂, g⟫_[R] := by + simp_rw [dL2Inner_eq_sum, Pi.add_apply, map_add, add_mul, sum_add_distrib] + +lemma dL2Inner_add_right (f g₁ g₂ : ι → R) : ⟪f, g₁ + g₂⟫_[R] = ⟪f, g₁⟫_[R] + ⟪f, g₂⟫_[R] := by + simp_rw [dL2Inner_eq_sum, Pi.add_apply, mul_add, sum_add_distrib] + +lemma dL2Inner_smul_left [Star S] [StarModule S R] [IsScalarTower S R R] (c : S) (f g : ι → R) : + ⟪c • f, g⟫_[R] = star c • ⟪f, g⟫_[R] := by + simp only [dL2Inner_eq_sum, Pi.smul_apply, smul_mul_assoc, smul_sum, starRingEnd_apply, star_smul] + +lemma dL2Inner_smul_right [Star S] [StarModule S R] [SMulCommClass S R R] (c : S) (f g : ι → R) : + ⟪f, c • g⟫_[R] = c • ⟪f, g⟫_[R] := by + simp only [dL2Inner_eq_sum, Pi.smul_apply, mul_smul_comm, smul_sum, starRingEnd_apply, star_smul] + +lemma smul_dL2Inner_left [InvolutiveStar S] [StarModule S R] [IsScalarTower S R R] (c : S) + (f g : ι → R) : c • ⟪f, g⟫_[R] = ⟪star c • f, g⟫_[R] := by rw [dL2Inner_smul_left, star_star] + +end CommSemiring + +section CommRing +variable [CommRing R] [StarRing R] + +@[simp] +lemma dL2Inner_neg_left (f g : ι → R) : ⟪-f, g⟫_[R] = -⟪f, g⟫_[R] := by simp [dL2Inner_eq_sum] + +@[simp] +lemma dL2Inner_neg_right (f g : ι → R) : ⟪f, -g⟫_[R] = -⟪f, g⟫_[R] := by simp [dL2Inner_eq_sum] + +lemma dL2Inner_sub_left (f₁ f₂ g : ι → R) : ⟪f₁ - f₂, g⟫_[R] = ⟪f₁, g⟫_[R] - ⟪f₂, g⟫_[R] := by + simp_rw [sub_eq_add_neg, dL2Inner_add_left, dL2Inner_neg_left] + +lemma dL2Inner_sub_right (f g₁ g₂ : ι → R) : ⟪f, g₁ - g₂⟫_[R] = ⟪f, g₁⟫_[R] - ⟪f, g₂⟫_[R] := by + simp_rw [sub_eq_add_neg, dL2Inner_add_right, dL2Inner_neg_right] + +end CommRing + +section OrderedCommSemiring +variable [OrderedCommSemiring R] [StarRing R] [StarOrderedRing R] {f g : ι → R} + +lemma dL2Inner_nonneg (hf : 0 ≤ f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫_[R] := + sum_nonneg fun _ _ ↦ mul_nonneg (star_nonneg_iff.2 $ hf _) $ hg _ + +end OrderedCommSemiring + +section LinearOrderedCommRing +variable [LinearOrderedCommRing R] [StarRing R] [TrivialStar R] {f g : ι → R} + +--TODO: Can we remove the `TrivialStar` assumption? +lemma abs_dL2Inner_le_dL2Inner_abs : |⟪f, g⟫_[R]| ≤ ⟪|f|, |g|⟫_[R] := + (abs_sum_le_sum_abs _ _).trans_eq $ + sum_congr rfl fun i _ ↦ by simp only [abs_mul, conj_trivial, Pi.abs_apply] + +end LinearOrderedCommRing + +section RCLike +variable {κ : Type*} [RCLike R] {f : ι → R} + +lemma dL2Inner_eq_inner (f g : ι → R) : + ⟪f, g⟫_[R] = inner ((WithLp.equiv 2 _).symm f) ((WithLp.equiv 2 _).symm g) := rfl + +lemma inner_eq_dL2Inner (f g : PiLp 2 fun _i : ι ↦ R) : + inner f g = ⟪WithLp.equiv 2 _ f, WithLp.equiv 2 _ g⟫_[R] := rfl + +@[simp] lemma dL2Inner_self {_ : MeasurableSpace ι} [DiscreteMeasurableSpace ι] (f : ι → R) : + ⟪f, f⟫_[R] = ((‖f‖_[2] : ℝ) : R) ^ 2 := by + simp_rw [← algebraMap.coe_pow, ← NNReal.coe_pow] + simp [dL2Norm_sq_eq_sum_nnnorm, dL2Inner_eq_sum, RCLike.conj_mul] + +lemma dL2Inner_self_of_norm_eq_one (hf : ∀ x, ‖f x‖ = 1) : ⟪f, f⟫_[R] = Fintype.card ι := by + simp [-dL2Inner_self, dL2Inner_eq_sum, RCLike.conj_mul, hf, card_univ] + +lemma linearIndependent_of_ne_zero_of_dL2Inner_eq_zero {v : κ → ι → R} (hz : ∀ k, v k ≠ 0) + (ho : Pairwise fun k l ↦ ⟪v k, v l⟫_[R] = 0) : LinearIndependent R v := by + simp_rw [dL2Inner_eq_inner] at ho + have := linearIndependent_of_ne_zero_of_inner_eq_zero ?_ ho + exacts [this, hz] + +variable {mι : MeasurableSpace ι} [DiscreteMeasurableSpace ι] + +lemma dL1Norm_mul (f g : ι → R) : ‖f * g‖_[1] = ⟪fun i ↦ ‖f i‖, fun i ↦ ‖g i‖⟫_[ℝ] := by + simp [dL2Inner_eq_sum, dL1Norm_eq_sum_nnnorm] + +end RCLike + +variable {mι : MeasurableSpace ι} [DiscreteMeasurableSpace ι] + +/-- **Cauchy-Schwarz inequality** -/ +lemma dL2Inner_le_dL2Norm_mul_dL2Norm (f g : ι → ℝ) : ⟪f, g⟫_[ℝ] ≤ ‖f‖_[2] * ‖g‖_[2] := by + simpa [dL2Norm_eq_sum_nnnorm, PiLp.norm_eq_of_L2, sqrt_eq_rpow] + using real_inner_le_norm ((WithLp.equiv 2 _).symm f) _ + +end MeasureTheory + +namespace Mathlib.Meta.Positivity +open Lean Meta Qq Function MeasureTheory + +section OrderedCommSemiring +variable [OrderedCommSemiring R] [StarRing R] [StarOrderedRing R] {f g : ι → R} + +private lemma dL2Inner_nonneg_of_pos_of_nonneg (hf : 0 < f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫_[R] := + dL2Inner_nonneg hf.le hg + +private lemma dL2Inner_nonneg_of_nonneg_of_pos (hf : 0 ≤ f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫_[R] := + dL2Inner_nonneg hf hg.le + +private lemma dL2Inner_nonneg_of_pos_of_pos (hf : 0 < f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫_[R] := + dL2Inner_nonneg hf.le hg.le + +end OrderedCommSemiring + +/-- The `positivity` extension which identifies expressions of the form `⟪f, g⟫_[R]`. -/ +@[positivity ⟪_, _⟫_[_]] def evalL2Inner : PositivityExt where eval {u R} _ _ e := do + match e with + | ~q(@dL2Inner $ι _ $instι $instring $inststar $f $g) => + let _p𝕜 ← synthInstanceQ q(OrderedCommSemiring $R) + let _p𝕜 ← synthInstanceQ q(StarOrderedRing $R) + assumeInstancesCommute + match ← core q(inferInstance) q(inferInstance) f, + ← core q(inferInstance) q(inferInstance) g with + | .positive pf, .positive pg => return .nonnegative q(dL2Inner_nonneg_of_pos_of_pos $pf $pg) + | .positive pf, .nonnegative pg => + return .nonnegative q(dL2Inner_nonneg_of_pos_of_nonneg $pf $pg) + | .nonnegative pf, .positive pg => + return .nonnegative q(dL2Inner_nonneg_of_nonneg_of_pos $pf $pg) + | .nonnegative pf, .nonnegative pg => return .nonnegative q(dL2Inner_nonneg $pf $pg) + | _, _ => return .none + | _ => throwError "not dL2Inner" + +section OrderedCommSemiring +variable [OrderedCommSemiring R] [StarRing R] [StarOrderedRing R] {f g : ι → R} + +example (hf : 0 < f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫_[R] := by positivity +example (hf : 0 < f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫_[R] := by positivity +example (hf : 0 ≤ f) (hg : 0 < g) : 0 ≤ ⟪f, g⟫_[R] := by positivity +example (hf : 0 ≤ f) (hg : 0 ≤ g) : 0 ≤ ⟪f, g⟫_[R] := by positivity + +end OrderedCommSemiring +end Mathlib.Meta.Positivity + +/-! ### Hölder inequality -/ + +namespace MeasureTheory +section Real +variable {α : Type*} {mα : MeasurableSpace α} [DiscreteMeasurableSpace α] [Fintype α] {p q : ℝ≥0} + {f g : α → ℝ} + +lemma dL1Norm_mul_of_nonneg (hf : 0 ≤ f) (hg : 0 ≤ g) : ‖f * g‖_[1] = ⟪f, g⟫_[ℝ] := by + convert dL1Norm_mul f g using 2 <;> ext a <;> refine (norm_of_nonneg ?_).symm; exacts [hf _, hg _] + +/-- **Hölder's inequality**, binary case. -/ +lemma dL2Inner_le_dLpNorm_mul_dLpNorm (hpq : p.IsConjExponent q) (f g : α → ℝ) : + ⟪f, g⟫_[ℝ] ≤ ‖f‖_[p] * ‖g‖_[q] := by + have hp := hpq.ne_zero + have hq := hpq.symm.ne_zero + norm_cast at hp hq + simpa [dL2Inner_eq_sum, dLpNorm_eq_sum_nnnorm, *] using inner_le_Lp_mul_Lq _ f g hpq.coe + +/-- **Hölder's inequality**, binary case. -/ +lemma abs_dL2Inner_le_dLpNorm_mul_dLpNorm (hpq : p.IsConjExponent q) (f g : α → ℝ) : + |⟪f, g⟫_[ℝ]| ≤ ‖f‖_[p] * ‖g‖_[q] := + abs_dL2Inner_le_dL2Inner_abs.trans $ + (dL2Inner_le_dLpNorm_mul_dLpNorm hpq _ _).trans_eq $ by simp_rw [dLpNorm_abs] + +end Real + +section Hoelder +variable {α : Type*} {mα : MeasurableSpace α} [DiscreteMeasurableSpace α] [Fintype α] [RCLike R] + {p q : ℝ≥0} {f g : α → R} + +lemma norm_dL2Inner_le (f g : α → R) : ‖⟪f, g⟫_[R]‖₊ ≤ ⟪fun a ↦ ‖f a‖, fun a ↦ ‖g a‖⟫_[ℝ] := + (norm_sum_le _ _).trans $ by simp [dL2Inner] + +/-- **Hölder's inequality**, binary case. -/ +lemma norm_dL2Inner_le_dLpNorm_mul_dLpNorm (hpq : p.IsConjExponent q) (f g : α → R) : + ‖⟪f, g⟫_[R]‖₊ ≤ ‖f‖_[p] * ‖g‖_[q] := + calc + _ ≤ ⟪fun a ↦ ‖f a‖, fun a ↦ ‖g a‖⟫_[ℝ] := norm_dL2Inner_le _ _ + _ ≤ ‖fun a ↦ ‖f a‖‖_[p] * ‖fun a ↦ ‖g a‖‖_[q] := dL2Inner_le_dLpNorm_mul_dLpNorm hpq _ _ + _ = ‖f‖_[p] * ‖g‖_[q] := by simp_rw [dLpNorm_norm] + +/-- **Hölder's inequality**, binary case. -/ +lemma dLpNorm_mul_le (hp : p ≠ 0) (hq : q ≠ 0) (r : ℝ≥0) (hpqr : p⁻¹ + q⁻¹ = r⁻¹) (f g : α → R) : + ‖f * g‖_[r] ≤ ‖f‖_[p] * ‖g‖_[q] := by + have hr : r ≠ 0 := by + rintro rfl + simp [hp] at hpqr + have : (‖(f * g) ·‖ ^ (r : ℝ)) = (‖f ·‖ ^ (r : ℝ)) * (‖g ·‖ ^ (r : ℝ)) := by + ext; simp [mul_rpow, abs_mul] + rw [dLpNorm_eq_dL1Norm_rpow, NNReal.rpow_inv_le_iff_of_pos, this, ← NNReal.coe_le_coe] + push_cast + rw [dL1Norm_mul_of_nonneg, mul_rpow, ← NNReal.coe_rpow, ← NNReal.coe_rpow, dLpNorm_rpow', + dLpNorm_rpow', ← ENNReal.coe_div, ← ENNReal.coe_div] + refine dL2Inner_le_dLpNorm_mul_dLpNorm ⟨?_, ?_⟩ _ _ + · norm_cast + rw [div_eq_mul_inv, ← hpqr, mul_add, mul_inv_cancel₀ hp] + exact lt_add_of_pos_right _ (by positivity) + · norm_cast + simp [div_eq_mul_inv, hpqr, ← mul_add, hr] + any_goals intro a; dsimp + all_goals positivity + +/-- **Hölder's inequality**, binary case. -/ +lemma dL1Norm_mul_le (hpq : p.IsConjExponent q) (f g : α → R) : + ‖f * g‖_[1] ≤ ‖f‖_[p] * ‖g‖_[q] := + dLpNorm_mul_le (mod_cast hpq.ne_zero) (mod_cast hpq.symm.ne_zero) _ + (by simpa using hpq.inv_add_inv_conj) _ _ + +/-- **Hölder's inequality**, finitary case. -/ +lemma dLpNorm_prod_le {ι : Type*} {s : Finset ι} (hs : s.Nonempty) {p : ι → ℝ≥0} + (hp : ∀ i, p i ≠ 0) (q : ℝ≥0) + (hpq : ∑ i in s, (p i)⁻¹ = q⁻¹) (f : ι → α → R) : + ‖∏ i in s, f i‖_[q] ≤ ∏ i in s, ‖f i‖_[p i] := by + induction' s using Finset.cons_induction with i s hi ih generalizing q + · cases not_nonempty_empty hs + obtain rfl | hs := s.eq_empty_or_nonempty + · simp only [sum_cons, sum_empty, add_zero, inv_inj] at hpq + simp [← hpq] + simp_rw [prod_cons] + rw [sum_cons, ← inv_inv (∑ _ ∈ _, _ : ℝ≥0)] at hpq + refine (dLpNorm_mul_le (hp _) (inv_ne_zero (sum_pos (fun _ _ ↦ ?_) hs).ne') _ hpq _ _).trans + (mul_le_mul_left' (ih hs _ (inv_inv _).symm) _) + exact pos_iff_ne_zero.2 (inv_ne_zero $ hp _) + +end Hoelder +end MeasureTheory diff --git a/LeanAPAP/Prereqs/LpNorm/Weighted.lean b/LeanAPAP/Prereqs/LpNorm/Weighted.lean index c959bfb103..2ef5a6ae1b 100644 --- a/LeanAPAP/Prereqs/LpNorm/Weighted.lean +++ b/LeanAPAP/Prereqs/LpNorm/Weighted.lean @@ -4,193 +4,141 @@ import LeanAPAP.Prereqs.LpNorm.Discrete.Basic # Lp norms -/ -open Finset Function Real +open Finset Function Real MeasureTheory open scoped ComplexConjugate ENNReal NNReal -variable {ι 𝕜 : Type*} [Fintype ι] +variable {α 𝕜 E : Type*} [MeasurableSpace α] /-! #### Weighted Lp norm -/ section NormedAddCommGroup -variable {α : ι → Type*} [∀ i, NormedAddCommGroup (α i)] {p q : ℝ≥0} {w : ι → ℝ≥0} - {f g h : ∀ i, α i} +variable [NormedAddCommGroup E] {p q : ℝ≥0∞} {w : α → ℝ≥0} {f g h : α → E} /-- The weighted Lp norm of a function. -/ -noncomputable def wlpNorm (p : ℝ≥0) (w : ι → ℝ≥0) (f : ∀ i, α i) : ℝ := - ‖fun i ↦ w i ^ (p⁻¹ : ℝ) • ‖f i‖‖_[p] +noncomputable def wLpNorm (p : ℝ≥0∞) (w : α → ℝ≥0) (f : α → E) : ℝ≥0 := + nnLpNorm f p $ .sum fun i ↦ w i • .dirac i -notation "‖" f "‖_[" p ", " w "]" => wlpNorm p w f +notation "‖" f "‖_[" p ", " w "]" => wLpNorm p w f -@[simp] -lemma wlpNorm_one_eq_lpNorm (p : ℝ≥0) (f : ∀ i, α i) : ‖f‖_[p, 1] = ‖f‖_[p] := by - simp [wlpNorm, l0Norm_eq_card, lpNorm_eq_sum, *] +@[simp] lemma wLpNorm_zero (w : α → ℝ≥0) : ‖(0 : α → E)‖_[p, w] = 0 := by simp [wLpNorm] -@[simp] -lemma wlpNorm_const_right (hp : 1 ≤ p) (w : ℝ≥0) (f : ∀ i, α i) : - ‖f‖_[p, const _ w] = w ^ (p⁻¹ : ℝ) * ‖f‖_[p] := by - simpa [wlpNorm, -norm_eq_abs, norm_of_nonneg, Pi.smul_def, NNReal.smul_def, rpow_nonneg] using - lpNorm_smul (ENNReal.one_le_coe_iff.2 hp) (w ^ (p⁻¹ : ℝ) : ℝ) fun i ↦ ‖f i‖ - -lemma wlpNorm_eq_sum (hp : p ≠ 0) (w : ι → ℝ≥0) (f : ∀ i, α i) : - ‖f‖_[p, w] = (∑ i, w i • ‖f i‖ ^ (p : ℝ)) ^ (p⁻¹ : ℝ) := by - have : (p : ℝ) ≠ 0 := by positivity - simp_rw [wlpNorm, lpNorm_eq_sum hp, NNReal.smul_def, norm_smul] - simp only [NNReal.coe_rpow, norm_norm, Algebra.id.smul_eq_mul, mul_rpow, norm_nonneg, - rpow_nonneg, hp, NNReal.coe_nonneg, norm_of_nonneg, rpow_inv_rpow _ this] - -lemma wlpNorm_eq_sum' {p : ℝ} (hp : 0 < p) (w : ι → ℝ≥0) (f : ∀ i, α i) : - ‖f‖_[p.toNNReal, w] = (∑ i, w i • ‖f i‖ ^ p) ^ p⁻¹ := by - rw [wlpNorm_eq_sum] <;> simp [hp, hp.ne', hp.le] +@[simp] lemma wLpNorm_neg (w : α → ℝ≥0) (f : α → E) : ‖-f‖_[p, w] = ‖f‖_[p, w] := by + simp [wLpNorm] -lemma wlpNorm_rpow_eq_sum {p : ℝ≥0} (hp : p ≠ 0) (w : ι → ℝ≥0) (f : ∀ i, α i) : - ‖f‖_[p, w] ^ (p : ℝ) = ∑ i, w i • ‖f i‖ ^ (p : ℝ) := by - rw [wlpNorm_eq_sum hp, rpow_inv_rpow] -- positivity - · exact sum_nonneg fun _ _ ↦ by positivity - · positivity +lemma wLpNorm_sub_comm (w : α → ℝ≥0) (f g : α → E) : ‖f - g‖_[p, w] = ‖g - f‖_[p, w] := by + simp [wLpNorm, nnLpNorm_sub_comm] -lemma wlpNorm_pow_eq_sum {p : ℕ} (hp : p ≠ 0) (w : ι → ℝ≥0) (f : ∀ i, α i) : - ‖f‖_[p, w] ^ p = ∑ i, w i • ‖f i‖ ^ p := by - simpa using wlpNorm_rpow_eq_sum (Nat.cast_ne_zero.2 hp) w f +@[simp] lemma wLpNorm_one_eq_dLpNorm (p : ℝ≥0) (f : α → E) : ‖f‖_[p, 1] = ‖f‖_[p] := by + simp [wLpNorm, dLpNorm, nnLpNorm, Measure.count] -lemma wl1Norm_eq_sum (w : ι → ℝ≥0) (f : ∀ i, α i) : ‖f‖_[1, w] = ∑ i, w i • ‖f i‖ := by - simp [wlpNorm_eq_sum] +@[simp] lemma wLpNorm_exponent_zero (w : α → ℝ≥0) (f : α → E) : ‖f‖_[0, w] = 0 := by simp [wLpNorm] -lemma wl0Norm_eq_card (w : ι → ℝ≥0) (f : ∀ i, α i) : - ‖f‖_[0, w] = {i | f i ≠ 0}.toFinite.toFinset.card := by simp [wlpNorm, l0Norm_eq_card] +@[simp] lemma wLpNorm_norm (w : α → ℝ≥0) (f : α → E) : ‖fun i ↦ ‖f i‖‖_[p, w] = ‖f‖_[p, w] := by + simp [wLpNorm] -@[simp] -lemma wlpNorm_zero (w : ι → ℝ≥0) : ‖(0 : ∀ i, α i)‖_[p, w] = 0 := by simp [wlpNorm, ← Pi.zero_def] +lemma wLpNorm_smul [NormedField 𝕜] [NormedSpace 𝕜 E] (c : 𝕜) (f : α → E) (p : ℝ≥0∞) (w : α → ℝ≥0) : + ‖c • f‖_[p, w] = ‖c‖₊ * ‖f‖_[p, w] := nnLpNorm_const_smul .. -@[simp] lemma wlpNorm_norm (w : ι → ℝ≥0) (f : ∀ i, α i) : ‖fun i ↦ ‖f i‖‖_[p, w] = ‖f‖_[p, w] := by - obtain rfl | hp := @eq_zero_or_pos _ _ p <;> simp [wl0Norm_eq_card, wlpNorm_eq_sum, *, ne_of_gt] +lemma wLpNorm_nsmul [NormedSpace ℝ E] (n : ℕ) (f : α → E) (p : ℝ≥0∞) (w : α → ℝ≥0) : + ‖n • f‖_[p, w] = n • ‖f‖_[p, w] := nnLpNorm_nsmul .. -@[simp]lemma wlpNorm_neg (w : ι → ℝ≥0) (f : ∀ i, α i) : ‖-f‖_[p, w] = ‖f‖_[p, w] := by - simp [← wlpNorm_norm _ (-f)] +section RCLike +variable {K : Type*} [RCLike K] -lemma wlpNorm_sub_comm (w : ι → ℝ≥0) (f g : ∀ i, α i) : ‖f - g‖_[p, w] = ‖g - f‖_[p, w] := by - simp [← wlpNorm_neg _ (f - g)] +@[simp] lemma wLpNorm_conj (f : α → K) : ‖conj f‖_[p, w] = ‖f‖_[p, w] := by simp [← wLpNorm_norm] -@[simp] lemma wlpNorm_nonneg : 0 ≤ ‖f‖_[p, w] := lpNorm_nonneg +variable [AddGroup α] -lemma wlpNorm_mono_right (hpq : p ≤ q) (w : ι → ℝ≥0) (f : ∀ i, α i) : ‖f‖_[p, w] ≤ ‖f‖_[q, w] := +@[simp] lemma wLpNorm_conjneg (f : α → K) : ‖conjneg f‖_[p, w] = ‖f‖_[p, w] := by + simp [← wLpNorm_norm] sorry -section one_le +end RCLike -lemma wlpNorm_add_le (hp : 1 ≤ p) (w : ι → ℝ≥0) (f g : ∀ i, α i) : - ‖f + g‖_[p, w] ≤ ‖f‖_[p, w] + ‖g‖_[p, w] := by - unfold wlpNorm - refine (lpNorm_add_le (by exact_mod_cast hp) _ _).trans' - (lpNorm_mono (fun i ↦ by dsimp; positivity) fun i ↦ ?_) - dsimp - rw [← smul_add] - exact smul_le_smul_of_nonneg_left (norm_add_le _ _) (zero_le _) +variable [Fintype α] -lemma wlpNorm_sub_le (hp : 1 ≤ p) (w : ι → ℝ≥0) (f g : ∀ i, α i) : - ‖f - g‖_[p, w] ≤ ‖f‖_[p, w] + ‖g‖_[p, w] := by - simpa [sub_eq_add_neg] using wlpNorm_add_le hp w f (-g) +@[simp] lemma wLpNorm_const_right (hp : p ≠ ∞) (w : ℝ≥0) (f : α → E) : + ‖f‖_[p, const _ w] = w ^ p.toReal⁻¹ * ‖f‖_[p] := by + simp [wLpNorm, dLpNorm, ← Finset.smul_sum, nnLpNorm_smul_measure_of_ne_top hp, Measure.count] -lemma wlpNorm_le_wlpNorm_add_wlpNorm_sub' (hp : 1 ≤ p) (w : ι → ℝ≥0) (f g : ∀ i, α i) : - ‖f‖_[p, w] ≤ ‖g‖_[p, w] + ‖f - g‖_[p, w] := by simpa using wlpNorm_add_le hp w g (f - g) +@[simp] lemma wLpNorm_smul_right (hp : p ≠ ⊤) (c : ℝ≥0) (f : α → E) : + ‖f‖_[p, c • w] = c ^ p.toReal⁻¹ * ‖f‖_[p, w] := by + simp [wLpNorm, mul_smul, ← Finset.smul_sum, nnLpNorm_smul_measure_of_ne_top hp] -lemma wlpNorm_le_wlpNorm_add_wlpNorm_sub (hp : 1 ≤ p) (w : ι → ℝ≥0) (f g : ∀ i, α i) : - ‖f‖_[p, w] ≤ ‖g‖_[p, w] + ‖g - f‖_[p, w] := by - rw [wlpNorm_sub_comm]; exact wlpNorm_le_wlpNorm_add_wlpNorm_sub' hp _ _ _ +variable [DiscreteMeasurableSpace α] -lemma wlpNorm_le_add_wlpNorm_add (hp : 1 ≤ p) (w : ι → ℝ≥0) (f g : ∀ i, α i) : - ‖f‖_[p, w] ≤ ‖f + g‖_[p, w] + ‖g‖_[p, w] := by simpa using wlpNorm_add_le hp w (f + g) (-g) +lemma wLpNorm_eq_sum (hp₀ : p ≠ 0) (hp : p ≠ ∞) (w : α → ℝ≥0) (f : α → E) : + ‖f‖_[p, w] = (∑ i, w i • ‖f i‖₊ ^ p.toReal) ^ p.toReal⁻¹ := by + ext + simp [wLpNorm, coe_nnLpNorm_eq_integral_norm_rpow_toReal hp₀ hp .of_discrete, one_div, + integral_finset_sum_measure, tsum_eq_sum' (s := univ) (by simp), ENNReal.coe_rpow_of_nonneg, + ENNReal.toReal_inv, NNReal.smul_def] -lemma wlpNorm_sub_le_wlpNorm_sub_add_wlpNorm_sub (hp : 1 ≤ p) (f g : ∀ i, α i) : - ‖f - h‖_[p, w] ≤ ‖f - g‖_[p, w] + ‖g - h‖_[p, w] := by - simpa using wlpNorm_add_le hp w (f - g) (g - h) +lemma wLpNorm_toNNReal_eq_sum {p : ℝ} (hp : 0 < p) (w : α → ℝ≥0) (f : α → E) : + ‖f‖_[p.toNNReal, w] = (∑ i, w i • ‖f i‖ ^ p) ^ p⁻¹ := by + rw [wLpNorm_eq_sum] <;> simp [hp, hp.ne', hp.le, NNReal.smul_def] -variable [NormedField 𝕜] [∀ i, NormedSpace 𝕜 (α i)] +lemma wLpNorm_rpow_eq_sum {p : ℝ≥0} (hp : p ≠ 0) (w : α → ℝ≥0) (f : α → E) : + ‖f‖_[p, w] ^ (p : ℝ) = ∑ i, w i • ‖f i‖₊ ^ (p : ℝ) := by simp [wLpNorm_eq_sum, hp] --- TODO: `p ≠ 0` is enough -lemma wlpNorm_smul (hp : 1 ≤ p) (c : 𝕜) (f : ∀ i, α i) : ‖c • f‖_[p, w] = ‖c‖ * ‖f‖_[p, w] := by - rw [wlpNorm, wlpNorm] - have : (1 : ℝ≥0∞) ≤ p := by exact_mod_cast hp - have := lpNorm_smul this ‖c‖ fun i ↦ w i ^ (p⁻¹ : ℝ) • ‖f i‖ - rw [norm_norm] at this - rw [← this] - congr with i : 1 - simp only [Pi.smul_apply, Algebra.id.smul_eq_mul, Algebra.mul_smul_comm, norm_smul] +lemma wLpNorm_pow_eq_sum {p : ℕ} (hp : p ≠ 0) (w : α → ℝ≥0) (f : α → E) : + ‖f‖_[p, w] ^ p = ∑ i, w i • ‖f i‖₊ ^ p := by + simpa using wLpNorm_rpow_eq_sum (Nat.cast_ne_zero.2 hp) w f -@[simp] -lemma wlpNorm_smul_right (hp : p ≠ 0) (c : ℝ≥0) (f : ∀ i, α i) : - ‖f‖_[p, c • w] = c ^ (p⁻¹ : ℝ) * ‖f‖_[p, w] := by - simp only [wlpNorm_eq_sum hp, NNReal.smul_def, Pi.smul_apply, Algebra.id.smul_eq_mul, - NNReal.coe_mul, mul_assoc, ← mul_sum] - exact mul_rpow (by positivity) (sum_nonneg fun _ _ ↦ by positivity) -- positivity +lemma wdL1Norm_eq_sum (w : α → ℝ≥0) (f : α → E) : ‖f‖_[1, w] = ∑ i, w i • ‖f i‖₊ := by + simp [wLpNorm_eq_sum] -variable [∀ i, NormedSpace ℝ (α i)] +lemma wLpNorm_mono_right (hpq : p ≤ q) (w : α → ℝ≥0) (f : α → E) : ‖f‖_[p, w] ≤ ‖f‖_[q, w] := + sorry -lemma wlpNorm_nsmul (hp : 1 ≤ p) (n : ℕ) (w : ι → ℝ≥0) (f : ∀ i, α i) : - ‖n • f‖_[p, w] = n • ‖f‖_[p, w] := by - rw [← Nat.cast_smul_eq_nsmul ℝ, wlpNorm_smul hp, RCLike.norm_natCast, nsmul_eq_mul] +section one_le -end one_le +lemma wLpNorm_add_le (hp : 1 ≤ p) (w : α → ℝ≥0) (f g : α → E) : + ‖f + g‖_[p, w] ≤ ‖f‖_[p, w] + ‖g‖_[p, w] := nnLpNorm_add_le .of_discrete .of_discrete hp -end NormedAddCommGroup +lemma wLpNorm_sub_le (hp : 1 ≤ p) (w : α → ℝ≥0) (f g : α → E) : + ‖f - g‖_[p, w] ≤ ‖f‖_[p, w] + ‖g‖_[p, w] := by + simpa [sub_eq_add_neg] using wLpNorm_add_le hp w f (-g) -section Real -variable {p : ℝ≥0} {w : ι → ℝ≥0} {f g : ι → ℝ} +lemma wLpNorm_le_wLpNorm_add_wLpNorm_sub' (hp : 1 ≤ p) (w : α → ℝ≥0) (f g : α → E) : + ‖f‖_[p, w] ≤ ‖g‖_[p, w] + ‖f - g‖_[p, w] := by simpa using wLpNorm_add_le hp w g (f - g) -@[simp] -lemma wlpNorm_one (hp : p ≠ 0) (w : ι → ℝ≥0) : - ‖(1 : ι → ℝ)‖_[p, w] = (∑ i, w i : ℝ) ^ (p⁻¹ : ℝ) := by - simp [wlpNorm_eq_sum hp, NNReal.smul_def] - -lemma wlpNorm_mono (hf : 0 ≤ f) (hfg : f ≤ g) : ‖f‖_[p, w] ≤ ‖g‖_[p, w] := - lpNorm_mono (fun i ↦ by dsimp; positivity) fun i ↦ smul_le_smul_of_nonneg_left - (by rw [norm_of_nonneg (hf _), norm_of_nonneg (hf.trans hfg _)]; exact hfg _) $ by positivity +lemma wLpNorm_le_wLpNorm_add_wLpNorm_sub (hp : 1 ≤ p) (w : α → ℝ≥0) (f g : α → E) : + ‖f‖_[p, w] ≤ ‖g‖_[p, w] + ‖g - f‖_[p, w] := by + rw [wLpNorm_sub_comm]; exact wLpNorm_le_wLpNorm_add_wLpNorm_sub' hp .. -end Real +lemma wLpNorm_le_add_wLpNorm_add (hp : 1 ≤ p) (w : α → ℝ≥0) (f g : α → E) : + ‖f‖_[p, w] ≤ ‖f + g‖_[p, w] + ‖g‖_[p, w] := by simpa using wLpNorm_add_le hp w (f + g) (-g) -section wlpNorm -variable {α β : Type*} [Fintype α] {p : ℝ≥0} {w : α → ℝ≥0} +lemma wLpNorm_sub_le_wLpNorm_sub_add_wLpNorm_sub (hp : 1 ≤ p) (f g : α → E) : + ‖f - h‖_[p, w] ≤ ‖f - g‖_[p, w] + ‖g - h‖_[p, w] := by + simpa using wLpNorm_add_le hp w (f - g) (g - h) -@[simp] -lemma wlpNorm_conj [RCLike β] (f : α → β) : ‖conj f‖_[p, w] = ‖f‖_[p, w] := by simp [wlpNorm] +end one_le -variable [AddCommGroup α] +end NormedAddCommGroup -@[simp] lemma wlpNorm_translate [NormedAddCommGroup β] (a : α) (f : α → β) : - ‖τ a f‖_[p, τ a w] = ‖f‖_[p, w] := - (lpNorm_translate a fun i ↦ w i ^ (p⁻¹ : ℝ) • ‖f i‖ : _) +section Real +variable [Fintype α] [DiscreteMeasurableSpace α] {p : ℝ≥0∞} {w : α → ℝ≥0} {f g : α → ℝ} @[simp] -lemma wlpNorm_conjneg [RCLike β] (f : α → β) : ‖conjneg f‖_[p] = ‖f‖_[p] := by simp [wlpNorm] - -end wlpNorm - -namespace Mathlib.Meta.Positivity -open Lean Meta Qq Function - -/-- The `positivity` extension which identifies expressions of the form `‖f‖_[p, w]`. -/ -@[positivity ‖_‖_[_, _]] def evalWLpNorm : PositivityExt where eval {u α} _ _ e := do - if let 0 := u then -- lean4#3060 means we can't combine this with the match below - match α, e with - | ~q(ℝ), ~q(@wlpNorm $ι $instι $α $instnorm $p $w $f) => - assumeInstancesCommute - return .nonnegative q(wlpNorm_nonneg) - | _ => throwError "not wlpNorm" - else - throwError "not wlpNorm" +lemma wLpNorm_one (hp₀ : p ≠ 0) (hp : p ≠ ∞) (w : α → ℝ≥0) : + ‖(1 : α → ℝ)‖_[p, w] = (∑ i, w i) ^ p.toReal⁻¹ := by simp [wLpNorm_eq_sum hp₀ hp] -section Examples -section NormedAddCommGroup -variable {α : ι → Type*} [∀ i, NormedAddCommGroup (α i)] {w : ι → ℝ≥0} {f : ∀ i, α i} +lemma wLpNorm_mono (hf : 0 ≤ f) (hfg : f ≤ g) : ‖f‖_[p, w] ≤ ‖g‖_[p, w] := + nnLpNorm_mono_real .of_discrete (by simpa [abs_of_nonneg (hf _)]) -example {p : ℝ≥0} : 0 ≤ ‖f‖_[p, w] := by positivity +end Real -end NormedAddCommGroup +section wLpNorm +variable [Fintype α] [DiscreteMeasurableSpace α] {p : ℝ≥0} {w : α → ℝ≥0} -section Complex -variable {w : ι → ℝ≥0} {f : ι → ℂ} +variable [AddCommGroup α] -example {p : ℝ≥0} : 0 ≤ ‖f‖_[p, w] := by positivity +@[simp] lemma wLpNorm_translate [NormedAddCommGroup E] (a : α) (f : α → E) : + ‖τ a f‖_[p, τ a w] = ‖f‖_[p, w] := by + obtain rfl | hp := eq_or_ne p 0 + · simp + · simp [wLpNorm_eq_sum, hp, ← sum_translate a fun x ↦ w x * ‖f x‖₊ ^ (p : ℝ)] -end Complex -end Examples -end Mathlib.Meta.Positivity +end wLpNorm diff --git a/LeanAPAP/Prereqs/NNLpNorm.lean b/LeanAPAP/Prereqs/NNLpNorm.lean new file mode 100644 index 0000000000..c4b588ed1b --- /dev/null +++ b/LeanAPAP/Prereqs/NNLpNorm.lean @@ -0,0 +1,195 @@ +import Mathlib.Analysis.Normed.Group.Constructions +import Mathlib.MeasureTheory.Integral.Bochner +import LeanAPAP.Mathlib.Analysis.SpecialFunctions.Pow.NNReal +import LeanAPAP.Mathlib.MeasureTheory.Function.LpSeminorm.Basic +import LeanAPAP.Prereqs.Expect.Basic + +open Filter +open scoped BigOperators ComplexConjugate ENNReal NNReal + +namespace MeasureTheory +variable {α E : Type*} {m : MeasurableSpace α} {p : ℝ≥0∞} {q : ℝ} {μ ν : Measure α} + [NormedAddCommGroup E] {f g h : α → E} + +/-- `ℒp` seminorm, equal to `0` for `p = 0`, to `(∫ ‖f a‖^p ∂μ) ^ (1/p)` for `0 < p < ∞` and to +`essSup ‖f‖ μ` for `p = ∞`. -/ +noncomputable def nnLpNorm (f : α → E) (p : ℝ≥0∞) (μ : Measure α) : ℝ≥0 := (eLpNorm f p μ).toNNReal + +lemma coe_nnLpNorm_eq_eLpNorm (hf : Memℒp f p μ) : nnLpNorm f p μ = eLpNorm f p μ := by + rw [nnLpNorm, ENNReal.coe_toNNReal hf.eLpNorm_ne_top] + +-- TODO: Rename `eLpNorm_eq_lintegral_rpow_nnnorm` to `eLpNorm_eq_lintegral_nnnorm_rpow_toReal` +lemma coe_nnLpNorm_eq_integral_norm_rpow_toReal (hp₀ : p ≠ 0) (hp : p ≠ ∞) + (hf : AEMeasurable (fun x ↦ (‖f x‖₊ : ℝ≥0∞) ^ p.toReal) μ) : + nnLpNorm f p μ = (∫ x, ‖f x‖ ^ p.toReal ∂μ) ^ p.toReal⁻¹ := by + rw [nnLpNorm, eLpNorm_eq_lintegral_rpow_nnnorm hp₀ hp, ENNReal.coe_toNNReal_eq_toReal, + ← ENNReal.toReal_rpow, ← integral_toReal hf] + simp [← ENNReal.toReal_rpow] + · exact .of_forall fun x ↦ ENNReal.rpow_lt_top_of_nonneg (by positivity) (by simp) + +-- TODO: Rename `coe_eLpNorm_nnreal_eq_lintegral_norm_rpow` +lemma coe_nnLpNorm_nnreal_eq_integral_norm_rpow {p : ℝ≥0} (hp : p ≠ 0) + (hf : AEMeasurable (fun x ↦ (‖f x‖₊ : ℝ≥0∞) ^ p.toReal) μ) : + nnLpNorm f p μ = (∫ x, ‖f x‖ ^ (p : ℝ) ∂μ) ^ (p⁻¹ : ℝ) := by + rw [coe_nnLpNorm_eq_integral_norm_rpow_toReal (by positivity) (by simp) hf]; simp + +lemma coe_nnLpNorm_one_eq_integral_norm (hf : AEMeasurable (fun x ↦ ‖f x‖₊) μ) : + nnLpNorm f 1 μ = ∫ x, ‖f x‖ ∂μ := by + rw [coe_nnLpNorm_eq_integral_norm_rpow_toReal one_ne_zero ENNReal.coe_ne_top (by simpa using hf)] + simp + +@[simp] lemma nnLpNorm_exponent_zero (f : α → E) : nnLpNorm f 0 μ = 0 := by simp [nnLpNorm] +@[simp] lemma nnLpNorm_measure_zero (f : α → E) : nnLpNorm f p (0 : Measure α) = 0 := by + simp [nnLpNorm] + +lemma ae_le_nnLinftyNorm (hf : Memℒp f ∞ μ) : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ nnLpNorm f ⊤ μ := by + simp_rw [← ENNReal.coe_le_coe, coe_nnLpNorm_eq_eLpNorm hf]; exact ae_le_eLpNormEssSup + +lemma nnLinftyNorm_eq_essSup (hf : Memℒp f ∞ μ) : nnLpNorm f ∞ μ = essSup (‖f ·‖₊) μ := by + refine ENNReal.coe_injective ?_ + rw [coe_nnLpNorm_eq_eLpNorm hf, ENNReal.coe_essSup, eLpNorm_exponent_top, eLpNormEssSup] + exact ⟨_, ae_le_nnLinftyNorm hf⟩ + +@[simp] lemma nnLpNorm_zero (p : ℝ≥0∞) (μ : Measure α) : + nnLpNorm (0 : α → E) p μ = 0 := by simp [nnLpNorm] + +@[simp] lemma nnLpNorm_zero' (p : ℝ≥0∞) (μ : Measure α) : + nnLpNorm (fun _ ↦ 0 : α → E) p μ = 0 := by simp [nnLpNorm] + +@[simp] +lemma nnLpNorm_eq_zero (hf : Memℒp f p μ) (hp : p ≠ 0) : nnLpNorm f p μ = 0 ↔ f =ᵐ[μ] 0 := by + simp [nnLpNorm, ENNReal.toNNReal_eq_zero_iff, hf.eLpNorm_ne_top, eLpNorm_eq_zero_iff hf.1 hp] + +@[simp] lemma nnLpNorm_of_isEmpty [IsEmpty α] (f : α → E) (p : ℝ≥0∞) : nnLpNorm f p μ = 0 := by + simp [Subsingleton.elim f 0] + +@[simp] lemma nnLpNorm_neg (f : α → E) (p : ℝ≥0∞) (μ : Measure α) : + nnLpNorm (-f) p μ = nnLpNorm f p μ := by simp [nnLpNorm] + +@[simp] lemma nnLpNorm_neg' (f : α → E) (p : ℝ≥0∞) (μ : Measure α) : + nnLpNorm (fun x ↦ -f x) p μ = nnLpNorm f p μ := nnLpNorm_neg .. + +lemma nnLpNorm_sub_comm (f g : α → E) (p : ℝ≥0∞) (μ : Measure α) : + nnLpNorm (f - g) p μ = nnLpNorm (g - f) p μ := by simp [nnLpNorm, eLpNorm_sub_comm] + +@[simp] lemma nnLpNorm_norm (f : α → E) (p : ℝ≥0∞) : + nnLpNorm (fun x ↦ ‖f x‖) p μ = nnLpNorm f p μ := by simp [nnLpNorm] + +@[simp] lemma nnLpNorm_abs (f : α → ℝ) (p : ℝ≥0∞) : nnLpNorm (|f|) p μ = nnLpNorm f p μ := + nnLpNorm_norm f p + +@[simp] lemma nnLpNorm_abs' (f : α → ℝ) (p : ℝ≥0∞) : + nnLpNorm (fun x ↦ |f x|) p μ = nnLpNorm f p μ := nnLpNorm_abs .. + +@[simp] lemma nnLpNorm_const (hp : p ≠ 0) (hμ : μ ≠ 0) (c : E) : + nnLpNorm (fun _x ↦ c) p μ = ‖c‖₊ * (μ Set.univ).toNNReal ^ p.toReal⁻¹ := by + simp [nnLpNorm, eLpNorm_const c hp hμ] + +@[simp] lemma nnLpNorm_const' (hp₀ : p ≠ 0) (hp : p ≠ ∞) (c : E) : + nnLpNorm (fun _x ↦ c) p μ = ‖c‖₊ * (μ Set.univ).toNNReal ^ p.toReal⁻¹ := by + simp [nnLpNorm, eLpNorm_const' c hp₀ hp] + +section NormedField +variable {𝕜 : Type*} [NormedField 𝕜] + +@[simp] lemma nnLpNorm_one (hp : p ≠ 0) (hμ : μ ≠ 0) : + nnLpNorm (1 : α → 𝕜) p μ = (μ Set.univ).toNNReal ^ (p.toReal⁻¹ : ℝ) := by + simp [Pi.one_def, nnLpNorm_const hp hμ] + +@[simp] lemma nnLpNorm_one' (hp₀ : p ≠ 0) (hp : p ≠ ∞) (μ : Measure α) : + nnLpNorm (1 : α → 𝕜) p μ = (μ Set.univ).toNNReal ^ (p.toReal⁻¹ : ℝ) := by + simp [Pi.one_def, nnLpNorm_const' hp₀ hp] + +lemma nnLpNorm_const_smul [Module 𝕜 E] [BoundedSMul 𝕜 E] (c : 𝕜) (f : α → E) (μ : Measure α) : + nnLpNorm (c • f) p μ = ‖c‖₊ * nnLpNorm f p μ := by simp [nnLpNorm, eLpNorm_const_smul] + +lemma nnLpNorm_nsmul [NormedSpace ℝ E] (n : ℕ) (f : α → E) (μ : Measure α) : + nnLpNorm (n • f) p μ = n • nnLpNorm f p μ := by simp [nnLpNorm, eLpNorm_nsmul] + +variable [NormedSpace ℝ 𝕜] + +lemma nnLpNorm_natCast_mul (n : ℕ) (f : α → 𝕜) (p : ℝ≥0∞) (μ : Measure α) : + nnLpNorm ((n : α → 𝕜) * f) p μ = n * nnLpNorm f p μ := by + simpa only [nsmul_eq_mul] using nnLpNorm_nsmul n f μ + +lemma nnLpNorm_natCast_mul' (n : ℕ) (f : α → 𝕜) (p : ℝ≥0∞) (μ : Measure α) : + nnLpNorm (n * f ·) p μ = n * nnLpNorm f p μ := nnLpNorm_natCast_mul .. + +lemma nnLpNorm_mul_natCast (f : α → 𝕜) (n : ℕ) (p : ℝ≥0∞) (μ : Measure α) : + nnLpNorm (f * (n : α → 𝕜)) p μ = nnLpNorm f p μ * n := by + simpa only [mul_comm] using nnLpNorm_natCast_mul n f p μ + +lemma nnLpNorm_mul_natCast' (f : α → 𝕜) (n : ℕ) (p : ℝ≥0∞) (μ : Measure α) : + nnLpNorm (f · * n) p μ = nnLpNorm f p μ * n := nnLpNorm_mul_natCast .. + +lemma nnLpNorm_div_natCast [CharZero 𝕜] {n : ℕ} (hn : n ≠ 0) (f : α → 𝕜) (p : ℝ≥0∞) + (μ : Measure α) : nnLpNorm (f / (n : α → 𝕜)) p μ = nnLpNorm f p μ / n := by + rw [eq_div_iff (by positivity), ← nnLpNorm_mul_natCast]; simp [Pi.mul_def, hn] + +lemma nnLpNorm_div_natCast' [CharZero 𝕜] {n : ℕ} (hn : n ≠ 0) (f : α → 𝕜) (p : ℝ≥0∞) + (μ : Measure α) : nnLpNorm (f · / n) p μ = nnLpNorm f p μ / n := nnLpNorm_div_natCast hn .. + +end NormedField + +lemma nnLpNorm_add_le (hf : Memℒp f p μ) (hg : Memℒp g p μ) (hp : 1 ≤ p) : + nnLpNorm (f + g) p μ ≤ nnLpNorm f p μ + nnLpNorm g p μ := by + unfold nnLpNorm + rw [← ENNReal.toNNReal_add hf.eLpNorm_ne_top hg.eLpNorm_ne_top] + gcongr + exacts [ENNReal.add_ne_top.2 ⟨hf.eLpNorm_ne_top, hg.eLpNorm_ne_top⟩, + eLpNorm_add_le hf.aestronglyMeasurable hg.aestronglyMeasurable hp] + +lemma nnLpNorm_sub_le (hf : Memℒp f p μ) (hg : Memℒp g p μ) (hp : 1 ≤ p) : + nnLpNorm (f - g) p μ ≤ nnLpNorm f p μ + nnLpNorm g p μ := by + simpa [sub_eq_add_neg] using nnLpNorm_add_le hf hg.neg hp + +lemma nnLpNorm_le_nnLpNorm_add_nnLpNorm_sub' (hf : Memℒp f p μ) (hg : Memℒp g p μ) (hp : 1 ≤ p) : + nnLpNorm f p μ ≤ nnLpNorm g p μ + nnLpNorm (f - g) p μ := by + simpa using nnLpNorm_add_le hg (hf.sub hg) hp + +lemma nnLpNorm_le_nnLpNorm_add_nnLpNorm_sub (hf : Memℒp f p μ) (hg : Memℒp g p μ) (hp : 1 ≤ p) : + nnLpNorm f p μ ≤ nnLpNorm g p μ + nnLpNorm (g - f) p μ := by + simpa [neg_add_eq_sub] using nnLpNorm_add_le hg.neg (hg.sub hf) hp + +lemma nnLpNorm_le_add_nnLpNorm_add (hf : Memℒp f p μ) (hg : Memℒp g p μ) (hp : 1 ≤ p) : + nnLpNorm f p μ ≤ nnLpNorm (f + g) p μ + nnLpNorm g p μ := by + simpa using nnLpNorm_add_le (hf.add hg) hg.neg hp + +lemma nnLpNorm_sub_le_nnLpNorm_sub_add_nnLpNorm_sub (hf : Memℒp f p μ) (hg : Memℒp g p μ) + (hh : Memℒp h p μ) (hp : 1 ≤ p) : + nnLpNorm (f - h) p μ ≤ nnLpNorm (f - g) p μ + nnLpNorm (g - h) p μ := by + simpa using nnLpNorm_add_le (hf.sub hg) (hg.sub hh) hp + +lemma nnLpNorm_sum_le {ι : Type*} {s : Finset ι} {f : ι → α → E} (hf : ∀ i ∈ s, Memℒp (f i) p μ) + (hp : 1 ≤ p) : nnLpNorm (∑ i ∈ s, f i) p μ ≤ ∑ i ∈ s, nnLpNorm (f i) p μ := by + unfold nnLpNorm + rw [← ENNReal.toNNReal_sum fun i hi ↦ (hf _ hi).eLpNorm_ne_top] + gcongr + exacts [ENNReal.sum_ne_top.2 fun i hi ↦ (hf _ hi).eLpNorm_ne_top, + eLpNorm_sum_le (fun i hi ↦ (hf _ hi).aestronglyMeasurable) hp] + +lemma nnLpNorm_expect_le [Module ℚ≥0 E] [NormedSpace ℝ E] {ι : Type*} {s : Finset ι} + {f : ι → α → E} (hf : ∀ i ∈ s, Memℒp (f i) p μ) (hp : 1 ≤ p) : + nnLpNorm (𝔼 i ∈ s, f i) p μ ≤ 𝔼 i ∈ s, nnLpNorm (f i) p μ := by + obtain rfl | hs := s.eq_empty_or_nonempty + · simp + refine (le_inv_smul_iff_of_pos $ by positivity).2 ?_ + rw [Nat.cast_smul_eq_nsmul, ← nnLpNorm_nsmul, Finset.card_smul_expect] + exact nnLpNorm_sum_le hf hp + +lemma nnLpNorm_mono_real {g : α → ℝ} (hg : Memℒp g p μ) (h : ∀ x, ‖f x‖ ≤ g x) : + nnLpNorm f p μ ≤ nnLpNorm g p μ := + ENNReal.toNNReal_mono (hg.eLpNorm_ne_top) (eLpNorm_mono_real h) + +lemma nnLpNorm_smul_measure_of_ne_zero {f : α → E} {c : ℝ≥0} (hc : c ≠ 0) : + nnLpNorm f p (c • μ) = c ^ p.toReal⁻¹ • nnLpNorm f p μ := by + simp [nnLpNorm, eLpNorm_smul_measure_of_ne_zero'' hc f p μ] + +lemma nnLpNorm_smul_measure_of_ne_top (hp : p ≠ ∞) {f : α → E} (c : ℝ≥0) : + nnLpNorm f p (c • μ) = c ^ p.toReal⁻¹ • nnLpNorm f p μ := by + simp [nnLpNorm, eLpNorm_smul_measure_of_ne_top'' hp] + +@[simp] lemma nnLpNorm_conj {K : Type*} [RCLike K] (f : α → K) (p : ℝ≥0∞) (μ : Measure α) : + nnLpNorm (conj f) p μ = nnLpNorm f p μ := by simp [← nnLpNorm_norm] + +end MeasureTheory diff --git a/LeanAPAP/Prereqs/Rudin.lean b/LeanAPAP/Prereqs/Rudin.lean index bb1d3b320b..a3ebe86ede 100644 --- a/LeanAPAP/Prereqs/Rudin.lean +++ b/LeanAPAP/Prereqs/Rudin.lean @@ -73,7 +73,7 @@ lemma rudin_exp_ineq (f : α → ℂ) (hf : AddDissociated $ support $ cft f) : AddDissociated.randomisation _ _ $ by simpa [-Complex.ofReal_sinh, hc₀] _ ≤ ∏ ψ, exp (‖cft f ψ‖ ^ 2 / 2) := prod_le_prod (fun _ _ ↦ by positivity) fun _ _ ↦ cosh_le_exp_half_sq _ - _ = _ := by simp_rw [← exp_sum, ← sum_div, ← l2Norm_cft, l2Norm_sq_eq_sum] + _ = _ := by simp_rw [← exp_sum, ← sum_div, ← dL2Norm_cft, dL2Norm_sq_eq_sum_nnnorm] /-- **Rudin's inequality**, exponential form with absolute values. -/ lemma rudin_exp_abs_ineq (f : α → ℂ) (hf : AddDissociated $ support $ cft f) : @@ -94,7 +94,7 @@ private lemma rudin_ineq_aux (hp : 2 ≤ p) (f : α → ℂ) (hf : AddDissociate · rwa [cft_smul, support_const_smul_of_ne_zero] positivity simp_rw [Function.comp_def, Pi.smul_apply, Complex.smul_re, ← Pi.smul_def] at H - rw [nlpNorm_smul, nlpNorm_smul, norm_div, norm_of_nonneg, norm_of_nonneg, mul_left_comm, + rw [cLpNorm_smul, cLpNorm_smul, norm_div, norm_of_nonneg, norm_of_nonneg, mul_left_comm, mul_le_mul_left] at H refine H ?_ rw [div_mul_cancel₀] @@ -106,7 +106,7 @@ private lemma rudin_ineq_aux (hp : 2 ≤ p) (f : α → ℂ) (hf : AddDissociate have : (‖re ∘ f‖ₙ_[↑p] / p) ^ p ≤ (2 * exp 2⁻¹) ^ p := by calc _ = 𝔼 a, |(f a).re| ^ p / p ^ p := by - simp [div_pow, nlpNorm_pow_eq_expect hp₀]; rw [expect_div] + simp [div_pow, cLpNorm_pow_eq_expect_nnnorm hp₀]; rw [expect_div] _ ≤ 𝔼 a, |(f a).re| ^ p / p ! := by gcongr; norm_cast; exact p.factorial_le_pow _ ≤ 𝔼 a, exp |(f a).re| := by gcongr; exact pow_div_factorial_le_exp _ (abs_nonneg _) _ _ ≤ _ := rudin_exp_abs_ineq f hf @@ -125,10 +125,10 @@ lemma rudin_ineq (hp : 2 ≤ p) (f : α → ℂ) (hf : AddDissociated $ support ‖f‖ₙ_[p] = ‖(fun a ↦ ((f a).re : ℂ)) + I • (fun a ↦ ((f a).im : ℂ))‖ₙ_[p] := by congr with a; simp [mul_comm I] _ ≤ ‖fun a ↦ ((f a).re : ℂ)‖ₙ_[p] + ‖I • (fun a ↦ ((f a).im : ℂ))‖ₙ_[p] - := nlpNorm_add_le hp₁ _ _ + := cLpNorm_add_le hp₁ _ _ _ = ‖re ∘ f‖ₙ_[p] + ‖re ∘ ((-I) • f)‖ₙ_[p] := by - rw [nlpNorm_smul hp₁, Complex.norm_I, one_mul, ← Complex.nlpNorm_coe_comp, - ← Complex.nlpNorm_coe_comp] + rw [cLpNorm_smul hp₁, Complex.norm_I, one_mul, ← Complex.cLpNorm_coe_comp, + ← Complex.cLpNorm_coe_comp] congr ext a : 1 simp @@ -136,4 +136,4 @@ lemma rudin_ineq (hp : 2 ≤ p) (f : α → ℂ) (hf : AddDissociated $ support := add_le_add (rudin_ineq_aux hp _ hf) $ rudin_ineq_aux hp _ $ by rwa [cft_smul, support_const_smul_of_ne_zero]; simp _ = 4 * exp 2⁻¹ * sqrt p * ‖f‖ₙ_[2] := by - rw [nlpNorm_smul one_le_two, norm_neg, Complex.norm_I, one_mul]; ring + rw [cLpNorm_smul one_le_two, norm_neg, Complex.norm_I, one_mul]; ring diff --git a/blueprint/src/chapter/ff.tex b/blueprint/src/chapter/ff.tex index 272719c18b..33dd7629ac 100644 --- a/blueprint/src/chapter/ff.tex +++ b/blueprint/src/chapter/ff.tex @@ -172,7 +172,7 @@ \chapter{Finite field model} \begin{lemma} \label{no3aps_inner_prod} -\lean{AddSalemSpencer.l2Inner_mu_conv_mu_mu_two_smul_mu} +\lean{AddSalemSpencer.dL2Inner_mu_conv_mu_mu_two_smul_mu} \leanok If $A\subseteq G$ has no non-trivial three-term arithmetic progressions and $G$ has odd order then \[\langle \mu_A\ast \mu_A,\mu_{2\cdot A}\rangle = 1/\abs{A}^2.\]