From 1149cc8757d647a13e883af44e0fd55f8be5ae74 Mon Sep 17 00:00:00 2001 From: kkytola <“kalle.kytola@aalto.fi”> Date: Mon, 26 Jun 2023 12:18:48 +0300 Subject: [PATCH 01/41] A file with an outline of the proof of le_liminf_open_implies_convergence. --- ...lForPortmanteauOpenImpliesConvergence.lean | 576 ++++++++++++++++++ 1 file changed, 576 insertions(+) create mode 100644 Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean new file mode 100644 index 0000000000000..58ceb81992cda --- /dev/null +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -0,0 +1,576 @@ +import Mathlib.MeasureTheory.Measure.Portmanteau +import Mathlib.MeasureTheory.Integral.Layercake +import Mathlib.Tactic.LibrarySearch +import Mathlib.Tactic.Propose +import Mathlib.Tactic.GCongr + +open MeasureTheory Filter +open scoped ENNReal NNReal BoundedContinuousFunction Topology + +-- This file contains steps needed to prove the portmanteau implication +-- `le_liminf_open_implies_convergence`. Some intermediate steps need to be generalized +-- and cleaned up, and are to be placed in appropriate files. The main part should go +-- to the file `Mathlib.MeasureTheory.Measure.Portmanteau`. + + + +section liminf_lemma + +variable {ι R S : Type _} {F : Filter ι} [NeBot F] + [ConditionallyCompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] + [ConditionallyCompleteLinearOrder S] [TopologicalSpace S] [OrderTopology S] + +-- NOTE: The result `Monotone.map_sSup_of_continuousAt'` in the library has a natural version +-- generalized to *conditionally* complete linear orders. One just needs a hypothesis `BddAbove s`. +/-- A monotone function continuous at the supremum of a nonempty set sends this supremum to +the supremum of the image of this set. -/ +theorem Monotone.map_sSup_of_continuousAt'' {f : R → S} {A : Set R} (Cf : ContinuousAt f (sSup A)) + (Mf : Monotone f) (A_nonemp : A.Nonempty) (A_bdd : BddAbove A) : + f (sSup A) = sSup (f '' A) := by + --This is a particular case of the more general `IsLUB.isLUB_of_tendsto` + refine ((@IsLUB.isLUB_of_tendsto R S _ _ _ _ _ _ f A (sSup A) (f (sSup A)) (fun _ _ _ _ xy => Mf xy) ?_ A_nonemp ?_).csSup_eq (Set.nonempty_image_iff.mpr A_nonemp)).symm + · exact isLUB_csSup A_nonemp A_bdd + · exact tendsto_nhdsWithin_of_tendsto_nhds Cf + +theorem Monotone.map_sInf_of_continuousAt'' {f : R → S} {A : Set R} (Cf : ContinuousAt f (sInf A)) + (Mf : Monotone f) (A_nonemp : A.Nonempty) (A_bdd : BddBelow A) : + f (sInf A) = sInf (f '' A) := by + exact @Monotone.map_sSup_of_continuousAt'' Rᵒᵈ Sᵒᵈ _ _ _ _ _ _ f A Cf Mf.dual A_nonemp A_bdd + +theorem Antitone.map_sInf_of_continuousAt'' {f : R → S} {A : Set R} (Cf : ContinuousAt f (sInf A)) + (Af : Antitone f) (A_nonemp : A.Nonempty) (A_bdd : BddBelow A) : + f (sInf A) = sSup (f '' A) := by + exact @Monotone.map_sInf_of_continuousAt'' R Sᵒᵈ _ _ _ _ _ _ f A Cf Af.dual_right A_nonemp A_bdd + +theorem Antitone.map_sSup_of_continuousAt'' {f : R → S} {A : Set R} (Cf : ContinuousAt f (sSup A)) + (Af : Antitone f) (A_nonemp : A.Nonempty) (A_bdd : BddAbove A) : + f (sSup A) = sInf (f '' A) := by + exact @Monotone.map_sSup_of_continuousAt'' R Sᵒᵈ _ _ _ _ _ _ f A Cf Af.dual_right A_nonemp A_bdd + +-- NOTE: Missing from Mathlib? Probably not, but what is the name...? +lemma Monotone.isCoboundedUnder_ge [Preorder X] [Preorder Y] {f : X → Y} (hf : Monotone f) {u : ι → X} + (F : Filter ι) [NeBot F] (bdd : F.IsBoundedUnder (· ≤ ·) u) : + F.IsCoboundedUnder (· ≥ ·) (f ∘ u) := by + obtain ⟨b, hb⟩ := bdd + refine ⟨f b, fun y hy ↦ ?_⟩ + have obs : ∀ᶠ _ in map u F, y ≤ f b := by + filter_upwards [hb, Filter.map_compose.symm ▸ hy] with x h₁ h₂ using h₂.trans (hf h₁) + exact eventually_const.mp obs + +-- NOTE: Missing from Mathlib? Probably not, but what is the name...? +lemma Monotone.isCoboundedUnder_le [Preorder X] [Preorder Y] {f : X → Y} (hf : Monotone f) {u : ι → X} + (F : Filter ι) [NeBot F] (bdd : F.IsBoundedUnder (· ≥ ·) u) : + F.IsCoboundedUnder (· ≤ ·) (f ∘ u) := + @Monotone.isCoboundedUnder_ge ι Xᵒᵈ Yᵒᵈ _ _ f hf.dual u F _ bdd + +-- NOTE: Missing from Mathlib? Probably not, but what is the name...? +lemma Antitone.isCoboundedUnder_le [Preorder X] [Preorder Y] {f : X → Y} (hf : Antitone f) {u : ι → X} + (F : Filter ι) [NeBot F] (bdd : F.IsBoundedUnder (· ≤ ·) u) : + F.IsCoboundedUnder (· ≤ ·) (f ∘ u) := + @Monotone.isCoboundedUnder_ge ι X Yᵒᵈ _ _ f hf u F _ bdd + +-- NOTE: Missing from Mathlib? Probably not, but what is the name...? +lemma Antitone.isCoboundedUnder_ge [Preorder X] [Preorder Y] {f : X → Y} (hf : Antitone f) {u : ι → X} + (F : Filter ι) [NeBot F] (bdd : F.IsBoundedUnder (· ≥ ·) u) : + F.IsCoboundedUnder (· ≥ ·) (f ∘ u) := + @Monotone.isCoboundedUnder_le ι X Yᵒᵈ _ _ f hf u F _ bdd + +-- NOTE: Missing from Mathlib? What would be a good generality? +lemma Filter.isBounded_le_map_of_isBounded_range (F : Filter ι) {f : ι → ℝ} + (h : Bornology.IsBounded (Set.range f)) : + (F.map f).IsBounded (· ≤ ·) := by + rw [Metric.isBounded_iff] at h + obtain ⟨c, hc⟩ := h + by_cases hι : Nonempty ι + · obtain ⟨j⟩ := hι + use f j + c + simp only [eventually_map] + apply eventually_of_forall + intro i + have obs := hc (Set.mem_range_self j) (Set.mem_range_self i) + rw [Real.dist_eq, abs_le] at obs + linarith [obs] + · simp only [not_nonempty_iff] at hι + simp only [filter_eq_bot_of_isEmpty F, map_bot] + exact Iff.mpr isBounded_bot (by infer_instance) + +-- NOTE: Missing from Mathlib? What would be a good generality? +lemma Filter.isBounded_ge_map_of_isBounded_range (F : Filter ι) {f : ι → ℝ} + (h : Bornology.IsBounded (Set.range f)) : + (F.map f).IsBounded (· ≥ ·) := by + rw [Metric.isBounded_iff] at h + obtain ⟨c, hc⟩ := h + by_cases hι : Nonempty ι + · obtain ⟨j⟩ := hι + use f j - c + simp only [eventually_map] + apply eventually_of_forall + intro i + have obs := hc (Set.mem_range_self j) (Set.mem_range_self i) + rw [Real.dist_eq, abs_le] at obs + linarith [obs] + · simp only [not_nonempty_iff] at hι + simp only [filter_eq_bot_of_isEmpty F, map_bot] + exact Iff.mpr isBounded_bot (by infer_instance) + +-- NOTE: Missing from Mathlib? What would be a good generality? +lemma Filter.isBounded_map_of_isBounded_range (F : Filter ι) {f : ι → ℝ} + (h : Bornology.IsBounded (Set.range f)) : + (F.map f).IsBounded (· ≤ ·) := by + rw [Metric.isBounded_iff] at h + obtain ⟨c, hc⟩ := h + by_cases hι : Nonempty ι + · obtain ⟨j⟩ := hι + use f j + c + simp only [eventually_map] + apply eventually_of_forall + intro i + have obs := hc (Set.mem_range_self j) (Set.mem_range_self i) + rw [Real.dist_eq, abs_le] at obs + linarith [obs] + · simp only [not_nonempty_iff] at hι + simp only [filter_eq_bot_of_isEmpty F, map_bot] + exact Iff.mpr isBounded_bot (by infer_instance) + +-- NOTE: This was the key lemma to generalize. +/-- An antitone function between conditionally complete linear ordered spaces sends +a `Filter.limsSup` to the `Filter.liminf` of the image if it is continuous at the `limsSup`. -/ +theorem Antitone.map_limsSup_of_continuousAt' {F : Filter R} [NeBot F] {f : R → S} + (bdd_above : F.IsBounded (· ≤ ·)) (bdd_below : F.IsBounded (· ≥ ·)) + (f_decr : Antitone f) (f_cont : ContinuousAt f F.limsSup) : f F.limsSup = F.liminf f := by + have cobdd : F.IsCobounded (· ≤ ·) := bdd_below.isCobounded_flip + apply le_antisymm + · rw [limsSup, f_decr.map_sInf_of_continuousAt'' f_cont bdd_above cobdd] + apply le_of_forall_lt + intro c hc + simp only [liminf, limsInf, eventually_map] at hc ⊢ + obtain ⟨d, hd, h'd⟩ := exists_lt_of_lt_csSup ((@Set.nonempty_image_iff R S f _).mpr bdd_above) hc + apply lt_csSup_of_lt ?_ ?_ h'd + · exact f_decr.isCoboundedUnder_ge F bdd_below + · rcases hd with ⟨e, ⟨he, fe_eq_d⟩⟩ + filter_upwards [he] with x hx using (fe_eq_d.symm ▸ f_decr hx) + · by_cases h' : ∃ c, c < F.limsSup ∧ Set.Ioo c F.limsSup = ∅ + · rcases h' with ⟨c, c_lt, hc⟩ + have B : ∃ᶠ n in F, F.limsSup ≤ n := by + apply (frequently_lt_of_lt_limsSup cobdd c_lt).mono + intro x hx + by_contra' + have : (Set.Ioo c F.limsSup).Nonempty := ⟨x, ⟨hx, this⟩⟩ + simp only [hc, Set.not_nonempty_empty] at this + apply @liminf_le_of_frequently_le R S _ F f (f (limsSup F)) (B.mono fun x hx ↦ f_decr hx) ?_ + obtain ⟨b, hb⟩ := bdd_above + use f b + simp only [ge_iff_le, eventually_map] + filter_upwards [hb] with t ht using f_decr ht + simp only [gt_iff_lt, not_lt, ge_iff_le, not_exists, not_and] at h' + by_contra' H + obtain ⟨l, l_lt, h'l⟩ : ∃ l < F.limsSup, Set.Ioc l F.limsSup ⊆ { x : R | f x < F.liminf f } + · apply exists_Ioc_subset_of_mem_nhds ((tendsto_order.1 f_cont.tendsto).2 _ H) + by_contra con + simp only [not_exists, not_lt] at con + simpa only [lt_self_iff_false] using lt_of_le_of_lt + (@liminf_le_of_frequently_le R S _ F f (f (limsSup F)) + (frequently_of_forall (fun r ↦ f_decr (con r))) (bdd_above.isBoundedUnder f_decr)) H + obtain ⟨m, l_m, m_lt⟩ : (Set.Ioo l F.limsSup).Nonempty := by + contrapose! h' + refine' ⟨l, l_lt, by rwa [Set.not_nonempty_iff_eq_empty] at h'⟩ + have B : F.liminf f ≤ f m := by + apply @liminf_le_of_frequently_le R S _ F f (f m) + · apply (frequently_lt_of_lt_limsSup cobdd m_lt).mono + exact fun x hx ↦ f_decr hx.le + · exact IsBounded.isBoundedUnder f_decr bdd_above + have I : f m < F.liminf f := h'l ⟨l_m, m_lt.le⟩ + exact lt_irrefl _ (B.trans_lt I) + +theorem Antitone.map_limsInf_of_continuousAt' {F : Filter R} [NeBot F] {f : R → S} + (bdd_above : F.IsBoundedUnder (· ≤ ·) id) (bdd_below : F.IsBoundedUnder (· ≥ ·) id) + (f_decr : Antitone f) (f_cont : ContinuousAt f F.limsInf) : f F.limsInf = F.limsup f := + @Antitone.map_limsSup_of_continuousAt' Rᵒᵈ Sᵒᵈ _ _ _ _ _ _ F _ f + bdd_below bdd_above f_decr.dual f_cont + +theorem Monotone.map_limsSup_of_continuousAt' {F : Filter R} [NeBot F] {f : R → S} + (bdd_above : F.IsBoundedUnder (· ≤ ·) id) (bdd_below : F.IsBoundedUnder (· ≥ ·) id) + (f_incr : Monotone f) (f_cont : ContinuousAt f F.limsSup) : f F.limsSup = F.limsup f := + @Antitone.map_limsSup_of_continuousAt' R Sᵒᵈ _ _ _ _ _ _ F _ f + bdd_above bdd_below f_incr f_cont + +theorem Monotone.map_limsInf_of_continuousAt' {F : Filter R} [NeBot F] {f : R → S} + (bdd_above : F.IsBoundedUnder (· ≤ ·) id) (bdd_below : F.IsBoundedUnder (· ≥ ·) id) + (f_incr : Monotone f) (f_cont : ContinuousAt f F.limsInf) : f F.limsInf = F.liminf f := + @Antitone.map_limsSup_of_continuousAt' Rᵒᵈ S _ _ _ _ _ _ F _ f + bdd_below bdd_above f_incr.dual f_cont + +lemma limsup_add_const (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.limsup (fun i ↦ f i + c) F = Filter.limsup f F + c := by + have key := @Monotone.map_limsSup_of_continuousAt' ℝ ℝ _ _ _ _ _ _ (F.map f) _ (fun x ↦ x + c) bdd_above bdd_below ?_ ?_ + simp only at key + convert key.symm + · intro x y hxy + simp only [add_le_add_iff_right, hxy] + · exact (continuous_add_right c).continuousAt + +lemma liminf_add_const (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by + have key := @Monotone.map_limsInf_of_continuousAt' ℝ ℝ _ _ _ _ _ _ (F.map f) _ (fun x ↦ x + c) bdd_above bdd_below ?_ ?_ + simp only at key + convert key.symm + · intro x y hxy + simp only [add_le_add_iff_right, hxy] + · exact (continuous_add_right c).continuousAt + +lemma limsup_const_sub (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.limsup (fun i ↦ c - f i ) F = c - Filter.liminf f F := by + have key := @Antitone.map_limsInf_of_continuousAt' ℝ ℝ _ _ _ _ _ _ (F.map f) _ (fun x ↦ c - x) bdd_above bdd_below ?_ ?_ + simp only at key + convert key.symm + · intro x y hxy + simp only [sub_le_sub_iff_left, hxy] + · exact (continuous_sub_left c).continuousAt + +lemma liminf_const_sub (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.liminf (fun i ↦ c - f i ) F = c - Filter.limsup f F := by + have key := @Antitone.map_limsSup_of_continuousAt' ℝ ℝ _ _ _ _ _ _ (F.map f) _ (fun x ↦ c - x) bdd_above bdd_below ?_ ?_ + simp only at key + convert key.symm + · intro x y hxy + simp only [sub_le_sub_iff_left, hxy] + · exact (continuous_sub_left c).continuousAt + +end liminf_lemma + + + +section boundedness_by_norm_bounds + +-- NOTE: Can this really be missing from Mathlib? +lemma Metric.isBounded_closedBall [PseudoMetricSpace X] (z : X) (r : ℝ) : + Bornology.IsBounded (Metric.closedBall z r) := by + rw [Metric.isBounded_iff] + use 2 * r + intro x hx y hy + simp only [closedBall, Set.mem_setOf_eq] at hx hy + calc dist x y + _ ≤ dist x z + dist z y := dist_triangle x z y + _ = dist x z + dist y z := by rw [dist_comm z y] + _ ≤ r + r := by gcongr + _ = 2 * r := by ring + +-- NOTE: Can this really be missing from Mathlib? +lemma Metric.isBounded_ball [PseudoMetricSpace X] (z : X) (r : ℝ) : + Bornology.IsBounded (Metric.ball z r) := + (Metric.isBounded_closedBall z r).subset ball_subset_closedBall + +-- NOTE: Should this be in Mathlib? +lemma isBounded_range_of_forall_norm_le [NormedAddGroup E] (f : ι → E) (c : ℝ) (h : ∀ i, ‖f i‖ ≤ c) : + Bornology.IsBounded (Set.range f) := by + apply (Metric.isBounded_closedBall 0 c).subset + intro x ⟨i, hi⟩ + simpa only [← hi, Metric.closedBall, dist_zero_right, Set.mem_setOf_eq, gt_iff_lt] using h i + +end boundedness_by_norm_bounds + + + +section layercake_for_integral + +-- TODO: Generalize from ℝ to the usual type classes. +-- NOTE: This is currently a mess, because of mixing Measurable and AEStronglyMeasurable. +lemma Integrable.measure_pos_le_norm_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] + {f : α → ℝ} --(f_nn : 0 ≤ f) + (f_intble : Integrable f μ) + {t : ℝ} (t_pos : 0 < t) : + μ {a : α | t ≤ ‖f a‖} < ∞ := by + have f_aemble := f_intble.1.aemeasurable + have norm_f_aemble : AEMeasurable (fun a ↦ ENNReal.ofReal ‖f a‖) μ := by + --have := ENNReal.measurable_ofReal.comp (@measurable_norm ℝ _ _ _) + exact (ENNReal.measurable_ofReal.comp measurable_norm).comp_aemeasurable f_aemble + obtain ⟨g, ⟨g_mble, ⟨g_nn, aeeq_g⟩⟩⟩ := @AEMeasurable.exists_measurable_nonneg α _ μ ℝ≥0∞ _ _ _ _ + norm_f_aemble (eventually_of_forall (fun x ↦ zero_le _)) + have foo : MeasurableSet {a : α | ENNReal.ofReal t < g a} := by + sorry + have aux := @lintegral_indicator_const _ _ μ _ foo (ENNReal.ofReal t) + have markov := @mul_meas_ge_le_lintegral₀ α _ μ + (fun a ↦ ENNReal.ofReal ‖f a‖) norm_f_aemble (ENNReal.ofReal t) + have same : ∀ a, ENNReal.ofReal t ≤ ENNReal.ofReal ‖f a‖ ↔ t ≤ ‖f a‖ := by sorry + have also : ∫⁻ (a : α), ENNReal.ofReal ‖f a‖ ∂μ = ∫⁻ (a : α), ‖f a‖₊ ∂μ := by + apply lintegral_congr + intro x + sorry + simp_rw [same, also] at markov + have almost := lt_of_le_of_lt markov f_intble.2 + have t_inv_ne_top : (ENNReal.ofReal t)⁻¹ ≠ ∞ := by + exact ENNReal.inv_ne_top.mpr (ENNReal.ofReal_pos.mpr t_pos).ne.symm + convert ENNReal.mul_lt_top t_inv_ne_top almost.ne + simp [← mul_assoc, + ENNReal.inv_mul_cancel (ENNReal.ofReal_pos.mpr t_pos).ne.symm ENNReal.ofReal_ne_top] + +lemma Integrable.measure_pos_lt_norm_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] + {f : α → ℝ} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : + μ {a : α | t < ‖f a‖} < ∞ := + lt_of_le_of_lt (measure_mono (fun _ h ↦ (Set.mem_setOf_eq ▸ h).le)) + (Integrable.measure_pos_le_norm_lt_top f_intble t_pos) + +lemma Integrable.measure_pos_le_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] + {f : α → ℝ} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : + μ {a : α | t ≤ f a} < ∞ := by + -- Need to do separately positive and negative parts? + sorry + +lemma Integrable.measure_pos_lt_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] + {f : α → ℝ} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : + μ {a : α | t < f a} < ∞ := by + apply lt_of_le_of_lt (measure_mono ?_) (Integrable.measure_pos_le_lt_top f_intble t_pos) + exact fun x hx ↦ (Set.mem_setOf_eq ▸ hx).le + +-- NOTE: This is a version of the basic "Layercake formula" for real-valued nonnegative integrands +-- and Bochner integral ∫ instead of ∫⁻. I don't know if the other (more general) versions of +-- layercake should be similarly generalized. The proofs are basically similar, but the statements +-- themselves become a bit unpleasant due to integrability requirements for something slightly +-- complicated. +-- TODO: Should remove `Measurable` assumption and just embrace the `AEStronglyMeasurable` +-- which comes from `Integrable`. This is not pleasant in the proof, but pays off for the user... +theorem integral_eq_integral_meas_lt [MeasurableSpace α] (μ : Measure α) [SigmaFinite μ] + {f : α → ℝ} (f_nn : 0 ≤ f) (f_mble : Measurable f) (f_intble : Integrable f μ) : + (∫ ω, f ω ∂μ) = ∫ t in Set.Ioi 0, ENNReal.toReal (μ {a : α | t < f a}) := by + have key := lintegral_eq_lintegral_meas_lt μ f_nn f_mble -- should use `Integrable` + have lhs_finite : ∫⁻ (ω : α), ENNReal.ofReal (f ω) ∂μ < ∞ := Integrable.lintegral_lt_top f_intble + have rhs_finite : ∫⁻ (t : ℝ) in Set.Ioi 0, μ {a | t < f a} < ∞ := by simp only [← key, lhs_finite] + have rhs_integrand_decr : Antitone (fun t ↦ (μ {a : α | t < f a})) := + fun _ _ hst ↦ measure_mono (fun _ h ↦ lt_of_le_of_lt hst h) + have rhs_integrand_finite : ∀ (t : ℝ), t > 0 → μ {a | t < f a} < ∞ := by + exact fun t ht ↦ Integrable.measure_pos_lt_lt_top f_intble ht + convert (ENNReal.toReal_eq_toReal lhs_finite.ne rhs_finite.ne).mpr key + · refine integral_eq_lintegral_of_nonneg_ae ?_ ?_ + · -- TODO: Maybe should relax the assumption to ae nonnegativity. + exact eventually_of_forall f_nn + · --exact f_mble.aestronglyMeasurable + exact f_intble.aestronglyMeasurable + · have aux := @integral_eq_lintegral_of_nonneg_ae _ _ ((volume : Measure ℝ).restrict (Set.Ioi 0)) + (fun t ↦ ENNReal.toReal (μ {a : α | t < f a})) ?_ ?_ + · rw [aux] + apply congrArg ENNReal.toReal + apply set_lintegral_congr_fun measurableSet_Ioi + apply eventually_of_forall + exact fun t t_pos ↦ ENNReal.ofReal_toReal (rhs_integrand_finite t t_pos).ne + · exact eventually_of_forall (fun x ↦ by simp only [Pi.zero_apply, ENNReal.toReal_nonneg]) + · apply Measurable.aestronglyMeasurable + refine Measurable.ennreal_toReal ?_ + apply Antitone.measurable + exact rhs_integrand_decr + +end layercake_for_integral + + + +section le_liminf_open_implies_convergence + +variable {Ω : Type} [MeasurableSpace Ω] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] + +/- +-- TODO: Is it possible to add a @[gcongr] attribute to `lintegral_mono`? + +attribute [gcongr] lintegral_mono -- @[gcongr] attribute only applies to lemmas proving x₁ ~₁ x₁' → ... xₙ ~ₙ xₙ' → f x₁ ... xₙ ∼ f x₁' ... xₙ', got ∀ {α : Type u_1} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} ⦃f g : α → ℝ≥0∞⦄, f ≤ g → ∫⁻ (a : α), f a ∂μ ≤ ∫⁻ (a : α), g a ∂μ + +lemma foo (μ : Measure Ω) {f g : Ω → ℝ≥0∞} (hfg : f ≤ g) : + ∫⁻ ω, f ω ∂μ ≤ ∫⁻ ω, g ω ∂μ := by + gcongr -- gcongr did not make progress + sorry + -/ + +-- NOTE: I think I will work with real-valued integrals, after all... +lemma fatou_argument_lintegral + {μ : Measure Ω} [SigmaFinite μ] {μs : ℕ → Measure Ω} [∀ i, SigmaFinite (μs i)] + {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) + (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : + ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫⁻ x, ENNReal.ofReal (f x) ∂ (μs i)) := by + simp_rw [lintegral_eq_lintegral_meas_lt _ f_nn f_cont.measurable] + calc ∫⁻ (t : ℝ) in Set.Ioi 0, μ {a | t < f a} + ≤ ∫⁻ (t : ℝ) in Set.Ioi 0, atTop.liminf (fun i ↦ (μs i) {a | t < f a}) + := (lintegral_mono (fun t ↦ h_opens _ (continuous_def.mp f_cont _ isOpen_Ioi))).trans ?_ + _ ≤ atTop.liminf (fun i ↦ ∫⁻ (t : ℝ) in Set.Ioi 0, (μs i) {a | t < f a}) + := lintegral_liminf_le (fun n ↦ Antitone.measurable + (fun s t hst ↦ measure_mono (fun ω hω ↦ lt_of_le_of_lt hst hω))) + rfl + +-- NOTE: I think this is the version I prefer to use, after all... +lemma fatou_argument_integral_nonneg + {μ : Measure Ω} [IsFiniteMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsFiniteMeasure (μs i)] + {f : Ω →ᵇ ℝ} (f_nn : 0 ≤ f) + (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : + ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by + sorry + +-- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? +lemma BoundedContinuousFunction.integrable (μ : Measure Ω) [IsFiniteMeasure μ] (f : Ω →ᵇ ℝ) : + Integrable f μ := by + refine ⟨f.continuous.measurable.aestronglyMeasurable, ?_⟩ + simp [HasFiniteIntegral] + calc ∫⁻ x, ‖f x‖₊ ∂μ + _ ≤ ∫⁻ _, ‖f‖₊ ∂μ := ?_ + _ = ‖f‖₊ * (μ Set.univ) := by rw [lintegral_const] + _ < ∞ := ENNReal.mul_lt_top + ENNReal.coe_ne_top (measure_ne_top μ Set.univ) + · apply lintegral_mono + exact fun x ↦ ENNReal.coe_le_coe.mpr (nnnorm_coe_le_nnnorm f x) + +-- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? +lemma BoundedContinuousFunction.norm_integral_le_mul_norm_of_isFiniteMeasure + (μ : Measure Ω) [IsFiniteMeasure μ] (f : Ω →ᵇ ℝ) : + ‖∫ x, (f x) ∂μ‖ ≤ ENNReal.toReal (μ Set.univ) * ‖f‖ := by + calc ‖∫ x, (f x) ∂μ‖ + _ ≤ ∫ x, ‖f x‖ ∂μ := by exact norm_integral_le_integral_norm _ + _ ≤ ∫ _, ‖f‖ ∂μ := ?_ + _ = ENNReal.toReal (μ Set.univ) • ‖f‖ := by rw [integral_const] + · apply integral_mono _ (integrable_const ‖f‖) (fun x ↦ f.norm_coe_le_norm x) + exact (integrable_norm_iff f.continuous.measurable.aestronglyMeasurable).mpr (f.integrable μ) + +-- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? +lemma BoundedContinuousFunction.norm_integral_le_norm_of_isProbabilityMeasure + (μ : Measure Ω) [IsProbabilityMeasure μ] (f : Ω →ᵇ ℝ) : + ‖∫ x, (f x) ∂μ‖ ≤ ‖f‖ := by + convert f.norm_integral_le_mul_norm_of_isFiniteMeasure μ + simp only [measure_univ, ENNReal.one_toReal, one_mul] + +-- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? +-- TODO: Should this be generalized to functions with values in Banach spaces? +lemma isBounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure + (μs : ι → Measure Ω) [∀ i, IsProbabilityMeasure (μs i)] (f : Ω →ᵇ ℝ) : + Bornology.IsBounded (Set.range (fun i ↦ ∫ x, (f x) ∂ (μs i))) := by + apply isBounded_range_of_forall_norm_le _ ‖f‖ + exact fun i ↦ f.norm_integral_le_norm_of_isProbabilityMeasure (μs i) + +lemma main_thing' + {μ : Measure Ω} [IsProbabilityMeasure μ] + {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] + {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) + (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : + ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫⁻ x, ENNReal.ofReal (f x) ∂ (μs i)) := by + simp_rw [lintegral_eq_lintegral_meas_lt _ f_nn f_cont.measurable] + have obs : ∀ t, IsOpen {a : Ω | t < f a} := fun t ↦ (continuous_def.mp f_cont) _ isOpen_Ioi + apply (lintegral_mono (fun t ↦ h_opens _ (obs t))).trans + refine lintegral_liminf_le ?_ + refine fun i ↦ Antitone.measurable (fun s t hst ↦ measure_mono ?_) + intro ω hω + simp only [Set.mem_setOf_eq] at * + linarith + +lemma main_thing + {μ : Measure Ω} [IsProbabilityMeasure μ] + {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] + {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) + (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : + ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by + sorry + +-- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? +lemma BoundedContinuousFunction.neg_norm_le [TopologicalSpace X] (f : X →ᵇ ℝ) (x : X) : + -‖f‖ ≤ f x := by + exact (abs_le.mp (norm_coe_le_norm f x)).1 + +-- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? +lemma BoundedContinuousFunction.le_norm [TopologicalSpace X] (f : X →ᵇ ℝ) (x : X): + f x ≤ ‖f‖ := by + exact (abs_le.mp (norm_coe_le_norm f x)).2 + +-- NOTE: Where should such things be placed? In the `Portmanteau`-file only? +lemma BoundedContinuousFunction.add_norm_nonneg [TopologicalSpace X] (f : X →ᵇ ℝ) : + 0 ≤ f + BoundedContinuousFunction.const _ ‖f‖ := by + intro x + dsimp + linarith [(abs_le.mp (norm_coe_le_norm f x)).1] + +-- NOTE: Where should such things be placed? In the `Portmanteau`-file only? +lemma BoundedContinuousFunction.norm_sub_nonneg [TopologicalSpace X] (f : X →ᵇ ℝ) : + 0 ≤ BoundedContinuousFunction.const _ ‖f‖ - f := by + intro x + dsimp + linarith [(abs_le.mp (norm_coe_le_norm f x)).2] + +-- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? +lemma BoundedContinuousFunction.integral_add_const {μ : Measure Ω} [IsFiniteMeasure μ] + (f : Ω →ᵇ ℝ) (c : ℝ) : + ∫ x, (f + BoundedContinuousFunction.const Ω c) x ∂μ = + ∫ x, f x ∂μ + ENNReal.toReal (μ (Set.univ)) • c := by + simp only [coe_add, const_toFun, Pi.add_apply, smul_eq_mul] + simp_rw [integral_add (FiniteMeasure.integrable_of_boundedContinuous_to_real _ f) + (integrable_const c)] + simp only [integral_const, smul_eq_mul] + +-- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? +lemma BoundedContinuousFunction.integral_const_sub {μ : Measure Ω} [IsFiniteMeasure μ] + (f : Ω →ᵇ ℝ) (c : ℝ) : + ∫ x, (BoundedContinuousFunction.const Ω c - f) x ∂μ = + ENNReal.toReal (μ (Set.univ)) • c - ∫ x, f x ∂μ := by + simp only [coe_sub, const_toFun, Pi.sub_apply, smul_eq_mul] + simp_rw [integral_sub (integrable_const c) + (FiniteMeasure.integrable_of_boundedContinuous_to_real _ f)] + simp only [integral_const, smul_eq_mul] + +lemma reduction_to_liminf {ι : Type} {L : Filter ι} [NeBot L] + {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ι → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] + (h : ∀ f : Ω →ᵇ ℝ, 0 ≤ f → ∫ x, (f x) ∂μ ≤ L.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i))) + (f : Ω →ᵇ ℝ) : + Tendsto (fun i ↦ ∫ x, (f x) ∂ (μs i)) L (𝓝 (∫ x, (f x) ∂μ)) := by + have obs := isBounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure μs f + have bdd_above : IsBoundedUnder (· ≤ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := by + apply isBounded_le_map_of_isBounded_range + apply isBounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure + have bdd_below : IsBoundedUnder (· ≥ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := by + apply isBounded_ge_map_of_isBounded_range + apply isBounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure + apply @tendsto_of_le_liminf_of_limsup_le ℝ ι _ _ _ L (fun i ↦ ∫ x, (f x) ∂ (μs i)) (∫ x, (f x) ∂μ) + · have key := h _ (f.add_norm_nonneg) + simp_rw [f.integral_add_const ‖f‖] at key + simp only [measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key + -- TODO: Should the case of ⊥ filter be treated separately and not included as an assumption? + have := liminf_add_const L (fun i ↦ ∫ x, (f x) ∂ (μs i)) ‖f‖ bdd_above bdd_below + rwa [this, add_le_add_iff_right] at key + · have key := h _ (f.norm_sub_nonneg) + simp_rw [f.integral_const_sub ‖f‖] at key + simp only [measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key + have := liminf_const_sub L (fun i ↦ ∫ x, (f x) ∂ (μs i)) ‖f‖ bdd_above bdd_below + rwa [this, sub_le_sub_iff_left] at key + · exact L.isBounded_le_map_of_isBounded_range obs + · exact L.isBounded_ge_map_of_isBounded_range obs + +/-- A characterization of weak convergence of probability measures by the condition that the +integrals of every continuous bounded nonnegative function converge to the integral of the function +against the limit measure. -/ +lemma ProbabilityMeasure.tendsto_iff_forall_nonneg_integral_tendsto {γ : Type _} {F : Filter γ} + {μs : γ → ProbabilityMeasure Ω} {μ : ProbabilityMeasure Ω} : + Tendsto μs F (𝓝 μ) ↔ + ∀ f : Ω →ᵇ ℝ, 0 ≤ f → + Tendsto (fun i ↦ ∫ ω, (f ω) ∂(μs i : Measure Ω)) F (𝓝 (∫ ω, (f ω) ∂(μ : Measure Ω))) := by + rw [ProbabilityMeasure.tendsto_iff_forall_integral_tendsto] + refine ⟨fun h f _ ↦ h f, fun h f ↦ ?_⟩ + set g := f + BoundedContinuousFunction.const _ ‖f‖ with g_def + have fx_eq : ∀ x, f x = g x - ‖f‖ := by + intro x + simp only [BoundedContinuousFunction.coe_add, BoundedContinuousFunction.const_toFun, Pi.add_apply, + add_sub_cancel] + have key := h g (f.add_norm_nonneg) + simp [g_def] at key + simp_rw [integral_add (FiniteMeasure.integrable_of_boundedContinuous_to_real _ f) + (integrable_const ‖f‖)] at key + simp only [integral_const, measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key + simp_rw [fx_eq] + convert tendsto_add.comp (Tendsto.prod_mk_nhds key (@tendsto_const_nhds _ _ _ (-‖f‖) F)) <;> simp + +theorem le_liminf_open_implies_convergence + {μ : ProbabilityMeasure Ω} {μs : ℕ → ProbabilityMeasure Ω} + (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : + atTop.Tendsto (fun i ↦ μs i) (𝓝 μ) := by + refine ProbabilityMeasure.tendsto_iff_forall_nonneg_integral_tendsto.mpr ?_ + intro g g_nn + apply reduction_to_liminf + intro f f_nn + have f_nn' : 0 ≤ (f : Ω → ℝ) := fun x ↦ by simpa using f_nn x + apply main_thing f.continuous f_nn' + -- Annoying coercions to reduce to `h_opens`... + sorry + +end le_liminf_open_implies_convergence From a9ddc6d68f91e252d86dd1456be3f9422d3569df Mon Sep 17 00:00:00 2001 From: kkytola <“kalle.kytola@aalto.fi”> Date: Thu, 29 Jun 2023 17:50:55 +0300 Subject: [PATCH 02/41] Thoughts. --- ...lForPortmanteauOpenImpliesConvergence.lean | 46 +++++++++++++++---- 1 file changed, 37 insertions(+), 9 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index 58ceb81992cda..242b899eb7f99 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -20,6 +20,8 @@ variable {ι R S : Type _} {F : Filter ι} [NeBot F] [ConditionallyCompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] [ConditionallyCompleteLinearOrder S] [TopologicalSpace S] [OrderTopology S] +#check Monotone.map_sSup_of_continuousAt' + -- NOTE: The result `Monotone.map_sSup_of_continuousAt'` in the library has a natural version -- generalized to *conditionally* complete linear orders. One just needs a hypothesis `BddAbove s`. /-- A monotone function continuous at the supremum of a nonempty set sends this supremum to @@ -28,29 +30,32 @@ theorem Monotone.map_sSup_of_continuousAt'' {f : R → S} {A : Set R} (Cf : Cont (Mf : Monotone f) (A_nonemp : A.Nonempty) (A_bdd : BddAbove A) : f (sSup A) = sSup (f '' A) := by --This is a particular case of the more general `IsLUB.isLUB_of_tendsto` - refine ((@IsLUB.isLUB_of_tendsto R S _ _ _ _ _ _ f A (sSup A) (f (sSup A)) (fun _ _ _ _ xy => Mf xy) ?_ A_nonemp ?_).csSup_eq (Set.nonempty_image_iff.mpr A_nonemp)).symm + refine ((@IsLUB.isLUB_of_tendsto R S _ _ _ _ _ _ f A (sSup A) (f (sSup A)) (Mf.monotoneOn _) ?_ A_nonemp ?_).csSup_eq (Set.nonempty_image_iff.mpr A_nonemp)).symm · exact isLUB_csSup A_nonemp A_bdd · exact tendsto_nhdsWithin_of_tendsto_nhds Cf theorem Monotone.map_sInf_of_continuousAt'' {f : R → S} {A : Set R} (Cf : ContinuousAt f (sInf A)) (Mf : Monotone f) (A_nonemp : A.Nonempty) (A_bdd : BddBelow A) : - f (sInf A) = sInf (f '' A) := by - exact @Monotone.map_sSup_of_continuousAt'' Rᵒᵈ Sᵒᵈ _ _ _ _ _ _ f A Cf Mf.dual A_nonemp A_bdd + f (sInf A) = sInf (f '' A) := + @Monotone.map_sSup_of_continuousAt'' Rᵒᵈ Sᵒᵈ _ _ _ _ _ _ f A Cf Mf.dual A_nonemp A_bdd theorem Antitone.map_sInf_of_continuousAt'' {f : R → S} {A : Set R} (Cf : ContinuousAt f (sInf A)) (Af : Antitone f) (A_nonemp : A.Nonempty) (A_bdd : BddBelow A) : - f (sInf A) = sSup (f '' A) := by - exact @Monotone.map_sInf_of_continuousAt'' R Sᵒᵈ _ _ _ _ _ _ f A Cf Af.dual_right A_nonemp A_bdd + f (sInf A) = sSup (f '' A) := + @Monotone.map_sInf_of_continuousAt'' R Sᵒᵈ _ _ _ _ _ _ f A Cf Af.dual_right A_nonemp A_bdd theorem Antitone.map_sSup_of_continuousAt'' {f : R → S} {A : Set R} (Cf : ContinuousAt f (sSup A)) (Af : Antitone f) (A_nonemp : A.Nonempty) (A_bdd : BddAbove A) : - f (sSup A) = sInf (f '' A) := by - exact @Monotone.map_sSup_of_continuousAt'' R Sᵒᵈ _ _ _ _ _ _ f A Cf Af.dual_right A_nonemp A_bdd + f (sSup A) = sInf (f '' A) := + @Monotone.map_sSup_of_continuousAt'' R Sᵒᵈ _ _ _ _ _ _ f A Cf Af.dual_right A_nonemp A_bdd + +#check Filter.IsBounded.isCobounded_flip -- NOTE: Missing from Mathlib? Probably not, but what is the name...? lemma Monotone.isCoboundedUnder_ge [Preorder X] [Preorder Y] {f : X → Y} (hf : Monotone f) {u : ι → X} (F : Filter ι) [NeBot F] (bdd : F.IsBoundedUnder (· ≤ ·) u) : F.IsCoboundedUnder (· ≥ ·) (f ∘ u) := by + --refine Filter.IsBounded.isCobounded_flip ?_ obtain ⟨b, hb⟩ := bdd refine ⟨f b, fun y hy ↦ ?_⟩ have obs : ∀ᶠ _ in map u F, y ≤ f b := by @@ -75,10 +80,16 @@ lemma Antitone.isCoboundedUnder_ge [Preorder X] [Preorder Y] {f : X → Y} (hf : F.IsCoboundedUnder (· ≥ ·) (f ∘ u) := @Monotone.isCoboundedUnder_le ι X Yᵒᵈ _ _ f hf u F _ bdd --- NOTE: Missing from Mathlib? What would be a good generality? +-- NOTE: Missing from Mathlib? +-- What would be a good generality? (Mixes order and metric, so typeclasses don't readily exist.) lemma Filter.isBounded_le_map_of_isBounded_range (F : Filter ι) {f : ι → ℝ} (h : Bornology.IsBounded (Set.range f)) : (F.map f).IsBounded (· ≤ ·) := by + rw [← Metric.bounded_iff_isBounded, Real.bounded_iff_bddBelow_bddAbove] at h + apply isBoundedUnder_of + sorry + +/- rw [Metric.isBounded_iff] at h obtain ⟨c, hc⟩ := h by_cases hι : Nonempty ι @@ -93,6 +104,7 @@ lemma Filter.isBounded_le_map_of_isBounded_range (F : Filter ι) {f : ι → ℝ · simp only [not_nonempty_iff] at hι simp only [filter_eq_bot_of_isEmpty F, map_bot] exact Iff.mpr isBounded_bot (by infer_instance) + -/ -- NOTE: Missing from Mathlib? What would be a good generality? lemma Filter.isBounded_ge_map_of_isBounded_range (F : Filter ι) {f : ι → ℝ} @@ -132,6 +144,8 @@ lemma Filter.isBounded_map_of_isBounded_range (F : Filter ι) {f : ι → ℝ} simp only [filter_eq_bot_of_isEmpty F, map_bot] exact Iff.mpr isBounded_bot (by infer_instance) +#check Antitone.map_limsSup_of_continuousAt + -- NOTE: This was the key lemma to generalize. /-- An antitone function between conditionally complete linear ordered spaces sends a `Filter.limsSup` to the `Filter.liminf` of the image if it is continuous at the `limsSup`. -/ @@ -246,6 +260,10 @@ end liminf_lemma section boundedness_by_norm_bounds +-- TODO: Should use `Metric.Bounded` +#check Metric.Bounded +#check Metric.bounded_closedBall + -- NOTE: Can this really be missing from Mathlib? lemma Metric.isBounded_closedBall [PseudoMetricSpace X] (z : X) (r : ℝ) : Bornology.IsBounded (Metric.closedBall z r) := by @@ -292,6 +310,8 @@ lemma Integrable.measure_pos_le_norm_lt_top [MeasurableSpace α] {μ : Measure norm_f_aemble (eventually_of_forall (fun x ↦ zero_le _)) have foo : MeasurableSet {a : α | ENNReal.ofReal t < g a} := by sorry + -- TODO: Generalize `lintegral_indicator_const` to null-measurable sets so there is no need + -- to use g instead of f. (Have already `lintegral_indicator₀` so easy!) have aux := @lintegral_indicator_const _ _ μ _ foo (ENNReal.ofReal t) have markov := @mul_meas_ge_le_lintegral₀ α _ μ (fun a ↦ ENNReal.ofReal ‖f a‖) norm_f_aemble (ENNReal.ofReal t) @@ -369,7 +389,6 @@ end layercake_for_integral section le_liminf_open_implies_convergence variable {Ω : Type} [MeasurableSpace Ω] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] - /- -- TODO: Is it possible to add a @[gcongr] attribute to `lintegral_mono`? @@ -379,6 +398,15 @@ lemma foo (μ : Measure Ω) {f g : Ω → ℝ≥0∞} (hfg : f ≤ g) : ∫⁻ ω, f ω ∂μ ≤ ∫⁻ ω, g ω ∂μ := by gcongr -- gcongr did not make progress sorry + +-- This would solve it! + +lemma MeasureTheory.lintegral_mono'' {α : Type} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} {f g : α → ℝ≥0∞} + (hfg : f ≤ g) : lintegral μ f ≤ lintegral μ g := by sorry + +#check lintegral_mono'' + +attribute [gcongr] lintegral_mono'' -- @[gcongr] attribute only applies to lemmas proving x₁ ~₁ x₁' → ... xₙ ~ₙ xₙ' → f x₁ ... xₙ ∼ f x₁' ... xₙ', got ∀ {α : Type u_1} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} ⦃f g : α → ℝ≥0∞⦄, f ≤ g → ∫⁻ (a : α), f a ∂μ ≤ ∫⁻ (a : α), g a ∂μ -/ -- NOTE: I think I will work with real-valued integrals, after all... From a10a67468e2049a7a9ff42358c6cba60cb7aba89 Mon Sep 17 00:00:00 2001 From: Kalle Date: Fri, 30 Jun 2023 22:40:30 +0300 Subject: [PATCH 03/41] Switched to Metric.Bounded instead of Bornology.IsBounded. --- ...lForPortmanteauOpenImpliesConvergence.lean | 117 +++++++++--------- 1 file changed, 56 insertions(+), 61 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index 242b899eb7f99..912b10fc3b400 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -80,69 +80,42 @@ lemma Antitone.isCoboundedUnder_ge [Preorder X] [Preorder Y] {f : X → Y} (hf : F.IsCoboundedUnder (· ≥ ·) (f ∘ u) := @Monotone.isCoboundedUnder_le ι X Yᵒᵈ _ _ f hf u F _ bdd +-- NOTE: Missing from Mathlib? +-- What would be a good generality? (Mixes order and metric, so typeclasses don't readily exist.) +lemma Filter.isBounded_le_map_of_bounded_range (F : Filter ι) {f : ι → ℝ} + (h : Metric.Bounded (Set.range f)) : + (F.map f).IsBounded (· ≤ ·) := by + sorry + --rw [← Metric.bounded_iff_isBounded, Real.bounded_iff_bddBelow_bddAbove] at h + --obtain ⟨c, hc⟩ := h.2 + --apply isBoundedUnder_of + --use c + --simpa [mem_upperBounds] using hc + +lemma Filter.isBounded_ge_map_of_bounded_range (F : Filter ι) {f : ι → ℝ} + (h : Metric.Bounded (Set.range f)) : + (F.map f).IsBounded (· ≥ ·) := by sorry + -- NOTE: Missing from Mathlib? -- What would be a good generality? (Mixes order and metric, so typeclasses don't readily exist.) lemma Filter.isBounded_le_map_of_isBounded_range (F : Filter ι) {f : ι → ℝ} (h : Bornology.IsBounded (Set.range f)) : (F.map f).IsBounded (· ≤ ·) := by rw [← Metric.bounded_iff_isBounded, Real.bounded_iff_bddBelow_bddAbove] at h + obtain ⟨c, hc⟩ := h.2 apply isBoundedUnder_of - sorry - -/- - rw [Metric.isBounded_iff] at h - obtain ⟨c, hc⟩ := h - by_cases hι : Nonempty ι - · obtain ⟨j⟩ := hι - use f j + c - simp only [eventually_map] - apply eventually_of_forall - intro i - have obs := hc (Set.mem_range_self j) (Set.mem_range_self i) - rw [Real.dist_eq, abs_le] at obs - linarith [obs] - · simp only [not_nonempty_iff] at hι - simp only [filter_eq_bot_of_isEmpty F, map_bot] - exact Iff.mpr isBounded_bot (by infer_instance) - -/ + use c + simpa [mem_upperBounds] using hc -- NOTE: Missing from Mathlib? What would be a good generality? lemma Filter.isBounded_ge_map_of_isBounded_range (F : Filter ι) {f : ι → ℝ} (h : Bornology.IsBounded (Set.range f)) : (F.map f).IsBounded (· ≥ ·) := by - rw [Metric.isBounded_iff] at h - obtain ⟨c, hc⟩ := h - by_cases hι : Nonempty ι - · obtain ⟨j⟩ := hι - use f j - c - simp only [eventually_map] - apply eventually_of_forall - intro i - have obs := hc (Set.mem_range_self j) (Set.mem_range_self i) - rw [Real.dist_eq, abs_le] at obs - linarith [obs] - · simp only [not_nonempty_iff] at hι - simp only [filter_eq_bot_of_isEmpty F, map_bot] - exact Iff.mpr isBounded_bot (by infer_instance) - --- NOTE: Missing from Mathlib? What would be a good generality? -lemma Filter.isBounded_map_of_isBounded_range (F : Filter ι) {f : ι → ℝ} - (h : Bornology.IsBounded (Set.range f)) : - (F.map f).IsBounded (· ≤ ·) := by - rw [Metric.isBounded_iff] at h - obtain ⟨c, hc⟩ := h - by_cases hι : Nonempty ι - · obtain ⟨j⟩ := hι - use f j + c - simp only [eventually_map] - apply eventually_of_forall - intro i - have obs := hc (Set.mem_range_self j) (Set.mem_range_self i) - rw [Real.dist_eq, abs_le] at obs - linarith [obs] - · simp only [not_nonempty_iff] at hι - simp only [filter_eq_bot_of_isEmpty F, map_bot] - exact Iff.mpr isBounded_bot (by infer_instance) + rw [← Metric.bounded_iff_isBounded, Real.bounded_iff_bddBelow_bddAbove] at h + obtain ⟨c, hc⟩ := h.1 + apply isBoundedUnder_of + use c + simpa [mem_lowerBounds] using hc #check Antitone.map_limsSup_of_continuousAt @@ -263,7 +236,17 @@ section boundedness_by_norm_bounds -- TODO: Should use `Metric.Bounded` #check Metric.Bounded #check Metric.bounded_closedBall +#check Metric.bounded_ball + +-- NOTE: Should this be in Mathlib? +lemma Metric.bounded_range_of_forall_norm_le [NormedAddGroup E] + (f : ι → E) (c : ℝ) (h : ∀ i, ‖f i‖ ≤ c) : + Metric.Bounded (Set.range f) := by + apply Metric.Bounded.mono _ (@Metric.bounded_closedBall _ _ 0 c) + intro x ⟨i, hi⟩ + simpa only [← hi, Metric.closedBall, dist_zero_right, Set.mem_setOf_eq, ge_iff_le] using h i +/- -- NOTE: Can this really be missing from Mathlib? lemma Metric.isBounded_closedBall [PseudoMetricSpace X] (z : X) (r : ℝ) : Bornology.IsBounded (Metric.closedBall z r) := by @@ -288,6 +271,7 @@ lemma isBounded_range_of_forall_norm_le [NormedAddGroup E] (f : ι → E) (c : apply (Metric.isBounded_closedBall 0 c).subset intro x ⟨i, hi⟩ simpa only [← hi, Metric.closedBall, dist_zero_right, Set.mem_setOf_eq, gt_iff_lt] using h i + -/ end boundedness_by_norm_bounds @@ -465,12 +449,23 @@ lemma BoundedContinuousFunction.norm_integral_le_norm_of_isProbabilityMeasure -- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? -- TODO: Should this be generalized to functions with values in Banach spaces? -lemma isBounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure +lemma bounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure (μs : ι → Measure Ω) [∀ i, IsProbabilityMeasure (μs i)] (f : Ω →ᵇ ℝ) : - Bornology.IsBounded (Set.range (fun i ↦ ∫ x, (f x) ∂ (μs i))) := by - apply isBounded_range_of_forall_norm_le _ ‖f‖ + Metric.Bounded (Set.range (fun i ↦ ∫ x, (f x) ∂ (μs i))) := by + apply Metric.bounded_range_of_forall_norm_le _ ‖f‖ exact fun i ↦ f.norm_integral_le_norm_of_isProbabilityMeasure (μs i) +/- +-- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? +-- TODO: Should this be generalized to functions with values in Banach spaces? +lemma isBounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure' + (μs : ι → Measure Ω) [∀ i, IsProbabilityMeasure (μs i)] (f : Ω →ᵇ ℝ) : + Bornology.IsBounded (Set.range (fun i ↦ ∫ x, (f x) ∂ (μs i))) := by + --apply isBounded_range_of_forall_norm_le _ ‖f‖ + --exact fun i ↦ f.norm_integral_le_norm_of_isProbabilityMeasure (μs i) + sorry + -/ + lemma main_thing' {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] @@ -543,13 +538,13 @@ lemma reduction_to_liminf {ι : Type} {L : Filter ι} [NeBot L] (h : ∀ f : Ω →ᵇ ℝ, 0 ≤ f → ∫ x, (f x) ∂μ ≤ L.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i))) (f : Ω →ᵇ ℝ) : Tendsto (fun i ↦ ∫ x, (f x) ∂ (μs i)) L (𝓝 (∫ x, (f x) ∂μ)) := by - have obs := isBounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure μs f + have obs := bounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure μs f have bdd_above : IsBoundedUnder (· ≤ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := by - apply isBounded_le_map_of_isBounded_range - apply isBounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure + apply isBounded_le_map_of_bounded_range + apply bounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure have bdd_below : IsBoundedUnder (· ≥ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := by - apply isBounded_ge_map_of_isBounded_range - apply isBounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure + apply isBounded_ge_map_of_bounded_range + apply bounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure apply @tendsto_of_le_liminf_of_limsup_le ℝ ι _ _ _ L (fun i ↦ ∫ x, (f x) ∂ (μs i)) (∫ x, (f x) ∂μ) · have key := h _ (f.add_norm_nonneg) simp_rw [f.integral_add_const ‖f‖] at key @@ -562,8 +557,8 @@ lemma reduction_to_liminf {ι : Type} {L : Filter ι} [NeBot L] simp only [measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key have := liminf_const_sub L (fun i ↦ ∫ x, (f x) ∂ (μs i)) ‖f‖ bdd_above bdd_below rwa [this, sub_le_sub_iff_left] at key - · exact L.isBounded_le_map_of_isBounded_range obs - · exact L.isBounded_ge_map_of_isBounded_range obs + · exact L.isBounded_le_map_of_bounded_range obs + · exact L.isBounded_ge_map_of_bounded_range obs /-- A characterization of weak convergence of probability measures by the condition that the integrals of every continuous bounded nonnegative function converge to the integral of the function From c47250ea89171fe8a98ceb0b6687998bab4e49fe Mon Sep 17 00:00:00 2001 From: Kalle Date: Fri, 30 Jun 2023 22:45:48 +0300 Subject: [PATCH 04/41] Continuing to switch. --- ...lForPortmanteauOpenImpliesConvergence.lean | 25 +++---------------- 1 file changed, 4 insertions(+), 21 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index 912b10fc3b400..16a7d1d8dd486 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -85,33 +85,16 @@ lemma Antitone.isCoboundedUnder_ge [Preorder X] [Preorder Y] {f : X → Y} (hf : lemma Filter.isBounded_le_map_of_bounded_range (F : Filter ι) {f : ι → ℝ} (h : Metric.Bounded (Set.range f)) : (F.map f).IsBounded (· ≤ ·) := by - sorry - --rw [← Metric.bounded_iff_isBounded, Real.bounded_iff_bddBelow_bddAbove] at h - --obtain ⟨c, hc⟩ := h.2 - --apply isBoundedUnder_of - --use c - --simpa [mem_upperBounds] using hc - -lemma Filter.isBounded_ge_map_of_bounded_range (F : Filter ι) {f : ι → ℝ} - (h : Metric.Bounded (Set.range f)) : - (F.map f).IsBounded (· ≥ ·) := by sorry - --- NOTE: Missing from Mathlib? --- What would be a good generality? (Mixes order and metric, so typeclasses don't readily exist.) -lemma Filter.isBounded_le_map_of_isBounded_range (F : Filter ι) {f : ι → ℝ} - (h : Bornology.IsBounded (Set.range f)) : - (F.map f).IsBounded (· ≤ ·) := by - rw [← Metric.bounded_iff_isBounded, Real.bounded_iff_bddBelow_bddAbove] at h + rw [Real.bounded_iff_bddBelow_bddAbove] at h obtain ⟨c, hc⟩ := h.2 apply isBoundedUnder_of use c simpa [mem_upperBounds] using hc --- NOTE: Missing from Mathlib? What would be a good generality? -lemma Filter.isBounded_ge_map_of_isBounded_range (F : Filter ι) {f : ι → ℝ} - (h : Bornology.IsBounded (Set.range f)) : +lemma Filter.isBounded_ge_map_of_bounded_range (F : Filter ι) {f : ι → ℝ} + (h : Metric.Bounded (Set.range f)) : (F.map f).IsBounded (· ≥ ·) := by - rw [← Metric.bounded_iff_isBounded, Real.bounded_iff_bddBelow_bddAbove] at h + rw [Real.bounded_iff_bddBelow_bddAbove] at h obtain ⟨c, hc⟩ := h.1 apply isBoundedUnder_of use c From 310fc41f805fc9a6a27be337861ecf041d7d40be Mon Sep 17 00:00:00 2001 From: Kalle Date: Fri, 30 Jun 2023 22:47:47 +0300 Subject: [PATCH 05/41] Removing commented out bornology lemmas. --- ...lForPortmanteauOpenImpliesConvergence.lean | 38 ------------------- 1 file changed, 38 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index 16a7d1d8dd486..7fa4ed3e8248f 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -229,33 +229,6 @@ lemma Metric.bounded_range_of_forall_norm_le [NormedAddGroup E] intro x ⟨i, hi⟩ simpa only [← hi, Metric.closedBall, dist_zero_right, Set.mem_setOf_eq, ge_iff_le] using h i -/- --- NOTE: Can this really be missing from Mathlib? -lemma Metric.isBounded_closedBall [PseudoMetricSpace X] (z : X) (r : ℝ) : - Bornology.IsBounded (Metric.closedBall z r) := by - rw [Metric.isBounded_iff] - use 2 * r - intro x hx y hy - simp only [closedBall, Set.mem_setOf_eq] at hx hy - calc dist x y - _ ≤ dist x z + dist z y := dist_triangle x z y - _ = dist x z + dist y z := by rw [dist_comm z y] - _ ≤ r + r := by gcongr - _ = 2 * r := by ring - --- NOTE: Can this really be missing from Mathlib? -lemma Metric.isBounded_ball [PseudoMetricSpace X] (z : X) (r : ℝ) : - Bornology.IsBounded (Metric.ball z r) := - (Metric.isBounded_closedBall z r).subset ball_subset_closedBall - --- NOTE: Should this be in Mathlib? -lemma isBounded_range_of_forall_norm_le [NormedAddGroup E] (f : ι → E) (c : ℝ) (h : ∀ i, ‖f i‖ ≤ c) : - Bornology.IsBounded (Set.range f) := by - apply (Metric.isBounded_closedBall 0 c).subset - intro x ⟨i, hi⟩ - simpa only [← hi, Metric.closedBall, dist_zero_right, Set.mem_setOf_eq, gt_iff_lt] using h i - -/ - end boundedness_by_norm_bounds @@ -438,17 +411,6 @@ lemma bounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure apply Metric.bounded_range_of_forall_norm_le _ ‖f‖ exact fun i ↦ f.norm_integral_le_norm_of_isProbabilityMeasure (μs i) -/- --- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? --- TODO: Should this be generalized to functions with values in Banach spaces? -lemma isBounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure' - (μs : ι → Measure Ω) [∀ i, IsProbabilityMeasure (μs i)] (f : Ω →ᵇ ℝ) : - Bornology.IsBounded (Set.range (fun i ↦ ∫ x, (f x) ∂ (μs i))) := by - --apply isBounded_range_of_forall_norm_le _ ‖f‖ - --exact fun i ↦ f.norm_integral_le_norm_of_isProbabilityMeasure (μs i) - sorry - -/ - lemma main_thing' {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] From 4949de91ccc8d64fa69c45b9d6cda0d14b57eaa1 Mon Sep 17 00:00:00 2001 From: Kalle Date: Wed, 9 Aug 2023 01:07:54 +0300 Subject: [PATCH 06/41] Messing around trying to find the optimal hypotheses for liminf_add_const and 7 friends. --- ...lForPortmanteauOpenImpliesConvergence.lean | 281 +++++++----------- 1 file changed, 110 insertions(+), 171 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index 7fa4ed3e8248f..a4740d8cebe27 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -1,87 +1,24 @@ import Mathlib.MeasureTheory.Measure.Portmanteau import Mathlib.MeasureTheory.Integral.Layercake -import Mathlib.Tactic.LibrarySearch -import Mathlib.Tactic.Propose -import Mathlib.Tactic.GCongr +import Mathlib.Tactic + +/-! +# Outline of portmanteau implication: liminf condition for open sets implies weak convergence + +This file contains steps needed to prove the portmanteau implication +`le_liminf_open_implies_convergence`. Some intermediate steps need to be generalized +and cleaned up, and are to be placed in appropriate files. The main part should go +to the file `Mathlib.MeasureTheory.Measure.Portmanteau`. +-/ open MeasureTheory Filter open scoped ENNReal NNReal BoundedContinuousFunction Topology --- This file contains steps needed to prove the portmanteau implication --- `le_liminf_open_implies_convergence`. Some intermediate steps need to be generalized --- and cleaned up, and are to be placed in appropriate files. The main part should go --- to the file `Mathlib.MeasureTheory.Measure.Portmanteau`. - - - -section liminf_lemma - -variable {ι R S : Type _} {F : Filter ι} [NeBot F] - [ConditionallyCompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] - [ConditionallyCompleteLinearOrder S] [TopologicalSpace S] [OrderTopology S] - -#check Monotone.map_sSup_of_continuousAt' - --- NOTE: The result `Monotone.map_sSup_of_continuousAt'` in the library has a natural version --- generalized to *conditionally* complete linear orders. One just needs a hypothesis `BddAbove s`. -/-- A monotone function continuous at the supremum of a nonempty set sends this supremum to -the supremum of the image of this set. -/ -theorem Monotone.map_sSup_of_continuousAt'' {f : R → S} {A : Set R} (Cf : ContinuousAt f (sSup A)) - (Mf : Monotone f) (A_nonemp : A.Nonempty) (A_bdd : BddAbove A) : - f (sSup A) = sSup (f '' A) := by - --This is a particular case of the more general `IsLUB.isLUB_of_tendsto` - refine ((@IsLUB.isLUB_of_tendsto R S _ _ _ _ _ _ f A (sSup A) (f (sSup A)) (Mf.monotoneOn _) ?_ A_nonemp ?_).csSup_eq (Set.nonempty_image_iff.mpr A_nonemp)).symm - · exact isLUB_csSup A_nonemp A_bdd - · exact tendsto_nhdsWithin_of_tendsto_nhds Cf - -theorem Monotone.map_sInf_of_continuousAt'' {f : R → S} {A : Set R} (Cf : ContinuousAt f (sInf A)) - (Mf : Monotone f) (A_nonemp : A.Nonempty) (A_bdd : BddBelow A) : - f (sInf A) = sInf (f '' A) := - @Monotone.map_sSup_of_continuousAt'' Rᵒᵈ Sᵒᵈ _ _ _ _ _ _ f A Cf Mf.dual A_nonemp A_bdd - -theorem Antitone.map_sInf_of_continuousAt'' {f : R → S} {A : Set R} (Cf : ContinuousAt f (sInf A)) - (Af : Antitone f) (A_nonemp : A.Nonempty) (A_bdd : BddBelow A) : - f (sInf A) = sSup (f '' A) := - @Monotone.map_sInf_of_continuousAt'' R Sᵒᵈ _ _ _ _ _ _ f A Cf Af.dual_right A_nonemp A_bdd - -theorem Antitone.map_sSup_of_continuousAt'' {f : R → S} {A : Set R} (Cf : ContinuousAt f (sSup A)) - (Af : Antitone f) (A_nonemp : A.Nonempty) (A_bdd : BddAbove A) : - f (sSup A) = sInf (f '' A) := - @Monotone.map_sSup_of_continuousAt'' R Sᵒᵈ _ _ _ _ _ _ f A Cf Af.dual_right A_nonemp A_bdd - -#check Filter.IsBounded.isCobounded_flip - --- NOTE: Missing from Mathlib? Probably not, but what is the name...? -lemma Monotone.isCoboundedUnder_ge [Preorder X] [Preorder Y] {f : X → Y} (hf : Monotone f) {u : ι → X} - (F : Filter ι) [NeBot F] (bdd : F.IsBoundedUnder (· ≤ ·) u) : - F.IsCoboundedUnder (· ≥ ·) (f ∘ u) := by - --refine Filter.IsBounded.isCobounded_flip ?_ - obtain ⟨b, hb⟩ := bdd - refine ⟨f b, fun y hy ↦ ?_⟩ - have obs : ∀ᶠ _ in map u F, y ≤ f b := by - filter_upwards [hb, Filter.map_compose.symm ▸ hy] with x h₁ h₂ using h₂.trans (hf h₁) - exact eventually_const.mp obs - --- NOTE: Missing from Mathlib? Probably not, but what is the name...? -lemma Monotone.isCoboundedUnder_le [Preorder X] [Preorder Y] {f : X → Y} (hf : Monotone f) {u : ι → X} - (F : Filter ι) [NeBot F] (bdd : F.IsBoundedUnder (· ≥ ·) u) : - F.IsCoboundedUnder (· ≤ ·) (f ∘ u) := - @Monotone.isCoboundedUnder_ge ι Xᵒᵈ Yᵒᵈ _ _ f hf.dual u F _ bdd - --- NOTE: Missing from Mathlib? Probably not, but what is the name...? -lemma Antitone.isCoboundedUnder_le [Preorder X] [Preorder Y] {f : X → Y} (hf : Antitone f) {u : ι → X} - (F : Filter ι) [NeBot F] (bdd : F.IsBoundedUnder (· ≤ ·) u) : - F.IsCoboundedUnder (· ≤ ·) (f ∘ u) := - @Monotone.isCoboundedUnder_ge ι X Yᵒᵈ _ _ f hf u F _ bdd - --- NOTE: Missing from Mathlib? Probably not, but what is the name...? -lemma Antitone.isCoboundedUnder_ge [Preorder X] [Preorder Y] {f : X → Y} (hf : Antitone f) {u : ι → X} - (F : Filter ι) [NeBot F] (bdd : F.IsBoundedUnder (· ≥ ·) u) : - F.IsCoboundedUnder (· ≥ ·) (f ∘ u) := - @Monotone.isCoboundedUnder_le ι X Yᵒᵈ _ _ f hf u F _ bdd + -- NOTE: Missing from Mathlib? --- What would be a good generality? (Mixes order and metric, so typeclasses don't readily exist.) +-- What would be a good generality? +-- (Mixes order-boundedness and metric-boundedness, so typeclasses don't readily exist.) lemma Filter.isBounded_le_map_of_bounded_range (F : Filter ι) {f : ι → ℝ} (h : Metric.Bounded (Set.range f)) : (F.map f).IsBounded (· ≤ ·) := by @@ -100,117 +37,119 @@ lemma Filter.isBounded_ge_map_of_bounded_range (F : Filter ι) {f : ι → ℝ} use c simpa [mem_lowerBounds] using hc -#check Antitone.map_limsSup_of_continuousAt - --- NOTE: This was the key lemma to generalize. -/-- An antitone function between conditionally complete linear ordered spaces sends -a `Filter.limsSup` to the `Filter.liminf` of the image if it is continuous at the `limsSup`. -/ -theorem Antitone.map_limsSup_of_continuousAt' {F : Filter R} [NeBot F] {f : R → S} - (bdd_above : F.IsBounded (· ≤ ·)) (bdd_below : F.IsBounded (· ≥ ·)) - (f_decr : Antitone f) (f_cont : ContinuousAt f F.limsSup) : f F.limsSup = F.liminf f := by - have cobdd : F.IsCobounded (· ≤ ·) := bdd_below.isCobounded_flip - apply le_antisymm - · rw [limsSup, f_decr.map_sInf_of_continuousAt'' f_cont bdd_above cobdd] - apply le_of_forall_lt - intro c hc - simp only [liminf, limsInf, eventually_map] at hc ⊢ - obtain ⟨d, hd, h'd⟩ := exists_lt_of_lt_csSup ((@Set.nonempty_image_iff R S f _).mpr bdd_above) hc - apply lt_csSup_of_lt ?_ ?_ h'd - · exact f_decr.isCoboundedUnder_ge F bdd_below - · rcases hd with ⟨e, ⟨he, fe_eq_d⟩⟩ - filter_upwards [he] with x hx using (fe_eq_d.symm ▸ f_decr hx) - · by_cases h' : ∃ c, c < F.limsSup ∧ Set.Ioo c F.limsSup = ∅ - · rcases h' with ⟨c, c_lt, hc⟩ - have B : ∃ᶠ n in F, F.limsSup ≤ n := by - apply (frequently_lt_of_lt_limsSup cobdd c_lt).mono - intro x hx - by_contra' - have : (Set.Ioo c F.limsSup).Nonempty := ⟨x, ⟨hx, this⟩⟩ - simp only [hc, Set.not_nonempty_empty] at this - apply @liminf_le_of_frequently_le R S _ F f (f (limsSup F)) (B.mono fun x hx ↦ f_decr hx) ?_ - obtain ⟨b, hb⟩ := bdd_above - use f b - simp only [ge_iff_le, eventually_map] - filter_upwards [hb] with t ht using f_decr ht - simp only [gt_iff_lt, not_lt, ge_iff_le, not_exists, not_and] at h' - by_contra' H - obtain ⟨l, l_lt, h'l⟩ : ∃ l < F.limsSup, Set.Ioc l F.limsSup ⊆ { x : R | f x < F.liminf f } - · apply exists_Ioc_subset_of_mem_nhds ((tendsto_order.1 f_cont.tendsto).2 _ H) - by_contra con - simp only [not_exists, not_lt] at con - simpa only [lt_self_iff_false] using lt_of_le_of_lt - (@liminf_le_of_frequently_le R S _ F f (f (limsSup F)) - (frequently_of_forall (fun r ↦ f_decr (con r))) (bdd_above.isBoundedUnder f_decr)) H - obtain ⟨m, l_m, m_lt⟩ : (Set.Ioo l F.limsSup).Nonempty := by - contrapose! h' - refine' ⟨l, l_lt, by rwa [Set.not_nonempty_iff_eq_empty] at h'⟩ - have B : F.liminf f ≤ f m := by - apply @liminf_le_of_frequently_le R S _ F f (f m) - · apply (frequently_lt_of_lt_limsSup cobdd m_lt).mono - exact fun x hx ↦ f_decr hx.le - · exact IsBounded.isBoundedUnder f_decr bdd_above - have I : f m < F.liminf f := h'l ⟨l_m, m_lt.le⟩ - exact lt_irrefl _ (B.trans_lt I) - -theorem Antitone.map_limsInf_of_continuousAt' {F : Filter R} [NeBot F] {f : R → S} - (bdd_above : F.IsBoundedUnder (· ≤ ·) id) (bdd_below : F.IsBoundedUnder (· ≥ ·) id) - (f_decr : Antitone f) (f_cont : ContinuousAt f F.limsInf) : f F.limsInf = F.limsup f := - @Antitone.map_limsSup_of_continuousAt' Rᵒᵈ Sᵒᵈ _ _ _ _ _ _ F _ f - bdd_below bdd_above f_decr.dual f_cont - -theorem Monotone.map_limsSup_of_continuousAt' {F : Filter R} [NeBot F] {f : R → S} - (bdd_above : F.IsBoundedUnder (· ≤ ·) id) (bdd_below : F.IsBoundedUnder (· ≥ ·) id) - (f_incr : Monotone f) (f_cont : ContinuousAt f F.limsSup) : f F.limsSup = F.limsup f := - @Antitone.map_limsSup_of_continuousAt' R Sᵒᵈ _ _ _ _ _ _ F _ f - bdd_above bdd_below f_incr f_cont - -theorem Monotone.map_limsInf_of_continuousAt' {F : Filter R} [NeBot F] {f : R → S} - (bdd_above : F.IsBoundedUnder (· ≤ ·) id) (bdd_below : F.IsBoundedUnder (· ≥ ·) id) - (f_incr : Monotone f) (f_cont : ContinuousAt f F.limsInf) : f F.limsInf = F.liminf f := - @Antitone.map_limsSup_of_continuousAt' Rᵒᵈ S _ _ _ _ _ _ F _ f - bdd_below bdd_above f_incr.dual f_cont +section limsup_liminf_add_sub + +example : ContinuousAdd ℝ := by exact LipschitzAdd.continuousAdd -- ok. +example : ContinuousAdd ℝ≥0 := by exact LipschitzAdd.continuousAdd -- ok. +example : ContinuousAdd ℝ≥0∞ := by infer_instance -- ok. + +variable {R : Type _} [ConditionallyCompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] + +#check add_le_add_iff_right + +lemma limsup_add_const (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] + [CovariantClass R R (Function.swap fun x x_1 ↦ x + x_1) fun x x_1 ↦ x ≤ x_1] + [ContravariantClass R R (Function.swap fun x x_1 ↦ x + x_1) fun x x_1 ↦ x ≤ x_1] + (f : ι → R) (c : R) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.limsup (fun i ↦ f i + c) F = Filter.limsup f F + c := by + convert (Monotone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x + c) ?_ + (continuous_add_right c).continuousAt bdd_above bdd_below).symm + intro x y hxy + simp only + simp only [add_le_add_iff_right, hxy] + +lemma liminf_add_const (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] + [CovariantClass R R (Function.swap fun x x_1 ↦ x + x_1) fun x x_1 ↦ x ≤ x_1] + [ContravariantClass R R (Function.swap fun x x_1 ↦ x + x_1) fun x x_1 ↦ x ≤ x_1] + (f : ι → R) (c : R) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by + convert (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x + c) ?_ + (continuous_add_right c).continuousAt bdd_above bdd_below).symm + intro x y hxy + simp only [add_le_add_iff_right, hxy] + +#check AddLECancellable +#check AddLECancellable.add_le_add_iff_right +#check AddLECancellable.add_le_add_iff_left + +example : AddLECancellable (15 : ℝ≥0) := by exact Contravariant.AddLECancellable +example : AddLECancellable (15 : ℝ) := by exact Contravariant.AddLECancellable +--example : AddLECancellable (1 : ℝ≥0∞) := by exact? + +lemma ENNReal.addLECancellable_of_ne_top {c : ℝ≥0∞} (c_ne_top : c ≠ ∞) : + AddLECancellable c := fun _ _ h ↦ (ENNReal.add_le_add_iff_left c_ne_top).mp h + +--#check StrictMono.add_monotone +#check add_le_add + +lemma liminf_const_add_improved (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] + [CovariantClass R R (fun x x_1 ↦ x + x_1) fun x x_1 ↦ x ≤ x_1] (f : ι → R) (c : R) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.liminf (fun i ↦ c + f i) F = c + Filter.liminf f F := by + convert (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ c + x) ?_ + (continuous_add_left c).continuousAt bdd_above bdd_below).symm + exact fun _ _ h ↦ add_le_add_left h c + +lemma liminf_add_const_improved (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] + [CovariantClass R R (Function.swap fun x x_1 ↦ x + x_1) fun x x_1 ↦ x ≤ x_1] (f : ι → R) (c : R) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by + convert (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x + c) ?_ + (continuous_add_right c).continuousAt bdd_above bdd_below).symm + exact fun _ _ h ↦ add_le_add_right h c + +lemma liminf_add_const_real (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by + exact liminf_add_const_improved F (fun i ↦ f i) c bdd_above bdd_below + +lemma liminf_add_const_nnReal (F : Filter ι) [NeBot F] (f : ι → ℝ≥0) (c : ℝ≥0) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by + exact liminf_add_const_improved F (fun i ↦ f i) c bdd_above bdd_below + +lemma liminf_add_const_ennReal (F : Filter ι) [NeBot F] (f : ι → ℝ≥0∞) (c : ℝ≥0∞) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by + exact liminf_add_const_improved F (fun i ↦ f i) c bdd_above bdd_below +/- lemma limsup_add_const (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : Filter.limsup (fun i ↦ f i + c) F = Filter.limsup f F + c := by - have key := @Monotone.map_limsSup_of_continuousAt' ℝ ℝ _ _ _ _ _ _ (F.map f) _ (fun x ↦ x + c) bdd_above bdd_below ?_ ?_ - simp only at key - convert key.symm - · intro x y hxy - simp only [add_le_add_iff_right, hxy] - · exact (continuous_add_right c).continuousAt + convert (Monotone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : ℝ) ↦ x + c) ?_ + (continuous_add_right c).continuousAt bdd_above bdd_below).symm + intro x y hxy + simp only [add_le_add_iff_right, hxy] lemma liminf_add_const (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by - have key := @Monotone.map_limsInf_of_continuousAt' ℝ ℝ _ _ _ _ _ _ (F.map f) _ (fun x ↦ x + c) bdd_above bdd_below ?_ ?_ - simp only at key - convert key.symm - · intro x y hxy - simp only [add_le_add_iff_right, hxy] - · exact (continuous_add_right c).continuousAt + convert (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : ℝ) ↦ x + c) ?_ + (continuous_add_right c).continuousAt bdd_above bdd_below).symm + intro x y hxy + simp only [add_le_add_iff_right, hxy] lemma limsup_const_sub (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : Filter.limsup (fun i ↦ c - f i ) F = c - Filter.liminf f F := by - have key := @Antitone.map_limsInf_of_continuousAt' ℝ ℝ _ _ _ _ _ _ (F.map f) _ (fun x ↦ c - x) bdd_above bdd_below ?_ ?_ - simp only at key - convert key.symm - · intro x y hxy - simp only [sub_le_sub_iff_left, hxy] - · exact (continuous_sub_left c).continuousAt + convert (Antitone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : ℝ) ↦ c - x) ?_ + (continuous_sub_left c).continuousAt bdd_above bdd_below).symm + intro x y hxy + simp only [sub_le_sub_iff_left, hxy] lemma liminf_const_sub (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : Filter.liminf (fun i ↦ c - f i ) F = c - Filter.limsup f F := by - have key := @Antitone.map_limsSup_of_continuousAt' ℝ ℝ _ _ _ _ _ _ (F.map f) _ (fun x ↦ c - x) bdd_above bdd_below ?_ ?_ - simp only at key - convert key.symm - · intro x y hxy - simp only [sub_le_sub_iff_left, hxy] - · exact (continuous_sub_left c).continuousAt - -end liminf_lemma + convert (Antitone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : ℝ) ↦ c - x) ?_ + (continuous_sub_left c).continuousAt bdd_above bdd_below).symm + intro x y hxy + simp only [sub_le_sub_iff_left, hxy] + -/ +end limsup_liminf_add_sub -- section + From 8da382ccbaad46f787343250675730eb56294eae Mon Sep 17 00:00:00 2001 From: Kalle Date: Wed, 9 Aug 2023 01:31:29 +0300 Subject: [PATCH 07/41] limsup_const_add and 7 friends with essentially satisfactory assumptions (except with subtraction counterparts have to assume AddCommGroup and one additional covariance property). --- ...lForPortmanteauOpenImpliesConvergence.lean | 139 ++++++++---------- 1 file changed, 63 insertions(+), 76 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index a4740d8cebe27..db74917ce6f8a 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -37,117 +37,104 @@ lemma Filter.isBounded_ge_map_of_bounded_range (F : Filter ι) {f : ι → ℝ} use c simpa [mem_lowerBounds] using hc -section limsup_liminf_add_sub -example : ContinuousAdd ℝ := by exact LipschitzAdd.continuousAdd -- ok. -example : ContinuousAdd ℝ≥0 := by exact LipschitzAdd.continuousAdd -- ok. -example : ContinuousAdd ℝ≥0∞ := by infer_instance -- ok. + +section limsup_liminf_add_sub variable {R : Type _} [ConditionallyCompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] -#check add_le_add_iff_right +lemma limsup_const_add (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] + [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.limsup (fun i ↦ c + f i) F = c + Filter.limsup f F := by + convert (Monotone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ c + x) ?_ + (continuous_add_left c).continuousAt bdd_above bdd_below).symm + exact fun _ _ h ↦ add_le_add_left h c lemma limsup_add_const (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] - [CovariantClass R R (Function.swap fun x x_1 ↦ x + x_1) fun x x_1 ↦ x ≤ x_1] - [ContravariantClass R R (Function.swap fun x x_1 ↦ x + x_1) fun x x_1 ↦ x ≤ x_1] - (f : ι → R) (c : R) + [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : Filter.limsup (fun i ↦ f i + c) F = Filter.limsup f F + c := by convert (Monotone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x + c) ?_ (continuous_add_right c).continuousAt bdd_above bdd_below).symm - intro x y hxy - simp only - simp only [add_le_add_iff_right, hxy] - -lemma liminf_add_const (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] - [CovariantClass R R (Function.swap fun x x_1 ↦ x + x_1) fun x x_1 ↦ x ≤ x_1] - [ContravariantClass R R (Function.swap fun x x_1 ↦ x + x_1) fun x x_1 ↦ x ≤ x_1] - (f : ι → R) (c : R) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by - convert (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x + c) ?_ - (continuous_add_right c).continuousAt bdd_above bdd_below).symm - intro x y hxy - simp only [add_le_add_iff_right, hxy] - -#check AddLECancellable -#check AddLECancellable.add_le_add_iff_right -#check AddLECancellable.add_le_add_iff_left - -example : AddLECancellable (15 : ℝ≥0) := by exact Contravariant.AddLECancellable -example : AddLECancellable (15 : ℝ) := by exact Contravariant.AddLECancellable ---example : AddLECancellable (1 : ℝ≥0∞) := by exact? - -lemma ENNReal.addLECancellable_of_ne_top {c : ℝ≥0∞} (c_ne_top : c ≠ ∞) : - AddLECancellable c := fun _ _ h ↦ (ENNReal.add_le_add_iff_left c_ne_top).mp h - ---#check StrictMono.add_monotone -#check add_le_add + exact fun _ _ h ↦ add_le_add_right h c -lemma liminf_const_add_improved (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] - [CovariantClass R R (fun x x_1 ↦ x + x_1) fun x x_1 ↦ x ≤ x_1] (f : ι → R) (c : R) +lemma liminf_const_add (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] + [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : Filter.liminf (fun i ↦ c + f i) F = c + Filter.liminf f F := by convert (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ c + x) ?_ (continuous_add_left c).continuousAt bdd_above bdd_below).symm exact fun _ _ h ↦ add_le_add_left h c -lemma liminf_add_const_improved (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] - [CovariantClass R R (Function.swap fun x x_1 ↦ x + x_1) fun x x_1 ↦ x ≤ x_1] (f : ι → R) (c : R) +lemma liminf_add_const (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] + [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by convert (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x + c) ?_ (continuous_add_right c).continuousAt bdd_above bdd_below).symm exact fun _ _ h ↦ add_le_add_right h c +lemma limsup_const_sub (F : Filter ι) [NeBot F] [AddGroup R] [ContinuousSub R] + [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] + [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.limsup (fun i ↦ c - f i) F = c - Filter.liminf f F := by + convert (Antitone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ c - x) ?_ + (continuous_sub_left c).continuousAt bdd_above bdd_below).symm + exact fun _ _ h ↦ sub_le_sub_left h c + +-- Ok, since we are currently assuming `AddGroup R`, this is actually a special case +-- of `limsup_add_const`. But it can be convenient in its own right, and it should +-- be generalized to something less than an `AddGroup R` (when `sub_le_sub_right` is generalized). +lemma limsup_sub_const (F : Filter ι) [NeBot F] [AddGroup R] [ContinuousSub R] + [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] + [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.limsup (fun i ↦ f i - c) F = Filter.limsup f F - c := by + convert (Monotone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x - c) ?_ + (continuous_sub_right c).continuousAt bdd_above bdd_below).symm + exact fun _ _ h ↦ sub_le_sub_right h c + +lemma liminf_const_sub (F : Filter ι) [NeBot F] [AddGroup R] [ContinuousSub R] + [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] + [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.liminf (fun i ↦ c - f i) F = c - Filter.limsup f F := by + convert (Antitone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ c - x) ?_ + (continuous_sub_left c).continuousAt bdd_above bdd_below).symm + exact fun _ _ h ↦ sub_le_sub_left h c + +-- Ok, since we are currently assuming `AddGroup R`, this is actually a special case +-- of `liminf_add_const`. But it can be convenient in its own right, and it should +-- be generalized to something less than an `AddGroup R` (when `sub_le_sub_right` is generalized). +lemma liminf_sub_const (F : Filter ι) [NeBot F] [AddGroup R] [ContinuousSub R] + [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] + [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) + (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : + Filter.liminf (fun i ↦ f i - c) F = Filter.liminf f F - c := by + convert (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x - c) ?_ + (continuous_sub_right c).continuousAt bdd_above bdd_below).symm + exact fun _ _ h ↦ sub_le_sub_right h c + + +-- ### TEST CASES (The lemmas should apply to ℝ, ℝ≥0, and ℝ≥0∞.) + lemma liminf_add_const_real (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by - exact liminf_add_const_improved F (fun i ↦ f i) c bdd_above bdd_below + exact liminf_add_const F (fun i ↦ f i) c bdd_above bdd_below lemma liminf_add_const_nnReal (F : Filter ι) [NeBot F] (f : ι → ℝ≥0) (c : ℝ≥0) (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by - exact liminf_add_const_improved F (fun i ↦ f i) c bdd_above bdd_below + exact liminf_add_const F (fun i ↦ f i) c bdd_above bdd_below lemma liminf_add_const_ennReal (F : Filter ι) [NeBot F] (f : ι → ℝ≥0∞) (c : ℝ≥0∞) (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by - exact liminf_add_const_improved F (fun i ↦ f i) c bdd_above bdd_below - -/- -lemma limsup_add_const (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.limsup (fun i ↦ f i + c) F = Filter.limsup f F + c := by - convert (Monotone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : ℝ) ↦ x + c) ?_ - (continuous_add_right c).continuousAt bdd_above bdd_below).symm - intro x y hxy - simp only [add_le_add_iff_right, hxy] + exact liminf_add_const F (fun i ↦ f i) c bdd_above bdd_below -lemma liminf_add_const (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by - convert (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : ℝ) ↦ x + c) ?_ - (continuous_add_right c).continuousAt bdd_above bdd_below).symm - intro x y hxy - simp only [add_le_add_iff_right, hxy] - -lemma limsup_const_sub (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.limsup (fun i ↦ c - f i ) F = c - Filter.liminf f F := by - convert (Antitone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : ℝ) ↦ c - x) ?_ - (continuous_sub_left c).continuousAt bdd_above bdd_below).symm - intro x y hxy - simp only [sub_le_sub_iff_left, hxy] - -lemma liminf_const_sub (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.liminf (fun i ↦ c - f i ) F = c - Filter.limsup f F := by - convert (Antitone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : ℝ) ↦ c - x) ?_ - (continuous_sub_left c).continuousAt bdd_above bdd_below).symm - intro x y hxy - simp only [sub_le_sub_iff_left, hxy] - -/ end limsup_liminf_add_sub -- section From 6ccc892625bbf0e5d6f8f424dc3d4cbced7fef5e Mon Sep 17 00:00:00 2001 From: Kalle Date: Sun, 13 Aug 2023 21:38:04 +0300 Subject: [PATCH 08/41] Think about cleanup. --- ...lForPortmanteauOpenImpliesConvergence.lean | 119 ++---------------- 1 file changed, 11 insertions(+), 108 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index db74917ce6f8a..c1e0cf5794521 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -15,6 +15,7 @@ open MeasureTheory Filter open scoped ENNReal NNReal BoundedContinuousFunction Topology +#check BddBelow -- NOTE: Missing from Mathlib? -- What would be a good generality? @@ -24,125 +25,19 @@ lemma Filter.isBounded_le_map_of_bounded_range (F : Filter ι) {f : ι → ℝ} (F.map f).IsBounded (· ≤ ·) := by rw [Real.bounded_iff_bddBelow_bddAbove] at h obtain ⟨c, hc⟩ := h.2 - apply isBoundedUnder_of - use c - simpa [mem_upperBounds] using hc + refine isBoundedUnder_of ⟨c, by simpa [mem_upperBounds] using hc⟩ lemma Filter.isBounded_ge_map_of_bounded_range (F : Filter ι) {f : ι → ℝ} (h : Metric.Bounded (Set.range f)) : (F.map f).IsBounded (· ≥ ·) := by rw [Real.bounded_iff_bddBelow_bddAbove] at h obtain ⟨c, hc⟩ := h.1 - apply isBoundedUnder_of - use c - simpa [mem_lowerBounds] using hc - - - -section limsup_liminf_add_sub - -variable {R : Type _} [ConditionallyCompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] - -lemma limsup_const_add (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] - [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.limsup (fun i ↦ c + f i) F = c + Filter.limsup f F := by - convert (Monotone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ c + x) ?_ - (continuous_add_left c).continuousAt bdd_above bdd_below).symm - exact fun _ _ h ↦ add_le_add_left h c - -lemma limsup_add_const (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] - [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.limsup (fun i ↦ f i + c) F = Filter.limsup f F + c := by - convert (Monotone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x + c) ?_ - (continuous_add_right c).continuousAt bdd_above bdd_below).symm - exact fun _ _ h ↦ add_le_add_right h c - -lemma liminf_const_add (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] - [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.liminf (fun i ↦ c + f i) F = c + Filter.liminf f F := by - convert (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ c + x) ?_ - (continuous_add_left c).continuousAt bdd_above bdd_below).symm - exact fun _ _ h ↦ add_le_add_left h c - -lemma liminf_add_const (F : Filter ι) [NeBot F] [Add R] [ContinuousAdd R] - [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by - convert (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x + c) ?_ - (continuous_add_right c).continuousAt bdd_above bdd_below).symm - exact fun _ _ h ↦ add_le_add_right h c - -lemma limsup_const_sub (F : Filter ι) [NeBot F] [AddGroup R] [ContinuousSub R] - [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] - [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.limsup (fun i ↦ c - f i) F = c - Filter.liminf f F := by - convert (Antitone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ c - x) ?_ - (continuous_sub_left c).continuousAt bdd_above bdd_below).symm - exact fun _ _ h ↦ sub_le_sub_left h c - --- Ok, since we are currently assuming `AddGroup R`, this is actually a special case --- of `limsup_add_const`. But it can be convenient in its own right, and it should --- be generalized to something less than an `AddGroup R` (when `sub_le_sub_right` is generalized). -lemma limsup_sub_const (F : Filter ι) [NeBot F] [AddGroup R] [ContinuousSub R] - [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] - [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.limsup (fun i ↦ f i - c) F = Filter.limsup f F - c := by - convert (Monotone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x - c) ?_ - (continuous_sub_right c).continuousAt bdd_above bdd_below).symm - exact fun _ _ h ↦ sub_le_sub_right h c - -lemma liminf_const_sub (F : Filter ι) [NeBot F] [AddGroup R] [ContinuousSub R] - [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] - [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.liminf (fun i ↦ c - f i) F = c - Filter.limsup f F := by - convert (Antitone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ c - x) ?_ - (continuous_sub_left c).continuousAt bdd_above bdd_below).symm - exact fun _ _ h ↦ sub_le_sub_left h c - --- Ok, since we are currently assuming `AddGroup R`, this is actually a special case --- of `liminf_add_const`. But it can be convenient in its own right, and it should --- be generalized to something less than an `AddGroup R` (when `sub_le_sub_right` is generalized). -lemma liminf_sub_const (F : Filter ι) [NeBot F] [AddGroup R] [ContinuousSub R] - [CovariantClass R R (fun x y ↦ x + y) fun x y ↦ x ≤ y] - [CovariantClass R R (Function.swap fun x y ↦ x + y) fun x y ↦ x ≤ y] (f : ι → R) (c : R) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.liminf (fun i ↦ f i - c) F = Filter.liminf f F - c := by - convert (Monotone.map_limsInf_of_continuousAt (F := F.map f) (f := fun (x : R) ↦ x - c) ?_ - (continuous_sub_right c).continuousAt bdd_above bdd_below).symm - exact fun _ _ h ↦ sub_le_sub_right h c - - --- ### TEST CASES (The lemmas should apply to ℝ, ℝ≥0, and ℝ≥0∞.) - -lemma liminf_add_const_real (F : Filter ι) [NeBot F] (f : ι → ℝ) (c : ℝ) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by - exact liminf_add_const F (fun i ↦ f i) c bdd_above bdd_below - -lemma liminf_add_const_nnReal (F : Filter ι) [NeBot F] (f : ι → ℝ≥0) (c : ℝ≥0) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by - exact liminf_add_const F (fun i ↦ f i) c bdd_above bdd_below - -lemma liminf_add_const_ennReal (F : Filter ι) [NeBot F] (f : ι → ℝ≥0∞) (c : ℝ≥0∞) - (bdd_above : F.IsBoundedUnder (· ≤ ·) f) (bdd_below : F.IsBoundedUnder (· ≥ ·) f) : - Filter.liminf (fun i ↦ f i + c) F = Filter.liminf f F + c := by - exact liminf_add_const F (fun i ↦ f i) c bdd_above bdd_below - -end limsup_liminf_add_sub -- section - + apply isBoundedUnder_of ⟨c, by simpa [mem_lowerBounds] using hc⟩ section boundedness_by_norm_bounds --- TODO: Should use `Metric.Bounded` #check Metric.Bounded #check Metric.bounded_closedBall #check Metric.bounded_ball @@ -155,6 +50,14 @@ lemma Metric.bounded_range_of_forall_norm_le [NormedAddGroup E] intro x ⟨i, hi⟩ simpa only [← hi, Metric.closedBall, dist_zero_right, Set.mem_setOf_eq, ge_iff_le] using h i +-- NOTE: Should this be in Mathlib? +lemma Metric.forall_norm_le_of_bounded_range [NormedAddGroup E] + (f : ι → E) (h : Metric.Bounded (Set.range f)) : + ∃ c, ∀ i, ‖f i‖ ≤ c := by + sorry + +-- I think there were some versions of this in Mathlib already... + end boundedness_by_norm_bounds From d64675806053d6697584847bb4c23adf1797d1ee Mon Sep 17 00:00:00 2001 From: Kalle Date: Mon, 28 Aug 2023 23:41:28 +0300 Subject: [PATCH 09/41] A more reasonable proof of a triviality (still not in the right generality). --- ...lForPortmanteauOpenImpliesConvergence.lean | 46 +++++++++---------- 1 file changed, 21 insertions(+), 25 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index c1e0cf5794521..133bb845a886f 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -38,6 +38,8 @@ lemma Filter.isBounded_ge_map_of_bounded_range (F : Filter ι) {f : ι → ℝ} section boundedness_by_norm_bounds +variable {ι : Type*} {E : Type*} + #check Metric.Bounded #check Metric.bounded_closedBall #check Metric.bounded_ball @@ -64,38 +66,32 @@ end boundedness_by_norm_bounds section layercake_for_integral +variable {α : Type*} + +#check lintegral_indicator₀ +#check lintegral_indicator_const₀ + +#check AEMeasurable.nullMeasurable +#check nullMeasurableSet_lt +#check NullMeasurable + -- TODO: Generalize from ℝ to the usual type classes. --- NOTE: This is currently a mess, because of mixing Measurable and AEStronglyMeasurable. lemma Integrable.measure_pos_le_norm_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] - {f : α → ℝ} --(f_nn : 0 ≤ f) - (f_intble : Integrable f μ) - {t : ℝ} (t_pos : 0 < t) : + {f : α → ℝ} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : μ {a : α | t ≤ ‖f a‖} < ∞ := by - have f_aemble := f_intble.1.aemeasurable - have norm_f_aemble : AEMeasurable (fun a ↦ ENNReal.ofReal ‖f a‖) μ := by - --have := ENNReal.measurable_ofReal.comp (@measurable_norm ℝ _ _ _) - exact (ENNReal.measurable_ofReal.comp measurable_norm).comp_aemeasurable f_aemble - obtain ⟨g, ⟨g_mble, ⟨g_nn, aeeq_g⟩⟩⟩ := @AEMeasurable.exists_measurable_nonneg α _ μ ℝ≥0∞ _ _ _ _ - norm_f_aemble (eventually_of_forall (fun x ↦ zero_le _)) - have foo : MeasurableSet {a : α | ENNReal.ofReal t < g a} := by - sorry - -- TODO: Generalize `lintegral_indicator_const` to null-measurable sets so there is no need - -- to use g instead of f. (Have already `lintegral_indicator₀` so easy!) - have aux := @lintegral_indicator_const _ _ μ _ foo (ENNReal.ofReal t) - have markov := @mul_meas_ge_le_lintegral₀ α _ μ - (fun a ↦ ENNReal.ofReal ‖f a‖) norm_f_aemble (ENNReal.ofReal t) - have same : ∀ a, ENNReal.ofReal t ≤ ENNReal.ofReal ‖f a‖ ↔ t ≤ ‖f a‖ := by sorry - have also : ∫⁻ (a : α), ENNReal.ofReal ‖f a‖ ∂μ = ∫⁻ (a : α), ‖f a‖₊ ∂μ := by + have norm_f_aemble : AEMeasurable (fun a ↦ ENNReal.ofReal ‖f a‖) μ := + (ENNReal.measurable_ofReal.comp measurable_norm).comp_aemeasurable f_intble.1.aemeasurable + have markov := mul_meas_ge_le_lintegral₀ (μ := μ) norm_f_aemble (ENNReal.ofReal t) + have obs : ∫⁻ (a : α), ENNReal.ofReal ‖f a‖ ∂μ = ∫⁻ (a : α), ‖f a‖₊ ∂μ := by apply lintegral_congr - intro x - sorry - simp_rw [same, also] at markov + exact fun x ↦ ofReal_norm_eq_coe_nnnorm (f x) + simp_rw [ENNReal.ofReal_le_ofReal_iff (norm_nonneg _), obs] at markov have almost := lt_of_le_of_lt markov f_intble.2 have t_inv_ne_top : (ENNReal.ofReal t)⁻¹ ≠ ∞ := by exact ENNReal.inv_ne_top.mpr (ENNReal.ofReal_pos.mpr t_pos).ne.symm - convert ENNReal.mul_lt_top t_inv_ne_top almost.ne - simp [← mul_assoc, - ENNReal.inv_mul_cancel (ENNReal.ofReal_pos.mpr t_pos).ne.symm ENNReal.ofReal_ne_top] + simpa [← mul_assoc, + ENNReal.inv_mul_cancel (ENNReal.ofReal_pos.mpr t_pos).ne.symm ENNReal.ofReal_ne_top] + using ENNReal.mul_lt_top t_inv_ne_top almost.ne lemma Integrable.measure_pos_lt_norm_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] {f : α → ℝ} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : From e1918ec84c6548a85f50f9ded1a2c60db3e8321a Mon Sep 17 00:00:00 2001 From: Kalle Date: Mon, 28 Aug 2023 23:53:33 +0300 Subject: [PATCH 10/41] Probably the right generality for the triviality. --- .../NewMaterialForPortmanteauOpenImpliesConvergence.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index 133bb845a886f..28facf2092d83 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -75,9 +75,9 @@ variable {α : Type*} #check nullMeasurableSet_lt #check NullMeasurable --- TODO: Generalize from ℝ to the usual type classes. lemma Integrable.measure_pos_le_norm_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] - {f : α → ℝ} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : + {E : Type*} [NormedAddCommGroup E] [MeasurableSpace E] [BorelSpace E] + {f : α → E} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : μ {a : α | t ≤ ‖f a‖} < ∞ := by have norm_f_aemble : AEMeasurable (fun a ↦ ENNReal.ofReal ‖f a‖) μ := (ENNReal.measurable_ofReal.comp measurable_norm).comp_aemeasurable f_intble.1.aemeasurable From 7c2f5c93710ec604dc2ec0f0b69cbcf21f74af11 Mon Sep 17 00:00:00 2001 From: Kalle Date: Tue, 29 Aug 2023 00:24:20 +0300 Subject: [PATCH 11/41] Should maybe generalize the layercake file to use AEMeasurable (or AEStronglyMeasurable?). --- ...lForPortmanteauOpenImpliesConvergence.lean | 38 +++++++++++++++++-- 1 file changed, 35 insertions(+), 3 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index 28facf2092d83..b00ef6670c356 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -94,7 +94,8 @@ lemma Integrable.measure_pos_le_norm_lt_top [MeasurableSpace α] {μ : Measure using ENNReal.mul_lt_top t_inv_ne_top almost.ne lemma Integrable.measure_pos_lt_norm_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] - {f : α → ℝ} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : + {E : Type*} [NormedAddCommGroup E] [MeasurableSpace E] [BorelSpace E] + {f : α → E} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : μ {a : α | t < ‖f a‖} < ∞ := lt_of_le_of_lt (measure_mono (fun _ h ↦ (Set.mem_setOf_eq ▸ h).le)) (Integrable.measure_pos_le_norm_lt_top f_intble t_pos) @@ -102,8 +103,10 @@ lemma Integrable.measure_pos_lt_norm_lt_top [MeasurableSpace α] {μ : Measure lemma Integrable.measure_pos_le_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] {f : α → ℝ} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : μ {a : α | t ≤ f a} < ∞ := by - -- Need to do separately positive and negative parts? - sorry + refine lt_of_le_of_lt (measure_mono ?_) (Integrable.measure_pos_le_norm_lt_top f_intble t_pos) + intro x hx + simp only [Real.norm_eq_abs, Set.mem_setOf_eq] at hx ⊢ + exact hx.trans (le_abs_self _) lemma Integrable.measure_pos_lt_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] {f : α → ℝ} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : @@ -147,6 +150,35 @@ theorem integral_eq_integral_meas_lt [MeasurableSpace α] (μ : Measure α) [Sig apply Antitone.measurable exact rhs_integrand_decr +theorem integral_eq_integral_meas_lt' [MeasurableSpace α] (μ : Measure α) [SigmaFinite μ] + {f : α → ℝ} (f_nn : 0 ≤ f) (f_mble : Measurable f) (f_intble : Integrable f μ) : + (∫ ω, f ω ∂μ) = ∫ t in Set.Ioi 0, ENNReal.toReal (μ {a : α | t < f a}) := by + have key := lintegral_eq_lintegral_meas_lt μ f_nn f_mble -- should use `Integrable` + have lhs_finite : ∫⁻ (ω : α), ENNReal.ofReal (f ω) ∂μ < ∞ := Integrable.lintegral_lt_top f_intble + have rhs_finite : ∫⁻ (t : ℝ) in Set.Ioi 0, μ {a | t < f a} < ∞ := by simp only [← key, lhs_finite] + have rhs_integrand_decr : Antitone (fun t ↦ (μ {a : α | t < f a})) := + fun _ _ hst ↦ measure_mono (fun _ h ↦ lt_of_le_of_lt hst h) + have rhs_integrand_finite : ∀ (t : ℝ), t > 0 → μ {a | t < f a} < ∞ := by + exact fun t ht ↦ Integrable.measure_pos_lt_lt_top f_intble ht + convert (ENNReal.toReal_eq_toReal lhs_finite.ne rhs_finite.ne).mpr key + · refine integral_eq_lintegral_of_nonneg_ae ?_ ?_ + · -- TODO: Maybe should relax the assumption to ae nonnegativity. + exact eventually_of_forall f_nn + · --exact f_mble.aestronglyMeasurable + exact f_intble.aestronglyMeasurable + · have aux := @integral_eq_lintegral_of_nonneg_ae _ _ ((volume : Measure ℝ).restrict (Set.Ioi 0)) + (fun t ↦ ENNReal.toReal (μ {a : α | t < f a})) ?_ ?_ + · rw [aux] + apply congrArg ENNReal.toReal + apply set_lintegral_congr_fun measurableSet_Ioi + apply eventually_of_forall + exact fun t t_pos ↦ ENNReal.ofReal_toReal (rhs_integrand_finite t t_pos).ne + · exact eventually_of_forall (fun x ↦ by simp only [Pi.zero_apply, ENNReal.toReal_nonneg]) + · apply Measurable.aestronglyMeasurable + refine Measurable.ennreal_toReal ?_ + apply Antitone.measurable + exact rhs_integrand_decr + end layercake_for_integral From 534bcde358b26a49ff7f89b4d6a5eee51876a8b9 Mon Sep 17 00:00:00 2001 From: Kalle Date: Wed, 13 Sep 2023 23:42:28 +0300 Subject: [PATCH 12/41] Some fixes but VS Code is stuck...k... --- ...lForPortmanteauOpenImpliesConvergence.lean | 31 +++++++------------ 1 file changed, 11 insertions(+), 20 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index b00ef6670c356..c9d343e6840a7 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -20,14 +20,14 @@ open scoped ENNReal NNReal BoundedContinuousFunction Topology -- NOTE: Missing from Mathlib? -- What would be a good generality? -- (Mixes order-boundedness and metric-boundedness, so typeclasses don't readily exist.) -lemma Filter.isBounded_le_map_of_bounded_range (F : Filter ι) {f : ι → ℝ} +lemma Filter.isBounded_le_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} (h : Metric.Bounded (Set.range f)) : (F.map f).IsBounded (· ≤ ·) := by rw [Real.bounded_iff_bddBelow_bddAbove] at h obtain ⟨c, hc⟩ := h.2 refine isBoundedUnder_of ⟨c, by simpa [mem_upperBounds] using hc⟩ -lemma Filter.isBounded_ge_map_of_bounded_range (F : Filter ι) {f : ι → ℝ} +lemma Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} (h : Metric.Bounded (Set.range f)) : (F.map f).IsBounded (· ≥ ·) := by rw [Real.bounded_iff_bddBelow_bddAbove] at h @@ -68,13 +68,6 @@ section layercake_for_integral variable {α : Type*} -#check lintegral_indicator₀ -#check lintegral_indicator_const₀ - -#check AEMeasurable.nullMeasurable -#check nullMeasurableSet_lt -#check NullMeasurable - lemma Integrable.measure_pos_le_norm_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] {E : Type*} [NormedAddCommGroup E] [MeasurableSpace E] [BorelSpace E] {f : α → E} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : @@ -119,12 +112,10 @@ lemma Integrable.measure_pos_lt_lt_top [MeasurableSpace α] {μ : Measure α} [S -- layercake should be similarly generalized. The proofs are basically similar, but the statements -- themselves become a bit unpleasant due to integrability requirements for something slightly -- complicated. --- TODO: Should remove `Measurable` assumption and just embrace the `AEStronglyMeasurable` --- which comes from `Integrable`. This is not pleasant in the proof, but pays off for the user... theorem integral_eq_integral_meas_lt [MeasurableSpace α] (μ : Measure α) [SigmaFinite μ] - {f : α → ℝ} (f_nn : 0 ≤ f) (f_mble : Measurable f) (f_intble : Integrable f μ) : + {f : α → ℝ} (f_nn : 0 ≤ᵐ[μ] f) (f_intble : Integrable f μ) : (∫ ω, f ω ∂μ) = ∫ t in Set.Ioi 0, ENNReal.toReal (μ {a : α | t < f a}) := by - have key := lintegral_eq_lintegral_meas_lt μ f_nn f_mble -- should use `Integrable` + have key := lintegral_eq_lintegral_meas_lt μ f_nn f_intble.aemeasurable have lhs_finite : ∫⁻ (ω : α), ENNReal.ofReal (f ω) ∂μ < ∞ := Integrable.lintegral_lt_top f_intble have rhs_finite : ∫⁻ (t : ℝ) in Set.Ioi 0, μ {a | t < f a} < ∞ := by simp only [← key, lhs_finite] have rhs_integrand_decr : Antitone (fun t ↦ (μ {a : α | t < f a})) := @@ -132,11 +123,7 @@ theorem integral_eq_integral_meas_lt [MeasurableSpace α] (μ : Measure α) [Sig have rhs_integrand_finite : ∀ (t : ℝ), t > 0 → μ {a | t < f a} < ∞ := by exact fun t ht ↦ Integrable.measure_pos_lt_lt_top f_intble ht convert (ENNReal.toReal_eq_toReal lhs_finite.ne rhs_finite.ne).mpr key - · refine integral_eq_lintegral_of_nonneg_ae ?_ ?_ - · -- TODO: Maybe should relax the assumption to ae nonnegativity. - exact eventually_of_forall f_nn - · --exact f_mble.aestronglyMeasurable - exact f_intble.aestronglyMeasurable + · exact integral_eq_lintegral_of_nonneg_ae f_nn f_intble.aestronglyMeasurable · have aux := @integral_eq_lintegral_of_nonneg_ae _ _ ((volume : Measure ℝ).restrict (Set.Ioi 0)) (fun t ↦ ENNReal.toReal (μ {a : α | t < f a})) ?_ ?_ · rw [aux] @@ -150,6 +137,7 @@ theorem integral_eq_integral_meas_lt [MeasurableSpace α] (μ : Measure α) [Sig apply Antitone.measurable exact rhs_integrand_decr +/- theorem integral_eq_integral_meas_lt' [MeasurableSpace α] (μ : Measure α) [SigmaFinite μ] {f : α → ℝ} (f_nn : 0 ≤ f) (f_mble : Measurable f) (f_intble : Integrable f μ) : (∫ ω, f ω ∂μ) = ∫ t in Set.Ioi 0, ENNReal.toReal (μ {a : α | t < f a}) := by @@ -178,6 +166,7 @@ theorem integral_eq_integral_meas_lt' [MeasurableSpace α] (μ : Measure α) [Si refine Measurable.ennreal_toReal ?_ apply Antitone.measurable exact rhs_integrand_decr + -/ end layercake_for_integral @@ -206,13 +195,15 @@ lemma MeasureTheory.lintegral_mono'' {α : Type} {m : MeasurableSpace α} {μ : attribute [gcongr] lintegral_mono'' -- @[gcongr] attribute only applies to lemmas proving x₁ ~₁ x₁' → ... xₙ ~ₙ xₙ' → f x₁ ... xₙ ∼ f x₁' ... xₙ', got ∀ {α : Type u_1} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} ⦃f g : α → ℝ≥0∞⦄, f ≤ g → ∫⁻ (a : α), f a ∂μ ≤ ∫⁻ (a : α), g a ∂μ -/ +--#check lintegral_liminf_le + -- NOTE: I think I will work with real-valued integrals, after all... lemma fatou_argument_lintegral {μ : Measure Ω} [SigmaFinite μ] {μs : ℕ → Measure Ω} [∀ i, SigmaFinite (μs i)] {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : - ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫⁻ x, ENNReal.ofReal (f x) ∂ (μs i)) := by - simp_rw [lintegral_eq_lintegral_meas_lt _ f_nn f_cont.measurable] + ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫⁻ x, ENNReal.ofReal (f x) ∂ (μs i)) := by + simp_rw [lintegral_eq_lintegral_meas_lt _ (eventually_of_forall f_nn) f_cont.aemeasurable] calc ∫⁻ (t : ℝ) in Set.Ioi 0, μ {a | t < f a} ≤ ∫⁻ (t : ℝ) in Set.Ioi 0, atTop.liminf (fun i ↦ (μs i) {a | t < f a}) := (lintegral_mono (fun t ↦ h_opens _ (continuous_def.mp f_cont _ isOpen_Ioi))).trans ?_ From 26baf407d59a2e5f7ac2a2d61e73217b3137a028 Mon Sep 17 00:00:00 2001 From: kkytola <“kalle.kytola@aalto.fi”> Date: Thu, 14 Sep 2023 20:31:19 +0300 Subject: [PATCH 13/41] Clean up by moving lemmas about integrals of bounded continuous functions to one file. Generalize to usual typeclasses. --- .../Integral/BoundedContinuousFunction.lean | 188 ++++++++++++++++++ .../MeasureTheory/Measure/FiniteMeasure.lean | 40 +--- ...lForPortmanteauOpenImpliesConvergence.lean | 115 +---------- 3 files changed, 192 insertions(+), 151 deletions(-) create mode 100644 Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean diff --git a/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean new file mode 100644 index 0000000000000..b8f1f69a2db09 --- /dev/null +++ b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean @@ -0,0 +1,188 @@ +/- +Copyright (c) 2023 Kalle Kytölä. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kalle Kytölä +-/ +import Mathlib.MeasureTheory.Integral.Layercake +import Mathlib.Tactic + +/-! +# Integration of bounded continuous functions + +This file contains some results about integrals of bounded continuous functions. +-/ + +open MeasureTheory Filter +open scoped ENNReal NNReal BoundedContinuousFunction Topology + +section moved_from_FiniteMeasure +-- These were moved (now in fact only copied) from `Mathlib.MeasureTheory.Measure.FiniteMeasure`. + +variable {X : Type*} [MeasurableSpace X] [TopologicalSpace X] [OpensMeasurableSpace X] + +theorem _root_.BoundedContinuousFunction.NNReal.coe_ennreal_comp_measurable (f : X →ᵇ ℝ≥0) : + Measurable fun X => (f X : ℝ≥0∞) := + measurable_coe_nnreal_ennreal.comp f.continuous.measurable +#align bounded_continuous_function.nnreal.to_ennreal_comp_measurable BoundedContinuousFunction.NNReal.coe_ennreal_comp_measurable + +namespace MeasureTheory + +variable (μ : Measure X) [IsFiniteMeasure μ] + +theorem lintegral_lt_top_of_boundedContinuous_to_nnreal (f : X →ᵇ ℝ≥0) : + (∫⁻ X, f X ∂μ) < ∞ := by + apply IsFiniteMeasure.lintegral_lt_top_of_bounded_to_ennreal + use nndist f 0 + intro x + have key := BoundedContinuousFunction.Nnreal.upper_bound f x + rw [ENNReal.coe_le_coe] + have eq : nndist f 0 = ⟨dist f 0, dist_nonneg⟩ := by + ext + simp only [Real.coe_toNNReal', max_eq_left_iff, NNReal.coe_mk, coe_nndist] + rwa [eq] at key +#align measure_theory.lintegral_lt_top_of_bounded_continuous_to_nnreal MeasureTheory.lintegral_lt_top_of_boundedContinuous_to_nnreal + +theorem FiniteMeasure.integrable_of_boundedContinuous_to_nnreal (f : X →ᵇ ℝ≥0) : + Integrable (((↑) : ℝ≥0 → ℝ) ∘ ⇑f) μ := by + refine' ⟨(NNReal.continuous_coe.comp f.continuous).measurable.aestronglyMeasurable, _⟩ + simp only [HasFiniteIntegral, Function.comp_apply, NNReal.nnnorm_eq] + exact lintegral_lt_top_of_boundedContinuous_to_nnreal _ f +#align measure_theory.finite_measure.integrable_of_bounded_continuous_to_nnreal MeasureTheory.FiniteMeasure.integrable_of_boundedContinuous_to_nnreal + +theorem FiniteMeasure.integrable_of_boundedContinuous_to_real (f : X →ᵇ ℝ) : + Integrable (⇑f) μ := by + refine' ⟨f.continuous.measurable.aestronglyMeasurable, _⟩ + have aux : ((↑) : ℝ≥0 → ℝ) ∘ ⇑f.nnnorm = fun x => ‖f x‖ := by + ext X + simp only [Function.comp_apply, BoundedContinuousFunction.nnnorm_coeFn_eq, coe_nnnorm] + apply (hasFiniteIntegral_iff_norm f).mpr + rw [← ofReal_integral_eq_lintegral_ofReal] + · exact ENNReal.ofReal_lt_top + · exact aux ▸ integrable_of_boundedContinuous_to_nnreal μ f.nnnorm + · exact eventually_of_forall fun X => norm_nonneg (f X) +#align measure_theory.finite_measure.integrable_of_bounded_continuous_to_real MeasureTheory.FiniteMeasure.integrable_of_boundedContinuous_to_real + +end MeasureTheory + +end moved_from_FiniteMeasure + + + +section boundedness_by_norm_bounds + +lemma Metric.bounded_range_of_forall_norm_le {ι : Type*} {E : Type*} [NormedAddGroup E] + {f : ι → E} {c : ℝ} (h : ∀ i, ‖f i‖ ≤ c) : + Metric.Bounded (Set.range f) := by + apply Metric.Bounded.mono _ (@Metric.bounded_closedBall _ _ 0 c) + intro x ⟨i, hi⟩ + simpa only [← hi, Metric.closedBall, dist_zero_right, Set.mem_setOf_eq, ge_iff_le] using h i + +end boundedness_by_norm_bounds + + + +section + +variable {X : Type} [MeasurableSpace X] [TopologicalSpace X] [OpensMeasurableSpace X] + +/- +-- TODO: Is it possible to add a @[gcongr] attribute to `lintegral_mono`? + +attribute [gcongr] lintegral_mono -- @[gcongr] attribute only applies to lemmas proving x₁ ~₁ x₁' → ... xₙ ~ₙ xₙ' → f x₁ ... xₙ ∼ f x₁' ... xₙ', got ∀ {α : Type u_1} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} ⦃f g : α → ℝ≥0∞⦄, f ≤ g → ∫⁻ (a : α), f a ∂μ ≤ ∫⁻ (a : α), g a ∂μ + +lemma foo (μ : Measure X) {f g : X → ℝ≥0∞} (hfg : f ≤ g) : + ∫⁻ X, f X ∂μ ≤ ∫⁻ X, g X ∂μ := by + gcongr -- gcongr did not make progress + sorry + +-- This would solve it! + +lemma MeasureTheory.lintegral_mono'' {α : Type} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} {f g : α → ℝ≥0∞} + (hfg : f ≤ g) : lintegral μ f ≤ lintegral μ g := by sorry + +#check lintegral_mono'' + +attribute [gcongr] lintegral_mono'' -- @[gcongr] attribute only applies to lemmas proving x₁ ~₁ x₁' → ... xₙ ~ₙ xₙ' → f x₁ ... xₙ ∼ f x₁' ... xₙ', got ∀ {α : Type u_1} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} ⦃f g : α → ℝ≥0∞⦄, f ≤ g → ∫⁻ (a : α), f a ∂μ ≤ ∫⁻ (a : α), g a ∂μ + -/ + +variable {E : Type*} [NormedAddCommGroup E] [TopologicalSpace.SecondCountableTopology E] +variable [MeasurableSpace E] [BorelSpace E] + +lemma BoundedContinuousFunction.integrable (μ : Measure X) [IsFiniteMeasure μ] (f : X →ᵇ E) : + Integrable f μ := by + refine ⟨f.continuous.measurable.aestronglyMeasurable, (hasFiniteIntegral_def _ _).mp ?_⟩ + calc ∫⁻ x, ‖f x‖₊ ∂μ + _ ≤ ∫⁻ _, ‖f‖₊ ∂μ := ?_ + _ = ‖f‖₊ * (μ Set.univ) := by rw [lintegral_const] + _ < ∞ := ENNReal.mul_lt_top + ENNReal.coe_ne_top (measure_ne_top μ Set.univ) + · apply lintegral_mono + exact fun x ↦ ENNReal.coe_le_coe.mpr (nnnorm_coe_le_nnnorm f x) + +variable [NormedSpace ℝ E] + +lemma BoundedContinuousFunction.norm_integral_le_mul_norm_of_isFiniteMeasure + (μ : Measure X) [IsFiniteMeasure μ] (f : X →ᵇ E) : + ‖∫ x, (f x) ∂μ‖ ≤ ENNReal.toReal (μ Set.univ) * ‖f‖ := by + calc ‖∫ x, (f x) ∂μ‖ + _ ≤ ∫ x, ‖f x‖ ∂μ := by exact norm_integral_le_integral_norm _ + _ ≤ ∫ _, ‖f‖ ∂μ := ?_ + _ = ENNReal.toReal (μ Set.univ) • ‖f‖ := by rw [integral_const] + · apply integral_mono _ (integrable_const ‖f‖) (fun x ↦ f.norm_coe_le_norm x) + exact (integrable_norm_iff f.continuous.measurable.aestronglyMeasurable).mpr (f.integrable μ) + +lemma BoundedContinuousFunction.norm_integral_le_norm_of_isProbabilityMeasure + (μ : Measure X) [IsProbabilityMeasure μ] (f : X →ᵇ E) : + ‖∫ x, (f x) ∂μ‖ ≤ ‖f‖ := by + convert f.norm_integral_le_mul_norm_of_isFiniteMeasure μ + simp only [measure_univ, ENNReal.one_toReal, one_mul] + +lemma bounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure + {ι : Type*} (μs : ι → Measure X) [∀ i, IsProbabilityMeasure (μs i)] (f : X →ᵇ E) : + Metric.Bounded (Set.range (fun i ↦ ∫ x, (f x) ∂ (μs i))) := by + apply Metric.bounded_range_of_forall_norm_le (c := ‖f‖) + exact fun i ↦ f.norm_integral_le_norm_of_isProbabilityMeasure (μs i) + +end + +section + +variable {X : Type*} [TopologicalSpace X] + +lemma BoundedContinuousFunction.neg_norm_le (f : X →ᵇ ℝ) (x : X) : + -‖f‖ ≤ f x := (abs_le.mp (norm_coe_le_norm f x)).1 + +lemma BoundedContinuousFunction.le_norm (f : X →ᵇ ℝ) (x : X): + f x ≤ ‖f‖ := (abs_le.mp (norm_coe_le_norm f x)).2 + +lemma BoundedContinuousFunction.add_norm_nonneg (f : X →ᵇ ℝ) : + 0 ≤ f + BoundedContinuousFunction.const _ ‖f‖ := by + intro x + dsimp + linarith [(abs_le.mp (norm_coe_le_norm f x)).1] + +lemma BoundedContinuousFunction.norm_sub_nonneg (f : X →ᵇ ℝ) : + 0 ≤ BoundedContinuousFunction.const _ ‖f‖ - f := by + intro x + dsimp + linarith [(abs_le.mp (norm_coe_le_norm f x)).2] + +variable [MeasurableSpace X] [OpensMeasurableSpace X] {μ : Measure X} [IsFiniteMeasure μ] + +lemma BoundedContinuousFunction.integral_add_const (f : X →ᵇ ℝ) (c : ℝ) : + ∫ x, (f + BoundedContinuousFunction.const X c) x ∂μ = + ∫ x, f x ∂μ + ENNReal.toReal (μ (Set.univ)) • c := by + simp only [coe_add, const_toFun, Pi.add_apply, smul_eq_mul] + simp_rw [integral_add (FiniteMeasure.integrable_of_boundedContinuous_to_real _ f) + (integrable_const c)] + simp only [integral_const, smul_eq_mul] + +lemma BoundedContinuousFunction.integral_const_sub (f : X →ᵇ ℝ) (c : ℝ) : + ∫ x, (BoundedContinuousFunction.const X c - f) x ∂μ = + ENNReal.toReal (μ (Set.univ)) • c - ∫ x, f x ∂μ := by + simp only [coe_sub, const_toFun, Pi.sub_apply, smul_eq_mul] + simp_rw [integral_sub (integrable_const c) + (FiniteMeasure.integrable_of_boundedContinuous_to_real _ f)] + simp only [integral_const, smul_eq_mul] + +end diff --git a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean index 85c3857294aeb..00c37e0cbbacf 100644 --- a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean @@ -6,6 +6,7 @@ Authors: Kalle Kytölä import Mathlib.Topology.ContinuousFunction.Bounded import Mathlib.Topology.Algebra.Module.WeakDual import Mathlib.MeasureTheory.Integral.Bochner +import Mathlib.MeasureTheory.Integral.BoundedContinuousFunction #align_import measure_theory.measure.finite_measure from "leanprover-community/mathlib"@"f2ce6086713c78a7f880485f7917ea547a215982" @@ -321,25 +322,6 @@ def testAgainstNN (μ : FiniteMeasure Ω) (f : Ω →ᵇ ℝ≥0) : ℝ≥0 := (∫⁻ ω, f ω ∂(μ : Measure Ω)).toNNReal #align measure_theory.finite_measure.test_against_nn MeasureTheory.FiniteMeasure.testAgainstNN -theorem _root_.BoundedContinuousFunction.NNReal.coe_ennreal_comp_measurable {Ω : Type*} - [TopologicalSpace Ω] [MeasurableSpace Ω] [OpensMeasurableSpace Ω] (f : Ω →ᵇ ℝ≥0) : - Measurable fun ω => (f ω : ℝ≥0∞) := - measurable_coe_nnreal_ennreal.comp f.continuous.measurable -#align bounded_continuous_function.nnreal.to_ennreal_comp_measurable BoundedContinuousFunction.NNReal.coe_ennreal_comp_measurable - -theorem _root_.MeasureTheory.lintegral_lt_top_of_boundedContinuous_to_nnreal (μ : Measure Ω) - [IsFiniteMeasure μ] (f : Ω →ᵇ ℝ≥0) : (∫⁻ ω, f ω ∂μ) < ∞ := by - apply IsFiniteMeasure.lintegral_lt_top_of_bounded_to_ennreal - use nndist f 0 - intro x - have key := BoundedContinuousFunction.Nnreal.upper_bound f x - rw [ENNReal.coe_le_coe] - have eq : nndist f 0 = ⟨dist f 0, dist_nonneg⟩ := by - ext - simp only [Real.coe_toNNReal', max_eq_left_iff, NNReal.coe_mk, coe_nndist] - rwa [eq] at key -#align measure_theory.lintegral_lt_top_of_bounded_continuous_to_nnreal MeasureTheory.lintegral_lt_top_of_boundedContinuous_to_nnreal - @[simp] theorem testAgainstNN_coe_eq {μ : FiniteMeasure Ω} {f : Ω →ᵇ ℝ≥0} : (μ.testAgainstNN f : ℝ≥0∞) = ∫⁻ ω, f ω ∂(μ : Measure Ω) := @@ -668,26 +650,6 @@ condition that the integrals of all bounded continuous real-valued functions con variable {Ω : Type*} [MeasurableSpace Ω] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] -theorem integrable_of_boundedContinuous_to_nnreal (μ : Measure Ω) [IsFiniteMeasure μ] - (f : Ω →ᵇ ℝ≥0) : Integrable (((↑) : ℝ≥0 → ℝ) ∘ ⇑f) μ := by - refine' ⟨(NNReal.continuous_coe.comp f.continuous).measurable.aestronglyMeasurable, _⟩ - simp only [HasFiniteIntegral, Function.comp_apply, NNReal.nnnorm_eq] - exact lintegral_lt_top_of_boundedContinuous_to_nnreal _ f -#align measure_theory.finite_measure.integrable_of_bounded_continuous_to_nnreal MeasureTheory.FiniteMeasure.integrable_of_boundedContinuous_to_nnreal - -theorem integrable_of_boundedContinuous_to_real (μ : Measure Ω) [IsFiniteMeasure μ] (f : Ω →ᵇ ℝ) : - Integrable (⇑f) μ := by - refine' ⟨f.continuous.measurable.aestronglyMeasurable, _⟩ - have aux : ((↑) : ℝ≥0 → ℝ) ∘ ⇑f.nnnorm = fun x => ‖f x‖ := by - ext ω - simp only [Function.comp_apply, BoundedContinuousFunction.nnnorm_coeFn_eq, coe_nnnorm] - apply (hasFiniteIntegral_iff_norm f).mpr - rw [← ofReal_integral_eq_lintegral_ofReal] - · exact ENNReal.ofReal_lt_top - · exact aux ▸ integrable_of_boundedContinuous_to_nnreal μ f.nnnorm - · exact eventually_of_forall fun ω => norm_nonneg (f ω) -#align measure_theory.finite_measure.integrable_of_bounded_continuous_to_real MeasureTheory.FiniteMeasure.integrable_of_boundedContinuous_to_real - theorem _root_.BoundedContinuousFunction.integral_eq_integral_nnrealPart_sub (μ : Measure Ω) [IsFiniteMeasure μ] (f : Ω →ᵇ ℝ) : ∫ ω, f ω ∂μ = (∫ ω, (f.nnrealPart ω : ℝ) ∂μ) - ∫ ω, ((-f).nnrealPart ω : ℝ) ∂μ := by diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index c9d343e6840a7..bf008be2f4c30 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -1,5 +1,6 @@ import Mathlib.MeasureTheory.Measure.Portmanteau import Mathlib.MeasureTheory.Integral.Layercake +import Mathlib.MeasureTheory.Integral.BoundedContinuousFunction import Mathlib.Tactic /-! @@ -36,34 +37,6 @@ lemma Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : -section boundedness_by_norm_bounds - -variable {ι : Type*} {E : Type*} - -#check Metric.Bounded -#check Metric.bounded_closedBall -#check Metric.bounded_ball - --- NOTE: Should this be in Mathlib? -lemma Metric.bounded_range_of_forall_norm_le [NormedAddGroup E] - (f : ι → E) (c : ℝ) (h : ∀ i, ‖f i‖ ≤ c) : - Metric.Bounded (Set.range f) := by - apply Metric.Bounded.mono _ (@Metric.bounded_closedBall _ _ 0 c) - intro x ⟨i, hi⟩ - simpa only [← hi, Metric.closedBall, dist_zero_right, Set.mem_setOf_eq, ge_iff_le] using h i - --- NOTE: Should this be in Mathlib? -lemma Metric.forall_norm_le_of_bounded_range [NormedAddGroup E] - (f : ι → E) (h : Metric.Bounded (Set.range f)) : - ∃ c, ∀ i, ‖f i‖ ≤ c := by - sorry - --- I think there were some versions of this in Mathlib already... - -end boundedness_by_norm_bounds - - - section layercake_for_integral variable {α : Type*} @@ -220,52 +193,13 @@ lemma fatou_argument_integral_nonneg ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by sorry --- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? -lemma BoundedContinuousFunction.integrable (μ : Measure Ω) [IsFiniteMeasure μ] (f : Ω →ᵇ ℝ) : - Integrable f μ := by - refine ⟨f.continuous.measurable.aestronglyMeasurable, ?_⟩ - simp [HasFiniteIntegral] - calc ∫⁻ x, ‖f x‖₊ ∂μ - _ ≤ ∫⁻ _, ‖f‖₊ ∂μ := ?_ - _ = ‖f‖₊ * (μ Set.univ) := by rw [lintegral_const] - _ < ∞ := ENNReal.mul_lt_top - ENNReal.coe_ne_top (measure_ne_top μ Set.univ) - · apply lintegral_mono - exact fun x ↦ ENNReal.coe_le_coe.mpr (nnnorm_coe_le_nnnorm f x) - --- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? -lemma BoundedContinuousFunction.norm_integral_le_mul_norm_of_isFiniteMeasure - (μ : Measure Ω) [IsFiniteMeasure μ] (f : Ω →ᵇ ℝ) : - ‖∫ x, (f x) ∂μ‖ ≤ ENNReal.toReal (μ Set.univ) * ‖f‖ := by - calc ‖∫ x, (f x) ∂μ‖ - _ ≤ ∫ x, ‖f x‖ ∂μ := by exact norm_integral_le_integral_norm _ - _ ≤ ∫ _, ‖f‖ ∂μ := ?_ - _ = ENNReal.toReal (μ Set.univ) • ‖f‖ := by rw [integral_const] - · apply integral_mono _ (integrable_const ‖f‖) (fun x ↦ f.norm_coe_le_norm x) - exact (integrable_norm_iff f.continuous.measurable.aestronglyMeasurable).mpr (f.integrable μ) - --- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? -lemma BoundedContinuousFunction.norm_integral_le_norm_of_isProbabilityMeasure - (μ : Measure Ω) [IsProbabilityMeasure μ] (f : Ω →ᵇ ℝ) : - ‖∫ x, (f x) ∂μ‖ ≤ ‖f‖ := by - convert f.norm_integral_le_mul_norm_of_isFiniteMeasure μ - simp only [measure_univ, ENNReal.one_toReal, one_mul] - --- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? --- TODO: Should this be generalized to functions with values in Banach spaces? -lemma bounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure - (μs : ι → Measure Ω) [∀ i, IsProbabilityMeasure (μs i)] (f : Ω →ᵇ ℝ) : - Metric.Bounded (Set.range (fun i ↦ ∫ x, (f x) ∂ (μs i))) := by - apply Metric.bounded_range_of_forall_norm_le _ ‖f‖ - exact fun i ↦ f.norm_integral_le_norm_of_isProbabilityMeasure (μs i) - lemma main_thing' {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫⁻ x, ENNReal.ofReal (f x) ∂ (μs i)) := by - simp_rw [lintegral_eq_lintegral_meas_lt _ f_nn f_cont.measurable] + simp_rw [lintegral_eq_lintegral_meas_lt _ (eventually_of_forall f_nn) f_cont.aemeasurable] have obs : ∀ t, IsOpen {a : Ω | t < f a} := fun t ↦ (continuous_def.mp f_cont) _ isOpen_Ioi apply (lintegral_mono (fun t ↦ h_opens _ (obs t))).trans refine lintegral_liminf_le ?_ @@ -282,50 +216,6 @@ lemma main_thing ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by sorry --- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? -lemma BoundedContinuousFunction.neg_norm_le [TopologicalSpace X] (f : X →ᵇ ℝ) (x : X) : - -‖f‖ ≤ f x := by - exact (abs_le.mp (norm_coe_le_norm f x)).1 - --- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? -lemma BoundedContinuousFunction.le_norm [TopologicalSpace X] (f : X →ᵇ ℝ) (x : X): - f x ≤ ‖f‖ := by - exact (abs_le.mp (norm_coe_le_norm f x)).2 - --- NOTE: Where should such things be placed? In the `Portmanteau`-file only? -lemma BoundedContinuousFunction.add_norm_nonneg [TopologicalSpace X] (f : X →ᵇ ℝ) : - 0 ≤ f + BoundedContinuousFunction.const _ ‖f‖ := by - intro x - dsimp - linarith [(abs_le.mp (norm_coe_le_norm f x)).1] - --- NOTE: Where should such things be placed? In the `Portmanteau`-file only? -lemma BoundedContinuousFunction.norm_sub_nonneg [TopologicalSpace X] (f : X →ᵇ ℝ) : - 0 ≤ BoundedContinuousFunction.const _ ‖f‖ - f := by - intro x - dsimp - linarith [(abs_le.mp (norm_coe_le_norm f x)).2] - --- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? -lemma BoundedContinuousFunction.integral_add_const {μ : Measure Ω} [IsFiniteMeasure μ] - (f : Ω →ᵇ ℝ) (c : ℝ) : - ∫ x, (f + BoundedContinuousFunction.const Ω c) x ∂μ = - ∫ x, f x ∂μ + ENNReal.toReal (μ (Set.univ)) • c := by - simp only [coe_add, const_toFun, Pi.add_apply, smul_eq_mul] - simp_rw [integral_add (FiniteMeasure.integrable_of_boundedContinuous_to_real _ f) - (integrable_const c)] - simp only [integral_const, smul_eq_mul] - --- NOTE: Maybe there should be a file for lemmas about integrals of `BoundedContinuousFunction`s? -lemma BoundedContinuousFunction.integral_const_sub {μ : Measure Ω} [IsFiniteMeasure μ] - (f : Ω →ᵇ ℝ) (c : ℝ) : - ∫ x, (BoundedContinuousFunction.const Ω c - f) x ∂μ = - ENNReal.toReal (μ (Set.univ)) • c - ∫ x, f x ∂μ := by - simp only [coe_sub, const_toFun, Pi.sub_apply, smul_eq_mul] - simp_rw [integral_sub (integrable_const c) - (FiniteMeasure.integrable_of_boundedContinuous_to_real _ f)] - simp only [integral_const, smul_eq_mul] - lemma reduction_to_liminf {ι : Type} {L : Filter ι} [NeBot L] {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ι → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] (h : ∀ f : Ω →ᵇ ℝ, 0 ≤ f → ∫ x, (f x) ∂μ ≤ L.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i))) @@ -340,6 +230,7 @@ lemma reduction_to_liminf {ι : Type} {L : Filter ι} [NeBot L] apply bounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure apply @tendsto_of_le_liminf_of_limsup_le ℝ ι _ _ _ L (fun i ↦ ∫ x, (f x) ∂ (μs i)) (∫ x, (f x) ∂μ) · have key := h _ (f.add_norm_nonneg) + --have := @BoundedContinuousFunction.integral_add_const simp_rw [f.integral_add_const ‖f‖] at key simp only [measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key -- TODO: Should the case of ⊥ filter be treated separately and not included as an assumption? From 8e42e1f823d971508436820bf3f9aac5dbbbb0ba Mon Sep 17 00:00:00 2001 From: kkytola <“kalle.kytola@aalto.fi”> Date: Thu, 14 Sep 2023 20:54:12 +0300 Subject: [PATCH 14/41] Small simplifications and cleaning up. --- .../Integral/BoundedContinuousFunction.lean | 39 +++++++++++-------- 1 file changed, 23 insertions(+), 16 deletions(-) diff --git a/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean index b8f1f69a2db09..18776bc3282d3 100644 --- a/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean +++ b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean @@ -49,6 +49,7 @@ theorem FiniteMeasure.integrable_of_boundedContinuous_to_nnreal (f : X →ᵇ exact lintegral_lt_top_of_boundedContinuous_to_nnreal _ f #align measure_theory.finite_measure.integrable_of_bounded_continuous_to_nnreal MeasureTheory.FiniteMeasure.integrable_of_boundedContinuous_to_nnreal +/- theorem FiniteMeasure.integrable_of_boundedContinuous_to_real (f : X →ᵇ ℝ) : Integrable (⇑f) μ := by refine' ⟨f.continuous.measurable.aestronglyMeasurable, _⟩ @@ -61,6 +62,7 @@ theorem FiniteMeasure.integrable_of_boundedContinuous_to_real (f : X →ᵇ ℝ) · exact aux ▸ integrable_of_boundedContinuous_to_nnreal μ f.nnnorm · exact eventually_of_forall fun X => norm_nonneg (f X) #align measure_theory.finite_measure.integrable_of_bounded_continuous_to_real MeasureTheory.FiniteMeasure.integrable_of_boundedContinuous_to_real + -/ end MeasureTheory @@ -83,28 +85,38 @@ end boundedness_by_norm_bounds section -variable {X : Type} [MeasurableSpace X] [TopologicalSpace X] [OpensMeasurableSpace X] +variable {X : Type*} [MeasurableSpace X] [TopologicalSpace X] [OpensMeasurableSpace X] /- -- TODO: Is it possible to add a @[gcongr] attribute to `lintegral_mono`? -attribute [gcongr] lintegral_mono -- @[gcongr] attribute only applies to lemmas proving x₁ ~₁ x₁' → ... xₙ ~ₙ xₙ' → f x₁ ... xₙ ∼ f x₁' ... xₙ', got ∀ {α : Type u_1} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} ⦃f g : α → ℝ≥0∞⦄, f ≤ g → ∫⁻ (a : α), f a ∂μ ≤ ∫⁻ (a : α), g a ∂μ - lemma foo (μ : Measure X) {f g : X → ℝ≥0∞} (hfg : f ≤ g) : ∫⁻ X, f X ∂μ ≤ ∫⁻ X, g X ∂μ := by gcongr -- gcongr did not make progress sorry +attribute [gcongr] lintegral_mono -- @[gcongr] attribute only applies to lemmas proving x₁ ~₁ x₁' → ... xₙ ~ₙ xₙ' → f x₁ ... xₙ ∼ f x₁' ... xₙ', got ∀ {α : Type u_1} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} ⦃f g : α → ℝ≥0∞⦄, f ≤ g → ∫⁻ (a : α), f a ∂μ ≤ ∫⁻ (a : α), g a ∂μ + -- This would solve it! -lemma MeasureTheory.lintegral_mono'' {α : Type} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} {f g : α → ℝ≥0∞} - (hfg : f ≤ g) : lintegral μ f ≤ lintegral μ g := by sorry +lemma MeasureTheory.lintegral_mono'' {α : Type} {m : MeasurableSpace α} + {μ : MeasureTheory.Measure α} {f g : α → ℝ≥0∞} (hfg : f ≤ g) : + lintegral μ f ≤ lintegral μ g := + lintegral_mono hfg -#check lintegral_mono'' +attribute [gcongr] MeasureTheory.lintegral_mono'' -attribute [gcongr] lintegral_mono'' -- @[gcongr] attribute only applies to lemmas proving x₁ ~₁ x₁' → ... xₙ ~ₙ xₙ' → f x₁ ... xₙ ∼ f x₁' ... xₙ', got ∀ {α : Type u_1} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} ⦃f g : α → ℝ≥0∞⦄, f ≤ g → ∫⁻ (a : α), f a ∂μ ≤ ∫⁻ (a : α), g a ∂μ +#check lintegral_mono +#check lintegral_mono'' -/ +lemma MeasureTheory.lintegral_mono'' {α : Type*} {m : MeasurableSpace α} + {μ : MeasureTheory.Measure α} {f g : α → ℝ≥0∞} (hfg : f ≤ g) : + lintegral μ f ≤ lintegral μ g := + lintegral_mono hfg + +attribute [gcongr] MeasureTheory.lintegral_mono'' + variable {E : Type*} [NormedAddCommGroup E] [TopologicalSpace.SecondCountableTopology E] variable [MeasurableSpace E] [BorelSpace E] @@ -116,7 +128,8 @@ lemma BoundedContinuousFunction.integrable (μ : Measure X) [IsFiniteMeasure μ] _ = ‖f‖₊ * (μ Set.univ) := by rw [lintegral_const] _ < ∞ := ENNReal.mul_lt_top ENNReal.coe_ne_top (measure_ne_top μ Set.univ) - · apply lintegral_mono + · --gcongr -- or `apply lintegral_mono` + apply lintegral_mono exact fun x ↦ ENNReal.coe_le_coe.mpr (nnnorm_coe_le_nnnorm f x) variable [NormedSpace ℝ E] @@ -172,17 +185,11 @@ variable [MeasurableSpace X] [OpensMeasurableSpace X] {μ : Measure X} [IsFinite lemma BoundedContinuousFunction.integral_add_const (f : X →ᵇ ℝ) (c : ℝ) : ∫ x, (f + BoundedContinuousFunction.const X c) x ∂μ = ∫ x, f x ∂μ + ENNReal.toReal (μ (Set.univ)) • c := by - simp only [coe_add, const_toFun, Pi.add_apply, smul_eq_mul] - simp_rw [integral_add (FiniteMeasure.integrable_of_boundedContinuous_to_real _ f) - (integrable_const c)] - simp only [integral_const, smul_eq_mul] + simp [integral_add (f.integrable _) (integrable_const c)] lemma BoundedContinuousFunction.integral_const_sub (f : X →ᵇ ℝ) (c : ℝ) : ∫ x, (BoundedContinuousFunction.const X c - f) x ∂μ = ENNReal.toReal (μ (Set.univ)) • c - ∫ x, f x ∂μ := by - simp only [coe_sub, const_toFun, Pi.sub_apply, smul_eq_mul] - simp_rw [integral_sub (integrable_const c) - (FiniteMeasure.integrable_of_boundedContinuous_to_real _ f)] - simp only [integral_const, smul_eq_mul] + simp [integral_sub (integrable_const c) (f.integrable _)] end From 3b293b77e645c65e24b36cea278486cb1db7ac34 Mon Sep 17 00:00:00 2001 From: kkytola <“kalle.kytola@aalto.fi”> Date: Fri, 15 Sep 2023 21:14:34 +0300 Subject: [PATCH 15/41] Small edits. --- ...lForPortmanteauOpenImpliesConvergence.lean | 29 +++++++++++++++++-- 1 file changed, 26 insertions(+), 3 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index bf008be2f4c30..528f2ae6a8501 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -216,7 +216,17 @@ lemma main_thing ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by sorry -lemma reduction_to_liminf {ι : Type} {L : Filter ι} [NeBot L] +-- Hmm... Fatou's lemma probably requires some countability properties of the filter here. +-- But it should be more general than just `atTop` in `ℕ`, right? +lemma main_thing'' {ι : Type*} (L : Filter ι) + {μ : Measure Ω} [IsProbabilityMeasure μ] + {μs : ι → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] + {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) + (h_opens : ∀ G, IsOpen G → μ G ≤ L.liminf (fun i ↦ μs i G)) : + ∫ x, (f x) ∂μ ≤ L.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by + sorry + +lemma reduction_to_liminf {ι : Type*} {L : Filter ι} [NeBot L] {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ι → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] (h : ∀ f : Ω →ᵇ ℝ, 0 ≤ f → ∫ x, (f x) ∂μ ≤ L.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i))) (f : Ω →ᵇ ℝ) : @@ -261,13 +271,26 @@ lemma ProbabilityMeasure.tendsto_iff_forall_nonneg_integral_tendsto {γ : Type _ add_sub_cancel] have key := h g (f.add_norm_nonneg) simp [g_def] at key - simp_rw [integral_add (FiniteMeasure.integrable_of_boundedContinuous_to_real _ f) - (integrable_const ‖f‖)] at key + simp_rw [integral_add (f.integrable _) (integrable_const ‖f‖)] at key simp only [integral_const, measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key simp_rw [fx_eq] convert tendsto_add.comp (Tendsto.prod_mk_nhds key (@tendsto_const_nhds _ _ _ (-‖f‖) F)) <;> simp theorem le_liminf_open_implies_convergence + {ι : Type*} {L : Filter ι} [NeBot L] {μ : ProbabilityMeasure Ω} {μs : ι → ProbabilityMeasure Ω} + (h_opens : ∀ G, IsOpen G → μ G ≤ L.liminf (fun i ↦ μs i G)) : + L.Tendsto (fun i ↦ μs i) (𝓝 μ) := by + refine ProbabilityMeasure.tendsto_iff_forall_nonneg_integral_tendsto.mpr ?_ + intro g g_nn + apply reduction_to_liminf + --have := @reduction_to_liminf Ω _ _ _ ι + intro f f_nn + have f_nn' : 0 ≤ (f : Ω → ℝ) := fun x ↦ by simpa using f_nn x + apply main_thing f.continuous f_nn' + -- Annoying coercions to reduce to `h_opens`... + sorry + +theorem le_liminf_open_implies_convergence' {μ : ProbabilityMeasure Ω} {μs : ℕ → ProbabilityMeasure Ω} (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : atTop.Tendsto (fun i ↦ μs i) (𝓝 μ) := by From 597a41c32f3c4372441eeb510ec27121925dc86c Mon Sep 17 00:00:00 2001 From: Kalle Date: Mon, 2 Oct 2023 01:28:31 +0300 Subject: [PATCH 16/41] Towards an acceptable proof. --- ...lForPortmanteauOpenImpliesConvergence.lean | 277 +++++------------- 1 file changed, 78 insertions(+), 199 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index 528f2ae6a8501..4a971718a4658 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -15,160 +15,28 @@ to the file `Mathlib.MeasureTheory.Measure.Portmanteau`. open MeasureTheory Filter open scoped ENNReal NNReal BoundedContinuousFunction Topology - -#check BddBelow - -- NOTE: Missing from Mathlib? -- What would be a good generality? -- (Mixes order-boundedness and metric-boundedness, so typeclasses don't readily exist.) lemma Filter.isBounded_le_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} - (h : Metric.Bounded (Set.range f)) : + (h : Bornology.IsBounded (Set.range f)) : (F.map f).IsBounded (· ≤ ·) := by - rw [Real.bounded_iff_bddBelow_bddAbove] at h + rw [Real.isBounded_iff_bddBelow_bddAbove] at h obtain ⟨c, hc⟩ := h.2 refine isBoundedUnder_of ⟨c, by simpa [mem_upperBounds] using hc⟩ lemma Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} - (h : Metric.Bounded (Set.range f)) : + (h : Bornology.IsBounded (Set.range f)) : (F.map f).IsBounded (· ≥ ·) := by - rw [Real.bounded_iff_bddBelow_bddAbove] at h + rw [Real.isBounded_iff_bddBelow_bddAbove] at h obtain ⟨c, hc⟩ := h.1 apply isBoundedUnder_of ⟨c, by simpa [mem_lowerBounds] using hc⟩ -section layercake_for_integral - -variable {α : Type*} - -lemma Integrable.measure_pos_le_norm_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] - {E : Type*} [NormedAddCommGroup E] [MeasurableSpace E] [BorelSpace E] - {f : α → E} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : - μ {a : α | t ≤ ‖f a‖} < ∞ := by - have norm_f_aemble : AEMeasurable (fun a ↦ ENNReal.ofReal ‖f a‖) μ := - (ENNReal.measurable_ofReal.comp measurable_norm).comp_aemeasurable f_intble.1.aemeasurable - have markov := mul_meas_ge_le_lintegral₀ (μ := μ) norm_f_aemble (ENNReal.ofReal t) - have obs : ∫⁻ (a : α), ENNReal.ofReal ‖f a‖ ∂μ = ∫⁻ (a : α), ‖f a‖₊ ∂μ := by - apply lintegral_congr - exact fun x ↦ ofReal_norm_eq_coe_nnnorm (f x) - simp_rw [ENNReal.ofReal_le_ofReal_iff (norm_nonneg _), obs] at markov - have almost := lt_of_le_of_lt markov f_intble.2 - have t_inv_ne_top : (ENNReal.ofReal t)⁻¹ ≠ ∞ := by - exact ENNReal.inv_ne_top.mpr (ENNReal.ofReal_pos.mpr t_pos).ne.symm - simpa [← mul_assoc, - ENNReal.inv_mul_cancel (ENNReal.ofReal_pos.mpr t_pos).ne.symm ENNReal.ofReal_ne_top] - using ENNReal.mul_lt_top t_inv_ne_top almost.ne - -lemma Integrable.measure_pos_lt_norm_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] - {E : Type*} [NormedAddCommGroup E] [MeasurableSpace E] [BorelSpace E] - {f : α → E} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : - μ {a : α | t < ‖f a‖} < ∞ := - lt_of_le_of_lt (measure_mono (fun _ h ↦ (Set.mem_setOf_eq ▸ h).le)) - (Integrable.measure_pos_le_norm_lt_top f_intble t_pos) - -lemma Integrable.measure_pos_le_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] - {f : α → ℝ} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : - μ {a : α | t ≤ f a} < ∞ := by - refine lt_of_le_of_lt (measure_mono ?_) (Integrable.measure_pos_le_norm_lt_top f_intble t_pos) - intro x hx - simp only [Real.norm_eq_abs, Set.mem_setOf_eq] at hx ⊢ - exact hx.trans (le_abs_self _) - -lemma Integrable.measure_pos_lt_lt_top [MeasurableSpace α] {μ : Measure α} [SigmaFinite μ] - {f : α → ℝ} (f_intble : Integrable f μ) {t : ℝ} (t_pos : 0 < t) : - μ {a : α | t < f a} < ∞ := by - apply lt_of_le_of_lt (measure_mono ?_) (Integrable.measure_pos_le_lt_top f_intble t_pos) - exact fun x hx ↦ (Set.mem_setOf_eq ▸ hx).le - --- NOTE: This is a version of the basic "Layercake formula" for real-valued nonnegative integrands --- and Bochner integral ∫ instead of ∫⁻. I don't know if the other (more general) versions of --- layercake should be similarly generalized. The proofs are basically similar, but the statements --- themselves become a bit unpleasant due to integrability requirements for something slightly --- complicated. -theorem integral_eq_integral_meas_lt [MeasurableSpace α] (μ : Measure α) [SigmaFinite μ] - {f : α → ℝ} (f_nn : 0 ≤ᵐ[μ] f) (f_intble : Integrable f μ) : - (∫ ω, f ω ∂μ) = ∫ t in Set.Ioi 0, ENNReal.toReal (μ {a : α | t < f a}) := by - have key := lintegral_eq_lintegral_meas_lt μ f_nn f_intble.aemeasurable - have lhs_finite : ∫⁻ (ω : α), ENNReal.ofReal (f ω) ∂μ < ∞ := Integrable.lintegral_lt_top f_intble - have rhs_finite : ∫⁻ (t : ℝ) in Set.Ioi 0, μ {a | t < f a} < ∞ := by simp only [← key, lhs_finite] - have rhs_integrand_decr : Antitone (fun t ↦ (μ {a : α | t < f a})) := - fun _ _ hst ↦ measure_mono (fun _ h ↦ lt_of_le_of_lt hst h) - have rhs_integrand_finite : ∀ (t : ℝ), t > 0 → μ {a | t < f a} < ∞ := by - exact fun t ht ↦ Integrable.measure_pos_lt_lt_top f_intble ht - convert (ENNReal.toReal_eq_toReal lhs_finite.ne rhs_finite.ne).mpr key - · exact integral_eq_lintegral_of_nonneg_ae f_nn f_intble.aestronglyMeasurable - · have aux := @integral_eq_lintegral_of_nonneg_ae _ _ ((volume : Measure ℝ).restrict (Set.Ioi 0)) - (fun t ↦ ENNReal.toReal (μ {a : α | t < f a})) ?_ ?_ - · rw [aux] - apply congrArg ENNReal.toReal - apply set_lintegral_congr_fun measurableSet_Ioi - apply eventually_of_forall - exact fun t t_pos ↦ ENNReal.ofReal_toReal (rhs_integrand_finite t t_pos).ne - · exact eventually_of_forall (fun x ↦ by simp only [Pi.zero_apply, ENNReal.toReal_nonneg]) - · apply Measurable.aestronglyMeasurable - refine Measurable.ennreal_toReal ?_ - apply Antitone.measurable - exact rhs_integrand_decr - -/- -theorem integral_eq_integral_meas_lt' [MeasurableSpace α] (μ : Measure α) [SigmaFinite μ] - {f : α → ℝ} (f_nn : 0 ≤ f) (f_mble : Measurable f) (f_intble : Integrable f μ) : - (∫ ω, f ω ∂μ) = ∫ t in Set.Ioi 0, ENNReal.toReal (μ {a : α | t < f a}) := by - have key := lintegral_eq_lintegral_meas_lt μ f_nn f_mble -- should use `Integrable` - have lhs_finite : ∫⁻ (ω : α), ENNReal.ofReal (f ω) ∂μ < ∞ := Integrable.lintegral_lt_top f_intble - have rhs_finite : ∫⁻ (t : ℝ) in Set.Ioi 0, μ {a | t < f a} < ∞ := by simp only [← key, lhs_finite] - have rhs_integrand_decr : Antitone (fun t ↦ (μ {a : α | t < f a})) := - fun _ _ hst ↦ measure_mono (fun _ h ↦ lt_of_le_of_lt hst h) - have rhs_integrand_finite : ∀ (t : ℝ), t > 0 → μ {a | t < f a} < ∞ := by - exact fun t ht ↦ Integrable.measure_pos_lt_lt_top f_intble ht - convert (ENNReal.toReal_eq_toReal lhs_finite.ne rhs_finite.ne).mpr key - · refine integral_eq_lintegral_of_nonneg_ae ?_ ?_ - · -- TODO: Maybe should relax the assumption to ae nonnegativity. - exact eventually_of_forall f_nn - · --exact f_mble.aestronglyMeasurable - exact f_intble.aestronglyMeasurable - · have aux := @integral_eq_lintegral_of_nonneg_ae _ _ ((volume : Measure ℝ).restrict (Set.Ioi 0)) - (fun t ↦ ENNReal.toReal (μ {a : α | t < f a})) ?_ ?_ - · rw [aux] - apply congrArg ENNReal.toReal - apply set_lintegral_congr_fun measurableSet_Ioi - apply eventually_of_forall - exact fun t t_pos ↦ ENNReal.ofReal_toReal (rhs_integrand_finite t t_pos).ne - · exact eventually_of_forall (fun x ↦ by simp only [Pi.zero_apply, ENNReal.toReal_nonneg]) - · apply Measurable.aestronglyMeasurable - refine Measurable.ennreal_toReal ?_ - apply Antitone.measurable - exact rhs_integrand_decr - -/ - -end layercake_for_integral - - - section le_liminf_open_implies_convergence variable {Ω : Type} [MeasurableSpace Ω] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] -/- --- TODO: Is it possible to add a @[gcongr] attribute to `lintegral_mono`? - -attribute [gcongr] lintegral_mono -- @[gcongr] attribute only applies to lemmas proving x₁ ~₁ x₁' → ... xₙ ~ₙ xₙ' → f x₁ ... xₙ ∼ f x₁' ... xₙ', got ∀ {α : Type u_1} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} ⦃f g : α → ℝ≥0∞⦄, f ≤ g → ∫⁻ (a : α), f a ∂μ ≤ ∫⁻ (a : α), g a ∂μ - -lemma foo (μ : Measure Ω) {f g : Ω → ℝ≥0∞} (hfg : f ≤ g) : - ∫⁻ ω, f ω ∂μ ≤ ∫⁻ ω, g ω ∂μ := by - gcongr -- gcongr did not make progress - sorry - --- This would solve it! - -lemma MeasureTheory.lintegral_mono'' {α : Type} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} {f g : α → ℝ≥0∞} - (hfg : f ≤ g) : lintegral μ f ≤ lintegral μ g := by sorry - -#check lintegral_mono'' - -attribute [gcongr] lintegral_mono'' -- @[gcongr] attribute only applies to lemmas proving x₁ ~₁ x₁' → ... xₙ ~ₙ xₙ' → f x₁ ... xₙ ∼ f x₁' ... xₙ', got ∀ {α : Type u_1} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} ⦃f g : α → ℝ≥0∞⦄, f ≤ g → ∫⁻ (a : α), f a ∂μ ≤ ∫⁻ (a : α), g a ∂μ - -/ - ---#check lintegral_liminf_le -- NOTE: I think I will work with real-valued integrals, after all... lemma fatou_argument_lintegral @@ -191,56 +59,57 @@ lemma fatou_argument_integral_nonneg {f : Ω →ᵇ ℝ} (f_nn : 0 ≤ f) (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by - sorry - -lemma main_thing' - {μ : Measure Ω} [IsProbabilityMeasure μ] - {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] - {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) - (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : - ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫⁻ x, ENNReal.ofReal (f x) ∂ (μs i)) := by - simp_rw [lintegral_eq_lintegral_meas_lt _ (eventually_of_forall f_nn) f_cont.aemeasurable] - have obs : ∀ t, IsOpen {a : Ω | t < f a} := fun t ↦ (continuous_def.mp f_cont) _ isOpen_Ioi - apply (lintegral_mono (fun t ↦ h_opens _ (obs t))).trans - refine lintegral_liminf_le ?_ - refine fun i ↦ Antitone.measurable (fun s t hst ↦ measure_mono ?_) - intro ω hω - simp only [Set.mem_setOf_eq] at * - linarith - -lemma main_thing - {μ : Measure Ω} [IsProbabilityMeasure μ] - {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] - {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) + have earlier := fatou_argument_lintegral f.continuous f_nn h_opens + rw [@integral_eq_lintegral_of_nonneg_ae Ω _ μ f (eventually_of_forall f_nn) + f.continuous.measurable.aestronglyMeasurable] + convert (ENNReal.toReal_le_toReal ?_ ?_).mpr earlier + · simp only [fun i ↦ @integral_eq_lintegral_of_nonneg_ae Ω _ (μs i) f (eventually_of_forall f_nn) + f.continuous.measurable.aestronglyMeasurable] + --have := @Monotone.map_liminf_of_continuousAt + sorry + · sorry + · sorry + +-- A direct proof attempt (should be discarded). +lemma fatou_argument_integral_nonneg' + {μ : Measure Ω} [IsFiniteMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsFiniteMeasure (μs i)] + {f : Ω →ᵇ ℝ} (f_nn : 0 ≤ f) (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by - sorry - --- Hmm... Fatou's lemma probably requires some countability properties of the filter here. --- But it should be more general than just `atTop` in `ℕ`, right? -lemma main_thing'' {ι : Type*} (L : Filter ι) - {μ : Measure Ω} [IsProbabilityMeasure μ] - {μs : ι → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] - {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) - (h_opens : ∀ G, IsOpen G → μ G ≤ L.liminf (fun i ↦ μs i G)) : - ∫ x, (f x) ∂μ ≤ L.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by - sorry + simp_rw [(f.integrable _).integral_eq_integral_meas_lt (eventually_of_forall f_nn)] + have rwr := @integral_eq_lintegral_of_nonneg_ae ℝ _ (volume.restrict (Set.Ioi 0)) + (fun t ↦ ENNReal.toReal (μ {x | t < f x})) (eventually_of_forall ?_) ?_ + have rwr' := fun i ↦ @integral_eq_lintegral_of_nonneg_ae ℝ _ (volume.restrict (Set.Ioi 0)) + (fun t ↦ ENNReal.toReal (μs i {x | t < f x})) (eventually_of_forall ?_) ?_ + simp only [rwr, ne_eq, rwr', ge_iff_le] + · --have aux : ∀ t, IsOpen {a : Ω | t < f a} := + -- fun t ↦ (continuous_def.mp f.continuous) _ isOpen_Ioi + --have obs := fun t ↦ h_opens _ (aux t) + --have := @lintegral_liminf_le + + --have earlier := fatou_argument_lintegral f.continuous f_nn h_opens + sorry + · exact fun _ ↦ by simp only [Pi.zero_apply, ENNReal.toReal_nonneg] + · apply Measurable.aestronglyMeasurable + apply ENNReal.measurable_toReal.comp (Antitone.measurable ?_) + exact fun x y hxy ↦ measure_mono (fun z hz ↦ lt_of_le_of_lt hxy hz) + · exact fun _ ↦ by simp only [Pi.zero_apply, ENNReal.toReal_nonneg] + · apply Measurable.aestronglyMeasurable + apply ENNReal.measurable_toReal.comp (Antitone.measurable ?_) + exact fun x y hxy ↦ measure_mono (fun z hz ↦ lt_of_le_of_lt hxy hz) lemma reduction_to_liminf {ι : Type*} {L : Filter ι} [NeBot L] {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ι → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] (h : ∀ f : Ω →ᵇ ℝ, 0 ≤ f → ∫ x, (f x) ∂μ ≤ L.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i))) (f : Ω →ᵇ ℝ) : Tendsto (fun i ↦ ∫ x, (f x) ∂ (μs i)) L (𝓝 (∫ x, (f x) ∂μ)) := by - have obs := bounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure μs f - have bdd_above : IsBoundedUnder (· ≤ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := by - apply isBounded_le_map_of_bounded_range - apply bounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure - have bdd_below : IsBoundedUnder (· ≥ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := by - apply isBounded_ge_map_of_bounded_range - apply bounded_range_integral_boundedContinuousFunction_of_isProbabilityMeasure + have obs := BoundedContinuousFunction.isBounded_range_integral μs f + have bdd_above : IsBoundedUnder (· ≤ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := + isBounded_le_map_of_bounded_range _ obs + have bdd_below : IsBoundedUnder (· ≥ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := + isBounded_ge_map_of_bounded_range _ obs apply @tendsto_of_le_liminf_of_limsup_le ℝ ι _ _ _ L (fun i ↦ ∫ x, (f x) ∂ (μs i)) (∫ x, (f x) ∂μ) · have key := h _ (f.add_norm_nonneg) - --have := @BoundedContinuousFunction.integral_add_const simp_rw [f.integral_add_const ‖f‖] at key simp only [measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key -- TODO: Should the case of ⊥ filter be treated separately and not included as an assumption? @@ -251,9 +120,10 @@ lemma reduction_to_liminf {ι : Type*} {L : Filter ι} [NeBot L] simp only [measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key have := liminf_const_sub L (fun i ↦ ∫ x, (f x) ∂ (μs i)) ‖f‖ bdd_above bdd_below rwa [this, sub_le_sub_iff_left] at key - · exact L.isBounded_le_map_of_bounded_range obs - · exact L.isBounded_ge_map_of_bounded_range obs + · exact bdd_above + · exact bdd_below +-- Not needed? /-- A characterization of weak convergence of probability measures by the condition that the integrals of every continuous bounded nonnegative function converge to the integral of the function against the limit measure. -/ @@ -277,30 +147,39 @@ lemma ProbabilityMeasure.tendsto_iff_forall_nonneg_integral_tendsto {γ : Type _ convert tendsto_add.comp (Tendsto.prod_mk_nhds key (@tendsto_const_nhds _ _ _ (-‖f‖) F)) <;> simp theorem le_liminf_open_implies_convergence - {ι : Type*} {L : Filter ι} [NeBot L] {μ : ProbabilityMeasure Ω} {μs : ι → ProbabilityMeasure Ω} - (h_opens : ∀ G, IsOpen G → μ G ≤ L.liminf (fun i ↦ μs i G)) : - L.Tendsto (fun i ↦ μs i) (𝓝 μ) := by - refine ProbabilityMeasure.tendsto_iff_forall_nonneg_integral_tendsto.mpr ?_ - intro g g_nn - apply reduction_to_liminf - --have := @reduction_to_liminf Ω _ _ _ ι - intro f f_nn - have f_nn' : 0 ≤ (f : Ω → ℝ) := fun x ↦ by simpa using f_nn x - apply main_thing f.continuous f_nn' - -- Annoying coercions to reduce to `h_opens`... - sorry - -theorem le_liminf_open_implies_convergence' - {μ : ProbabilityMeasure Ω} {μs : ℕ → ProbabilityMeasure Ω} + {ι : Type*} {μ : ProbabilityMeasure Ω} {μs : ℕ → ProbabilityMeasure Ω} (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : atTop.Tendsto (fun i ↦ μs i) (𝓝 μ) := by - refine ProbabilityMeasure.tendsto_iff_forall_nonneg_integral_tendsto.mpr ?_ - intro g g_nn + refine ProbabilityMeasure.tendsto_iff_forall_integral_tendsto.mpr ?_ apply reduction_to_liminf intro f f_nn - have f_nn' : 0 ≤ (f : Ω → ℝ) := fun x ↦ by simpa using f_nn x - apply main_thing f.continuous f_nn' - -- Annoying coercions to reduce to `h_opens`... - sorry + apply fatou_argument_integral_nonneg (f := f) f_nn + intro G G_open + specialize h_opens G G_open + simp only at h_opens + simp only [liminf, limsInf, eventually_map, eventually_atTop, ge_iff_le, le_sSup_iff] at * + intro b b_ub + by_cases b_eq_top : b = ∞ + · simp only [b_eq_top, le_top] + · have foo := (@le_csSup_iff ℝ≥0 _ {a | ∃ a_1, ∀ (b : ℕ), a_1 ≤ b → a ≤ ENNReal.toNNReal (μs b G)} + (ENNReal.toNNReal (μ G)) ?_ ?_).mp h_opens (ENNReal.toNNReal b) ?_ + · simp only [ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure] at foo + convert ENNReal.coe_le_coe.mpr foo + · simp only [ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure] + · simp only [ne_eq, b_eq_top, not_false_eq_true, ENNReal.coe_toNNReal] + · use 1 + simp [mem_upperBounds] + intro x n hn + exact (hn n (le_refl n)).trans (ProbabilityMeasure.apply_le_one _ _) + · refine ⟨0, by simp⟩ + · simp only [mem_upperBounds, Set.mem_setOf_eq, forall_exists_index, ne_eq, + ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure] at b_ub ⊢ + intro x n hn + specialize b_ub x n ?_ + · intro m hm + convert ENNReal.coe_le_coe.mpr (hn m hm) + simp only [ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure] + · rw [(ENNReal.coe_toNNReal b_eq_top).symm] at b_ub + exact ENNReal.coe_le_coe.mp b_ub end le_liminf_open_implies_convergence From 8d1aa9d9f6b09ec0e457d9d43472eb7f5c4510e5 Mon Sep 17 00:00:00 2001 From: Kalle Date: Wed, 4 Oct 2023 00:46:28 +0300 Subject: [PATCH 17/41] Mostly reduced to a MonotoneOn version of map_liminf lemma. --- ...lForPortmanteauOpenImpliesConvergence.lean | 113 +++++++++++++++++- 1 file changed, 108 insertions(+), 5 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index 4a971718a4658..d22247b7f2bee 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -15,6 +15,89 @@ to the file `Mathlib.MeasureTheory.Measure.Portmanteau`. open MeasureTheory Filter open scoped ENNReal NNReal BoundedContinuousFunction Topology + +section yet_another_map_liminf_lemma + +variable {ι R S : Type*} {F : Filter ι} [NeBot F] + [ConditionallyCompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] + [ConditionallyCompleteLinearOrder S] [TopologicalSpace S] [OrderTopology S] + +/-- An antitone function between (conditionally) complete linear ordered spaces sends a +`Filter.limsSup` to the `Filter.liminf` of the image if the function is continuous at the `limsSup` +(and the filter is bounded from above and below). -/ +theorem AntitoneOn.map_limsSup_of_continuousAt {F : Filter R} [NeBot F] {f : R → S} + {R' : Set R} (f_decr : AntitoneOn f R') (f_cont : ContinuousAt f F.limsSup) + (hF : Filter.principal R' ≤ F) + (bdd_above : F.IsBounded (· ≤ ·) := by isBoundedDefault) + (bdd_below : F.IsBounded (· ≥ ·) := by isBoundedDefault) : + f F.limsSup = F.liminf f := by + sorry +/- + have cobdd : F.IsCobounded (· ≤ ·) := bdd_below.isCobounded_flip + apply le_antisymm + · rw [limsSup, f_decr.map_sInf_of_continuousAt' f_cont bdd_above cobdd] + apply le_of_forall_lt + intro c hc + simp only [liminf, limsInf, eventually_map] at hc ⊢ + obtain ⟨d, hd, h'd⟩ := + exists_lt_of_lt_csSup (bdd_above.recOn fun x hx ↦ ⟨f x, Set.mem_image_of_mem f hx⟩) hc + apply lt_csSup_of_lt ?_ ?_ h'd + · exact (Antitone.isBoundedUnder_le_comp f_decr bdd_below).isCoboundedUnder_flip + · rcases hd with ⟨e, ⟨he, fe_eq_d⟩⟩ + filter_upwards [he] with x hx using (fe_eq_d.symm ▸ f_decr hx) + · by_cases h' : ∃ c, c < F.limsSup ∧ Set.Ioo c F.limsSup = ∅ + · rcases h' with ⟨c, c_lt, hc⟩ + have B : ∃ᶠ n in F, F.limsSup ≤ n := by + apply (frequently_lt_of_lt_limsSup cobdd c_lt).mono + intro x hx + by_contra' + have : (Set.Ioo c F.limsSup).Nonempty := ⟨x, ⟨hx, this⟩⟩ + simp only [hc, Set.not_nonempty_empty] at this + apply liminf_le_of_frequently_le _ (bdd_above.isBoundedUnder f_decr) + exact (B.mono fun x hx ↦ f_decr hx) + push_neg at h' + by_contra' H + have not_bot : ¬ IsBot F.limsSup := fun maybe_bot ↦ + lt_irrefl (F.liminf f) <| lt_of_le_of_lt + (liminf_le_of_frequently_le (frequently_of_forall (fun r ↦ f_decr (maybe_bot r))) + (bdd_above.isBoundedUnder f_decr)) H + obtain ⟨l, l_lt, h'l⟩ : ∃ l < F.limsSup, Set.Ioc l F.limsSup ⊆ { x : R | f x < F.liminf f } + · apply exists_Ioc_subset_of_mem_nhds ((tendsto_order.1 f_cont.tendsto).2 _ H) + simpa [IsBot] using not_bot + obtain ⟨m, l_m, m_lt⟩ : (Set.Ioo l F.limsSup).Nonempty := by + contrapose! h' + refine' ⟨l, l_lt, by rwa [Set.not_nonempty_iff_eq_empty] at h'⟩ + have B : F.liminf f ≤ f m := by + apply liminf_le_of_frequently_le _ _ + · apply (frequently_lt_of_lt_limsSup cobdd m_lt).mono + exact fun x hx ↦ f_decr hx.le + · exact IsBounded.isBoundedUnder f_decr bdd_above + have I : f m < F.liminf f := h'l ⟨l_m, m_lt.le⟩ + exact lt_irrefl _ (B.trans_lt I) + -/ + +theorem MonotoneOn.map_limsInf_of_continuousAt {F : Filter R} [NeBot F] {f : R → S} + {R' : Set R} (f_incr : MonotoneOn f R') (f_cont : ContinuousAt f F.limsInf) + (hF : Filter.principal R' ≤ F) + (bdd_above : F.IsBounded (· ≤ ·) := by isBoundedDefault) + (bdd_below : F.IsBounded (· ≥ ·) := by isBoundedDefault) : + f F.limsInf = F.liminf f := by + exact @AntitoneOn.map_limsSup_of_continuousAt Rᵒᵈ S _ _ _ _ F _ f R' + (antitone_on_dual_iff.mp f_incr) f_cont hF bdd_below bdd_above + +theorem MonotoneOn.map_liminf_of_continuousAt {F : Filter ι} [NeBot F] + {f : R → S} {R' : Set R} (f_incr : MonotoneOn f R') + (a : ι → R) (f_cont : ContinuousAt f (F.liminf a)) (hF : Filter.principal R' ≤ F.map a) + (bdd_above : F.IsBoundedUnder (· ≤ ·) a := by isBoundedDefault) + (bdd_below : F.IsBoundedUnder (· ≥ ·) a := by isBoundedDefault) : + f (F.liminf a) = F.liminf (f ∘ a) := by + apply @MonotoneOn.map_limsInf_of_continuousAt R S _ _ _ _ (F.map a) _ f R' f_incr f_cont hF + bdd_above bdd_below + +end yet_another_map_liminf_lemma + + + -- NOTE: Missing from Mathlib? -- What would be a good generality? -- (Mixes order-boundedness and metric-boundedness, so typeclasses don't readily exist.) @@ -55,7 +138,7 @@ lemma fatou_argument_lintegral -- NOTE: I think this is the version I prefer to use, after all... lemma fatou_argument_integral_nonneg - {μ : Measure Ω} [IsFiniteMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsFiniteMeasure (μs i)] + {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] {f : Ω →ᵇ ℝ} (f_nn : 0 ≤ f) (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by @@ -65,10 +148,30 @@ lemma fatou_argument_integral_nonneg convert (ENNReal.toReal_le_toReal ?_ ?_).mpr earlier · simp only [fun i ↦ @integral_eq_lintegral_of_nonneg_ae Ω _ (μs i) f (eventually_of_forall f_nn) f.continuous.measurable.aestronglyMeasurable] - --have := @Monotone.map_liminf_of_continuousAt - sorry - · sorry - · sorry + have aux : MonotoneOn ENNReal.toReal {∞}ᶜ := by sorry + apply (@MonotoneOn.map_liminf_of_continuousAt ℕ ℝ≥0∞ ℝ _ _ _ _ atTop _ ENNReal.toReal {∞}ᶜ aux + (fun i ↦ ∫⁻ (x : Ω), ENNReal.ofReal (f x) ∂(μs i)) ?_ ?_ ?_ ?_).symm + · --have := ENNReal.continuousOn_toNNReal + sorry + · sorry + · exact ⟨∞, eventually_of_forall (fun x ↦ by simp only [le_top])⟩ + · exact ⟨0, eventually_of_forall (fun x ↦ by simp only [ge_iff_le, zero_le])⟩ + · exact (f.lintegral_of_real_lt_top μ).ne + · apply ne_of_lt + have obs := fun (i : ℕ) ↦ @BoundedContinuousFunction.lintegral_nnnorm_le Ω _ _ (μs i) ℝ _ f + simp at obs + apply lt_of_le_of_lt _ (show (‖f‖₊ : ℝ≥0∞) < ∞ from ENNReal.coe_lt_top) + apply liminf_le_of_le + · refine ⟨0, eventually_of_forall (by simp only [ge_iff_le, zero_le, forall_const])⟩ + · intro x hx + obtain ⟨i, hi⟩ := hx.exists + apply le_trans hi + convert obs i with x + have aux := ENNReal.ofReal_eq_coe_nnreal (f_nn x) + simp only [ContinuousMap.toFun_eq_coe, BoundedContinuousFunction.coe_to_continuous_fun] at aux + rw [aux] + congr + exact (Real.norm_of_nonneg (f_nn x)).symm -- A direct proof attempt (should be discarded). lemma fatou_argument_integral_nonneg' From 38a8c70967f095c8cef09842138cda6cc8eaf521 Mon Sep 17 00:00:00 2001 From: Kalle Date: Sat, 21 Oct 2023 21:34:30 +0300 Subject: [PATCH 18/41] Filling in stuff (coercion heavy). Oops with filter le. --- ...lForPortmanteauOpenImpliesConvergence.lean | 29 +++++++++++++++++-- 1 file changed, 27 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index d22247b7f2bee..994dd56b20e56 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -136,6 +136,16 @@ lemma fatou_argument_lintegral (fun s t hst ↦ measure_mono (fun ω hω ↦ lt_of_le_of_lt hst hω))) rfl +theorem BoundedContinuousFunction.lintegral_le_edist_mul + {μ : Measure Ω} [IsFiniteMeasure μ] (f : Ω →ᵇ ℝ≥0) : + (∫⁻ x, f x ∂μ) ≤ edist 0 f * (μ Set.univ) := by + have bound : ∀ x, f x ≤ nndist 0 f := by + intro x + convert nndist_coe_le_nndist x + simp only [coe_zero, Pi.zero_apply, NNReal.nndist_zero_eq_val] + apply le_trans (lintegral_mono (fun x ↦ ENNReal.coe_le_coe.mpr (bound x))) + simp + -- NOTE: I think this is the version I prefer to use, after all... lemma fatou_argument_integral_nonneg {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] @@ -151,9 +161,24 @@ lemma fatou_argument_integral_nonneg have aux : MonotoneOn ENNReal.toReal {∞}ᶜ := by sorry apply (@MonotoneOn.map_liminf_of_continuousAt ℕ ℝ≥0∞ ℝ _ _ _ _ atTop _ ENNReal.toReal {∞}ᶜ aux (fun i ↦ ∫⁻ (x : Ω), ENNReal.ofReal (f x) ∂(μs i)) ?_ ?_ ?_ ?_).symm - · --have := ENNReal.continuousOn_toNNReal + have lip : LipschitzWith 1 Real.toNNReal := by sorry + let g := BoundedContinuousFunction.comp _ lip f + have bound : ∀ i, ∫⁻ x, ENNReal.ofReal (f x) ∂(μs i) ≤ nndist 0 g := fun i ↦ by + simpa only [coe_nnreal_ennreal_nndist, measure_univ, mul_one, ge_iff_le] using + BoundedContinuousFunction.lintegral_le_edist_mul (μ := μs i) g + · apply (NNReal.continuous_coe.comp_continuousOn ENNReal.continuousOn_toNNReal).continuousAt + have obs : {x : ℝ≥0∞ | x ≠ ∞} = Set.Iio ∞ := by sorry + rw [obs] + apply Iio_mem_nhds + have lip : LipschitzWith 1 Real.toNNReal := by sorry + let g := BoundedContinuousFunction.comp _ lip f + apply lt_of_le_of_lt (liminf_le_of_frequently_le (frequently_of_forall bound)) + exact ENNReal.coe_lt_top + · --rw [le_map_iff] + --intro s hs + --simp + --rw [principal_le_iff] sorry - · sorry · exact ⟨∞, eventually_of_forall (fun x ↦ by simp only [le_top])⟩ · exact ⟨0, eventually_of_forall (fun x ↦ by simp only [ge_iff_le, zero_le])⟩ · exact (f.lintegral_of_real_lt_top μ).ne From 1d1c447726a4a9d274249a395df86cb5188afef2 Mon Sep 17 00:00:00 2001 From: Kalle Date: Sat, 21 Oct 2023 21:55:15 +0300 Subject: [PATCH 19/41] Fix the filter le and separate two lemmas that should exist. --- ...lForPortmanteauOpenImpliesConvergence.lean | 40 +++++++++++-------- 1 file changed, 24 insertions(+), 16 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index 994dd56b20e56..ddb5b263df0ea 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -27,7 +27,7 @@ variable {ι R S : Type*} {F : Filter ι} [NeBot F] (and the filter is bounded from above and below). -/ theorem AntitoneOn.map_limsSup_of_continuousAt {F : Filter R} [NeBot F] {f : R → S} {R' : Set R} (f_decr : AntitoneOn f R') (f_cont : ContinuousAt f F.limsSup) - (hF : Filter.principal R' ≤ F) + (hF : F ≤ Filter.principal R') (bdd_above : F.IsBounded (· ≤ ·) := by isBoundedDefault) (bdd_below : F.IsBounded (· ≥ ·) := by isBoundedDefault) : f F.limsSup = F.liminf f := by @@ -78,7 +78,7 @@ theorem AntitoneOn.map_limsSup_of_continuousAt {F : Filter R} [NeBot F] {f : R theorem MonotoneOn.map_limsInf_of_continuousAt {F : Filter R} [NeBot F] {f : R → S} {R' : Set R} (f_incr : MonotoneOn f R') (f_cont : ContinuousAt f F.limsInf) - (hF : Filter.principal R' ≤ F) + (hF : F ≤ Filter.principal R') (bdd_above : F.IsBounded (· ≤ ·) := by isBoundedDefault) (bdd_below : F.IsBounded (· ≥ ·) := by isBoundedDefault) : f F.limsInf = F.liminf f := by @@ -87,7 +87,7 @@ theorem MonotoneOn.map_limsInf_of_continuousAt {F : Filter R} [NeBot F] {f : R theorem MonotoneOn.map_liminf_of_continuousAt {F : Filter ι} [NeBot F] {f : R → S} {R' : Set R} (f_incr : MonotoneOn f R') - (a : ι → R) (f_cont : ContinuousAt f (F.liminf a)) (hF : Filter.principal R' ≤ F.map a) + (a : ι → R) (f_cont : ContinuousAt f (F.liminf a)) (hF : F.map a ≤ Filter.principal R') (bdd_above : F.IsBoundedUnder (· ≤ ·) a := by isBoundedDefault) (bdd_below : F.IsBoundedUnder (· ≥ ·) a := by isBoundedDefault) : f (F.liminf a) = F.liminf (f ∘ a) := by @@ -146,6 +146,17 @@ theorem BoundedContinuousFunction.lintegral_le_edist_mul apply le_trans (lintegral_mono (fun x ↦ ENNReal.coe_le_coe.mpr (bound x))) simp +lemma ENNReal.monotoneOn_toReal : MonotoneOn ENNReal.toReal {∞}ᶜ := + fun _ _ _ hy x_le_y ↦ toReal_mono hy x_le_y + +-- Argh. :| +lemma Real.lipschitzWith_toNNReal : LipschitzWith 1 Real.toNNReal := by + apply lipschitzWith_iff_dist_le_mul.mpr + intro x y + simp only [NNReal.coe_one, one_mul, NNReal.dist_eq, coe_toNNReal', ge_iff_le, Real.dist_eq] + simpa only [ge_iff_le, NNReal.coe_one, dist_prod_same_right, one_mul, Real.dist_eq] using + lipschitzWith_iff_dist_le_mul.mp lipschitzWith_max (x, 0) (y, 0) + -- NOTE: I think this is the version I prefer to use, after all... lemma fatou_argument_integral_nonneg {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] @@ -158,27 +169,24 @@ lemma fatou_argument_integral_nonneg convert (ENNReal.toReal_le_toReal ?_ ?_).mpr earlier · simp only [fun i ↦ @integral_eq_lintegral_of_nonneg_ae Ω _ (μs i) f (eventually_of_forall f_nn) f.continuous.measurable.aestronglyMeasurable] - have aux : MonotoneOn ENNReal.toReal {∞}ᶜ := by sorry - apply (@MonotoneOn.map_liminf_of_continuousAt ℕ ℝ≥0∞ ℝ _ _ _ _ atTop _ ENNReal.toReal {∞}ᶜ aux - (fun i ↦ ∫⁻ (x : Ω), ENNReal.ofReal (f x) ∂(μs i)) ?_ ?_ ?_ ?_).symm - have lip : LipschitzWith 1 Real.toNNReal := by sorry - let g := BoundedContinuousFunction.comp _ lip f + let g := BoundedContinuousFunction.comp _ Real.lipschitzWith_toNNReal f have bound : ∀ i, ∫⁻ x, ENNReal.ofReal (f x) ∂(μs i) ≤ nndist 0 g := fun i ↦ by simpa only [coe_nnreal_ennreal_nndist, measure_univ, mul_one, ge_iff_le] using BoundedContinuousFunction.lintegral_le_edist_mul (μ := μs i) g + apply (@MonotoneOn.map_liminf_of_continuousAt ℕ ℝ≥0∞ ℝ _ _ _ _ atTop _ ENNReal.toReal {∞}ᶜ + ENNReal.monotoneOn_toReal + (fun i ↦ ∫⁻ (x : Ω), ENNReal.ofReal (f x) ∂(μs i)) ?_ ?_ ?_ ?_).symm · apply (NNReal.continuous_coe.comp_continuousOn ENNReal.continuousOn_toNNReal).continuousAt - have obs : {x : ℝ≥0∞ | x ≠ ∞} = Set.Iio ∞ := by sorry + have obs : {x : ℝ≥0∞ | x ≠ ∞} = Set.Iio ∞ := by + ext x + simp only [ne_eq, Set.mem_setOf_eq, Set.mem_Iio, ne_iff_lt_iff_le, le_top] rw [obs] apply Iio_mem_nhds - have lip : LipschitzWith 1 Real.toNNReal := by sorry - let g := BoundedContinuousFunction.comp _ lip f apply lt_of_le_of_lt (liminf_le_of_frequently_le (frequently_of_forall bound)) exact ENNReal.coe_lt_top - · --rw [le_map_iff] - --intro s hs - --simp - --rw [principal_le_iff] - sorry + · simp only [le_principal_iff, mem_map, Set.preimage_compl, mem_atTop_sets, ge_iff_le, + Set.mem_compl_iff, Set.mem_preimage, Set.mem_singleton_iff] + exact ⟨0, fun i _ ↦ (lt_of_le_of_lt (bound i) ENNReal.coe_lt_top).ne⟩ · exact ⟨∞, eventually_of_forall (fun x ↦ by simp only [le_top])⟩ · exact ⟨0, eventually_of_forall (fun x ↦ by simp only [ge_iff_le, zero_le])⟩ · exact (f.lintegral_of_real_lt_top μ).ne From af71e6b79b52a279158c5a45d5c47611e30fa071 Mon Sep 17 00:00:00 2001 From: Kalle Date: Sat, 21 Oct 2023 21:56:35 +0300 Subject: [PATCH 20/41] Delete a stupid direct proof attempt. --- ...lForPortmanteauOpenImpliesConvergence.lean | 28 ------------------- 1 file changed, 28 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index ddb5b263df0ea..e06d1d7eb1820 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -206,34 +206,6 @@ lemma fatou_argument_integral_nonneg congr exact (Real.norm_of_nonneg (f_nn x)).symm --- A direct proof attempt (should be discarded). -lemma fatou_argument_integral_nonneg' - {μ : Measure Ω} [IsFiniteMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsFiniteMeasure (μs i)] - {f : Ω →ᵇ ℝ} (f_nn : 0 ≤ f) - (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : - ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by - simp_rw [(f.integrable _).integral_eq_integral_meas_lt (eventually_of_forall f_nn)] - have rwr := @integral_eq_lintegral_of_nonneg_ae ℝ _ (volume.restrict (Set.Ioi 0)) - (fun t ↦ ENNReal.toReal (μ {x | t < f x})) (eventually_of_forall ?_) ?_ - have rwr' := fun i ↦ @integral_eq_lintegral_of_nonneg_ae ℝ _ (volume.restrict (Set.Ioi 0)) - (fun t ↦ ENNReal.toReal (μs i {x | t < f x})) (eventually_of_forall ?_) ?_ - simp only [rwr, ne_eq, rwr', ge_iff_le] - · --have aux : ∀ t, IsOpen {a : Ω | t < f a} := - -- fun t ↦ (continuous_def.mp f.continuous) _ isOpen_Ioi - --have obs := fun t ↦ h_opens _ (aux t) - --have := @lintegral_liminf_le - - --have earlier := fatou_argument_lintegral f.continuous f_nn h_opens - sorry - · exact fun _ ↦ by simp only [Pi.zero_apply, ENNReal.toReal_nonneg] - · apply Measurable.aestronglyMeasurable - apply ENNReal.measurable_toReal.comp (Antitone.measurable ?_) - exact fun x y hxy ↦ measure_mono (fun z hz ↦ lt_of_le_of_lt hxy hz) - · exact fun _ ↦ by simp only [Pi.zero_apply, ENNReal.toReal_nonneg] - · apply Measurable.aestronglyMeasurable - apply ENNReal.measurable_toReal.comp (Antitone.measurable ?_) - exact fun x y hxy ↦ measure_mono (fun z hz ↦ lt_of_le_of_lt hxy hz) - lemma reduction_to_liminf {ι : Type*} {L : Filter ι} [NeBot L] {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ι → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] (h : ∀ f : Ω →ᵇ ℝ, 0 ≤ f → ∫ x, (f x) ∂μ ≤ L.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i))) From 6eaf1344fcd094ba7a3bebf0729ad5ada4a4f688 Mon Sep 17 00:00:00 2001 From: Kalle Date: Sat, 21 Oct 2023 23:13:49 +0300 Subject: [PATCH 21/41] Hmmm... What is a reasonable way to prove the mapping of liminfs under AntitoneOn functions? --- ...lForPortmanteauOpenImpliesConvergence.lean | 91 ++++++++++++++++++- 1 file changed, 87 insertions(+), 4 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index e06d1d7eb1820..e025ea2b504cb 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -22,6 +22,39 @@ variable {ι R S : Type*} {F : Filter ι} [NeBot F] [ConditionallyCompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] [ConditionallyCompleteLinearOrder S] [TopologicalSpace S] [OrderTopology S] +--#check Antitone.map_sInf_of_continuousAt' + +/-- A function that is continuous at the supremum of a nonempty set and monotone on the set +sends this supremum to the supremum of the image of this set. -/ +theorem MonotoneOn.map_sSup_of_continuousAt' {α β : Type*} + [ConditionallyCompleteLinearOrder α] [TopologicalSpace α] [OrderTopology α] + [ConditionallyCompleteLinearOrder β] [TopologicalSpace β] [OrderClosedTopology β] + {f : α → β} {A A' : Set α} (hA : A ⊆ A') (Cf : ContinuousAt f (sSup A)) + (Mf : MonotoneOn f A') (A_nonemp : A.Nonempty) (A_bdd : BddAbove A := by bddDefault) : + f (sSup A) = sSup (f '' A) := + --This is a particular case of the more general `IsLUB.isLUB_of_tendsto` + .symm <| ((isLUB_csSup A_nonemp A_bdd).isLUB_of_tendsto (Mf.mono hA) A_nonemp <| + Cf.mono_left inf_le_left).csSup_eq (A_nonemp.image f) + +/-- A function that is continuous at the supremum of a nonempty set and monotone on the set +sends this supremum to the supremum of the image of this set. -/ +theorem AntitoneOn.map_sInf_of_continuousAt' {α β : Type*} + [ConditionallyCompleteLinearOrder α] [TopologicalSpace α] [OrderTopology α] + [ConditionallyCompleteLinearOrder β] [TopologicalSpace β] [OrderClosedTopology β] + {f : α → β} {A A' : Set α} (hA : A ⊆ A') (Cf : ContinuousAt f (sInf A)) + (Af : AntitoneOn f A') (A_nonemp : A.Nonempty) (A_bdd : BddBelow A := by bddDefault) : + f (sInf A) = sSup (f '' A) := by + sorry + +theorem AntitoneOn.map_limsSup_of_continuousAt₀ {F : Filter R} [NeBot F] {f : R → S} + {R' : Set R} (f_decr : AntitoneOn f R') (f_cont : ContinuousAt f F.limsSup) + (hF : F ≤ Filter.principal R') + (bdd_above : F.IsBounded (· ≤ ·) := by isBoundedDefault) + (bdd_below : F.IsBounded (· ≥ ·) := by isBoundedDefault) : + f F.limsSup = F.liminf f := by + --simp only [le_principal_iff] at hF + sorry + /-- An antitone function between (conditionally) complete linear ordered spaces sends a `Filter.limsSup` to the `Filter.liminf` of the image if the function is continuous at the `limsSup` (and the filter is bounded from above and below). -/ @@ -31,7 +64,58 @@ theorem AntitoneOn.map_limsSup_of_continuousAt {F : Filter R} [NeBot F] {f : R (bdd_above : F.IsBounded (· ≤ ·) := by isBoundedDefault) (bdd_below : F.IsBounded (· ≥ ·) := by isBoundedDefault) : f F.limsSup = F.liminf f := by - sorry + have cobdd : F.IsCobounded (· ≤ ·) := bdd_below.isCobounded_flip + apply le_antisymm + · rw [limsSup] + have := @AntitoneOn.map_sInf_of_continuousAt' R S _ _ _ _ _ _ _ f + {a | ∀ᶠ (n : R) in F, n ≤ a} R' _ f_cont f_decr bdd_above cobdd + rw [this] + --rw [limsSup, f_decr.map_sInf_of_continuousAt' f_cont bdd_above cobdd] + apply le_of_forall_lt + intro c hc + simp only [liminf, limsInf, eventually_map] at hc ⊢ + obtain ⟨d, hd, h'd⟩ := + exists_lt_of_lt_csSup (bdd_above.recOn fun x hx ↦ ⟨f x, Set.mem_image_of_mem f hx⟩) hc + apply lt_csSup_of_lt ?_ ?_ h'd + · --exact (Antitone.isBoundedUnder_le_comp f_decr bdd_below).isCoboundedUnder_flip + sorry + · rcases hd with ⟨e, ⟨he, fe_eq_d⟩⟩ + filter_upwards [he] with x hx --using (fe_eq_d.symm ▸ f_decr hx) + sorry + · by_cases h' : ∃ c, c < F.limsSup ∧ Set.Ioo c F.limsSup = ∅ + · rcases h' with ⟨c, c_lt, hc⟩ + have B : ∃ᶠ n in F, F.limsSup ≤ n := by + apply (frequently_lt_of_lt_limsSup cobdd c_lt).mono + intro x hx + by_contra' + have : (Set.Ioo c F.limsSup).Nonempty := ⟨x, ⟨hx, this⟩⟩ + simp only [hc, Set.not_nonempty_empty] at this + --apply liminf_le_of_frequently_le _ (bdd_above.isBoundedUnder f_decr) + --exact (B.mono fun x hx ↦ f_decr hx) + sorry + push_neg at h' + by_contra' H + have not_bot : ¬ IsBot F.limsSup := fun maybe_bot ↦ + --lt_irrefl (F.liminf f) <| lt_of_le_of_lt + -- (liminf_le_of_frequently_le (frequently_of_forall (fun r ↦ f_decr (maybe_bot r))) + -- (bdd_above.isBoundedUnder f_decr)) H + sorry + obtain ⟨l, l_lt, h'l⟩ : ∃ l < F.limsSup, Set.Ioc l F.limsSup ⊆ { x : R | f x < F.liminf f } + · apply exists_Ioc_subset_of_mem_nhds ((tendsto_order.1 f_cont.tendsto).2 _ H) + simpa [IsBot] using not_bot + obtain ⟨m, l_m, m_lt⟩ : (Set.Ioo l F.limsSup).Nonempty := by + contrapose! h' + refine' ⟨l, l_lt, by rwa [Set.not_nonempty_iff_eq_empty] at h'⟩ + have B : F.liminf f ≤ f m := by + apply liminf_le_of_frequently_le _ _ + · apply (frequently_lt_of_lt_limsSup cobdd m_lt).mono + --exact fun x hx ↦ f_decr hx.le + sorry + · --exact IsBounded.isBoundedUnder f_decr bdd_above + sorry + have I : f m < F.liminf f := h'l ⟨l_m, m_lt.le⟩ + exact lt_irrefl _ (B.trans_lt I) + /- have cobdd : F.IsCobounded (· ≤ ·) := bdd_below.isCobounded_flip apply le_antisymm @@ -254,9 +338,8 @@ lemma ProbabilityMeasure.tendsto_iff_forall_nonneg_integral_tendsto {γ : Type _ simp_rw [fx_eq] convert tendsto_add.comp (Tendsto.prod_mk_nhds key (@tendsto_const_nhds _ _ _ (-‖f‖) F)) <;> simp -theorem le_liminf_open_implies_convergence - {ι : Type*} {μ : ProbabilityMeasure Ω} {μs : ℕ → ProbabilityMeasure Ω} - (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : +theorem le_liminf_open_implies_convergence {μ : ProbabilityMeasure Ω} + {μs : ℕ → ProbabilityMeasure Ω} (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : atTop.Tendsto (fun i ↦ μs i) (𝓝 μ) := by refine ProbabilityMeasure.tendsto_iff_forall_integral_tendsto.mpr ?_ apply reduction_to_liminf From 261221d766db9f1ebbb2452b4e12ea22abdb62ed Mon Sep 17 00:00:00 2001 From: kkytola <“kalle.kytola@aalto.fi”> Date: Mon, 30 Oct 2023 19:33:06 +0200 Subject: [PATCH 22/41] Let us take an ad hoc route and avoind map_liminf for MonotoneOn. --- ...lForPortmanteauOpenImpliesConvergence.lean | 67 +++++++++++++++++++ 1 file changed, 67 insertions(+) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index e025ea2b504cb..5d287c5d76ccf 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -15,6 +15,73 @@ to the file `Mathlib.MeasureTheory.Measure.Portmanteau`. open MeasureTheory Filter open scoped ENNReal NNReal BoundedContinuousFunction Topology +noncomputable def ENNReal.truncateToReal (t x : ℝ≥0∞) : ℝ := + (min t x).toReal + +lemma ENNReal.truncateToReal_eq_toReal {t x : ℝ≥0∞} (t_ne_top : t ≠ ∞) (x_le : x ≤ t) : + ENNReal.truncateToReal t x = x.toReal := by + have x_lt_top : x < ∞ := lt_of_le_of_lt x_le t_ne_top.lt_top + have obs : min t x ≠ ∞ := by + simp_all only [ne_eq, ge_iff_le, min_eq_top, false_and, not_false_eq_true] + exact (ENNReal.toReal_eq_toReal obs x_lt_top.ne).mpr (min_eq_right x_le) + +lemma ENNReal.truncateToReal_le {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) {x : ℝ≥0∞} : + ENNReal.truncateToReal t x ≤ t.toReal := by + rw [ENNReal.truncateToReal] + apply (toReal_le_toReal _ t_ne_top).mpr (min_le_left t x) + simp_all only [ne_eq, ge_iff_le, min_eq_top, false_and, not_false_eq_true] + +lemma ENNReal.truncateToReal_nonneg {t x : ℝ≥0∞} : + 0 ≤ ENNReal.truncateToReal t x := toReal_nonneg + +lemma ENNReal.monotone_truncateToReal {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) : + Monotone (ENNReal.truncateToReal t) := by + intro x y x_le_y + have obs_x : min t x ≠ ∞ := by + simp_all only [ne_eq, ge_iff_le, min_eq_top, false_and, not_false_eq_true] + have obs_y : min t y ≠ ∞ := by + simp_all only [ne_eq, ge_iff_le, min_eq_top, false_and, not_false_eq_true] + exact (ENNReal.toReal_le_toReal obs_x obs_y).mpr (min_le_min_left t x_le_y) + +-- Q: Missing? +lemma ENNReal.continuousOn_toReal : + ContinuousOn (fun (x : ℝ≥0∞) ↦ x.toReal) { x | x ≠ ∞ } := by + change ContinuousOn (((↑) : ℝ≥0 → ℝ) ∘ (fun (x : ℝ≥0∞) ↦ x.toNNReal)) { x | x ≠ ∞ } + apply NNReal.continuous_coe.comp_continuousOn continuousOn_toNNReal + +lemma ENNReal.continuous_truncateToReal {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) : + Continuous (ENNReal.truncateToReal t) := by + apply ENNReal.continuousOn_toReal.comp_continuous (continuous_min.comp (Continuous.Prod.mk t)) + simp [t_ne_top] + +lemma bar {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} (b_ne_top : b ≠ ∞) {xs : ι → ℝ≥0∞} + (le_b : ∀ i, xs i ≤ b) : + F.liminf (fun i ↦ (xs i).toReal) = (F.liminf xs).toReal := by + have liminf_le : F.liminf xs ≤ b := by + apply liminf_le_of_le + · use 0 + simp only [ge_iff_le, zero_le, eventually_map, eventually_true] + · intro y h + obtain ⟨i, hi⟩ := h.exists + exact hi.trans (le_b i) + have key : ∀ i, (xs i).toReal = ENNReal.truncateToReal b (xs i) := by + intro i + rw [ENNReal.truncateToReal_eq_toReal b_ne_top (le_b i)] + have key' : (F.liminf xs).toReal = ENNReal.truncateToReal b (F.liminf xs) := by + rw [ENNReal.truncateToReal_eq_toReal b_ne_top liminf_le] + simp_rw [key, key'] + have := @Monotone.map_liminf_of_continuousAt ι ℝ≥0∞ ℝ F _ _ _ _ _ _ _ _ + (ENNReal.monotone_truncateToReal b_ne_top) xs + (ENNReal.continuous_truncateToReal b_ne_top).continuousAt ?_ ?_ + · rw [this] + rfl + · use b + simp only [eventually_map] + exact eventually_of_forall le_b + · use 0 + apply eventually_of_forall + intro y + simp only [ge_iff_le, zero_le] section yet_another_map_liminf_lemma From f7def1b238e0481f225977886affb893a0ba160f Mon Sep 17 00:00:00 2001 From: kkytola <“kalle.kytola@aalto.fi”> Date: Wed, 1 Nov 2023 18:46:19 +0200 Subject: [PATCH 23/41] Sorry-free avoiding MonotoneOn via ad hoc truncation. --- ...lForPortmanteauOpenImpliesConvergence.lean | 208 ++---------------- 1 file changed, 24 insertions(+), 184 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index 5d287c5d76ccf..bb03a7b3729d8 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -83,172 +83,6 @@ lemma bar {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} (b_ne_top : b intro y simp only [ge_iff_le, zero_le] -section yet_another_map_liminf_lemma - -variable {ι R S : Type*} {F : Filter ι} [NeBot F] - [ConditionallyCompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] - [ConditionallyCompleteLinearOrder S] [TopologicalSpace S] [OrderTopology S] - ---#check Antitone.map_sInf_of_continuousAt' - -/-- A function that is continuous at the supremum of a nonempty set and monotone on the set -sends this supremum to the supremum of the image of this set. -/ -theorem MonotoneOn.map_sSup_of_continuousAt' {α β : Type*} - [ConditionallyCompleteLinearOrder α] [TopologicalSpace α] [OrderTopology α] - [ConditionallyCompleteLinearOrder β] [TopologicalSpace β] [OrderClosedTopology β] - {f : α → β} {A A' : Set α} (hA : A ⊆ A') (Cf : ContinuousAt f (sSup A)) - (Mf : MonotoneOn f A') (A_nonemp : A.Nonempty) (A_bdd : BddAbove A := by bddDefault) : - f (sSup A) = sSup (f '' A) := - --This is a particular case of the more general `IsLUB.isLUB_of_tendsto` - .symm <| ((isLUB_csSup A_nonemp A_bdd).isLUB_of_tendsto (Mf.mono hA) A_nonemp <| - Cf.mono_left inf_le_left).csSup_eq (A_nonemp.image f) - -/-- A function that is continuous at the supremum of a nonempty set and monotone on the set -sends this supremum to the supremum of the image of this set. -/ -theorem AntitoneOn.map_sInf_of_continuousAt' {α β : Type*} - [ConditionallyCompleteLinearOrder α] [TopologicalSpace α] [OrderTopology α] - [ConditionallyCompleteLinearOrder β] [TopologicalSpace β] [OrderClosedTopology β] - {f : α → β} {A A' : Set α} (hA : A ⊆ A') (Cf : ContinuousAt f (sInf A)) - (Af : AntitoneOn f A') (A_nonemp : A.Nonempty) (A_bdd : BddBelow A := by bddDefault) : - f (sInf A) = sSup (f '' A) := by - sorry - -theorem AntitoneOn.map_limsSup_of_continuousAt₀ {F : Filter R} [NeBot F] {f : R → S} - {R' : Set R} (f_decr : AntitoneOn f R') (f_cont : ContinuousAt f F.limsSup) - (hF : F ≤ Filter.principal R') - (bdd_above : F.IsBounded (· ≤ ·) := by isBoundedDefault) - (bdd_below : F.IsBounded (· ≥ ·) := by isBoundedDefault) : - f F.limsSup = F.liminf f := by - --simp only [le_principal_iff] at hF - sorry - -/-- An antitone function between (conditionally) complete linear ordered spaces sends a -`Filter.limsSup` to the `Filter.liminf` of the image if the function is continuous at the `limsSup` -(and the filter is bounded from above and below). -/ -theorem AntitoneOn.map_limsSup_of_continuousAt {F : Filter R} [NeBot F] {f : R → S} - {R' : Set R} (f_decr : AntitoneOn f R') (f_cont : ContinuousAt f F.limsSup) - (hF : F ≤ Filter.principal R') - (bdd_above : F.IsBounded (· ≤ ·) := by isBoundedDefault) - (bdd_below : F.IsBounded (· ≥ ·) := by isBoundedDefault) : - f F.limsSup = F.liminf f := by - have cobdd : F.IsCobounded (· ≤ ·) := bdd_below.isCobounded_flip - apply le_antisymm - · rw [limsSup] - have := @AntitoneOn.map_sInf_of_continuousAt' R S _ _ _ _ _ _ _ f - {a | ∀ᶠ (n : R) in F, n ≤ a} R' _ f_cont f_decr bdd_above cobdd - rw [this] - --rw [limsSup, f_decr.map_sInf_of_continuousAt' f_cont bdd_above cobdd] - apply le_of_forall_lt - intro c hc - simp only [liminf, limsInf, eventually_map] at hc ⊢ - obtain ⟨d, hd, h'd⟩ := - exists_lt_of_lt_csSup (bdd_above.recOn fun x hx ↦ ⟨f x, Set.mem_image_of_mem f hx⟩) hc - apply lt_csSup_of_lt ?_ ?_ h'd - · --exact (Antitone.isBoundedUnder_le_comp f_decr bdd_below).isCoboundedUnder_flip - sorry - · rcases hd with ⟨e, ⟨he, fe_eq_d⟩⟩ - filter_upwards [he] with x hx --using (fe_eq_d.symm ▸ f_decr hx) - sorry - · by_cases h' : ∃ c, c < F.limsSup ∧ Set.Ioo c F.limsSup = ∅ - · rcases h' with ⟨c, c_lt, hc⟩ - have B : ∃ᶠ n in F, F.limsSup ≤ n := by - apply (frequently_lt_of_lt_limsSup cobdd c_lt).mono - intro x hx - by_contra' - have : (Set.Ioo c F.limsSup).Nonempty := ⟨x, ⟨hx, this⟩⟩ - simp only [hc, Set.not_nonempty_empty] at this - --apply liminf_le_of_frequently_le _ (bdd_above.isBoundedUnder f_decr) - --exact (B.mono fun x hx ↦ f_decr hx) - sorry - push_neg at h' - by_contra' H - have not_bot : ¬ IsBot F.limsSup := fun maybe_bot ↦ - --lt_irrefl (F.liminf f) <| lt_of_le_of_lt - -- (liminf_le_of_frequently_le (frequently_of_forall (fun r ↦ f_decr (maybe_bot r))) - -- (bdd_above.isBoundedUnder f_decr)) H - sorry - obtain ⟨l, l_lt, h'l⟩ : ∃ l < F.limsSup, Set.Ioc l F.limsSup ⊆ { x : R | f x < F.liminf f } - · apply exists_Ioc_subset_of_mem_nhds ((tendsto_order.1 f_cont.tendsto).2 _ H) - simpa [IsBot] using not_bot - obtain ⟨m, l_m, m_lt⟩ : (Set.Ioo l F.limsSup).Nonempty := by - contrapose! h' - refine' ⟨l, l_lt, by rwa [Set.not_nonempty_iff_eq_empty] at h'⟩ - have B : F.liminf f ≤ f m := by - apply liminf_le_of_frequently_le _ _ - · apply (frequently_lt_of_lt_limsSup cobdd m_lt).mono - --exact fun x hx ↦ f_decr hx.le - sorry - · --exact IsBounded.isBoundedUnder f_decr bdd_above - sorry - have I : f m < F.liminf f := h'l ⟨l_m, m_lt.le⟩ - exact lt_irrefl _ (B.trans_lt I) - -/- - have cobdd : F.IsCobounded (· ≤ ·) := bdd_below.isCobounded_flip - apply le_antisymm - · rw [limsSup, f_decr.map_sInf_of_continuousAt' f_cont bdd_above cobdd] - apply le_of_forall_lt - intro c hc - simp only [liminf, limsInf, eventually_map] at hc ⊢ - obtain ⟨d, hd, h'd⟩ := - exists_lt_of_lt_csSup (bdd_above.recOn fun x hx ↦ ⟨f x, Set.mem_image_of_mem f hx⟩) hc - apply lt_csSup_of_lt ?_ ?_ h'd - · exact (Antitone.isBoundedUnder_le_comp f_decr bdd_below).isCoboundedUnder_flip - · rcases hd with ⟨e, ⟨he, fe_eq_d⟩⟩ - filter_upwards [he] with x hx using (fe_eq_d.symm ▸ f_decr hx) - · by_cases h' : ∃ c, c < F.limsSup ∧ Set.Ioo c F.limsSup = ∅ - · rcases h' with ⟨c, c_lt, hc⟩ - have B : ∃ᶠ n in F, F.limsSup ≤ n := by - apply (frequently_lt_of_lt_limsSup cobdd c_lt).mono - intro x hx - by_contra' - have : (Set.Ioo c F.limsSup).Nonempty := ⟨x, ⟨hx, this⟩⟩ - simp only [hc, Set.not_nonempty_empty] at this - apply liminf_le_of_frequently_le _ (bdd_above.isBoundedUnder f_decr) - exact (B.mono fun x hx ↦ f_decr hx) - push_neg at h' - by_contra' H - have not_bot : ¬ IsBot F.limsSup := fun maybe_bot ↦ - lt_irrefl (F.liminf f) <| lt_of_le_of_lt - (liminf_le_of_frequently_le (frequently_of_forall (fun r ↦ f_decr (maybe_bot r))) - (bdd_above.isBoundedUnder f_decr)) H - obtain ⟨l, l_lt, h'l⟩ : ∃ l < F.limsSup, Set.Ioc l F.limsSup ⊆ { x : R | f x < F.liminf f } - · apply exists_Ioc_subset_of_mem_nhds ((tendsto_order.1 f_cont.tendsto).2 _ H) - simpa [IsBot] using not_bot - obtain ⟨m, l_m, m_lt⟩ : (Set.Ioo l F.limsSup).Nonempty := by - contrapose! h' - refine' ⟨l, l_lt, by rwa [Set.not_nonempty_iff_eq_empty] at h'⟩ - have B : F.liminf f ≤ f m := by - apply liminf_le_of_frequently_le _ _ - · apply (frequently_lt_of_lt_limsSup cobdd m_lt).mono - exact fun x hx ↦ f_decr hx.le - · exact IsBounded.isBoundedUnder f_decr bdd_above - have I : f m < F.liminf f := h'l ⟨l_m, m_lt.le⟩ - exact lt_irrefl _ (B.trans_lt I) - -/ - -theorem MonotoneOn.map_limsInf_of_continuousAt {F : Filter R} [NeBot F] {f : R → S} - {R' : Set R} (f_incr : MonotoneOn f R') (f_cont : ContinuousAt f F.limsInf) - (hF : F ≤ Filter.principal R') - (bdd_above : F.IsBounded (· ≤ ·) := by isBoundedDefault) - (bdd_below : F.IsBounded (· ≥ ·) := by isBoundedDefault) : - f F.limsInf = F.liminf f := by - exact @AntitoneOn.map_limsSup_of_continuousAt Rᵒᵈ S _ _ _ _ F _ f R' - (antitone_on_dual_iff.mp f_incr) f_cont hF bdd_below bdd_above - -theorem MonotoneOn.map_liminf_of_continuousAt {F : Filter ι} [NeBot F] - {f : R → S} {R' : Set R} (f_incr : MonotoneOn f R') - (a : ι → R) (f_cont : ContinuousAt f (F.liminf a)) (hF : F.map a ≤ Filter.principal R') - (bdd_above : F.IsBoundedUnder (· ≤ ·) a := by isBoundedDefault) - (bdd_below : F.IsBoundedUnder (· ≥ ·) a := by isBoundedDefault) : - f (F.liminf a) = F.liminf (f ∘ a) := by - apply @MonotoneOn.map_limsInf_of_continuousAt R S _ _ _ _ (F.map a) _ f R' f_incr f_cont hF - bdd_above bdd_below - -end yet_another_map_liminf_lemma - - - -- NOTE: Missing from Mathlib? -- What would be a good generality? -- (Mixes order-boundedness and metric-boundedness, so typeclasses don't readily exist.) @@ -267,7 +101,6 @@ lemma Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : apply isBoundedUnder_of ⟨c, by simpa [mem_lowerBounds] using hc⟩ - section le_liminf_open_implies_convergence variable {Ω : Type} [MeasurableSpace Ω] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] @@ -321,25 +154,32 @@ lemma fatou_argument_integral_nonneg · simp only [fun i ↦ @integral_eq_lintegral_of_nonneg_ae Ω _ (μs i) f (eventually_of_forall f_nn) f.continuous.measurable.aestronglyMeasurable] let g := BoundedContinuousFunction.comp _ Real.lipschitzWith_toNNReal f - have bound : ∀ i, ∫⁻ x, ENNReal.ofReal (f x) ∂(μs i) ≤ nndist 0 g := fun i ↦ by + let c := nndist 0 g + have c_ne_top : (c : ℝ≥0∞) ≠ ∞ := ENNReal.coe_ne_top + have bound : ∀ i, ∫⁻ x, ENNReal.ofReal (f x) ∂(μs i) ≤ c := fun i ↦ by simpa only [coe_nnreal_ennreal_nndist, measure_univ, mul_one, ge_iff_le] using BoundedContinuousFunction.lintegral_le_edist_mul (μ := μs i) g - apply (@MonotoneOn.map_liminf_of_continuousAt ℕ ℝ≥0∞ ℝ _ _ _ _ atTop _ ENNReal.toReal {∞}ᶜ - ENNReal.monotoneOn_toReal - (fun i ↦ ∫⁻ (x : Ω), ENNReal.ofReal (f x) ∂(μs i)) ?_ ?_ ?_ ?_).symm - · apply (NNReal.continuous_coe.comp_continuousOn ENNReal.continuousOn_toNNReal).continuousAt - have obs : {x : ℝ≥0∞ | x ≠ ∞} = Set.Iio ∞ := by - ext x - simp only [ne_eq, Set.mem_setOf_eq, Set.mem_Iio, ne_iff_lt_iff_le, le_top] - rw [obs] - apply Iio_mem_nhds - apply lt_of_le_of_lt (liminf_le_of_frequently_le (frequently_of_forall bound)) - exact ENNReal.coe_lt_top - · simp only [le_principal_iff, mem_map, Set.preimage_compl, mem_atTop_sets, ge_iff_le, - Set.mem_compl_iff, Set.mem_preimage, Set.mem_singleton_iff] - exact ⟨0, fun i _ ↦ (lt_of_le_of_lt (bound i) ENNReal.coe_lt_top).ne⟩ - · exact ⟨∞, eventually_of_forall (fun x ↦ by simp only [le_top])⟩ - · exact ⟨0, eventually_of_forall (fun x ↦ by simp only [ge_iff_le, zero_le])⟩ + have obs₁ := fun i ↦ @ENNReal.truncateToReal_eq_toReal c _ c_ne_top (bound i) + simp only [← obs₁] + have := @Monotone.map_liminf_of_continuousAt ℕ ℝ≥0∞ ℝ atTop _ _ _ _ _ _ _ + (ENNReal.truncateToReal c) (ENNReal.monotone_truncateToReal (t := c) c_ne_top) + (fun i ↦ ∫⁻ (a : Ω), ENNReal.ofReal (f a) ∂μs i) + (ENNReal.continuous_truncateToReal (t := c) c_ne_top).continuousAt ?_ ?_ + · convert this.symm using 1 + apply (ENNReal.truncateToReal_eq_toReal c_ne_top _).symm + apply liminf_le_of_le ?_ (fun b h ↦ ?_) + · use 0 + simp only [ge_iff_le, zero_le, eventually_map, eventually_atTop, implies_true, forall_const, + exists_const] + obtain ⟨i, hi⟩ := h.exists + exact hi.trans (bound i) + · use c + simp only [coe_nnreal_ennreal_nndist, eventually_map, eventually_atTop, ge_iff_le] + use 0 + exact fun j _ ↦ bound j + · use 0 + simp only [ge_iff_le, zero_le, eventually_map, eventually_atTop, implies_true, forall_const, + exists_const] · exact (f.lintegral_of_real_lt_top μ).ne · apply ne_of_lt have obs := fun (i : ℕ) ↦ @BoundedContinuousFunction.lintegral_nnnorm_le Ω _ _ (μs i) ℝ _ f From a58a27b3723f3a1184bd887ef1753bf6245a136f Mon Sep 17 00:00:00 2001 From: kkytola <“kalle.kytola@aalto.fi”> Date: Wed, 1 Nov 2023 19:37:07 +0200 Subject: [PATCH 24/41] Some cleanup. --- ...lForPortmanteauOpenImpliesConvergence.lean | 109 +++++++----------- 1 file changed, 41 insertions(+), 68 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean index bb03a7b3729d8..e0548475d6b29 100644 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean @@ -15,6 +15,8 @@ to the file `Mathlib.MeasureTheory.Measure.Portmanteau`. open MeasureTheory Filter open scoped ENNReal NNReal BoundedContinuousFunction Topology +/-- With truncation level `t`, the truncated cast `ℝ≥0∞ → ℝ` is given by `x ↦ (min t x).toReal`. +Unlike `ENNReal.toReal`, this cast is continuous and monotone when `t ≠ ∞`. -/ noncomputable def ENNReal.truncateToReal (t x : ℝ≥0∞) : ℝ := (min t x).toReal @@ -34,6 +36,7 @@ lemma ENNReal.truncateToReal_le {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) {x : ℝ lemma ENNReal.truncateToReal_nonneg {t x : ℝ≥0∞} : 0 ≤ ENNReal.truncateToReal t x := toReal_nonneg +/-- The truncated cast `ENNReal.truncateToReal t : ℝ≥0∞ → ℝ` is monotone when `t ≠ ∞`. -/ lemma ENNReal.monotone_truncateToReal {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) : Monotone (ENNReal.truncateToReal t) := by intro x y x_le_y @@ -49,13 +52,15 @@ lemma ENNReal.continuousOn_toReal : change ContinuousOn (((↑) : ℝ≥0 → ℝ) ∘ (fun (x : ℝ≥0∞) ↦ x.toNNReal)) { x | x ≠ ∞ } apply NNReal.continuous_coe.comp_continuousOn continuousOn_toNNReal +/-- The truncated cast `ENNReal.truncateToReal t : ℝ≥0∞ → ℝ` is continuous when `t ≠ ∞`. -/ lemma ENNReal.continuous_truncateToReal {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) : Continuous (ENNReal.truncateToReal t) := by apply ENNReal.continuousOn_toReal.comp_continuous (continuous_min.comp (Continuous.Prod.mk t)) simp [t_ne_top] -lemma bar {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} (b_ne_top : b ≠ ∞) {xs : ι → ℝ≥0∞} - (le_b : ∀ i, xs i ≤ b) : +/-- If `xs : ι → ℝ≥0∞` is bounded, then we have `liminf (toReal ∘ xs) = toReal (liminf xs)`. -/ +lemma ENNReal.liminf_toReal_eq {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} (b_ne_top : b ≠ ∞) + {xs : ι → ℝ≥0∞} (le_b : ∀ i, xs i ≤ b) : F.liminf (fun i ↦ (xs i).toReal) = (F.liminf xs).toReal := by have liminf_le : F.liminf xs ≤ b := by apply liminf_le_of_le @@ -64,16 +69,16 @@ lemma bar {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} (b_ne_top : b · intro y h obtain ⟨i, hi⟩ := h.exists exact hi.trans (le_b i) - have key : ∀ i, (xs i).toReal = ENNReal.truncateToReal b (xs i) := by + have aux : ∀ i, (xs i).toReal = ENNReal.truncateToReal b (xs i) := by intro i rw [ENNReal.truncateToReal_eq_toReal b_ne_top (le_b i)] - have key' : (F.liminf xs).toReal = ENNReal.truncateToReal b (F.liminf xs) := by + have aux' : (F.liminf xs).toReal = ENNReal.truncateToReal b (F.liminf xs) := by rw [ENNReal.truncateToReal_eq_toReal b_ne_top liminf_le] - simp_rw [key, key'] - have := @Monotone.map_liminf_of_continuousAt ι ℝ≥0∞ ℝ F _ _ _ _ _ _ _ _ + simp_rw [aux, aux'] + have key := Monotone.map_liminf_of_continuousAt (F := F) (ENNReal.monotone_truncateToReal b_ne_top) xs (ENNReal.continuous_truncateToReal b_ne_top).continuousAt ?_ ?_ - · rw [this] + · rw [key] rfl · use b simp only [eventually_map] @@ -100,13 +105,11 @@ lemma Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : obtain ⟨c, hc⟩ := h.1 apply isBoundedUnder_of ⟨c, by simpa [mem_lowerBounds] using hc⟩ - section le_liminf_open_implies_convergence variable {Ω : Type} [MeasurableSpace Ω] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] --- NOTE: I think I will work with real-valued integrals, after all... -lemma fatou_argument_lintegral +lemma lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure {μ : Measure Ω} [SigmaFinite μ] {μs : ℕ → Measure Ω} [∀ i, SigmaFinite (μs i)] {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : @@ -130,6 +133,7 @@ theorem BoundedContinuousFunction.lintegral_le_edist_mul apply le_trans (lintegral_mono (fun x ↦ ENNReal.coe_le_coe.mpr (bound x))) simp +-- Missing? lemma ENNReal.monotoneOn_toReal : MonotoneOn ENNReal.toReal {∞}ᶜ := fun _ _ _ hy x_le_y ↦ toReal_mono hy x_le_y @@ -141,49 +145,27 @@ lemma Real.lipschitzWith_toNNReal : LipschitzWith 1 Real.toNNReal := by simpa only [ge_iff_le, NNReal.coe_one, dist_prod_same_right, one_mul, Real.dist_eq] using lipschitzWith_iff_dist_le_mul.mp lipschitzWith_max (x, 0) (y, 0) --- NOTE: I think this is the version I prefer to use, after all... -lemma fatou_argument_integral_nonneg +lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] {f : Ω →ᵇ ℝ} (f_nn : 0 ≤ f) (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by - have earlier := fatou_argument_lintegral f.continuous f_nn h_opens + have same := lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure + f.continuous f_nn h_opens rw [@integral_eq_lintegral_of_nonneg_ae Ω _ μ f (eventually_of_forall f_nn) f.continuous.measurable.aestronglyMeasurable] - convert (ENNReal.toReal_le_toReal ?_ ?_).mpr earlier + convert (ENNReal.toReal_le_toReal ?_ ?_).mpr same · simp only [fun i ↦ @integral_eq_lintegral_of_nonneg_ae Ω _ (μs i) f (eventually_of_forall f_nn) f.continuous.measurable.aestronglyMeasurable] let g := BoundedContinuousFunction.comp _ Real.lipschitzWith_toNNReal f - let c := nndist 0 g - have c_ne_top : (c : ℝ≥0∞) ≠ ∞ := ENNReal.coe_ne_top - have bound : ∀ i, ∫⁻ x, ENNReal.ofReal (f x) ∂(μs i) ≤ c := fun i ↦ by + have bound : ∀ i, ∫⁻ x, ENNReal.ofReal (f x) ∂(μs i) ≤ nndist 0 g := fun i ↦ by simpa only [coe_nnreal_ennreal_nndist, measure_univ, mul_one, ge_iff_le] using BoundedContinuousFunction.lintegral_le_edist_mul (μ := μs i) g - have obs₁ := fun i ↦ @ENNReal.truncateToReal_eq_toReal c _ c_ne_top (bound i) - simp only [← obs₁] - have := @Monotone.map_liminf_of_continuousAt ℕ ℝ≥0∞ ℝ atTop _ _ _ _ _ _ _ - (ENNReal.truncateToReal c) (ENNReal.monotone_truncateToReal (t := c) c_ne_top) - (fun i ↦ ∫⁻ (a : Ω), ENNReal.ofReal (f a) ∂μs i) - (ENNReal.continuous_truncateToReal (t := c) c_ne_top).continuousAt ?_ ?_ - · convert this.symm using 1 - apply (ENNReal.truncateToReal_eq_toReal c_ne_top _).symm - apply liminf_le_of_le ?_ (fun b h ↦ ?_) - · use 0 - simp only [ge_iff_le, zero_le, eventually_map, eventually_atTop, implies_true, forall_const, - exists_const] - obtain ⟨i, hi⟩ := h.exists - exact hi.trans (bound i) - · use c - simp only [coe_nnreal_ennreal_nndist, eventually_map, eventually_atTop, ge_iff_le] - use 0 - exact fun j _ ↦ bound j - · use 0 - simp only [ge_iff_le, zero_le, eventually_map, eventually_atTop, implies_true, forall_const, - exists_const] + apply ENNReal.liminf_toReal_eq ENNReal.coe_ne_top bound · exact (f.lintegral_of_real_lt_top μ).ne · apply ne_of_lt have obs := fun (i : ℕ) ↦ @BoundedContinuousFunction.lintegral_nnnorm_le Ω _ _ (μs i) ℝ _ f - simp at obs + simp only [measure_univ, mul_one] at obs apply lt_of_le_of_lt _ (show (‖f‖₊ : ℝ≥0∞) < ∞ from ENNReal.coe_lt_top) apply liminf_le_of_le · refine ⟨0, eventually_of_forall (by simp only [ge_iff_le, zero_le, forall_const])⟩ @@ -197,11 +179,14 @@ lemma fatou_argument_integral_nonneg congr exact (Real.norm_of_nonneg (f_nn x)).symm -lemma reduction_to_liminf {ι : Type*} {L : Filter ι} [NeBot L] +lemma tendsto_integral_of_forall_integral_le_liminf_integral {ι : Type*} {L : Filter ι} {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ι → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] (h : ∀ f : Ω →ᵇ ℝ, 0 ≤ f → ∫ x, (f x) ∂μ ≤ L.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i))) (f : Ω →ᵇ ℝ) : Tendsto (fun i ↦ ∫ x, (f x) ∂ (μs i)) L (𝓝 (∫ x, (f x) ∂μ)) := by + by_cases L_bot : L = ⊥ + · simp only [L_bot, tendsto_bot] + have : NeBot L := ⟨L_bot⟩ have obs := BoundedContinuousFunction.isBounded_range_integral μs f have bdd_above : IsBoundedUnder (· ≤ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := isBounded_le_map_of_bounded_range _ obs @@ -211,7 +196,6 @@ lemma reduction_to_liminf {ι : Type*} {L : Filter ι} [NeBot L] · have key := h _ (f.add_norm_nonneg) simp_rw [f.integral_add_const ‖f‖] at key simp only [measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key - -- TODO: Should the case of ⊥ filter be treated separately and not included as an assumption? have := liminf_add_const L (fun i ↦ ∫ x, (f x) ∂ (μs i)) ‖f‖ bdd_above bdd_below rwa [this, add_le_add_iff_right] at key · have key := h _ (f.norm_sub_nonneg) @@ -245,39 +229,28 @@ lemma ProbabilityMeasure.tendsto_iff_forall_nonneg_integral_tendsto {γ : Type _ simp_rw [fx_eq] convert tendsto_add.comp (Tendsto.prod_mk_nhds key (@tendsto_const_nhds _ _ _ (-‖f‖) F)) <;> simp -theorem le_liminf_open_implies_convergence {μ : ProbabilityMeasure Ω} +/-- One implication of the portmanteau theorem. -/ +theorem ProbabilityMeasure.tendsto_of_forall_isOpen_le_liminf {μ : ProbabilityMeasure Ω} {μs : ℕ → ProbabilityMeasure Ω} (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : atTop.Tendsto (fun i ↦ μs i) (𝓝 μ) := by refine ProbabilityMeasure.tendsto_iff_forall_integral_tendsto.mpr ?_ - apply reduction_to_liminf + apply tendsto_integral_of_forall_integral_le_liminf_integral intro f f_nn - apply fatou_argument_integral_nonneg (f := f) f_nn + apply integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure (f := f) f_nn intro G G_open specialize h_opens G G_open simp only at h_opens - simp only [liminf, limsInf, eventually_map, eventually_atTop, ge_iff_le, le_sSup_iff] at * - intro b b_ub - by_cases b_eq_top : b = ∞ - · simp only [b_eq_top, le_top] - · have foo := (@le_csSup_iff ℝ≥0 _ {a | ∃ a_1, ∀ (b : ℕ), a_1 ≤ b → a ≤ ENNReal.toNNReal (μs b G)} - (ENNReal.toNNReal (μ G)) ?_ ?_).mp h_opens (ENNReal.toNNReal b) ?_ - · simp only [ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure] at foo - convert ENNReal.coe_le_coe.mpr foo - · simp only [ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure] - · simp only [ne_eq, b_eq_top, not_false_eq_true, ENNReal.coe_toNNReal] - · use 1 - simp [mem_upperBounds] - intro x n hn - exact (hn n (le_refl n)).trans (ProbabilityMeasure.apply_le_one _ _) - · refine ⟨0, by simp⟩ - · simp only [mem_upperBounds, Set.mem_setOf_eq, forall_exists_index, ne_eq, - ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure] at b_ub ⊢ - intro x n hn - specialize b_ub x n ?_ - · intro m hm - convert ENNReal.coe_le_coe.mpr (hn m hm) - simp only [ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure] - · rw [(ENNReal.coe_toNNReal b_eq_top).symm] at b_ub - exact ENNReal.coe_le_coe.mp b_ub + have aux := Monotone.map_liminf_of_continuousAt (F := atTop) ENNReal.coe_mono (μs · G) ?_ ?_ ?_ + · have obs := ENNReal.coe_mono h_opens + simp only [ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure, aux] at obs + convert obs + simp only [Function.comp_apply, ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure] + · apply ENNReal.continuous_coe.continuousAt + · use 1 + simp only [eventually_map, ProbabilityMeasure.apply_le_one, eventually_atTop, ge_iff_le, + implies_true, forall_const, exists_const] + · use 0 + simp only [zero_le, eventually_map, eventually_atTop, ge_iff_le, implies_true, forall_const, + exists_const] end le_liminf_open_implies_convergence From 9b397ab16578b2790cc0892a358e3c8eae267b39 Mon Sep 17 00:00:00 2001 From: kkytola <“kalle.kytola@aalto.fi”> Date: Wed, 1 Nov 2023 20:08:58 +0200 Subject: [PATCH 25/41] Move some ENNReal lemmas to an appropriate file. --- ...lForPortmanteauOpenImpliesConvergence.lean | 256 ------------------ Mathlib/Topology/Instances/ENNReal.lean | 83 +++++- 2 files changed, 81 insertions(+), 258 deletions(-) delete mode 100644 Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean diff --git a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean b/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean deleted file mode 100644 index e0548475d6b29..0000000000000 --- a/Mathlib/MeasureTheory/Measure/NewMaterialForPortmanteauOpenImpliesConvergence.lean +++ /dev/null @@ -1,256 +0,0 @@ -import Mathlib.MeasureTheory.Measure.Portmanteau -import Mathlib.MeasureTheory.Integral.Layercake -import Mathlib.MeasureTheory.Integral.BoundedContinuousFunction -import Mathlib.Tactic - -/-! -# Outline of portmanteau implication: liminf condition for open sets implies weak convergence - -This file contains steps needed to prove the portmanteau implication -`le_liminf_open_implies_convergence`. Some intermediate steps need to be generalized -and cleaned up, and are to be placed in appropriate files. The main part should go -to the file `Mathlib.MeasureTheory.Measure.Portmanteau`. --/ - -open MeasureTheory Filter -open scoped ENNReal NNReal BoundedContinuousFunction Topology - -/-- With truncation level `t`, the truncated cast `ℝ≥0∞ → ℝ` is given by `x ↦ (min t x).toReal`. -Unlike `ENNReal.toReal`, this cast is continuous and monotone when `t ≠ ∞`. -/ -noncomputable def ENNReal.truncateToReal (t x : ℝ≥0∞) : ℝ := - (min t x).toReal - -lemma ENNReal.truncateToReal_eq_toReal {t x : ℝ≥0∞} (t_ne_top : t ≠ ∞) (x_le : x ≤ t) : - ENNReal.truncateToReal t x = x.toReal := by - have x_lt_top : x < ∞ := lt_of_le_of_lt x_le t_ne_top.lt_top - have obs : min t x ≠ ∞ := by - simp_all only [ne_eq, ge_iff_le, min_eq_top, false_and, not_false_eq_true] - exact (ENNReal.toReal_eq_toReal obs x_lt_top.ne).mpr (min_eq_right x_le) - -lemma ENNReal.truncateToReal_le {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) {x : ℝ≥0∞} : - ENNReal.truncateToReal t x ≤ t.toReal := by - rw [ENNReal.truncateToReal] - apply (toReal_le_toReal _ t_ne_top).mpr (min_le_left t x) - simp_all only [ne_eq, ge_iff_le, min_eq_top, false_and, not_false_eq_true] - -lemma ENNReal.truncateToReal_nonneg {t x : ℝ≥0∞} : - 0 ≤ ENNReal.truncateToReal t x := toReal_nonneg - -/-- The truncated cast `ENNReal.truncateToReal t : ℝ≥0∞ → ℝ` is monotone when `t ≠ ∞`. -/ -lemma ENNReal.monotone_truncateToReal {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) : - Monotone (ENNReal.truncateToReal t) := by - intro x y x_le_y - have obs_x : min t x ≠ ∞ := by - simp_all only [ne_eq, ge_iff_le, min_eq_top, false_and, not_false_eq_true] - have obs_y : min t y ≠ ∞ := by - simp_all only [ne_eq, ge_iff_le, min_eq_top, false_and, not_false_eq_true] - exact (ENNReal.toReal_le_toReal obs_x obs_y).mpr (min_le_min_left t x_le_y) - --- Q: Missing? -lemma ENNReal.continuousOn_toReal : - ContinuousOn (fun (x : ℝ≥0∞) ↦ x.toReal) { x | x ≠ ∞ } := by - change ContinuousOn (((↑) : ℝ≥0 → ℝ) ∘ (fun (x : ℝ≥0∞) ↦ x.toNNReal)) { x | x ≠ ∞ } - apply NNReal.continuous_coe.comp_continuousOn continuousOn_toNNReal - -/-- The truncated cast `ENNReal.truncateToReal t : ℝ≥0∞ → ℝ` is continuous when `t ≠ ∞`. -/ -lemma ENNReal.continuous_truncateToReal {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) : - Continuous (ENNReal.truncateToReal t) := by - apply ENNReal.continuousOn_toReal.comp_continuous (continuous_min.comp (Continuous.Prod.mk t)) - simp [t_ne_top] - -/-- If `xs : ι → ℝ≥0∞` is bounded, then we have `liminf (toReal ∘ xs) = toReal (liminf xs)`. -/ -lemma ENNReal.liminf_toReal_eq {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} (b_ne_top : b ≠ ∞) - {xs : ι → ℝ≥0∞} (le_b : ∀ i, xs i ≤ b) : - F.liminf (fun i ↦ (xs i).toReal) = (F.liminf xs).toReal := by - have liminf_le : F.liminf xs ≤ b := by - apply liminf_le_of_le - · use 0 - simp only [ge_iff_le, zero_le, eventually_map, eventually_true] - · intro y h - obtain ⟨i, hi⟩ := h.exists - exact hi.trans (le_b i) - have aux : ∀ i, (xs i).toReal = ENNReal.truncateToReal b (xs i) := by - intro i - rw [ENNReal.truncateToReal_eq_toReal b_ne_top (le_b i)] - have aux' : (F.liminf xs).toReal = ENNReal.truncateToReal b (F.liminf xs) := by - rw [ENNReal.truncateToReal_eq_toReal b_ne_top liminf_le] - simp_rw [aux, aux'] - have key := Monotone.map_liminf_of_continuousAt (F := F) - (ENNReal.monotone_truncateToReal b_ne_top) xs - (ENNReal.continuous_truncateToReal b_ne_top).continuousAt ?_ ?_ - · rw [key] - rfl - · use b - simp only [eventually_map] - exact eventually_of_forall le_b - · use 0 - apply eventually_of_forall - intro y - simp only [ge_iff_le, zero_le] - --- NOTE: Missing from Mathlib? --- What would be a good generality? --- (Mixes order-boundedness and metric-boundedness, so typeclasses don't readily exist.) -lemma Filter.isBounded_le_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} - (h : Bornology.IsBounded (Set.range f)) : - (F.map f).IsBounded (· ≤ ·) := by - rw [Real.isBounded_iff_bddBelow_bddAbove] at h - obtain ⟨c, hc⟩ := h.2 - refine isBoundedUnder_of ⟨c, by simpa [mem_upperBounds] using hc⟩ - -lemma Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} - (h : Bornology.IsBounded (Set.range f)) : - (F.map f).IsBounded (· ≥ ·) := by - rw [Real.isBounded_iff_bddBelow_bddAbove] at h - obtain ⟨c, hc⟩ := h.1 - apply isBoundedUnder_of ⟨c, by simpa [mem_lowerBounds] using hc⟩ - -section le_liminf_open_implies_convergence - -variable {Ω : Type} [MeasurableSpace Ω] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] - -lemma lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure - {μ : Measure Ω} [SigmaFinite μ] {μs : ℕ → Measure Ω} [∀ i, SigmaFinite (μs i)] - {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) - (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : - ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫⁻ x, ENNReal.ofReal (f x) ∂ (μs i)) := by - simp_rw [lintegral_eq_lintegral_meas_lt _ (eventually_of_forall f_nn) f_cont.aemeasurable] - calc ∫⁻ (t : ℝ) in Set.Ioi 0, μ {a | t < f a} - ≤ ∫⁻ (t : ℝ) in Set.Ioi 0, atTop.liminf (fun i ↦ (μs i) {a | t < f a}) - := (lintegral_mono (fun t ↦ h_opens _ (continuous_def.mp f_cont _ isOpen_Ioi))).trans ?_ - _ ≤ atTop.liminf (fun i ↦ ∫⁻ (t : ℝ) in Set.Ioi 0, (μs i) {a | t < f a}) - := lintegral_liminf_le (fun n ↦ Antitone.measurable - (fun s t hst ↦ measure_mono (fun ω hω ↦ lt_of_le_of_lt hst hω))) - rfl - -theorem BoundedContinuousFunction.lintegral_le_edist_mul - {μ : Measure Ω} [IsFiniteMeasure μ] (f : Ω →ᵇ ℝ≥0) : - (∫⁻ x, f x ∂μ) ≤ edist 0 f * (μ Set.univ) := by - have bound : ∀ x, f x ≤ nndist 0 f := by - intro x - convert nndist_coe_le_nndist x - simp only [coe_zero, Pi.zero_apply, NNReal.nndist_zero_eq_val] - apply le_trans (lintegral_mono (fun x ↦ ENNReal.coe_le_coe.mpr (bound x))) - simp - --- Missing? -lemma ENNReal.monotoneOn_toReal : MonotoneOn ENNReal.toReal {∞}ᶜ := - fun _ _ _ hy x_le_y ↦ toReal_mono hy x_le_y - --- Argh. :| -lemma Real.lipschitzWith_toNNReal : LipschitzWith 1 Real.toNNReal := by - apply lipschitzWith_iff_dist_le_mul.mpr - intro x y - simp only [NNReal.coe_one, one_mul, NNReal.dist_eq, coe_toNNReal', ge_iff_le, Real.dist_eq] - simpa only [ge_iff_le, NNReal.coe_one, dist_prod_same_right, one_mul, Real.dist_eq] using - lipschitzWith_iff_dist_le_mul.mp lipschitzWith_max (x, 0) (y, 0) - -lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure - {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] - {f : Ω →ᵇ ℝ} (f_nn : 0 ≤ f) - (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : - ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by - have same := lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure - f.continuous f_nn h_opens - rw [@integral_eq_lintegral_of_nonneg_ae Ω _ μ f (eventually_of_forall f_nn) - f.continuous.measurable.aestronglyMeasurable] - convert (ENNReal.toReal_le_toReal ?_ ?_).mpr same - · simp only [fun i ↦ @integral_eq_lintegral_of_nonneg_ae Ω _ (μs i) f (eventually_of_forall f_nn) - f.continuous.measurable.aestronglyMeasurable] - let g := BoundedContinuousFunction.comp _ Real.lipschitzWith_toNNReal f - have bound : ∀ i, ∫⁻ x, ENNReal.ofReal (f x) ∂(μs i) ≤ nndist 0 g := fun i ↦ by - simpa only [coe_nnreal_ennreal_nndist, measure_univ, mul_one, ge_iff_le] using - BoundedContinuousFunction.lintegral_le_edist_mul (μ := μs i) g - apply ENNReal.liminf_toReal_eq ENNReal.coe_ne_top bound - · exact (f.lintegral_of_real_lt_top μ).ne - · apply ne_of_lt - have obs := fun (i : ℕ) ↦ @BoundedContinuousFunction.lintegral_nnnorm_le Ω _ _ (μs i) ℝ _ f - simp only [measure_univ, mul_one] at obs - apply lt_of_le_of_lt _ (show (‖f‖₊ : ℝ≥0∞) < ∞ from ENNReal.coe_lt_top) - apply liminf_le_of_le - · refine ⟨0, eventually_of_forall (by simp only [ge_iff_le, zero_le, forall_const])⟩ - · intro x hx - obtain ⟨i, hi⟩ := hx.exists - apply le_trans hi - convert obs i with x - have aux := ENNReal.ofReal_eq_coe_nnreal (f_nn x) - simp only [ContinuousMap.toFun_eq_coe, BoundedContinuousFunction.coe_to_continuous_fun] at aux - rw [aux] - congr - exact (Real.norm_of_nonneg (f_nn x)).symm - -lemma tendsto_integral_of_forall_integral_le_liminf_integral {ι : Type*} {L : Filter ι} - {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ι → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] - (h : ∀ f : Ω →ᵇ ℝ, 0 ≤ f → ∫ x, (f x) ∂μ ≤ L.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i))) - (f : Ω →ᵇ ℝ) : - Tendsto (fun i ↦ ∫ x, (f x) ∂ (μs i)) L (𝓝 (∫ x, (f x) ∂μ)) := by - by_cases L_bot : L = ⊥ - · simp only [L_bot, tendsto_bot] - have : NeBot L := ⟨L_bot⟩ - have obs := BoundedContinuousFunction.isBounded_range_integral μs f - have bdd_above : IsBoundedUnder (· ≤ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := - isBounded_le_map_of_bounded_range _ obs - have bdd_below : IsBoundedUnder (· ≥ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := - isBounded_ge_map_of_bounded_range _ obs - apply @tendsto_of_le_liminf_of_limsup_le ℝ ι _ _ _ L (fun i ↦ ∫ x, (f x) ∂ (μs i)) (∫ x, (f x) ∂μ) - · have key := h _ (f.add_norm_nonneg) - simp_rw [f.integral_add_const ‖f‖] at key - simp only [measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key - have := liminf_add_const L (fun i ↦ ∫ x, (f x) ∂ (μs i)) ‖f‖ bdd_above bdd_below - rwa [this, add_le_add_iff_right] at key - · have key := h _ (f.norm_sub_nonneg) - simp_rw [f.integral_const_sub ‖f‖] at key - simp only [measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key - have := liminf_const_sub L (fun i ↦ ∫ x, (f x) ∂ (μs i)) ‖f‖ bdd_above bdd_below - rwa [this, sub_le_sub_iff_left] at key - · exact bdd_above - · exact bdd_below - --- Not needed? -/-- A characterization of weak convergence of probability measures by the condition that the -integrals of every continuous bounded nonnegative function converge to the integral of the function -against the limit measure. -/ -lemma ProbabilityMeasure.tendsto_iff_forall_nonneg_integral_tendsto {γ : Type _} {F : Filter γ} - {μs : γ → ProbabilityMeasure Ω} {μ : ProbabilityMeasure Ω} : - Tendsto μs F (𝓝 μ) ↔ - ∀ f : Ω →ᵇ ℝ, 0 ≤ f → - Tendsto (fun i ↦ ∫ ω, (f ω) ∂(μs i : Measure Ω)) F (𝓝 (∫ ω, (f ω) ∂(μ : Measure Ω))) := by - rw [ProbabilityMeasure.tendsto_iff_forall_integral_tendsto] - refine ⟨fun h f _ ↦ h f, fun h f ↦ ?_⟩ - set g := f + BoundedContinuousFunction.const _ ‖f‖ with g_def - have fx_eq : ∀ x, f x = g x - ‖f‖ := by - intro x - simp only [BoundedContinuousFunction.coe_add, BoundedContinuousFunction.const_toFun, Pi.add_apply, - add_sub_cancel] - have key := h g (f.add_norm_nonneg) - simp [g_def] at key - simp_rw [integral_add (f.integrable _) (integrable_const ‖f‖)] at key - simp only [integral_const, measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key - simp_rw [fx_eq] - convert tendsto_add.comp (Tendsto.prod_mk_nhds key (@tendsto_const_nhds _ _ _ (-‖f‖) F)) <;> simp - -/-- One implication of the portmanteau theorem. -/ -theorem ProbabilityMeasure.tendsto_of_forall_isOpen_le_liminf {μ : ProbabilityMeasure Ω} - {μs : ℕ → ProbabilityMeasure Ω} (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : - atTop.Tendsto (fun i ↦ μs i) (𝓝 μ) := by - refine ProbabilityMeasure.tendsto_iff_forall_integral_tendsto.mpr ?_ - apply tendsto_integral_of_forall_integral_le_liminf_integral - intro f f_nn - apply integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure (f := f) f_nn - intro G G_open - specialize h_opens G G_open - simp only at h_opens - have aux := Monotone.map_liminf_of_continuousAt (F := atTop) ENNReal.coe_mono (μs · G) ?_ ?_ ?_ - · have obs := ENNReal.coe_mono h_opens - simp only [ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure, aux] at obs - convert obs - simp only [Function.comp_apply, ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure] - · apply ENNReal.continuous_coe.continuousAt - · use 1 - simp only [eventually_map, ProbabilityMeasure.apply_le_one, eventually_atTop, ge_iff_le, - implies_true, forall_const, exists_const] - · use 0 - simp only [zero_le, eventually_map, eventually_atTop, ge_iff_le, implies_true, forall_const, - exists_const] - -end le_liminf_open_implies_convergence diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 688949454900e..ffbfd835ee890 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -136,6 +136,9 @@ theorem tendsto_toReal {a : ℝ≥0∞} (ha : a ≠ ⊤) : Tendsto ENNReal.toRea NNReal.tendsto_coe.2 <| tendsto_toNNReal ha #align ennreal.tendsto_to_real ENNReal.tendsto_toReal +lemma continuousOn_toReal : ContinuousOn ENNReal.toReal { a | a ≠ ∞ } := + NNReal.continuous_coe.comp_continuousOn continuousOn_toNNReal + /-- The set of finite `ℝ≥0∞` numbers is homeomorphic to `ℝ≥0`. -/ def neTopHomeomorphNNReal : { a | a ≠ ∞ } ≃ₜ ℝ≥0 where toEquiv := neTopEquivNNReal @@ -1603,9 +1606,47 @@ theorem edist_le_tsum_of_edist_le_of_tendsto₀ {f : ℕ → α} (d : ℕ → end +namespace ENNReal + +section truncateToReal + +/-- With truncation level `t`, the truncated cast `ℝ≥0∞ → ℝ` is given by `x ↦ (min t x).toReal`. +Unlike `ENNReal.toReal`, this cast is continuous and monotone when `t ≠ ∞`. -/ +noncomputable def truncateToReal (t x : ℝ≥0∞) : ℝ := (min t x).toReal + +lemma truncateToReal_eq_toReal {t x : ℝ≥0∞} (t_ne_top : t ≠ ∞) (x_le : x ≤ t) : + truncateToReal t x = x.toReal := by + have x_lt_top : x < ∞ := lt_of_le_of_lt x_le t_ne_top.lt_top + have obs : min t x ≠ ∞ := by + simp_all only [ne_eq, ge_iff_le, min_eq_top, false_and, not_false_eq_true] + exact (ENNReal.toReal_eq_toReal obs x_lt_top.ne).mpr (min_eq_right x_le) + +lemma truncateToReal_le {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) {x : ℝ≥0∞} : + truncateToReal t x ≤ t.toReal := by + rw [truncateToReal] + apply (toReal_le_toReal _ t_ne_top).mpr (min_le_left t x) + simp_all only [ne_eq, ge_iff_le, min_eq_top, false_and, not_false_eq_true] + +lemma truncateToReal_nonneg {t x : ℝ≥0∞} : 0 ≤ truncateToReal t x := toReal_nonneg + +/-- The truncated cast `ENNReal.truncateToReal t : ℝ≥0∞ → ℝ` is monotone when `t ≠ ∞`. -/ +lemma monotone_truncateToReal {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) : Monotone (truncateToReal t) := by + intro x y x_le_y + have obs_x : min t x ≠ ∞ := by + simp_all only [ne_eq, ge_iff_le, min_eq_top, false_and, not_false_eq_true] + have obs_y : min t y ≠ ∞ := by + simp_all only [ne_eq, ge_iff_le, min_eq_top, false_and, not_false_eq_true] + exact (ENNReal.toReal_le_toReal obs_x obs_y).mpr (min_le_min_left t x_le_y) + +/-- The truncated cast `ENNReal.truncateToReal t : ℝ≥0∞ → ℝ` is continuous when `t ≠ ∞`. -/ +lemma continuous_truncateToReal {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) : Continuous (truncateToReal t) := by + apply continuousOn_toReal.comp_continuous (continuous_min.comp (Continuous.Prod.mk t)) + simp [t_ne_top] + +end truncateToReal + section LimsupLiminf -namespace ENNReal set_option autoImplicit true lemma limsup_sub_const (F : Filter ι) [NeBot F] (f : ι → ℝ≥0∞) (c : ℝ≥0∞) : @@ -1630,6 +1671,44 @@ lemma liminf_const_sub (F : Filter ι) [NeBot F] (f : ι → ℝ≥0∞) (Antitone.map_limsSup_of_continuousAt (F := F.map f) (f := fun (x : ℝ≥0∞) ↦ c - x) (fun _ _ h ↦ tsub_le_tsub_left h c) (continuous_sub_left c_ne_top).continuousAt).symm -end ENNReal -- namespace +/-- If `xs : ι → ℝ≥0∞` is bounded, then we have `liminf (toReal ∘ xs) = toReal (liminf xs)`. -/ +lemma liminf_toReal_eq {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} (b_ne_top : b ≠ ∞) + {xs : ι → ℝ≥0∞} (le_b : ∀ i, xs i ≤ b) : + F.liminf (fun i ↦ (xs i).toReal) = (F.liminf xs).toReal := by + have liminf_le : F.liminf xs ≤ b := by + apply liminf_le_of_le ⟨0, by simp⟩ + intro y h + obtain ⟨i, hi⟩ := h.exists + exact hi.trans (le_b i) + have aux : ∀ i, (xs i).toReal = ENNReal.truncateToReal b (xs i) := by + simp only [truncateToReal_eq_toReal b_ne_top (le_b _), implies_true] + have aux' : (F.liminf xs).toReal = ENNReal.truncateToReal b (F.liminf xs) := by + rw [truncateToReal_eq_toReal b_ne_top liminf_le] + simp_rw [aux, aux'] + have key := Monotone.map_liminf_of_continuousAt (F := F) (monotone_truncateToReal b_ne_top) xs + (continuous_truncateToReal b_ne_top).continuousAt + ⟨b, by simpa only [eventually_map] using eventually_of_forall le_b⟩ + ⟨0, eventually_of_forall (by simp)⟩ + rw [key] + rfl + +/-- If `xs : ι → ℝ≥0∞` is bounded, then we have `liminf (toReal ∘ xs) = toReal (liminf xs)`. -/ +lemma limsup_toReal_eq {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} (b_ne_top : b ≠ ∞) + {xs : ι → ℝ≥0∞} (le_b : ∀ i, xs i ≤ b) : + F.limsup (fun i ↦ (xs i).toReal) = (F.limsup xs).toReal := by + have aux : ∀ i, (xs i).toReal = ENNReal.truncateToReal b (xs i) := by + simp only [truncateToReal_eq_toReal b_ne_top (le_b _), implies_true] + have aux' : (F.limsup xs).toReal = ENNReal.truncateToReal b (F.limsup xs) := by + rw [truncateToReal_eq_toReal b_ne_top + (limsup_le_of_le ⟨0, by simp⟩ (eventually_of_forall le_b))] + simp_rw [aux, aux'] + have key := Monotone.map_limsup_of_continuousAt (F := F) (monotone_truncateToReal b_ne_top) xs + (continuous_truncateToReal b_ne_top).continuousAt + ⟨b, by simpa only [eventually_map] using eventually_of_forall le_b⟩ + ⟨0, eventually_of_forall (by simp)⟩ + rw [key] + rfl end LimsupLiminf + +end ENNReal -- namespace From cf94e140b8c50b3b27344a7e9d956d3e28cb4ae0 Mon Sep 17 00:00:00 2001 From: kkytola <“kalle.kytola@aalto.fi”> Date: Wed, 1 Nov 2023 20:20:12 +0200 Subject: [PATCH 26/41] Add Lipschitzness of Real.toNNReal. --- Mathlib/Topology/MetricSpace/Lipschitz.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Topology/MetricSpace/Lipschitz.lean b/Mathlib/Topology/MetricSpace/Lipschitz.lean index c81b6bc3bafa6..4f9721a3ec2ae 100644 --- a/Mathlib/Topology/MetricSpace/Lipschitz.lean +++ b/Mathlib/Topology/MetricSpace/Lipschitz.lean @@ -451,6 +451,11 @@ theorem _root_.lipschitzWith_min : LipschitzWith 1 fun p : ℝ × ℝ => min p.1 (le_abs_self _).trans (abs_min_sub_min_le_max _ _ _ _) #align lipschitz_with_min lipschitzWith_min +lemma _root_.Real.lipschitzWith_toNNReal : LipschitzWith 1 Real.toNNReal := by + refine lipschitzWith_iff_dist_le_mul.mpr (fun x y ↦ ?_) + simpa only [ge_iff_le, NNReal.coe_one, dist_prod_same_right, one_mul, Real.dist_eq] using + lipschitzWith_iff_dist_le_mul.mp lipschitzWith_max (x, 0) (y, 0) + end Metric section EMetric From fd6e7e6642537d8e5208518ee9c4873bba0ed2f3 Mon Sep 17 00:00:00 2001 From: kkytola <“kalle.kytola@aalto.fi”> Date: Wed, 1 Nov 2023 20:28:32 +0200 Subject: [PATCH 27/41] Indent correctly. --- Mathlib/Topology/Instances/ENNReal.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index ffbfd835ee890..04120a9a24c6d 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -1673,7 +1673,7 @@ lemma liminf_const_sub (F : Filter ι) [NeBot F] (f : ι → ℝ≥0∞) /-- If `xs : ι → ℝ≥0∞` is bounded, then we have `liminf (toReal ∘ xs) = toReal (liminf xs)`. -/ lemma liminf_toReal_eq {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} (b_ne_top : b ≠ ∞) - {xs : ι → ℝ≥0∞} (le_b : ∀ i, xs i ≤ b) : + {xs : ι → ℝ≥0∞} (le_b : ∀ i, xs i ≤ b) : F.liminf (fun i ↦ (xs i).toReal) = (F.liminf xs).toReal := by have liminf_le : F.liminf xs ≤ b := by apply liminf_le_of_le ⟨0, by simp⟩ @@ -1694,7 +1694,7 @@ lemma liminf_toReal_eq {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} ( /-- If `xs : ι → ℝ≥0∞` is bounded, then we have `liminf (toReal ∘ xs) = toReal (liminf xs)`. -/ lemma limsup_toReal_eq {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} (b_ne_top : b ≠ ∞) - {xs : ι → ℝ≥0∞} (le_b : ∀ i, xs i ≤ b) : + {xs : ι → ℝ≥0∞} (le_b : ∀ i, xs i ≤ b) : F.limsup (fun i ↦ (xs i).toReal) = (F.limsup xs).toReal := by have aux : ∀ i, (xs i).toReal = ENNReal.truncateToReal b (xs i) := by simp only [truncateToReal_eq_toReal b_ne_top (le_b _), implies_true] From cedbc623b51f323cd34fa671ee4070fa9c822be3 Mon Sep 17 00:00:00 2001 From: kkytola <“kalle.kytola@aalto.fi”> Date: Wed, 1 Nov 2023 20:31:48 +0200 Subject: [PATCH 28/41] Why did this material disappear when creating the PR? --- .../MeasureTheory/Measure/Portmanteau.lean | 129 ++++++++++++++++++ 1 file changed, 129 insertions(+) diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index bd7d49acc3d2e..3ae50f5b3d38b 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -539,4 +539,133 @@ lemma le_liminf_measure_open_of_forall_tendsto_measure end LimitBorelImpliesLimsupClosedLE --section +lemma Filter.isBounded_le_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} + (h : Bornology.IsBounded (Set.range f)) : + (F.map f).IsBounded (· ≤ ·) := by + rw [Real.isBounded_iff_bddBelow_bddAbove] at h + obtain ⟨c, hc⟩ := h.2 + refine isBoundedUnder_of ⟨c, by simpa [mem_upperBounds] using hc⟩ + +lemma Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} + (h : Bornology.IsBounded (Set.range f)) : + (F.map f).IsBounded (· ≥ ·) := by + rw [Real.isBounded_iff_bddBelow_bddAbove] at h + obtain ⟨c, hc⟩ := h.1 + apply isBoundedUnder_of ⟨c, by simpa [mem_lowerBounds] using hc⟩ + +section le_liminf_open_implies_convergence + +variable {Ω : Type} [MeasurableSpace Ω] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] + +lemma lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure + {μ : Measure Ω} [SigmaFinite μ] {μs : ℕ → Measure Ω} [∀ i, SigmaFinite (μs i)] + {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) + (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : + ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫⁻ x, ENNReal.ofReal (f x) ∂ (μs i)) := by + simp_rw [lintegral_eq_lintegral_meas_lt _ (eventually_of_forall f_nn) f_cont.aemeasurable] + calc ∫⁻ (t : ℝ) in Set.Ioi 0, μ {a | t < f a} + ≤ ∫⁻ (t : ℝ) in Set.Ioi 0, atTop.liminf (fun i ↦ (μs i) {a | t < f a}) + := (lintegral_mono (fun t ↦ h_opens _ (continuous_def.mp f_cont _ isOpen_Ioi))).trans ?_ + _ ≤ atTop.liminf (fun i ↦ ∫⁻ (t : ℝ) in Set.Ioi 0, (μs i) {a | t < f a}) + := lintegral_liminf_le (fun n ↦ Antitone.measurable + (fun s t hst ↦ measure_mono (fun ω hω ↦ lt_of_le_of_lt hst hω))) + rfl + +theorem BoundedContinuousFunction.lintegral_le_edist_mul + {μ : Measure Ω} [IsFiniteMeasure μ] (f : Ω →ᵇ ℝ≥0) : + (∫⁻ x, f x ∂μ) ≤ edist 0 f * (μ Set.univ) := by + have bound : ∀ x, f x ≤ nndist 0 f := by + intro x + convert nndist_coe_le_nndist x + simp only [coe_zero, Pi.zero_apply, NNReal.nndist_zero_eq_val] + apply le_trans (lintegral_mono (fun x ↦ ENNReal.coe_le_coe.mpr (bound x))) + simp + +lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure + {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] + {f : Ω →ᵇ ℝ} (f_nn : 0 ≤ f) + (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : + ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by + have same := lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure + f.continuous f_nn h_opens + rw [@integral_eq_lintegral_of_nonneg_ae Ω _ μ f (eventually_of_forall f_nn) + f.continuous.measurable.aestronglyMeasurable] + convert (ENNReal.toReal_le_toReal ?_ ?_).mpr same + · simp only [fun i ↦ @integral_eq_lintegral_of_nonneg_ae Ω _ (μs i) f (eventually_of_forall f_nn) + f.continuous.measurable.aestronglyMeasurable] + let g := BoundedContinuousFunction.comp _ Real.lipschitzWith_toNNReal f + have bound : ∀ i, ∫⁻ x, ENNReal.ofReal (f x) ∂(μs i) ≤ nndist 0 g := fun i ↦ by + simpa only [coe_nnreal_ennreal_nndist, measure_univ, mul_one, ge_iff_le] using + BoundedContinuousFunction.lintegral_le_edist_mul (μ := μs i) g + apply ENNReal.liminf_toReal_eq ENNReal.coe_ne_top bound + · exact (f.lintegral_of_real_lt_top μ).ne + · apply ne_of_lt + have obs := fun (i : ℕ) ↦ @BoundedContinuousFunction.lintegral_nnnorm_le Ω _ _ (μs i) ℝ _ f + simp only [measure_univ, mul_one] at obs + apply lt_of_le_of_lt _ (show (‖f‖₊ : ℝ≥0∞) < ∞ from ENNReal.coe_lt_top) + apply liminf_le_of_le + · refine ⟨0, eventually_of_forall (by simp only [ge_iff_le, zero_le, forall_const])⟩ + · intro x hx + obtain ⟨i, hi⟩ := hx.exists + apply le_trans hi + convert obs i with x + have aux := ENNReal.ofReal_eq_coe_nnreal (f_nn x) + simp only [ContinuousMap.toFun_eq_coe, BoundedContinuousFunction.coe_to_continuous_fun] at aux + rw [aux] + congr + exact (Real.norm_of_nonneg (f_nn x)).symm + +lemma tendsto_integral_of_forall_integral_le_liminf_integral {ι : Type*} {L : Filter ι} + {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ι → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] + (h : ∀ f : Ω →ᵇ ℝ, 0 ≤ f → ∫ x, (f x) ∂μ ≤ L.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i))) + (f : Ω →ᵇ ℝ) : + Tendsto (fun i ↦ ∫ x, (f x) ∂ (μs i)) L (𝓝 (∫ x, (f x) ∂μ)) := by + by_cases L_bot : L = ⊥ + · simp only [L_bot, tendsto_bot] + have : NeBot L := ⟨L_bot⟩ + have obs := BoundedContinuousFunction.isBounded_range_integral μs f + have bdd_above : IsBoundedUnder (· ≤ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := + isBounded_le_map_of_bounded_range _ obs + have bdd_below : IsBoundedUnder (· ≥ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := + isBounded_ge_map_of_bounded_range _ obs + apply @tendsto_of_le_liminf_of_limsup_le ℝ ι _ _ _ L (fun i ↦ ∫ x, (f x) ∂ (μs i)) (∫ x, (f x) ∂μ) + · have key := h _ (f.add_norm_nonneg) + simp_rw [f.integral_add_const ‖f‖] at key + simp only [measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key + have := liminf_add_const L (fun i ↦ ∫ x, (f x) ∂ (μs i)) ‖f‖ bdd_above bdd_below + rwa [this, add_le_add_iff_right] at key + · have key := h _ (f.norm_sub_nonneg) + simp_rw [f.integral_const_sub ‖f‖] at key + simp only [measure_univ, ENNReal.one_toReal, smul_eq_mul, one_mul] at key + have := liminf_const_sub L (fun i ↦ ∫ x, (f x) ∂ (μs i)) ‖f‖ bdd_above bdd_below + rwa [this, sub_le_sub_iff_left] at key + · exact bdd_above + · exact bdd_below + +/-- One implication of the portmanteau theorem. -/ +theorem ProbabilityMeasure.tendsto_of_forall_isOpen_le_liminf {μ : ProbabilityMeasure Ω} + {μs : ℕ → ProbabilityMeasure Ω} (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : + atTop.Tendsto (fun i ↦ μs i) (𝓝 μ) := by + refine ProbabilityMeasure.tendsto_iff_forall_integral_tendsto.mpr ?_ + apply tendsto_integral_of_forall_integral_le_liminf_integral + intro f f_nn + apply integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure (f := f) f_nn + intro G G_open + specialize h_opens G G_open + simp only at h_opens + have aux := Monotone.map_liminf_of_continuousAt (F := atTop) ENNReal.coe_mono (μs · G) ?_ ?_ ?_ + · have obs := ENNReal.coe_mono h_opens + simp only [ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure, aux] at obs + convert obs + simp only [Function.comp_apply, ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure] + · apply ENNReal.continuous_coe.continuousAt + · use 1 + simp only [eventually_map, ProbabilityMeasure.apply_le_one, eventually_atTop, ge_iff_le, + implies_true, forall_const, exists_const] + · use 0 + simp only [zero_le, eventually_map, eventually_atTop, ge_iff_le, implies_true, forall_const, + exists_const] + +end le_liminf_open_implies_convergence + end MeasureTheory --namespace From 552d0bc4b4c6b0e8cbcb8ff4fc9e1c3fe4206de1 Mon Sep 17 00:00:00 2001 From: Kalle Date: Thu, 2 Nov 2023 08:46:37 +0200 Subject: [PATCH 29/41] Add a necessary import (Layercake) and fix broken things. --- Mathlib/MeasureTheory/Measure/Portmanteau.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 3ae50f5b3d38b..e4dc0fd704b29 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -5,6 +5,7 @@ Authors: Kalle Kytölä -/ import Mathlib.MeasureTheory.Measure.ProbabilityMeasure import Mathlib.MeasureTheory.Measure.Lebesgue.Basic +import Mathlib.MeasureTheory.Integral.Layercake #align_import measure_theory.measure.portmanteau from "leanprover-community/mathlib"@"fd5edc43dc4f10b85abfe544b88f82cf13c5f844" @@ -539,14 +540,14 @@ lemma le_liminf_measure_open_of_forall_tendsto_measure end LimitBorelImpliesLimsupClosedLE --section -lemma Filter.isBounded_le_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} +lemma _root_.Filter.isBounded_le_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} (h : Bornology.IsBounded (Set.range f)) : (F.map f).IsBounded (· ≤ ·) := by rw [Real.isBounded_iff_bddBelow_bddAbove] at h obtain ⟨c, hc⟩ := h.2 refine isBoundedUnder_of ⟨c, by simpa [mem_upperBounds] using hc⟩ -lemma Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} +lemma _root_.Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} (h : Bornology.IsBounded (Set.range f)) : (F.map f).IsBounded (· ≥ ·) := by rw [Real.isBounded_iff_bddBelow_bddAbove] at h From 86e01b758c711f97c7cbc0199db50675417d43ff Mon Sep 17 00:00:00 2001 From: kkytola <39528102+kkytola@users.noreply.github.com> Date: Thu, 2 Nov 2023 21:12:33 +0200 Subject: [PATCH 30/41] Apply suggestions from code review Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/MeasureTheory/Measure/Portmanteau.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index e4dc0fd704b29..f3c9c47e236f3 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -573,7 +573,7 @@ lemma lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure rfl theorem BoundedContinuousFunction.lintegral_le_edist_mul - {μ : Measure Ω} [IsFiniteMeasure μ] (f : Ω →ᵇ ℝ≥0) : + {μ : Measure Ω} [IsFiniteMeasure μ] (f : Ω →ᵇ ℝ≥0) : (∫⁻ x, f x ∂μ) ≤ edist 0 f * (μ Set.univ) := by have bound : ∀ x, f x ≤ nndist 0 f := by intro x @@ -645,7 +645,7 @@ lemma tendsto_integral_of_forall_integral_le_liminf_integral {ι : Type*} {L : F /-- One implication of the portmanteau theorem. -/ theorem ProbabilityMeasure.tendsto_of_forall_isOpen_le_liminf {μ : ProbabilityMeasure Ω} - {μs : ℕ → ProbabilityMeasure Ω} (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : + {μs : ℕ → ProbabilityMeasure Ω} (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : atTop.Tendsto (fun i ↦ μs i) (𝓝 μ) := by refine ProbabilityMeasure.tendsto_iff_forall_integral_tendsto.mpr ?_ apply tendsto_integral_of_forall_integral_le_liminf_integral From e4636d8467c32230c6925307e7ab6f9a1bca0d11 Mon Sep 17 00:00:00 2001 From: Kalle Date: Thu, 2 Nov 2023 21:46:05 +0200 Subject: [PATCH 31/41] Moved a lemma and added docstrings and rearranged some. --- .../Integral/BoundedContinuousFunction.lean | 10 +++ .../MeasureTheory/Measure/Portmanteau.lean | 70 ++++++++++--------- 2 files changed, 46 insertions(+), 34 deletions(-) diff --git a/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean index 2826e4916bb1e..080af3b6fe849 100644 --- a/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean +++ b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean @@ -22,8 +22,18 @@ namespace BoundedContinuousFunction section NNRealValued +lemma apply_le_nndist_zero {X : Type*} [TopologicalSpace X] (f : X →ᵇ ℝ≥0) (x : X) : + f x ≤ nndist 0 f := by + convert nndist_coe_le_nndist x + simp only [coe_zero, Pi.zero_apply, NNReal.nndist_zero_eq_val] + variable {X : Type*} [MeasurableSpace X] [TopologicalSpace X] [OpensMeasurableSpace X] +lemma lintegral_le_edist_mul (f : X →ᵇ ℝ≥0) (μ : Measure X) : + (∫⁻ x, f x ∂μ) ≤ edist 0 f * (μ Set.univ) := by + apply le_trans (lintegral_mono (fun x ↦ ENNReal.coe_le_coe.mpr (f.apply_le_nndist_zero x))) + simp + theorem measurable_coe_ennreal_comp (f : X →ᵇ ℝ≥0) : Measurable fun x ↦ (f x : ℝ≥0∞) := measurable_coe_nnreal_ennreal.comp f.continuous.measurable diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index f3c9c47e236f3..9e2c9d7dd7e72 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -6,6 +6,7 @@ Authors: Kalle Kytölä import Mathlib.MeasureTheory.Measure.ProbabilityMeasure import Mathlib.MeasureTheory.Measure.Lebesgue.Basic import Mathlib.MeasureTheory.Integral.Layercake +import Mathlib.MeasureTheory.Integral.BoundedContinuousFunction #align_import measure_theory.measure.portmanteau from "leanprover-community/mathlib"@"fd5edc43dc4f10b85abfe544b88f82cf13c5f844" @@ -40,9 +41,8 @@ The separate implications are: * `MeasureTheory.limsup_measure_closed_le_iff_liminf_measure_open_ge` is the equivalence (C) ↔ (O). * `MeasureTheory.tendsto_measure_of_null_frontier` is the implication (O) → (B). * `MeasureTheory.limsup_measure_closed_le_of_forall_tendsto_measure` is the implication (B) → (C). - -TODO: - * Prove the remaining implication (O) → (T) to complete the proof of equivalence of the conditions. + * `MeasureTheory.tendsto_of_forall_isOpen_le_liminf` gives the implication (O) → (T) for + any sequence of Borel probability measures. ## Implementation notes @@ -78,13 +78,24 @@ probability measure noncomputable section -open MeasureTheory +open MeasureTheory Set Filter BoundedContinuousFunction -open Set +-- Q: Where do these belong? (Tried `Topology.Instances.Real` and `Order.LiminfLimsup` and +-- `Topology.Algebra.Order.LiminfLimsup` but none worked. `#find_home` suggests the present file.) -open Filter +lemma Filter.isBounded_le_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} + (h : Bornology.IsBounded (Set.range f)) : + (F.map f).IsBounded (· ≤ ·) := by + rw [Real.isBounded_iff_bddBelow_bddAbove] at h + obtain ⟨c, hc⟩ := h.2 + refine isBoundedUnder_of ⟨c, by simpa [mem_upperBounds] using hc⟩ -open BoundedContinuousFunction +lemma Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} + (h : Bornology.IsBounded (Set.range f)) : + (F.map f).IsBounded (· ≥ ·) := by + rw [Real.isBounded_iff_bddBelow_bddAbove] at h + obtain ⟨c, hc⟩ := h.1 + apply isBoundedUnder_of ⟨c, by simpa [mem_lowerBounds] using hc⟩ open scoped Topology ENNReal NNReal BoundedContinuousFunction @@ -540,27 +551,26 @@ lemma le_liminf_measure_open_of_forall_tendsto_measure end LimitBorelImpliesLimsupClosedLE --section -lemma _root_.Filter.isBounded_le_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} - (h : Bornology.IsBounded (Set.range f)) : - (F.map f).IsBounded (· ≤ ·) := by - rw [Real.isBounded_iff_bddBelow_bddAbove] at h - obtain ⟨c, hc⟩ := h.2 - refine isBoundedUnder_of ⟨c, by simpa [mem_upperBounds] using hc⟩ +section le_liminf_open_implies_convergence -lemma _root_.Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} - (h : Bornology.IsBounded (Set.range f)) : - (F.map f).IsBounded (· ≥ ·) := by - rw [Real.isBounded_iff_bddBelow_bddAbove] at h - obtain ⟨c, hc⟩ := h.1 - apply isBoundedUnder_of ⟨c, by simpa [mem_lowerBounds] using hc⟩ +/-! ### Portmanteau implication: liminf condition for open sets implies weak convergence -section le_liminf_open_implies_convergence + +In this section we prove for a sequence (μsₙ)ₙ Borel probability measures that + + (O) For any open set G, the liminf of the measures of G under μsₙ is at least + the measure of G under μ, i.e., μ(G) ≤ liminfₙ μsₙ(G). + +implies + + (T) The measures μsₙ converge weakly to the measure μ. + +-/ variable {Ω : Type} [MeasurableSpace Ω] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] lemma lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure - {μ : Measure Ω} [SigmaFinite μ] {μs : ℕ → Measure Ω} [∀ i, SigmaFinite (μs i)] - {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) + {μ : Measure Ω} {μs : ℕ → Measure Ω} {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫⁻ x, ENNReal.ofReal (f x) ∂ (μs i)) := by simp_rw [lintegral_eq_lintegral_meas_lt _ (eventually_of_forall f_nn) f_cont.aemeasurable] @@ -572,16 +582,6 @@ lemma lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure (fun s t hst ↦ measure_mono (fun ω hω ↦ lt_of_le_of_lt hst hω))) rfl -theorem BoundedContinuousFunction.lintegral_le_edist_mul - {μ : Measure Ω} [IsFiniteMeasure μ] (f : Ω →ᵇ ℝ≥0) : - (∫⁻ x, f x ∂μ) ≤ edist 0 f * (μ Set.univ) := by - have bound : ∀ x, f x ≤ nndist 0 f := by - intro x - convert nndist_coe_le_nndist x - simp only [coe_zero, Pi.zero_apply, NNReal.nndist_zero_eq_val] - apply le_trans (lintegral_mono (fun x ↦ ENNReal.coe_le_coe.mpr (bound x))) - simp - lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] {f : Ω →ᵇ ℝ} (f_nn : 0 ≤ f) @@ -643,8 +643,10 @@ lemma tendsto_integral_of_forall_integral_le_liminf_integral {ι : Type*} {L : F · exact bdd_above · exact bdd_below -/-- One implication of the portmanteau theorem. -/ -theorem ProbabilityMeasure.tendsto_of_forall_isOpen_le_liminf {μ : ProbabilityMeasure Ω} +/-- One implication of the portmanteau theorem: +If for all open sets G we have the limsup condition μ(G) ≤ liminf μsᵢ(G), then the measures +μsₙ converge weakly to the measure μ. -/ +theorem tendsto_of_forall_isOpen_le_liminf {μ : ProbabilityMeasure Ω} {μs : ℕ → ProbabilityMeasure Ω} (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : atTop.Tendsto (fun i ↦ μs i) (𝓝 μ) := by refine ProbabilityMeasure.tendsto_iff_forall_integral_tendsto.mpr ?_ From 13070df0a65efd955c557328da104560bad88e75 Mon Sep 17 00:00:00 2001 From: Kalle Date: Thu, 2 Nov 2023 22:55:28 +0200 Subject: [PATCH 32/41] Lints. --- .../MeasureTheory/Measure/Portmanteau.lean | 19 +++++++++++-------- 1 file changed, 11 insertions(+), 8 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 9e2c9d7dd7e72..a7938abd3efb2 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -575,18 +575,20 @@ lemma lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫⁻ x, ENNReal.ofReal (f x) ∂ (μs i)) := by simp_rw [lintegral_eq_lintegral_meas_lt _ (eventually_of_forall f_nn) f_cont.aemeasurable] calc ∫⁻ (t : ℝ) in Set.Ioi 0, μ {a | t < f a} - ≤ ∫⁻ (t : ℝ) in Set.Ioi 0, atTop.liminf (fun i ↦ (μs i) {a | t < f a}) - := (lintegral_mono (fun t ↦ h_opens _ (continuous_def.mp f_cont _ isOpen_Ioi))).trans ?_ - _ ≤ atTop.liminf (fun i ↦ ∫⁻ (t : ℝ) in Set.Ioi 0, (μs i) {a | t < f a}) - := lintegral_liminf_le (fun n ↦ Antitone.measurable - (fun s t hst ↦ measure_mono (fun ω hω ↦ lt_of_le_of_lt hst hω))) - rfl + ≤ ∫⁻ (t : ℝ) in Set.Ioi 0, atTop.liminf (fun i ↦ (μs i) {a | t < f a}) := ?_ -- (i) + _ ≤ atTop.liminf (fun i ↦ ∫⁻ (t : ℝ) in Set.Ioi 0, (μs i) {a | t < f a}) := ?_ -- (ii) + · -- (i) + exact (lintegral_mono (fun t ↦ h_opens _ (continuous_def.mp f_cont _ isOpen_Ioi))).trans + (le_refl _) + · -- (ii) + exact lintegral_liminf_le (fun n ↦ Antitone.measurable (fun s t hst ↦ + measure_mono (fun ω hω ↦ lt_of_le_of_lt hst hω))) lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] {f : Ω →ᵇ ℝ} (f_nn : 0 ≤ f) (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : - ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by + ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by have same := lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure f.continuous f_nn h_opens rw [@integral_eq_lintegral_of_nonneg_ae Ω _ μ f (eventually_of_forall f_nn) @@ -647,7 +649,8 @@ lemma tendsto_integral_of_forall_integral_le_liminf_integral {ι : Type*} {L : F If for all open sets G we have the limsup condition μ(G) ≤ liminf μsᵢ(G), then the measures μsₙ converge weakly to the measure μ. -/ theorem tendsto_of_forall_isOpen_le_liminf {μ : ProbabilityMeasure Ω} - {μs : ℕ → ProbabilityMeasure Ω} (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : + {μs : ℕ → ProbabilityMeasure Ω} + (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : atTop.Tendsto (fun i ↦ μs i) (𝓝 μ) := by refine ProbabilityMeasure.tendsto_iff_forall_integral_tendsto.mpr ?_ apply tendsto_integral_of_forall_integral_le_liminf_integral From 40cc65220b79e5a89cdf11bd0bd7899e996d4008 Mon Sep 17 00:00:00 2001 From: kkytola <39528102+kkytola@users.noreply.github.com> Date: Fri, 10 Nov 2023 21:51:02 +0200 Subject: [PATCH 33/41] Apply suggestions from code review Co-authored-by: sgouezel --- .../MeasureTheory/Integral/BoundedContinuousFunction.lean | 5 ++--- Mathlib/MeasureTheory/Measure/Portmanteau.lean | 5 ++--- 2 files changed, 4 insertions(+), 6 deletions(-) diff --git a/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean index 080af3b6fe849..fda4b0db8e29d 100644 --- a/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean +++ b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean @@ -30,9 +30,8 @@ lemma apply_le_nndist_zero {X : Type*} [TopologicalSpace X] (f : X →ᵇ ℝ≥ variable {X : Type*} [MeasurableSpace X] [TopologicalSpace X] [OpensMeasurableSpace X] lemma lintegral_le_edist_mul (f : X →ᵇ ℝ≥0) (μ : Measure X) : - (∫⁻ x, f x ∂μ) ≤ edist 0 f * (μ Set.univ) := by - apply le_trans (lintegral_mono (fun x ↦ ENNReal.coe_le_coe.mpr (f.apply_le_nndist_zero x))) - simp + (∫⁻ x, f x ∂μ) ≤ edist 0 f * (μ Set.univ) := + le_trans (lintegral_mono (fun x ↦ ENNReal.coe_le_coe.mpr (f.apply_le_nndist_zero x))) (by simp) theorem measurable_coe_ennreal_comp (f : X →ᵇ ℝ≥0) : Measurable fun x ↦ (f x : ℝ≥0∞) := diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index a7938abd3efb2..a30983b10979a 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -623,9 +623,8 @@ lemma tendsto_integral_of_forall_integral_le_liminf_integral {ι : Type*} {L : F (h : ∀ f : Ω →ᵇ ℝ, 0 ≤ f → ∫ x, (f x) ∂μ ≤ L.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i))) (f : Ω →ᵇ ℝ) : Tendsto (fun i ↦ ∫ x, (f x) ∂ (μs i)) L (𝓝 (∫ x, (f x) ∂μ)) := by - by_cases L_bot : L = ⊥ - · simp only [L_bot, tendsto_bot] - have : NeBot L := ⟨L_bot⟩ + rcases eq_or_neBot L with rfl|hL + · simp only [tendsto_bot] have obs := BoundedContinuousFunction.isBounded_range_integral μs f have bdd_above : IsBoundedUnder (· ≤ ·) L (fun i ↦ ∫ (x : Ω), f x ∂μs i) := isBounded_le_map_of_bounded_range _ obs From 49178d322b2d263bfbedc77fce166621b720dbe9 Mon Sep 17 00:00:00 2001 From: Kalle Date: Sat, 11 Nov 2023 00:04:24 +0200 Subject: [PATCH 34/41] Clarify proof as requested. --- .../MeasureTheory/Measure/Portmanteau.lean | 26 ++++++++++--------- 1 file changed, 14 insertions(+), 12 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index a30983b10979a..5de2f1deb1447 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -658,18 +658,20 @@ theorem tendsto_of_forall_isOpen_le_liminf {μ : ProbabilityMeasure Ω} intro G G_open specialize h_opens G G_open simp only at h_opens - have aux := Monotone.map_liminf_of_continuousAt (F := atTop) ENNReal.coe_mono (μs · G) ?_ ?_ ?_ - · have obs := ENNReal.coe_mono h_opens - simp only [ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure, aux] at obs - convert obs - simp only [Function.comp_apply, ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure] - · apply ENNReal.continuous_coe.continuousAt - · use 1 - simp only [eventually_map, ProbabilityMeasure.apply_le_one, eventually_atTop, ge_iff_le, - implies_true, forall_const, exists_const] - · use 0 - simp only [zero_le, eventually_map, eventually_atTop, ge_iff_le, implies_true, forall_const, - exists_const] + have aux : ENNReal.some (liminf (fun i ↦ ENNReal.toNNReal ((μs i : Measure Ω) G)) atTop) = + liminf (ENNReal.some ∘ fun i ↦ (ENNReal.toNNReal ((μs i : Measure Ω) G))) atTop := by + refine Monotone.map_liminf_of_continuousAt (F := atTop) ENNReal.coe_mono (μs · G) ?_ ?_ ?_ + · apply ENNReal.continuous_coe.continuousAt + · use 1 + simp only [eventually_map, ProbabilityMeasure.apply_le_one, eventually_atTop, ge_iff_le, + implies_true, forall_const, exists_const] + · use 0 + simp only [zero_le, eventually_map, eventually_atTop, ge_iff_le, implies_true, forall_const, + exists_const] + have obs := ENNReal.coe_mono h_opens + simp only [ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure, aux] at obs + convert obs + simp only [Function.comp_apply, ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure] end le_liminf_open_implies_convergence From 0f4f62c35bb996dba3bd849227e80766c86c04ed Mon Sep 17 00:00:00 2001 From: Kalle Date: Sat, 11 Nov 2023 00:23:38 +0200 Subject: [PATCH 35/41] Generalize two lemmas to assume the hypotheses only eventually along the filter. --- Mathlib/Topology/Instances/ENNReal.lean | 31 ++++++++++++------------- 1 file changed, 15 insertions(+), 16 deletions(-) diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 04120a9a24c6d..8efea62033c9c 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -1673,39 +1673,38 @@ lemma liminf_const_sub (F : Filter ι) [NeBot F] (f : ι → ℝ≥0∞) /-- If `xs : ι → ℝ≥0∞` is bounded, then we have `liminf (toReal ∘ xs) = toReal (liminf xs)`. -/ lemma liminf_toReal_eq {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} (b_ne_top : b ≠ ∞) - {xs : ι → ℝ≥0∞} (le_b : ∀ i, xs i ≤ b) : + {xs : ι → ℝ≥0∞} (le_b : ∀ᶠ i in F, xs i ≤ b) : F.liminf (fun i ↦ (xs i).toReal) = (F.liminf xs).toReal := by have liminf_le : F.liminf xs ≤ b := by apply liminf_le_of_le ⟨0, by simp⟩ intro y h - obtain ⟨i, hi⟩ := h.exists - exact hi.trans (le_b i) - have aux : ∀ i, (xs i).toReal = ENNReal.truncateToReal b (xs i) := by - simp only [truncateToReal_eq_toReal b_ne_top (le_b _), implies_true] + obtain ⟨i, hi⟩ := (Eventually.and h le_b).exists + exact hi.1.trans hi.2 + have aux : ∀ᶠ i in F, (xs i).toReal = ENNReal.truncateToReal b (xs i) := by + filter_upwards [le_b] with i i_le_b + simp only [truncateToReal_eq_toReal b_ne_top i_le_b, implies_true] have aux' : (F.liminf xs).toReal = ENNReal.truncateToReal b (F.liminf xs) := by rw [truncateToReal_eq_toReal b_ne_top liminf_le] - simp_rw [aux, aux'] + simp_rw [liminf_congr aux, aux'] have key := Monotone.map_liminf_of_continuousAt (F := F) (monotone_truncateToReal b_ne_top) xs (continuous_truncateToReal b_ne_top).continuousAt - ⟨b, by simpa only [eventually_map] using eventually_of_forall le_b⟩ - ⟨0, eventually_of_forall (by simp)⟩ + ⟨b, by simpa only [eventually_map] using le_b⟩ ⟨0, eventually_of_forall (by simp)⟩ rw [key] rfl /-- If `xs : ι → ℝ≥0∞` is bounded, then we have `liminf (toReal ∘ xs) = toReal (liminf xs)`. -/ lemma limsup_toReal_eq {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} (b_ne_top : b ≠ ∞) - {xs : ι → ℝ≥0∞} (le_b : ∀ i, xs i ≤ b) : + {xs : ι → ℝ≥0∞} (le_b : ∀ᶠ i in F, xs i ≤ b) : F.limsup (fun i ↦ (xs i).toReal) = (F.limsup xs).toReal := by - have aux : ∀ i, (xs i).toReal = ENNReal.truncateToReal b (xs i) := by - simp only [truncateToReal_eq_toReal b_ne_top (le_b _), implies_true] + have aux : ∀ᶠ i in F, (xs i).toReal = ENNReal.truncateToReal b (xs i) := by + filter_upwards [le_b] with i i_le_b + simp only [truncateToReal_eq_toReal b_ne_top i_le_b, implies_true] have aux' : (F.limsup xs).toReal = ENNReal.truncateToReal b (F.limsup xs) := by - rw [truncateToReal_eq_toReal b_ne_top - (limsup_le_of_le ⟨0, by simp⟩ (eventually_of_forall le_b))] - simp_rw [aux, aux'] + rw [truncateToReal_eq_toReal b_ne_top (limsup_le_of_le ⟨0, by simp⟩ le_b)] + simp_rw [limsup_congr aux, aux'] have key := Monotone.map_limsup_of_continuousAt (F := F) (monotone_truncateToReal b_ne_top) xs (continuous_truncateToReal b_ne_top).continuousAt - ⟨b, by simpa only [eventually_map] using eventually_of_forall le_b⟩ - ⟨0, eventually_of_forall (by simp)⟩ + ⟨b, by simpa only [eventually_map] using le_b⟩ ⟨0, eventually_of_forall (by simp)⟩ rw [key] rfl From 049625c452a7530d6d631a8430058835e62294e2 Mon Sep 17 00:00:00 2001 From: Kalle Date: Sat, 11 Nov 2023 00:38:01 +0200 Subject: [PATCH 36/41] Move two lemmas to their own file. --- .../MeasureTheory/Measure/Portmanteau.lean | 19 +-------- Mathlib/Topology/Order/Bounded.lean | 40 +++++++++++++++++++ 2 files changed, 41 insertions(+), 18 deletions(-) create mode 100644 Mathlib/Topology/Order/Bounded.lean diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 5de2f1deb1447..76920dd080036 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -7,6 +7,7 @@ import Mathlib.MeasureTheory.Measure.ProbabilityMeasure import Mathlib.MeasureTheory.Measure.Lebesgue.Basic import Mathlib.MeasureTheory.Integral.Layercake import Mathlib.MeasureTheory.Integral.BoundedContinuousFunction +import Mathlib.Topology.Order.Bounded #align_import measure_theory.measure.portmanteau from "leanprover-community/mathlib"@"fd5edc43dc4f10b85abfe544b88f82cf13c5f844" @@ -79,24 +80,6 @@ probability measure noncomputable section open MeasureTheory Set Filter BoundedContinuousFunction - --- Q: Where do these belong? (Tried `Topology.Instances.Real` and `Order.LiminfLimsup` and --- `Topology.Algebra.Order.LiminfLimsup` but none worked. `#find_home` suggests the present file.) - -lemma Filter.isBounded_le_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} - (h : Bornology.IsBounded (Set.range f)) : - (F.map f).IsBounded (· ≤ ·) := by - rw [Real.isBounded_iff_bddBelow_bddAbove] at h - obtain ⟨c, hc⟩ := h.2 - refine isBoundedUnder_of ⟨c, by simpa [mem_upperBounds] using hc⟩ - -lemma Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} - (h : Bornology.IsBounded (Set.range f)) : - (F.map f).IsBounded (· ≥ ·) := by - rw [Real.isBounded_iff_bddBelow_bddAbove] at h - obtain ⟨c, hc⟩ := h.1 - apply isBoundedUnder_of ⟨c, by simpa [mem_lowerBounds] using hc⟩ - open scoped Topology ENNReal NNReal BoundedContinuousFunction namespace MeasureTheory diff --git a/Mathlib/Topology/Order/Bounded.lean b/Mathlib/Topology/Order/Bounded.lean new file mode 100644 index 0000000000000..2a4f3962dc4c7 --- /dev/null +++ b/Mathlib/Topology/Order/Bounded.lean @@ -0,0 +1,40 @@ +/- +Copyright (c) 2023 Kalle Kytölä. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kalle Kytölä +-/ +import Mathlib.Topology.Bornology.Basic +import Mathlib.Topology.Instances.Real +import Mathlib.Order.LiminfLimsup + +/-! +# Relating order and metric boundedness + +In spaces equipped with both an order and a metric, there are separate notions of boundedness +associated with each of the two structures. In specific cases such as ℝ, there are results which +relate the two notions. + +## Tags + +bounded, bornology, order, metric +-/ + +open Set Filter + +section Real + +lemma Filter.isBounded_le_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} + (h : Bornology.IsBounded (Set.range f)) : + (F.map f).IsBounded (· ≤ ·) := by + rw [Real.isBounded_iff_bddBelow_bddAbove] at h + obtain ⟨c, hc⟩ := h.2 + refine isBoundedUnder_of ⟨c, by simpa [mem_upperBounds] using hc⟩ + +lemma Filter.isBounded_ge_map_of_bounded_range {ι : Type*} (F : Filter ι) {f : ι → ℝ} + (h : Bornology.IsBounded (Set.range f)) : + (F.map f).IsBounded (· ≥ ·) := by + rw [Real.isBounded_iff_bddBelow_bddAbove] at h + obtain ⟨c, hc⟩ := h.1 + apply isBoundedUnder_of ⟨c, by simpa [mem_lowerBounds] using hc⟩ + +end Real From e81404b4afb6f889d512651216a8c383652df2be Mon Sep 17 00:00:00 2001 From: Kalle Date: Sat, 11 Nov 2023 14:55:09 +0200 Subject: [PATCH 37/41] Fix a broken proof. --- Mathlib/MeasureTheory/Measure/Portmanteau.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 76920dd080036..27ccb8361e1f8 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -583,7 +583,7 @@ lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure have bound : ∀ i, ∫⁻ x, ENNReal.ofReal (f x) ∂(μs i) ≤ nndist 0 g := fun i ↦ by simpa only [coe_nnreal_ennreal_nndist, measure_univ, mul_one, ge_iff_le] using BoundedContinuousFunction.lintegral_le_edist_mul (μ := μs i) g - apply ENNReal.liminf_toReal_eq ENNReal.coe_ne_top bound + apply ENNReal.liminf_toReal_eq ENNReal.coe_ne_top (eventually_of_forall bound) · exact (f.lintegral_of_real_lt_top μ).ne · apply ne_of_lt have obs := fun (i : ℕ) ↦ @BoundedContinuousFunction.lintegral_nnnorm_le Ω _ _ (μs i) ℝ _ f From 1731ccf28b5b5c4b5735dd429f7a8b2325f3d8cb Mon Sep 17 00:00:00 2001 From: Kalle Date: Sat, 11 Nov 2023 14:55:27 +0200 Subject: [PATCH 38/41] Import the new file in Mathlib.lean. --- Mathlib.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib.lean b/Mathlib.lean index 62387c973b352..f6f5a27180fac 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3523,6 +3523,7 @@ import Mathlib.Topology.NoetherianSpace import Mathlib.Topology.OmegaCompletePartialOrder import Mathlib.Topology.Order import Mathlib.Topology.Order.Basic +import Mathlib.Topology.Order.Bounded import Mathlib.Topology.Order.Category.AlexDisc import Mathlib.Topology.Order.Category.FrameAdjunction import Mathlib.Topology.Order.Hom.Basic From a92fa4f991ce6751bacc057b79ede72335a307bb Mon Sep 17 00:00:00 2001 From: kkytola <39528102+kkytola@users.noreply.github.com> Date: Sat, 11 Nov 2023 21:49:16 +0200 Subject: [PATCH 39/41] Update Mathlib/MeasureTheory/Measure/Portmanteau.lean Co-authored-by: sgouezel --- Mathlib/MeasureTheory/Measure/Portmanteau.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 27ccb8361e1f8..e7237c453d69f 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -628,7 +628,7 @@ lemma tendsto_integral_of_forall_integral_le_liminf_integral {ι : Type*} {L : F · exact bdd_below /-- One implication of the portmanteau theorem: -If for all open sets G we have the limsup condition μ(G) ≤ liminf μsᵢ(G), then the measures +If for all open sets G we have the liminf condition `μ(G) ≤ liminf μsᵢ(G)`, then the measures μsₙ converge weakly to the measure μ. -/ theorem tendsto_of_forall_isOpen_le_liminf {μ : ProbabilityMeasure Ω} {μs : ℕ → ProbabilityMeasure Ω} From 31776b1050605141895d70d90def6d633d5be77f Mon Sep 17 00:00:00 2001 From: Kalle Date: Sat, 11 Nov 2023 21:53:40 +0200 Subject: [PATCH 40/41] Correct one subscript in a docstring. --- Mathlib/MeasureTheory/Measure/Portmanteau.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index e7237c453d69f..cd98d262d863e 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -628,7 +628,7 @@ lemma tendsto_integral_of_forall_integral_le_liminf_integral {ι : Type*} {L : F · exact bdd_below /-- One implication of the portmanteau theorem: -If for all open sets G we have the liminf condition `μ(G) ≤ liminf μsᵢ(G)`, then the measures +If for all open sets G we have the liminf condition `μ(G) ≤ liminf μsₙ(G)`, then the measures μsₙ converge weakly to the measure μ. -/ theorem tendsto_of_forall_isOpen_le_liminf {μ : ProbabilityMeasure Ω} {μs : ℕ → ProbabilityMeasure Ω} From 98debd0556c188bc9d373fe10207eaaf6fefe798 Mon Sep 17 00:00:00 2001 From: Kalle Date: Sat, 11 Nov 2023 23:23:47 +0200 Subject: [PATCH 41/41] Replace ENNReal.some by ENNReal.ofNNReal. --- Mathlib/MeasureTheory/Measure/Portmanteau.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index cd98d262d863e..16f7861445864 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -641,8 +641,8 @@ theorem tendsto_of_forall_isOpen_le_liminf {μ : ProbabilityMeasure Ω} intro G G_open specialize h_opens G G_open simp only at h_opens - have aux : ENNReal.some (liminf (fun i ↦ ENNReal.toNNReal ((μs i : Measure Ω) G)) atTop) = - liminf (ENNReal.some ∘ fun i ↦ (ENNReal.toNNReal ((μs i : Measure Ω) G))) atTop := by + have aux : ENNReal.ofNNReal (liminf (fun i ↦ ENNReal.toNNReal ((μs i : Measure Ω) G)) atTop) = + liminf (ENNReal.ofNNReal ∘ fun i ↦ (ENNReal.toNNReal ((μs i : Measure Ω) G))) atTop := by refine Monotone.map_liminf_of_continuousAt (F := atTop) ENNReal.coe_mono (μs · G) ?_ ?_ ?_ · apply ENNReal.continuous_coe.continuousAt · use 1