Skip to content

Commit

Permalink
feat(measure_theory/lp_space): add snorm_le_snorm_mul_rpow_measure_un…
Browse files Browse the repository at this point in the history
…iv (#7926)

There were already versions of this lemma for `snorm'` and `snorm_ess_sup`. The new lemma collates these into a statement about `snorm`.
  • Loading branch information
RemyDegenne committed Jun 14, 2021
1 parent e041dbe commit 8377a1f
Showing 1 changed file with 29 additions and 30 deletions.
59 changes: 29 additions & 30 deletions src/measure_theory/lp_space.lean
Expand Up @@ -526,6 +526,34 @@ begin
rwa lintegral_const at h_le,
end

lemma snorm_le_snorm_mul_rpow_measure_univ {p q : ℝ≥0∞} (hpq : p ≤ q) {f : α → E}
(hf : ae_measurable f μ) :
snorm f p μ ≤ snorm f q μ * (μ set.univ) ^ (1/p.to_real - 1/q.to_real) :=
begin
by_cases hp0 : p = 0,
{ simp [hp0, zero_le], },
rw ← ne.def at hp0,
have hp0_lt : 0 < p, from lt_of_le_of_ne (zero_le _) hp0.symm,
have hq0_lt : 0 < q, from lt_of_lt_of_le hp0_lt hpq,
by_cases hq_top : q = ∞,
{ simp only [hq_top, div_zero, one_div, ennreal.top_to_real, sub_zero, snorm_exponent_top,
inv_zero],
by_cases hp_top : p = ∞,
{ simp only [hp_top, ennreal.rpow_zero, mul_one, ennreal.top_to_real, sub_zero, inv_zero,
snorm_exponent_top],
exact le_rfl, },
rw snorm_eq_snorm' hp0 hp_top,
have hp_pos : 0 < p.to_real, from ennreal.to_real_pos_iff.mpr ⟨hp0_lt, hp_top⟩,
refine (snorm'_le_snorm_ess_sup_mul_rpow_measure_univ hp_pos).trans (le_of_eq _),
congr,
exact one_div _, },
have hp_lt_top : p < ∞, from hpq.trans_lt (lt_top_iff_ne_top.mpr hq_top),
have hp_pos : 0 < p.to_real, from ennreal.to_real_pos_iff.mpr ⟨hp0_lt, hp_lt_top.ne⟩,
rw [snorm_eq_snorm' hp0_lt.ne.symm hp_lt_top.ne, snorm_eq_snorm' hq0_lt.ne.symm hq_top],
have hpq_real : p.to_real ≤ q.to_real, by rwa ennreal.to_real_le_to_real hp_lt_top.ne hq_top,
exact snorm'_le_snorm'_mul_rpow_measure_univ hp_pos hpq_real hf,
end

lemma snorm'_le_snorm'_of_exponent_le {p q : ℝ} (hp0_lt : 0 < p) (hpq : p ≤ q) (μ : measure α)
[probability_measure μ] {f : α → E} (hf : ae_measurable f μ) :
snorm' f p μ ≤ snorm' f q μ :=
Expand All @@ -541,36 +569,7 @@ le_trans (snorm'_le_snorm_ess_sup_mul_rpow_measure_univ hq_pos) (le_of_eq (by si
lemma snorm_le_snorm_of_exponent_le {p q : ℝ≥0∞} (hpq : p ≤ q) [probability_measure μ]
{f : α → E} (hf : ae_measurable f μ) :
snorm f p μ ≤ snorm f q μ :=
begin
by_cases hp0 : p = 0,
{ simp [hp0], },
rw ←ne.def at hp0,
by_cases hq_top : q = ∞,
{ by_cases hp_top : p = ∞,
{ rw [hq_top, hp_top],
exact le_refl _, },
{ have hp_pos : 0 < p.to_real,
from ennreal.to_real_pos_iff.mpr ⟨lt_of_le_of_ne (zero_le _) hp0.symm, hp_top⟩,
rw [snorm_eq_snorm' hp0 hp_top, hq_top, snorm_exponent_top],
refine le_trans (snorm'_le_snorm_ess_sup_mul_rpow_measure_univ hp_pos) (le_of_eq _),
simp [measure_univ], }, },
{ have hp_top : p ≠ ∞,
{ by_contra hp_eq_top,
push_neg at hp_eq_top,
refine hq_top _,
rwa [hp_eq_top, top_le_iff] at hpq, },
have hp_pos : 0 < p.to_real,
from ennreal.to_real_pos_iff.mpr ⟨lt_of_le_of_ne (zero_le _) hp0.symm, hp_top⟩,
have hq0 : q ≠ 0,
{ by_contra hq_eq_zero,
push_neg at hq_eq_zero,
have hp_eq_zero : p = 0, from le_antisymm (by rwa hq_eq_zero at hpq) (zero_le _),
rw [hp_eq_zero, ennreal.zero_to_real] at hp_pos,
exact (lt_irrefl _) hp_pos, },
have hpq_real : p.to_real ≤ q.to_real, by rwa ennreal.to_real_le_to_real hp_top hq_top,
rw [snorm_eq_snorm' hp0 hp_top, snorm_eq_snorm' hq0 hq_top],
exact snorm'_le_snorm'_of_exponent_le hp_pos hpq_real _ hf, },
end
(snorm_le_snorm_mul_rpow_measure_univ hpq hf).trans (le_of_eq (by simp [measure_univ]))

lemma snorm'_lt_top_of_snorm'_lt_top_of_exponent_le {p q : ℝ} [finite_measure μ] {f : α → E}
(hf : ae_measurable f μ) (hfq_lt_top : snorm' f q μ < ∞) (hp_nonneg : 0 ≤ p) (hpq : p ≤ q) :
Expand Down

0 comments on commit 8377a1f

Please sign in to comment.