Skip to content

Commit

Permalink
feat: add Integrable.piecewise (#8080)
Browse files Browse the repository at this point in the history
  • Loading branch information
RemyDegenne committed Nov 7, 2023
1 parent 1c4e5ae commit 18d2099
Show file tree
Hide file tree
Showing 4 changed files with 86 additions and 0 deletions.
29 changes: 29 additions & 0 deletions Mathlib/MeasureTheory/Function/LpSeminorm.lean
Expand Up @@ -176,6 +176,7 @@ theorem snorm'_exponent_zero {f : α → F} : snorm' f 0 μ = 1 := by
theorem snorm_exponent_zero {f : α → F} : snorm f 0 μ = 0 := by simp [snorm]
#align measure_theory.snorm_exponent_zero MeasureTheory.snorm_exponent_zero

@[simp]
theorem memℒp_zero_iff_aestronglyMeasurable {f : α → E} : Memℒp f 0 μ ↔ AEStronglyMeasurable f μ :=
by simp [Memℒp, snorm_exponent_zero]
#align measure_theory.mem_ℒp_zero_iff_ae_strongly_measurable MeasureTheory.memℒp_zero_iff_aestronglyMeasurable
Expand Down Expand Up @@ -908,6 +909,34 @@ theorem meas_snormEssSup_lt {f : α → F} : μ { y | snormEssSup f μ < ‖f y
meas_essSup_lt
#align measure_theory.meas_snorm_ess_sup_lt MeasureTheory.meas_snormEssSup_lt

lemma snormEssSup_piecewise_le {s : Set α} (f g : α → E) [DecidablePred (· ∈ s)]
(hs : MeasurableSet s) :
snormEssSup (Set.piecewise s f g) μ
≤ max (snormEssSup f (μ.restrict s)) (snormEssSup g (μ.restrict sᶜ)) := by
refine essSup_le_of_ae_le (max (snormEssSup f (μ.restrict s)) (snormEssSup g (μ.restrict sᶜ))) ?_
have hf : ∀ᵐ y ∂(μ.restrict s), ↑‖f y‖₊ ≤ snormEssSup f (μ.restrict s) :=
ae_le_snormEssSup (μ := μ.restrict s) (f := f)
have hg : ∀ᵐ y ∂(μ.restrict sᶜ), ↑‖g y‖₊ ≤ snormEssSup g (μ.restrict sᶜ) :=
ae_le_snormEssSup (μ := μ.restrict sᶜ) (f := g)
refine ae_of_ae_restrict_of_ae_restrict_compl s ?_ ?_
· rw [ae_restrict_iff' hs] at hf ⊢
filter_upwards [hf] with x hx
intro hx_mem
simp only [hx_mem, Set.piecewise_eq_of_mem]
exact (hx hx_mem).trans (le_max_left _ _)
· rw [ae_restrict_iff' hs.compl] at hg ⊢
filter_upwards [hg] with x hx
intro hx_mem
rw [Set.mem_compl_iff] at hx_mem
simp only [hx_mem, not_false_eq_true, Set.piecewise_eq_of_not_mem]
exact (hx hx_mem).trans (le_max_right _ _)

lemma snorm_top_piecewise_le {s : Set α} (f g : α → E) [DecidablePred (· ∈ s)]
(hs : MeasurableSet s) :
snorm (Set.piecewise s f g) ∞ μ
≤ max (snorm f ∞ (μ.restrict s)) (snorm g ∞ (μ.restrict sᶜ)) :=
snormEssSup_piecewise_le f g hs

section MapMeasure

variable {β : Type*} {mβ : MeasurableSpace β} {f : α → β} {g : β → E}
Expand Down
29 changes: 29 additions & 0 deletions Mathlib/MeasureTheory/Function/LpSpace.lean
Expand Up @@ -732,6 +732,35 @@ theorem exists_snorm_indicator_le (hp : p ≠ ∞) (c : E) {ε : ℝ≥0∞} (h
exact mul_le_mul_left' (ENNReal.rpow_le_rpow hs hp₀') _
#align measure_theory.exists_snorm_indicator_le MeasureTheory.exists_snorm_indicator_le

lemma Memℒp.piecewise [DecidablePred (· ∈ s)]
(hs : MeasurableSet s) (hf : Memℒp f p (μ.restrict s)) (hg : Memℒp g p (μ.restrict sᶜ)) :
Memℒp (s.piecewise f g) p μ := by
by_cases hp_zero : p = 0
· simp only [hp_zero, memℒp_zero_iff_aestronglyMeasurable]
exact AEStronglyMeasurable.piecewise hs hf.1 hg.1
refine ⟨AEStronglyMeasurable.piecewise hs hf.1 hg.1, ?_⟩
by_cases hp_top : p = ∞
· have hf2 := hf.2
have hg2 := hg.2
simp only [hp_top] at hf2 hg2 ⊢
exact (snorm_top_piecewise_le f g hs).trans_lt (max_lt_iff.mpr ⟨hf2, hg2⟩)
rw [snorm_lt_top_iff_lintegral_rpow_nnnorm_lt_top hp_zero hp_top, ← lintegral_add_compl _ hs,
ENNReal.add_lt_top]
constructor
· have h : ∀ᵐ (x : α) ∂μ, x ∈ s →
(‖Set.piecewise s f g x‖₊ : ℝ≥0∞) ^ p.toReal = (‖f x‖₊ : ℝ≥0∞) ^ p.toReal := by
refine ae_of_all _ (fun a ha ↦ ?_)
simp [ha]
rw [set_lintegral_congr_fun hs h]
exact lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top hp_zero hp_top hf.2
· have h : ∀ᵐ (x : α) ∂μ, x ∈ sᶜ →
(‖Set.piecewise s f g x‖₊ : ℝ≥0∞) ^ p.toReal = (‖g x‖₊ : ℝ≥0∞) ^ p.toReal := by
refine ae_of_all _ (fun a ha ↦ ?_)
have ha' : a ∉ s := ha
simp [ha']
rw [set_lintegral_congr_fun hs.compl h]
exact lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top hp_zero hp_top hg.2

end Indicator

section IndicatorConstLp
Expand Down
21 changes: 21 additions & 0 deletions Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean
Expand Up @@ -1720,6 +1720,27 @@ theorem _root_.exists_stronglyMeasurable_limit_of_tendsto_ae [PseudoMetrizableSp
rwa [h'x] at hx
#align exists_strongly_measurable_limit_of_tendsto_ae exists_stronglyMeasurable_limit_of_tendsto_ae

theorem piecewise {s : Set α} [DecidablePred (· ∈ s)]
(hs : MeasurableSet s) (hf : AEStronglyMeasurable f (μ.restrict s))
(hg : AEStronglyMeasurable g (μ.restrict sᶜ)) :
AEStronglyMeasurable (s.piecewise f g) μ := by
refine ⟨s.piecewise (hf.mk f) (hg.mk g),
StronglyMeasurable.piecewise hs hf.stronglyMeasurable_mk hg.stronglyMeasurable_mk, ?_⟩
refine ae_of_ae_restrict_of_ae_restrict_compl s ?_ ?_
· have h := hf.ae_eq_mk
rw [Filter.EventuallyEq, ae_restrict_iff' hs] at h
rw [ae_restrict_iff' hs]
filter_upwards [h] with x hx
intro hx_mem
simp only [hx_mem, Set.piecewise_eq_of_mem, hx hx_mem]
· have h := hg.ae_eq_mk
rw [Filter.EventuallyEq, ae_restrict_iff' hs.compl] at h
rw [ae_restrict_iff' hs.compl]
filter_upwards [h] with x hx
intro hx_mem
rw [Set.mem_compl_iff] at hx_mem
simp only [hx_mem, not_false_eq_true, Set.piecewise_eq_of_not_mem, hx hx_mem]

theorem sum_measure [PseudoMetrizableSpace β] {m : MeasurableSpace α} {μ : ι → Measure α}
(h : ∀ i, AEStronglyMeasurable f (μ i)) : AEStronglyMeasurable f (Measure.sum μ) := by
borelize β
Expand Down
7 changes: 7 additions & 0 deletions Mathlib/MeasureTheory/Integral/IntegrableOn.lean
Expand Up @@ -169,6 +169,13 @@ theorem IntegrableOn.restrict (h : IntegrableOn f s μ) (hs : MeasurableSet s) :
rw [IntegrableOn, Measure.restrict_restrict hs]; exact h.mono_set (inter_subset_left _ _)
#align measure_theory.integrable_on.restrict MeasureTheory.IntegrableOn.restrict

lemma Integrable.piecewise [DecidablePred (· ∈ s)]
(hs : MeasurableSet s) (hf : IntegrableOn f s μ) (hg : IntegrableOn g sᶜ μ) :
Integrable (s.piecewise f g) μ := by
rw [IntegrableOn] at hf hg
rw [← memℒp_one_iff_integrable] at hf hg ⊢
exact Memℒp.piecewise hs hf hg

theorem IntegrableOn.left_of_union (h : IntegrableOn f (s ∪ t) μ) : IntegrableOn f s μ :=
h.mono_set <| subset_union_left _ _
#align measure_theory.integrable_on.left_of_union MeasureTheory.IntegrableOn.left_of_union
Expand Down

0 comments on commit 18d2099

Please sign in to comment.