Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat: tendsto_of_integral_tendsto_of_monotone #11167

Closed
wants to merge 8 commits into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
81 changes: 81 additions & 0 deletions Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2162,6 +2162,61 @@ instance : MeasurableSMul ℝ≥0 ℝ≥0∞ where
simp_rw [ENNReal.smul_def]
exact measurable_coe_nnreal_ennreal.mul_const _

/-- A limit (over a general filter) of measurable `ℝ≥0∞` valued functions is measurable. -/
theorem measurable_of_tendsto' {ι : Type*} {f : ι → α → ℝ≥0∞} {g : α → ℝ≥0∞} (u : Filter ι)
[NeBot u] [IsCountablyGenerated u] (hf : ∀ i, Measurable (f i)) (lim : Tendsto f u (𝓝 g)) :
Measurable g := by
rcases u.exists_seq_tendsto with ⟨x, hx⟩
rw [tendsto_pi_nhds] at lim
have : (fun y => liminf (fun n => (f (x n) y : ℝ≥0∞)) atTop) = g := by
ext1 y
exact ((lim y).comp hx).liminf_eq
rw [← this]
show Measurable fun y => liminf (fun n => (f (x n) y : ℝ≥0∞)) atTop
exact measurable_liminf fun n => hf (x n)
#align measurable_of_tendsto_ennreal' ENNReal.measurable_of_tendsto'

-- 2024-03-09
@[deprecated] alias _root_.measurable_of_tendsto_ennreal' := ENNReal.measurable_of_tendsto'
RemyDegenne marked this conversation as resolved.
Show resolved Hide resolved

/-- A sequential limit of measurable `ℝ≥0∞` valued functions is measurable. -/
theorem measurable_of_tendsto {f : ℕ → α → ℝ≥0∞} {g : α → ℝ≥0∞} (hf : ∀ i, Measurable (f i))
(lim : Tendsto f atTop (𝓝 g)) : Measurable g :=
measurable_of_tendsto' atTop hf lim
#align measurable_of_tendsto_ennreal ENNReal.measurable_of_tendsto

-- 2024-03-09
@[deprecated] alias _root_.measurable_of_tendsto_ennreal := ENNReal.measurable_of_tendsto

/-- A limit (over a general filter) of a.e.-measurable `ℝ≥0∞` valued functions is
a.e.-measurable. -/
lemma aemeasurable_of_tendsto' {ι : Type*} {f : ι → α → ℝ≥0∞} {g : α → ℝ≥0∞}
{μ : Measure α} (u : Filter ι) [NeBot u] [IsCountablyGenerated u]
(hf : ∀ i, AEMeasurable (f i) μ) (hlim : ∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) u (𝓝 (g a))) :
AEMeasurable g μ := by
rcases u.exists_seq_tendsto with ⟨v, hv⟩
have h'f : ∀ n, AEMeasurable (f (v n)) μ := fun n ↦ hf (v n)
set p : α → (ℕ → ℝ≥0∞) → Prop := fun x f' ↦ Tendsto f' atTop (𝓝 (g x))
have hp : ∀ᵐ x ∂μ, p x fun n ↦ f (v n) x := by
filter_upwards [hlim] with x hx using hx.comp hv
set aeSeqLim := fun x ↦ ite (x ∈ aeSeqSet h'f p) (g x) (⟨f (v 0) x⟩ : Nonempty ℝ≥0∞).some
refine ⟨aeSeqLim, measurable_of_tendsto' atTop (aeSeq.measurable h'f p)
(tendsto_pi_nhds.mpr fun x ↦ ?_), ?_⟩
· unfold_let aeSeqLim
simp_rw [aeSeq]
split_ifs with hx
· simp_rw [aeSeq.mk_eq_fun_of_mem_aeSeqSet h'f hx]
exact aeSeq.fun_prop_of_mem_aeSeqSet h'f hx
· exact tendsto_const_nhds
· exact (ite_ae_eq_of_measure_compl_zero g (fun x ↦ (⟨f (v 0) x⟩ : Nonempty ℝ≥0∞).some)
(aeSeqSet h'f p) (aeSeq.measure_compl_aeSeqSet_eq_zero h'f hp)).symm

/-- A limit of a.e.-measurable `ℝ≥0∞` valued functions is a.e.-measurable. -/
lemma aemeasurable_of_tendsto {f : ℕ → α → ℝ≥0∞} {g : α → ℝ≥0∞} {μ : Measure α}
(hf : ∀ i, AEMeasurable (f i) μ) (hlim : ∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) atTop (𝓝 (g a))) :
AEMeasurable g μ :=
aemeasurable_of_tendsto' atTop hf hlim

end ENNReal

@[measurability]
Expand Down Expand Up @@ -2301,6 +2356,32 @@ theorem AEMeasurable.coe_ereal_ennreal {f : α → ℝ≥0∞} {μ : Measure α}
measurable_coe_ennreal_ereal.comp_aemeasurable hf
#align ae_measurable.coe_ereal_ennreal AEMeasurable.coe_ereal_ennreal

namespace NNReal

/-- A limit (over a general filter) of measurable `ℝ≥0` valued functions is measurable. -/
theorem measurable_of_tendsto' {ι} {f : ι → α → ℝ≥0} {g : α → ℝ≥0} (u : Filter ι) [NeBot u]
[IsCountablyGenerated u] (hf : ∀ i, Measurable (f i)) (lim : Tendsto f u (𝓝 g)) :
Measurable g := by
simp_rw [← measurable_coe_nnreal_ennreal_iff] at hf ⊢
refine' ENNReal.measurable_of_tendsto' u hf _
rw [tendsto_pi_nhds] at lim ⊢
exact fun x => (ENNReal.continuous_coe.tendsto (g x)).comp (lim x)
#align measurable_of_tendsto_nnreal' NNReal.measurable_of_tendsto'

-- 2024-03-09
@[deprecated] alias _root_.measurable_of_tendsto_nnreal' := NNReal.measurable_of_tendsto'

/-- A sequential limit of measurable `ℝ≥0` valued functions is measurable. -/
theorem measurable_of_tendsto {f : ℕ → α → ℝ≥0} {g : α → ℝ≥0} (hf : ∀ i, Measurable (f i))
(lim : Tendsto f atTop (𝓝 g)) : Measurable g :=
measurable_of_tendsto' atTop hf lim
#align measurable_of_tendsto_nnreal NNReal.measurable_of_tendsto

-- 2024-03-09
@[deprecated] alias _root_.measurable_of_tendsto_nnreal := NNReal.measurable_of_tendsto

end NNReal

/-- If a function `f : α → ℝ≥0` is measurable and the measure is σ-finite, then there exists
spanning measurable sets with finite measure on which `f` is bounded.
See also `StronglyMeasurable.exists_spanning_measurableSet_norm_le` for functions into normed
Expand Down
40 changes: 2 additions & 38 deletions Mathlib/MeasureTheory/Constructions/BorelSpace/Metrizable.lean
Original file line number Diff line number Diff line change
Expand Up @@ -25,42 +25,6 @@ variable [TopologicalSpace β] [PseudoMetrizableSpace β] [MeasurableSpace β] [

open Metric

/-- A limit (over a general filter) of measurable `ℝ≥0∞` valued functions is measurable. -/
theorem measurable_of_tendsto_ennreal' {ι} {f : ι → α → ℝ≥0∞} {g : α → ℝ≥0∞} (u : Filter ι)
RemyDegenne marked this conversation as resolved.
Show resolved Hide resolved
[NeBot u] [IsCountablyGenerated u] (hf : ∀ i, Measurable (f i)) (lim : Tendsto f u (𝓝 g)) :
Measurable g := by
rcases u.exists_seq_tendsto with ⟨x, hx⟩
rw [tendsto_pi_nhds] at lim
have : (fun y => liminf (fun n => (f (x n) y : ℝ≥0∞)) atTop) = g := by
ext1 y
exact ((lim y).comp hx).liminf_eq
rw [← this]
show Measurable fun y => liminf (fun n => (f (x n) y : ℝ≥0∞)) atTop
exact measurable_liminf fun n => hf (x n)
#align measurable_of_tendsto_ennreal' measurable_of_tendsto_ennreal'

/-- A sequential limit of measurable `ℝ≥0∞` valued functions is measurable. -/
theorem measurable_of_tendsto_ennreal {f : ℕ → α → ℝ≥0∞} {g : α → ℝ≥0∞} (hf : ∀ i, Measurable (f i))
(lim : Tendsto f atTop (𝓝 g)) : Measurable g :=
measurable_of_tendsto_ennreal' atTop hf lim
#align measurable_of_tendsto_ennreal measurable_of_tendsto_ennreal

/-- A limit (over a general filter) of measurable `ℝ≥0` valued functions is measurable. -/
theorem measurable_of_tendsto_nnreal' {ι} {f : ι → α → ℝ≥0} {g : α → ℝ≥0} (u : Filter ι) [NeBot u]
[IsCountablyGenerated u] (hf : ∀ i, Measurable (f i)) (lim : Tendsto f u (𝓝 g)) :
Measurable g := by
simp_rw [← measurable_coe_nnreal_ennreal_iff] at hf ⊢
refine' measurable_of_tendsto_ennreal' u hf _
rw [tendsto_pi_nhds] at lim ⊢
exact fun x => (ENNReal.continuous_coe.tendsto (g x)).comp (lim x)
#align measurable_of_tendsto_nnreal' measurable_of_tendsto_nnreal'

/-- A sequential limit of measurable `ℝ≥0` valued functions is measurable. -/
theorem measurable_of_tendsto_nnreal {f : ℕ → α → ℝ≥0} {g : α → ℝ≥0} (hf : ∀ i, Measurable (f i))
(lim : Tendsto f atTop (𝓝 g)) : Measurable g :=
measurable_of_tendsto_nnreal' atTop hf lim
#align measurable_of_tendsto_nnreal measurable_of_tendsto_nnreal

/-- A limit (over a general filter) of measurable functions valued in a (pseudo) metrizable space is
measurable. -/
theorem measurable_of_tendsto_metrizable' {ι} {f : ι → α → β} {g : α → β} (u : Filter ι) [NeBot u]
Expand All @@ -71,7 +35,7 @@ theorem measurable_of_tendsto_metrizable' {ι} {f : ι → α → β} {g : α
intro s h1s h2s h3s
have : Measurable fun x => infNndist (g x) s := by
suffices Tendsto (fun i x => infNndist (f i x) s) u (𝓝 fun x => infNndist (g x) s) from
measurable_of_tendsto_nnreal' u (fun i => (hf i).infNndist) this
NNReal.measurable_of_tendsto' u (fun i => (hf i).infNndist) this
rw [tendsto_pi_nhds] at lim ⊢
intro x
exact ((continuous_infNndist_pt s).tendsto (g x)).comp (lim x)
Expand Down Expand Up @@ -184,7 +148,7 @@ lemma measurableSet_of_tendsto_indicator [NeBot L] (As_mble : ∀ i, MeasurableS
(h_lim : ∀ x, ∀ᶠ i in L, x ∈ As i ↔ x ∈ A) :
MeasurableSet A := by
simp_rw [← measurable_indicator_const_iff (1 : ℝ≥0∞)] at As_mble ⊢
exact measurable_of_tendsto_ennreal' L As_mble
exact ENNReal.measurable_of_tendsto' L As_mble
((tendsto_indicator_const_iff_forall_eventually L (1 : ℝ≥0∞)).mpr h_lim)

/-- If the indicator functions of a.e.-measurable sets `Aᵢ` converge a.e. to the indicator function
Expand Down
4 changes: 4 additions & 0 deletions Mathlib/MeasureTheory/Function/AEMeasurableSequence.lean
Original file line number Diff line number Diff line change
Expand Up @@ -137,4 +137,8 @@ theorem iSup [CompleteLattice β] [Countable ι] (hf : ∀ i, AEMeasurable (f i)
exact measure_mono_null (Set.compl_subset_compl.mpr h_ss) (measure_compl_aeSeqSet_eq_zero hf hp)
#align ae_seq.supr aeSeq.iSup

theorem iInf [CompleteLattice β] [Countable ι] (hf : ∀ i, AEMeasurable (f i) μ)
(hp : ∀ᵐ x ∂μ, p x fun n ↦ f n x) : ⨅ n, aeSeq hf p n =ᵐ[μ] ⨅ n, f n :=
iSup (β := βᵒᵈ) hf hp

end aeSeq
83 changes: 83 additions & 0 deletions Mathlib/MeasureTheory/Integral/Bochner.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1328,6 +1328,89 @@ lemma integral_tendsto_of_tendsto_of_antitone {μ : Measure α} {f : ℕ → α
· filter_upwards [h_mono] with x hx n m hnm using neg_le_neg_iff.mpr <| hx hnm
· filter_upwards [h_tendsto] with x hx using hx.neg

/-- If a monotone sequence of functions has an upper bound and the sequence of integrals of these
functions tends to the integral of the upper bound, then the sequence of functions converges
almost everywhere to the upper bound. -/
lemma tendsto_of_integral_tendsto_of_monotone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ}
(hf_int : ∀ n, Integrable (f n) μ) (hF_int : Integrable F μ)
(hf_tendsto : Tendsto (fun i ↦ ∫ a, f i a ∂μ) atTop (𝓝 (∫ a, F a ∂μ)))
(hf_mono : ∀ᵐ a ∂μ, Monotone (fun i ↦ f i a))
(hf_bound : ∀ᵐ a ∂μ, ∀ i, f i a ≤ F a) :
∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) atTop (𝓝 (F a)) := by
-- reduce to the `ℝ≥0∞` case
let f' : ℕ → α → ℝ≥0∞ := fun n a ↦ ENNReal.ofReal (f n a - f 0 a)
let F' : α → ℝ≥0∞ := fun a ↦ ENNReal.ofReal (F a - f 0 a)
have hf'_int_eq : ∀ i, ∫⁻ a, f' i a ∂μ = ENNReal.ofReal (∫ a, f i a ∂μ - ∫ a, f 0 a ∂μ) := by
intro i
unfold_let f'
rw [← ofReal_integral_eq_lintegral_ofReal, integral_sub (hf_int i) (hf_int 0)]
· exact (hf_int i).sub (hf_int 0)
· filter_upwards [hf_mono] with a h_mono
simp [h_mono (zero_le i)]
have hF'_int_eq : ∫⁻ a, F' a ∂μ = ENNReal.ofReal (∫ a, F a ∂μ - ∫ a, f 0 a ∂μ) := by
unfold_let F'
rw [← ofReal_integral_eq_lintegral_ofReal, integral_sub hF_int (hf_int 0)]
· exact hF_int.sub (hf_int 0)
· filter_upwards [hf_bound] with a h_bound
simp [h_bound 0]
have h_tendsto : Tendsto (fun i ↦ ∫⁻ a, f' i a ∂μ) atTop (𝓝 (∫⁻ a, F' a ∂μ)) := by
simp_rw [hf'_int_eq, hF'_int_eq]
refine (ENNReal.continuous_ofReal.tendsto _).comp ?_
rwa [tendsto_sub_const_iff]
have h_mono : ∀ᵐ a ∂μ, Monotone (fun i ↦ f' i a) := by
filter_upwards [hf_mono] with a ha_mono i j hij
refine ENNReal.ofReal_le_ofReal ?_
simp [ha_mono hij]
have h_bound : ∀ᵐ a ∂μ, ∀ i, f' i a ≤ F' a := by
filter_upwards [hf_bound] with a ha_bound i
refine ENNReal.ofReal_le_ofReal ?_
simp only [tsub_le_iff_right, sub_add_cancel, ha_bound i]
-- use the corresponding lemma for `ℝ≥0∞`
have h := tendsto_of_lintegral_tendsto_of_monotone ?_ h_tendsto h_mono h_bound ?_
rotate_left
· exact (hF_int.1.aemeasurable.sub (hf_int 0).1.aemeasurable).ennreal_ofReal
· exact ((lintegral_ofReal_le_lintegral_nnnorm _).trans_lt (hF_int.sub (hf_int 0)).2).ne
filter_upwards [h, hf_mono, hf_bound] with a ha ha_mono ha_bound
have h1 : (fun i ↦ f i a) = fun i ↦ (f' i a).toReal + f 0 a := by
unfold_let f'
ext i
rw [ENNReal.toReal_ofReal]
· abel
· simp [ha_mono (zero_le i)]
have h2 : F a = (F' a).toReal + f 0 a := by
unfold_let F'
rw [ENNReal.toReal_ofReal]
· abel
· simp [ha_bound 0]
rw [h1, h2]
refine Filter.Tendsto.add ?_ tendsto_const_nhds
exact (ENNReal.continuousAt_toReal ENNReal.ofReal_ne_top).tendsto.comp ha

/-- If an antitone sequence of functions has a lower bound and the sequence of integrals of these
functions tends to the integral of the lower bound, then the sequence of functions converges
almost everywhere to the lower bound. -/
lemma tendsto_of_integral_tendsto_of_antitone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ}
(hf_int : ∀ n, Integrable (f n) μ) (hF_int : Integrable F μ)
(hf_tendsto : Tendsto (fun i ↦ ∫ a, f i a ∂μ) atTop (𝓝 (∫ a, F a ∂μ)))
(hf_mono : ∀ᵐ a ∂μ, Antitone (fun i ↦ f i a))
(hf_bound : ∀ᵐ a ∂μ, ∀ i, F a ≤ f i a) :
∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) atTop (𝓝 (F a)) := by
let f' : ℕ → α → ℝ := fun i a ↦ - f i a
let F' : α → ℝ := fun a ↦ - F a
suffices ∀ᵐ a ∂μ, Tendsto (fun i ↦ f' i a) atTop (𝓝 (F' a)) by
filter_upwards [this] with a ha_tendsto
convert ha_tendsto.neg
· simp [f']
· simp [F']
refine tendsto_of_integral_tendsto_of_monotone (fun n ↦ (hf_int n).neg) hF_int.neg ?_ ?_ ?_
· convert hf_tendsto.neg
· rw [integral_neg]
· rw [integral_neg]
· filter_upwards [hf_mono] with a ha i j hij
simp [f', ha hij]
· filter_upwards [hf_bound] with a ha i
simp [f', F', ha i]

section NormedAddCommGroup

variable {H : Type*} [NormedAddCommGroup H]
Expand Down
Loading
Loading