Skip to content

Commit

Permalink
feat(measure_theory/integral/lebesgue): lintegral is strictly monoton…
Browse files Browse the repository at this point in the history
…e under some conditions (#9373)
  • Loading branch information
JasonKYi committed Sep 25, 2021
1 parent eba2b2e commit 168806c
Show file tree
Hide file tree
Showing 2 changed files with 54 additions and 0 deletions.
46 changes: 46 additions & 0 deletions src/measure_theory/integral/lebesgue.lean
Expand Up @@ -1586,6 +1586,52 @@ begin
exact ennreal.sub_add_cancel_of_le hx
end

lemma lintegral_sub_le (f g : α → ℝ≥0∞)
(hf : measurable f) (hg : measurable g) (h : f ≤ᵐ[μ] g) :
∫⁻ x, g x ∂μ - ∫⁻ x, f x ∂μ ≤ ∫⁻ x, g x - f x ∂μ :=
begin
by_cases hfi : ∫⁻ x, f x ∂μ = ∞,
{ rw [hfi, ennreal.sub_top],
exact bot_le },
{ rw lintegral_sub hg hf hfi h,
refl' }
end

lemma lintegral_strict_mono_of_ae_le_of_ae_lt_on {f g : α → ℝ≥0∞}
(hf : measurable f) (hg : measurable g) (hfi : ∫⁻ x, f x ∂μ ≠ ∞) (h_le : f ≤ᵐ[μ] g)
{s : set α} (hμs : μ s ≠ 0) (h : ∀ᵐ x ∂μ, x ∈ s → f x < g x) :
∫⁻ x, f x ∂μ < ∫⁻ x, g x ∂μ :=
begin
rw [← ennreal.sub_pos, ← lintegral_sub hg hf hfi h_le],
by_contra hnlt,
rw [not_lt, nonpos_iff_eq_zero, lintegral_eq_zero_iff (hg.sub hf), filter.eventually_eq] at hnlt,
simp only [ae_iff, ennreal.sub_eq_zero_iff_le, pi.zero_apply, not_lt, not_le] at hnlt h,
refine hμs _,
push_neg at h,
have hs_eq : s = {a : α | a ∈ s ∧ g a ≤ f a} ∪ {a : α | a ∈ s ∧ f a < g a},
{ ext1 x,
simp_rw [set.mem_union, set.mem_set_of_eq, ← not_le],
tauto, },
rw hs_eq,
refine measure_union_null h (measure_mono_null _ hnlt),
simp,
end

lemma lintegral_strict_mono {f g : α → ℝ≥0∞} (hμ : μ ≠ 0)
(hf : measurable f) (hg : measurable g) (hfi : ∫⁻ x, f x ∂μ ≠ ∞) (h : ∀ᵐ x ∂μ, f x < g x) :
∫⁻ x, f x ∂μ < ∫⁻ x, g x ∂μ :=
begin
rw [ne.def, ← measure.measure_univ_eq_zero] at hμ,
refine lintegral_strict_mono_of_ae_le_of_ae_lt_on hf hg hfi (ae_le_of_ae_lt h) hμ _,
simpa using h,
end

lemma set_lintegral_strict_mono {f g : α → ℝ≥0∞} {s : set α}
(hsm : measurable_set s) (hs : μ s ≠ 0) (hf : measurable f) (hg : measurable g)
(hfi : ∫⁻ x in s, f x ∂μ ≠ ∞) (h : ∀ᵐ x ∂μ, x ∈ s → f x < g x) :
∫⁻ x in s, f x ∂μ < ∫⁻ x in s, g x ∂μ :=
lintegral_strict_mono (by simp [hs]) hf hg hfi ((ae_restrict_iff' hsm).mpr h)

/-- Monotone convergence theorem for nonincreasing sequences of functions -/
lemma lintegral_infi_ae
{f : ℕ → α → ℝ≥0∞} (h_meas : ∀n, measurable (f n))
Expand Down
8 changes: 8 additions & 0 deletions src/measure_theory/measure/measure_space_def.lean
Expand Up @@ -347,6 +347,14 @@ lemma ae_eq_trans {f g h: α → δ} (h₁ : f =ᵐ[μ] g) (h₂ : g =ᵐ[μ] h)
f =ᵐ[μ] h :=
h₁.trans h₂

lemma ae_le_of_ae_lt {f g : α → ℝ≥0∞} (h : ∀ᵐ x ∂μ, f x < g x) : f ≤ᵐ[μ] g :=
begin
rw [filter.eventually_le, ae_iff],
rw ae_iff at h,
refine measure_mono_null (λ x hx, _) h,
exact not_lt.2 (le_of_lt (not_le.1 hx)),
end

@[simp] lemma ae_eq_empty : s =ᵐ[μ] (∅ : set α) ↔ μ s = 0 :=
eventually_eq_empty.trans $ by simp [ae_iff]

Expand Down

0 comments on commit 168806c

Please sign in to comment.