From d5f79f023a80c97eb202176ce0ca820d2bb53336 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 1 Feb 2024 09:18:21 +0100 Subject: [PATCH 001/129] partial work towards kernel disintegration --- Mathlib/Probability/Kernel/BuildKernel.lean | 28 + Mathlib/Probability/Kernel/CondCdf.lean | 189 +---- Mathlib/Probability/Kernel/StieltjesReal.lean | 209 +++++ Mathlib/Probability/Kernel/draft.lean | 718 ++++++++++++++++++ 4 files changed, 993 insertions(+), 151 deletions(-) create mode 100644 Mathlib/Probability/Kernel/BuildKernel.lean create mode 100644 Mathlib/Probability/Kernel/StieltjesReal.lean create mode 100644 Mathlib/Probability/Kernel/draft.lean diff --git a/Mathlib/Probability/Kernel/BuildKernel.lean b/Mathlib/Probability/Kernel/BuildKernel.lean new file mode 100644 index 0000000000000..2d3c36e85bc89 --- /dev/null +++ b/Mathlib/Probability/Kernel/BuildKernel.lean @@ -0,0 +1,28 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Probability.Kernel.StieltjesReal +import Mathlib.Probability.Kernel.Basic + +/-! + + +-/ + + +open MeasureTheory Set Filter TopologicalSpace + +open scoped NNReal ENNReal MeasureTheory Topology + +namespace ProbabilityTheory + +variable {α : Type*} [MeasurableSpace α] + +noncomputable +def IsCDFLike.kernel {f : α → ℚ → ℝ} (hf : IsCDFLike f) : kernel α ℝ where + val (a : α) := (todo2 hf a).measure + property := measurable_measure_todo2 hf + +end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index 49eaad73a699e..475eded804785 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -4,9 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.Data.Set.Lattice -import Mathlib.MeasureTheory.Measure.Stieltjes -import Mathlib.MeasureTheory.Decomposition.RadonNikodym -import Mathlib.MeasureTheory.Constructions.Prod.Basic +import Mathlib.Probability.Kernel.StieltjesReal #align_import probability.kernel.cond_cdf from "leanprover-community/mathlib"@"3b88f4005dc2e28d42f974cc1ce838f0dafb39b8" @@ -695,139 +693,56 @@ theorem inf_gt_condCDFRat (ρ : Measure (α × ℝ)) (a : α) (t : ℚ) : exact h.trans (mem_Ioi.mp x.prop).le #align probability_theory.inf_gt_cond_cdf_rat ProbabilityTheory.inf_gt_condCDFRat -/-- Conditional cdf of the measure given the value on `α`, as a plain function. This is an auxiliary -definition used to define `cond_cdf`. -/ -noncomputable irreducible_def condCDF' (ρ : Measure (α × ℝ)) : α → ℝ → ℝ := fun a t => - ⨅ r : { r' : ℚ // t < r' }, condCDFRat ρ a r -#align probability_theory.cond_cdf' ProbabilityTheory.condCDF' - -theorem condCDF'_def' {ρ : Measure (α × ℝ)} {a : α} {x : ℝ} : - condCDF' ρ a x = ⨅ r : { r : ℚ // x < r }, condCDFRat ρ a r := by rw [condCDF'] -#align probability_theory.cond_cdf'_def ProbabilityTheory.condCDF'_def' - -theorem condCDF'_eq_condCDFRat (ρ : Measure (α × ℝ)) (a : α) (r : ℚ) : - condCDF' ρ a r = condCDFRat ρ a r := by - rw [← inf_gt_condCDFRat ρ a r, condCDF'] - refine' Equiv.iInf_congr _ _ - · exact - { toFun := fun t => ⟨t.1, mod_cast t.2⟩ - invFun := fun t => ⟨t.1, mod_cast t.2⟩ - left_inv := fun t => by simp only [Subtype.coe_eta] - right_inv := fun t => by simp only [Subtype.coe_eta] } - · intro t - simp only [Equiv.coe_fn_mk, Subtype.coe_mk] -#align probability_theory.cond_cdf'_eq_cond_cdf_rat ProbabilityTheory.condCDF'_eq_condCDFRat - -theorem condCDF'_nonneg (ρ : Measure (α × ℝ)) (a : α) (r : ℝ) : 0 ≤ condCDF' ρ a r := by - have : Nonempty { r' : ℚ // r < ↑r' } := by - obtain ⟨r, hrx⟩ := exists_rat_gt r - exact ⟨⟨r, hrx⟩⟩ - rw [condCDF'_def] - exact le_ciInf fun r' => condCDFRat_nonneg ρ a _ -#align probability_theory.cond_cdf'_nonneg ProbabilityTheory.condCDF'_nonneg - -theorem bddBelow_range_condCDFRat_gt (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : - BddBelow (range fun r : { r' : ℚ // x < ↑r' } => condCDFRat ρ a r) := by - refine' ⟨0, fun z => _⟩; rintro ⟨u, rfl⟩; exact condCDFRat_nonneg ρ a _ -#align probability_theory.bdd_below_range_cond_cdf_rat_gt ProbabilityTheory.bddBelow_range_condCDFRat_gt - -theorem monotone_condCDF' (ρ : Measure (α × ℝ)) (a : α) : Monotone (condCDF' ρ a) := by - intro x y hxy - have : Nonempty { r' : ℚ // y < ↑r' } := by - obtain ⟨r, hrx⟩ := exists_rat_gt y - exact ⟨⟨r, hrx⟩⟩ - simp_rw [condCDF'_def] - refine' le_ciInf fun r => (ciInf_le _ _).trans_eq _ - · exact bddBelow_range_condCDFRat_gt ρ a x - · exact ⟨r.1, hxy.trans_lt r.prop⟩ - · rfl -#align probability_theory.monotone_cond_cdf' ProbabilityTheory.monotone_condCDF' - -theorem continuousWithinAt_condCDF'_Ici (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : - ContinuousWithinAt (condCDF' ρ a) (Ici x) x := by - rw [← continuousWithinAt_Ioi_iff_Ici] - convert Monotone.tendsto_nhdsWithin_Ioi (monotone_condCDF' ρ a) x - rw [sInf_image'] - have h' : ⨅ r : Ioi x, condCDF' ρ a r = ⨅ r : { r' : ℚ // x < r' }, condCDF' ρ a r := by - refine' Real.iInf_Ioi_eq_iInf_rat_gt x _ (monotone_condCDF' ρ a) - refine' ⟨0, fun z => _⟩ - rintro ⟨u, -, rfl⟩ - exact condCDF'_nonneg ρ a u - have h'' : - ⨅ r : { r' : ℚ // x < r' }, condCDF' ρ a r = - ⨅ r : { r' : ℚ // x < r' }, condCDFRat ρ a r := by - congr with r - exact condCDF'_eq_condCDFRat ρ a r - rw [h', h'', ContinuousWithinAt] - congr! - exact condCDF'_def' -#align probability_theory.continuous_within_at_cond_cdf'_Ici ProbabilityTheory.continuousWithinAt_condCDF'_Ici +lemma isCDFLike_condCDFRat (ρ : Measure (α × ℝ)) : IsCDFLike (condCDFRat ρ) where + mono := monotone_condCDFRat ρ + nonneg := condCDFRat_nonneg ρ + le_one := condCDFRat_le_one ρ + tendsto_atTop_one := tendsto_condCDFRat_atTop ρ + tendsto_atBot_zero := tendsto_condCDFRat_atBot ρ + iInf_rat_gt_eq := inf_gt_condCDFRat ρ + measurable := measurable_condCDFRat ρ + +#noalign probability_theory.cond_cdf' +#noalign probability_theory.cond_cdf'_def +#noalign probability_theory.cond_cdf'_eq_cond_cdf_rat +#noalign probability_theory.cond_cdf'_nonneg +#noalign probability_theory.bdd_below_range_cond_cdf_rat_gt +#noalign probability_theory.monotone_cond_cdf' +#noalign probability_theory.continuous_within_at_cond_cdf'_Ici /-! ### Conditional cdf -/ /-- Conditional cdf of the measure given the value on `α`, as a Stieltjes function. -/ -noncomputable def condCDF (ρ : Measure (α × ℝ)) (a : α) : StieltjesFunction where - toFun := condCDF' ρ a - mono' := monotone_condCDF' ρ a - right_continuous' x := continuousWithinAt_condCDF'_Ici ρ a x +noncomputable def condCDF (ρ : Measure (α × ℝ)) (a : α) : StieltjesFunction := + todo2 (isCDFLike_condCDFRat ρ) a #align probability_theory.cond_cdf ProbabilityTheory.condCDF theorem condCDF_eq_condCDFRat (ρ : Measure (α × ℝ)) (a : α) (r : ℚ) : condCDF ρ a r = condCDFRat ρ a r := - condCDF'_eq_condCDFRat ρ a r + todo2_eq _ _ r #align probability_theory.cond_cdf_eq_cond_cdf_rat ProbabilityTheory.condCDF_eq_condCDFRat /-- The conditional cdf is non-negative for all `a : α`. -/ theorem condCDF_nonneg (ρ : Measure (α × ℝ)) (a : α) (r : ℝ) : 0 ≤ condCDF ρ a r := - condCDF'_nonneg ρ a r + todo2_nonneg _ a r #align probability_theory.cond_cdf_nonneg ProbabilityTheory.condCDF_nonneg /-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ -theorem condCDF_le_one (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : condCDF ρ a x ≤ 1 := by - obtain ⟨r, hrx⟩ := exists_rat_gt x - rw [← StieltjesFunction.iInf_rat_gt_eq] - simp_rw [condCDF_eq_condCDFRat] - refine' ciInf_le_of_le (bddBelow_range_condCDFRat_gt ρ a x) _ (condCDFRat_le_one _ _ _) - exact ⟨r, hrx⟩ +theorem condCDF_le_one (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : condCDF ρ a x ≤ 1 := + todo2_le_one _ _ _ #align probability_theory.cond_cdf_le_one ProbabilityTheory.condCDF_le_one /-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ theorem tendsto_condCDF_atBot (ρ : Measure (α × ℝ)) (a : α) : - Tendsto (condCDF ρ a) atBot (𝓝 0) := by - have h_exists : ∀ x : ℝ, ∃ q : ℚ, x < q ∧ ↑q < x + 1 := fun x => exists_rat_btwn (lt_add_one x) - let qs : ℝ → ℚ := fun x => (h_exists x).choose - have hqs_tendsto : Tendsto qs atBot atBot := by - rw [tendsto_atBot_atBot] - refine' fun q => ⟨q - 1, fun y hy => _⟩ - have h_le : ↑(qs y) ≤ (q : ℝ) - 1 + 1 := - (h_exists y).choose_spec.2.le.trans (add_le_add hy le_rfl) - rw [sub_add_cancel] at h_le - exact mod_cast h_le - refine' - tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds - ((tendsto_condCDFRat_atBot ρ a).comp hqs_tendsto) (condCDF_nonneg ρ a) fun x => _ - rw [Function.comp_apply, ← condCDF_eq_condCDFRat] - exact (condCDF ρ a).mono (h_exists x).choose_spec.1.le + Tendsto (condCDF ρ a) atBot (𝓝 0) := + tendsto_todo2_atBot _ _ #align probability_theory.tendsto_cond_cdf_at_bot ProbabilityTheory.tendsto_condCDF_atBot /-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ theorem tendsto_condCDF_atTop (ρ : Measure (α × ℝ)) (a : α) : - Tendsto (condCDF ρ a) atTop (𝓝 1) := by - have h_exists : ∀ x : ℝ, ∃ q : ℚ, x - 1 < q ∧ ↑q < x := fun x => exists_rat_btwn (sub_one_lt x) - let qs : ℝ → ℚ := fun x => (h_exists x).choose - have hqs_tendsto : Tendsto qs atTop atTop := by - rw [tendsto_atTop_atTop] - refine' fun q => ⟨q + 1, fun y hy => _⟩ - have h_le : y - 1 ≤ qs y := (h_exists y).choose_spec.1.le - rw [sub_le_iff_le_add] at h_le - exact_mod_cast le_of_add_le_add_right (hy.trans h_le) - refine' - tendsto_of_tendsto_of_tendsto_of_le_of_le ((tendsto_condCDFRat_atTop ρ a).comp hqs_tendsto) - tendsto_const_nhds _ (condCDF_le_one ρ a) - intro x - rw [Function.comp_apply, ← condCDF_eq_condCDFRat] - exact (condCDF ρ a).mono (le_of_lt (h_exists x).choose_spec.2) + Tendsto (condCDF ρ a) atTop (𝓝 1) := + tendsto_todo2_atTop _ _ #align probability_theory.tendsto_cond_cdf_at_top ProbabilityTheory.tendsto_condCDF_atTop theorem condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) : @@ -844,14 +759,8 @@ theorem ofReal_condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r #align probability_theory.of_real_cond_cdf_ae_eq ProbabilityTheory.ofReal_condCDF_ae_eq /-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ -theorem measurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : Measurable fun a => condCDF ρ a x := by - have : (fun a => condCDF ρ a x) = fun a => ⨅ r : { r' : ℚ // x < r' }, condCDFRat ρ a ↑r := by - ext1 a - rw [← StieltjesFunction.iInf_rat_gt_eq] - congr with q - rw [condCDF_eq_condCDFRat] - rw [this] - exact measurable_iInf (fun q => measurable_condCDFRat ρ q) +theorem measurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : Measurable fun a => condCDF ρ a x := + measurable_todo2 _ _ #align probability_theory.measurable_cond_cdf ProbabilityTheory.measurable_condCDF /-- Auxiliary lemma for `set_lintegral_cond_cdf`. -/ @@ -918,7 +827,7 @@ theorem lintegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : /-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ theorem stronglyMeasurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : StronglyMeasurable fun a => condCDF ρ a x := - (measurable_condCDF ρ x).stronglyMeasurable + stronglyMeasurable_todo2 _ _ #align probability_theory.strongly_measurable_cond_cdf ProbabilityTheory.stronglyMeasurable_condCDF theorem integrable_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) : @@ -956,44 +865,22 @@ theorem integral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : section Measure theorem measure_condCDF_Iic (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : - (condCDF ρ a).measure (Iic x) = ENNReal.ofReal (condCDF ρ a x) := by - rw [← sub_zero (condCDF ρ a x)] - exact (condCDF ρ a).measure_Iic (tendsto_condCDF_atBot ρ a) _ + (condCDF ρ a).measure (Iic x) = ENNReal.ofReal (condCDF ρ a x) := + measure_todo2_Iic _ _ _ #align probability_theory.measure_cond_cdf_Iic ProbabilityTheory.measure_condCDF_Iic -theorem measure_condCDF_univ (ρ : Measure (α × ℝ)) (a : α) : (condCDF ρ a).measure univ = 1 := by - rw [← ENNReal.ofReal_one, ← sub_zero (1 : ℝ)] - exact StieltjesFunction.measure_univ _ (tendsto_condCDF_atBot ρ a) (tendsto_condCDF_atTop ρ a) +theorem measure_condCDF_univ (ρ : Measure (α × ℝ)) (a : α) : (condCDF ρ a).measure univ = 1 := + measure_todo2_univ _ _ #align probability_theory.measure_cond_cdf_univ ProbabilityTheory.measure_condCDF_univ instance instIsProbabilityMeasure (ρ : Measure (α × ℝ)) (a : α) : - IsProbabilityMeasure (condCDF ρ a).measure := - ⟨measure_condCDF_univ ρ a⟩ + IsProbabilityMeasure (condCDF ρ a).measure := by + rw [condCDF]; infer_instance /-- The function `a ↦ (condCDF ρ a).measure` is measurable. -/ theorem measurable_measure_condCDF (ρ : Measure (α × ℝ)) : - Measurable fun a => (condCDF ρ a).measure := by - rw [Measure.measurable_measure] - refine' fun s hs => ?_ - -- Porting note: supplied `C` - refine' MeasurableSpace.induction_on_inter - (C := fun s => Measurable fun b ↦ StieltjesFunction.measure (condCDF ρ b) s) - (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic _ _ _ _ hs - · simp only [measure_empty, measurable_const] - · rintro S ⟨u, rfl⟩ - simp_rw [measure_condCDF_Iic ρ _ u] - exact (measurable_condCDF ρ u).ennreal_ofReal - · intro t ht ht_cd_meas - have : - (fun a => (condCDF ρ a).measure tᶜ) = - (fun a => (condCDF ρ a).measure univ) - fun a => (condCDF ρ a).measure t := by - ext1 a - rw [measure_compl ht (measure_ne_top (condCDF ρ a).measure _), Pi.sub_apply] - simp_rw [this, measure_condCDF_univ ρ] - exact Measurable.sub measurable_const ht_cd_meas - · intro f hf_disj hf_meas hf_cd_meas - simp_rw [measure_iUnion hf_disj hf_meas] - exact Measurable.ennreal_tsum hf_cd_meas + Measurable fun a => (condCDF ρ a).measure := + measurable_measure_todo2 _ #align probability_theory.measurable_measure_cond_cdf ProbabilityTheory.measurable_measure_condCDF end Measure diff --git a/Mathlib/Probability/Kernel/StieltjesReal.lean b/Mathlib/Probability/Kernel/StieltjesReal.lean new file mode 100644 index 0000000000000..a8e9742432c03 --- /dev/null +++ b/Mathlib/Probability/Kernel/StieltjesReal.lean @@ -0,0 +1,209 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.MeasureTheory.Measure.Stieltjes +import Mathlib.MeasureTheory.Decomposition.RadonNikodym +import Mathlib.MeasureTheory.Constructions.Prod.Basic + +/-! + + +-/ + + +open MeasureTheory Set Filter TopologicalSpace + +open scoped NNReal ENNReal MeasureTheory Topology + +namespace ProbabilityTheory + +variable {α β ι : Type*} [MeasurableSpace α] + +section IsCDFLike + +structure IsCDFLike (f : α → ℚ → ℝ) : Prop where + mono : ∀ a, Monotone (f a) + nonneg : ∀ a q, 0 ≤ f a q + le_one : ∀ a q, f a q ≤ 1 + tendsto_atTop_one : ∀ a, Tendsto (f a) atTop (𝓝 1) + tendsto_atBot_zero : ∀ a, Tendsto (f a) atBot (𝓝 0) + iInf_rat_gt_eq : ∀ a, ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t + measurable : ∀ q, Measurable (fun a ↦ f a q) + +end IsCDFLike + +variable {f : α → ℚ → ℝ} (hf : IsCDFLike f) + +/-- Conditional cdf of the measure given the value on `α`, as a plain function. This is an auxiliary +definition used to define `cond_cdf`. -/ +noncomputable irreducible_def todo1 (f : α → ℚ → ℝ) : α → ℝ → ℝ := + fun a t ↦ ⨅ r : { r' : ℚ // t < r' }, f a r + +lemma todo1_eq (a : α) (r : ℚ) : + todo1 f a r = f a r := by + rw [← hf.iInf_rat_gt_eq a r, todo1] + refine' Equiv.iInf_congr _ _ + · exact + { toFun := fun t ↦ ⟨t.1, mod_cast t.2⟩ + invFun := fun t ↦ ⟨t.1, mod_cast t.2⟩ + left_inv := fun t ↦ by simp only [Subtype.coe_eta] + right_inv := fun t ↦ by simp only [Subtype.coe_eta] } + · intro t + simp only [Equiv.coe_fn_mk, Subtype.coe_mk] + +theorem todo1_nonneg (a : α) (r : ℝ) : 0 ≤ todo1 f a r := by + have : Nonempty { r' : ℚ // r < ↑r' } := by + obtain ⟨r, hrx⟩ := exists_rat_gt r + exact ⟨⟨r, hrx⟩⟩ + rw [todo1_def] + exact le_ciInf fun r' ↦ hf.nonneg a _ + +theorem bddBelow_range_gt (a : α) (x : ℝ) : + BddBelow (range fun r : { r' : ℚ // x < ↑r' } ↦ f a r) := by + refine' ⟨0, fun z ↦ _⟩; rintro ⟨u, rfl⟩; exact hf.nonneg a _ + +theorem monotone_todo1 (a : α) : Monotone (todo1 f a) := by + intro x y hxy + have : Nonempty { r' : ℚ // y < ↑r' } := by + obtain ⟨r, hrx⟩ := exists_rat_gt y + exact ⟨⟨r, hrx⟩⟩ + simp_rw [todo1_def] + refine' le_ciInf fun r ↦ (ciInf_le _ _).trans_eq _ + · exact bddBelow_range_gt hf a x + · exact ⟨r.1, hxy.trans_lt r.prop⟩ + · rfl + +theorem continuousWithinAt_todo1_Ici (a : α) (x : ℝ) : + ContinuousWithinAt (todo1 f a) (Ici x) x := by + rw [← continuousWithinAt_Ioi_iff_Ici] + convert Monotone.tendsto_nhdsWithin_Ioi (monotone_todo1 hf a) x + rw [sInf_image'] + have h' : ⨅ r : Ioi x, todo1 f a r = ⨅ r : { r' : ℚ // x < r' }, todo1 f a r := by + refine' Real.iInf_Ioi_eq_iInf_rat_gt x _ (monotone_todo1 hf a) + refine' ⟨0, fun z ↦ _⟩ + rintro ⟨u, -, rfl⟩ + exact todo1_nonneg hf a u + have h'' : + ⨅ r : { r' : ℚ // x < r' }, todo1 f a r = + ⨅ r : { r' : ℚ // x < r' }, f a r := by + congr with r + exact todo1_eq hf a r + rw [h', h'', ContinuousWithinAt] + congr! + rw [todo1_def] + +/-! ### Conditional cdf -/ + + +/-- Conditional cdf of the measure given the value on `α`, as a Stieltjes function. -/ +noncomputable def todo2 (a : α) : StieltjesFunction where + toFun := todo1 f a + mono' := monotone_todo1 hf a + right_continuous' x := continuousWithinAt_todo1_Ici hf a x + +theorem todo2_eq (a : α) (r : ℚ) : todo2 hf a r = f a r := todo1_eq hf a r + +/-- The conditional cdf is non-negative for all `a : α`. -/ +theorem todo2_nonneg (a : α) (r : ℝ) : 0 ≤ todo2 hf a r := todo1_nonneg hf a r + +/-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ +theorem todo2_le_one (a : α) (x : ℝ) : todo2 hf a x ≤ 1 := by + obtain ⟨r, hrx⟩ := exists_rat_gt x + rw [← StieltjesFunction.iInf_rat_gt_eq] + simp_rw [todo2_eq] + refine ciInf_le_of_le (bddBelow_range_gt hf a x) ?_ (hf.le_one _ _) + exact ⟨r, hrx⟩ + +/-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ +theorem tendsto_todo2_atBot (a : α) : + Tendsto (todo2 hf a) atBot (𝓝 0) := by + have h_exists : ∀ x : ℝ, ∃ q : ℚ, x < q ∧ ↑q < x + 1 := fun x ↦ exists_rat_btwn (lt_add_one x) + let qs : ℝ → ℚ := fun x ↦ (h_exists x).choose + have hqs_tendsto : Tendsto qs atBot atBot := by + rw [tendsto_atBot_atBot] + refine fun q ↦ ⟨q - 1, fun y hy ↦ ?_⟩ + have h_le : ↑(qs y) ≤ (q : ℝ) - 1 + 1 := + (h_exists y).choose_spec.2.le.trans (add_le_add hy le_rfl) + rw [sub_add_cancel] at h_le + exact mod_cast h_le + refine tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds + ((hf.tendsto_atBot_zero a).comp hqs_tendsto) (todo2_nonneg hf a) fun x ↦ ?_ + rw [Function.comp_apply, ← todo2_eq hf] + exact (todo2 hf a).mono (h_exists x).choose_spec.1.le + +/-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ +theorem tendsto_todo2_atTop (a : α) : + Tendsto (todo2 hf a) atTop (𝓝 1) := by + have h_exists : ∀ x : ℝ, ∃ q : ℚ, x - 1 < q ∧ ↑q < x := fun x ↦ exists_rat_btwn (sub_one_lt x) + let qs : ℝ → ℚ := fun x ↦ (h_exists x).choose + have hqs_tendsto : Tendsto qs atTop atTop := by + rw [tendsto_atTop_atTop] + refine fun q ↦ ⟨q + 1, fun y hy ↦ ?_⟩ + have h_le : y - 1 ≤ qs y := (h_exists y).choose_spec.1.le + rw [sub_le_iff_le_add] at h_le + exact_mod_cast le_of_add_le_add_right (hy.trans h_le) + refine tendsto_of_tendsto_of_tendsto_of_le_of_le ((hf.tendsto_atTop_one a).comp hqs_tendsto) + tendsto_const_nhds ?_ (todo2_le_one hf a) + intro x + rw [Function.comp_apply, ← todo2_eq hf] + exact (todo2 hf a).mono (le_of_lt (h_exists x).choose_spec.2) + +/-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ +theorem measurable_todo2 (x : ℝ) : Measurable fun a ↦ todo2 hf a x := by + have : (fun a ↦ todo2 hf a x) = fun a ↦ ⨅ r : { r' : ℚ // x < r' }, f a ↑r := by + ext1 a + rw [← StieltjesFunction.iInf_rat_gt_eq] + congr with q + rw [todo2_eq] + rw [this] + exact measurable_iInf (fun q ↦ hf.measurable q) + +/-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ +theorem stronglyMeasurable_todo2 (x : ℝ) : + StronglyMeasurable fun a ↦ todo2 hf a x := + (measurable_todo2 hf x).stronglyMeasurable + +section Measure + +theorem measure_todo2_Iic (a : α) (x : ℝ) : + (todo2 hf a).measure (Iic x) = ENNReal.ofReal (todo2 hf a x) := by + rw [← sub_zero (todo2 hf a x)] + exact (todo2 hf a).measure_Iic (tendsto_todo2_atBot hf a) _ + +theorem measure_todo2_univ (a : α) : (todo2 hf a).measure univ = 1 := by + rw [← ENNReal.ofReal_one, ← sub_zero (1 : ℝ)] + exact StieltjesFunction.measure_univ _ (tendsto_todo2_atBot hf a) (tendsto_todo2_atTop hf a) + +instance instIsProbabilityMeasure_todo2 (a : α) : + IsProbabilityMeasure (todo2 hf a).measure := + ⟨measure_todo2_univ hf a⟩ + +/-- The function `a ↦ (condCDF ρ a).measure` is measurable. -/ +theorem measurable_measure_todo2 : + Measurable fun a ↦ (todo2 hf a).measure := by + rw [Measure.measurable_measure] + refine' fun s hs ↦ ?_ + -- Porting note: supplied `C` + refine' MeasurableSpace.induction_on_inter + (C := fun s ↦ Measurable fun b ↦ StieltjesFunction.measure (todo2 hf b) s) + (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic _ _ _ _ hs + · simp only [measure_empty, measurable_const] + · rintro S ⟨u, rfl⟩ + simp_rw [measure_todo2_Iic hf _ u] + exact (measurable_todo2 hf u).ennreal_ofReal + · intro t ht ht_cd_meas + have : (fun a ↦ (todo2 hf a).measure tᶜ) = + (fun a ↦ (todo2 hf a).measure univ) - fun a ↦ (todo2 hf a).measure t := by + ext1 a + rw [measure_compl ht (measure_ne_top (todo2 hf a).measure _), Pi.sub_apply] + simp_rw [this, measure_todo2_univ hf] + exact Measurable.sub measurable_const ht_cd_meas + · intro f hf_disj hf_meas hf_cd_meas + simp_rw [measure_iUnion hf_disj hf_meas] + exact Measurable.ennreal_tsum hf_cd_meas + +end Measure + +end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/draft.lean b/Mathlib/Probability/Kernel/draft.lean new file mode 100644 index 0000000000000..cf82cfb80d778 --- /dev/null +++ b/Mathlib/Probability/Kernel/draft.lean @@ -0,0 +1,718 @@ +import Mathlib.Probability.Kernel.Disintegration +import Mathlib.Probability.Martingale.Convergence +import Mathlib.Probability.Kernel.BuildKernel + +open MeasureTheory Set Filter + +open scoped ENNReal MeasureTheory Topology ProbabilityTheory + +namespace ProbabilityTheory + +variable {α Ω : Type*} {mα : MeasurableSpace α} + [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] + +lemma measurableSet_tendsto_nhds {β γ ι : Type*} [MeasurableSpace β] + [TopologicalSpace γ] [PolishSpace γ] [MeasurableSpace γ] + [hγ : OpensMeasurableSpace γ] [Countable ι] {l : Filter ι} + [l.IsCountablyGenerated] {f : ι → β → γ} (hf : ∀ i, Measurable (f i)) (c : γ) : + MeasurableSet { x | Tendsto (fun n ↦ f n x) l (𝓝 c) } := sorry + +lemma measurableSet_tendsto_fun {β γ ι : Type*} [MeasurableSpace β] + [TopologicalSpace γ] [PolishSpace γ] [MeasurableSpace γ] + [hγ : OpensMeasurableSpace γ] [Countable ι] {l : Filter ι} + [l.IsCountablyGenerated] {f : ι → β → γ} (hf : ∀ i, Measurable (f i)) {g : β → γ} + (hg : Measurable g) : + MeasurableSet { x | Tendsto (fun n ↦ f n x) l (𝓝 (g x)) } := by + letI := upgradePolishSpace γ + have : { x | Tendsto (fun n ↦ f n x) l (𝓝 (g x)) } + = { x | Tendsto (fun n ↦ dist (f n x) (g x)) l (𝓝 0) } := by + ext x + simp only [mem_setOf_eq] + rw [tendsto_iff_dist_tendsto_zero] + rw [this] + exact measurableSet_tendsto_nhds (fun n ↦ (hf n).dist hg) 0 + +section Real + +section dissection_system + +def I (n : ℕ) (k : ℤ) : Set ℝ := Set.Ico (k * (2⁻¹ : ℝ) ^ n) ((k + 1) * ((2 : ℝ) ^ n)⁻¹) + +lemma measurableSet_I (n : ℕ) (k : ℤ) : MeasurableSet (I n k) := measurableSet_Ico + +lemma pairwise_disjoint_I (n : ℕ) : Pairwise (Disjoint on fun k ↦ I n k) := by + sorry + +lemma I_succ_union (n : ℕ) (k : ℤ) : I (n+1) (2 * k) ∪ I (n+1) (2 * k + 1) = I n k := by + ext x + cases lt_or_le x ((2 * k + 1) * ((2 : ℝ) ^ (n+1))⁻¹) with + | inl h=> + simp only [I, inv_pow, mem_Ico, Int.cast_mul, Int.int_cast_ofNat, Int.cast_add, + Int.cast_one, mem_union, h, and_true, not_le.mpr h, false_and, or_false] + sorry + | inr h => + simp only [I, inv_pow, mem_Ico, Int.cast_mul, Int.int_cast_ofNat, Int.cast_add, + Int.cast_one, mem_union, not_lt.mpr h, and_false, h, true_and, false_or] + sorry + +-- todo : `Filtration` should be renamed to `filtration` +def ℱ : Filtration ℕ (borel ℝ) where + seq := fun n ↦ MeasurableSpace.generateFrom {s | ∃ k, s = I n k} + mono' := by + refine monotone_nat_of_le_succ ?_ + intro n + refine MeasurableSpace.generateFrom_le fun s ⟨k, hs⟩ ↦ ?_ + rw [hs, ← I_succ_union n k] + refine MeasurableSet.union ?_ ?_ + · exact MeasurableSpace.measurableSet_generateFrom ⟨2 * k, rfl⟩ + · exact MeasurableSpace.measurableSet_generateFrom ⟨2 * k + 1, rfl⟩ + le' := fun n ↦ by + refine MeasurableSpace.generateFrom_le fun s ⟨k, hs⟩ ↦ ?_ + rw [hs] + exact measurableSet_I n k + +lemma measurableSet_ℱ_I (n : ℕ) (k : ℤ) : MeasurableSet[ℱ n] (I n k) := + MeasurableSpace.measurableSet_generateFrom ⟨k, rfl⟩ + +noncomputable def indexI (n : ℕ) (t : ℝ) : ℤ := Int.floor (t * 2^n) + +lemma mem_I_indexI (n : ℕ) (t : ℝ) : t ∈ I n (indexI n t) := by + rw [indexI, I] + simp only [inv_pow, mem_Ico] + constructor + · rw [← div_eq_mul_inv, div_le_iff] + · exact Int.floor_le (t * 2 ^ n) + · positivity + · rw [← div_eq_mul_inv, lt_div_iff] + · exact Int.lt_floor_add_one (t * 2 ^ n) + · positivity + +lemma indexI_of_mem (n : ℕ) (k : ℤ) (t : ℝ) (ht : t ∈ I n k) : indexI n t = k := by + rw [indexI] + simp only [I, inv_pow, mem_Ico, ← div_eq_mul_inv] at ht + rw [div_le_iff, lt_div_iff] at ht + · rw [Int.floor_eq_iff] + exact ht + · positivity + · positivity + +lemma mem_I_iff_indexI (n : ℕ) (k : ℤ) (t : ℝ) : t ∈ I n k ↔ indexI n t = k := + ⟨fun h ↦ indexI_of_mem n k t h, fun h ↦ h ▸ mem_I_indexI n t⟩ + +lemma measurable_indexI (n : ℕ) : Measurable[ℱ n] (indexI n) := by + unfold indexI + refine @measurable_to_countable' ℤ ℝ _ _ (ℱ n) _ (fun k ↦ ?_) + have : (fun t ↦ ⌊t * (2 : ℝ) ^ n⌋) ⁻¹' {k} = I n k := by + ext t + simp only [mem_preimage, mem_singleton_iff, I, inv_pow, mem_Ico] + rw [Int.floor_eq_iff] + refine ⟨fun ⟨h1, h2⟩ ↦ ⟨?_, ?_⟩, fun ⟨h1, h2⟩ ↦ ⟨?_, ?_⟩⟩ + · rw [mul_inv_le_iff, mul_comm] + · exact h1 + · positivity + · rw [← div_eq_mul_inv, lt_div_iff] + · exact h2 + · positivity + · rw [mul_inv_le_iff, mul_comm] at h1 + · exact h1 + · positivity + · rw [← div_eq_mul_inv, lt_div_iff] at h2 + · exact h2 + · positivity + rw [this] + exact measurableSet_ℱ_I n k + +lemma iUnion_I (n : ℕ) : ⋃ k, I n k = univ := by + ext x + simp only [mem_iUnion, mem_univ, iff_true] + exact ⟨indexI n x, mem_I_indexI n x⟩ + +lemma indexI_le_indexI_iff (n : ℕ) (t x : ℝ) : + indexI n t ≤ indexI n x ↔ ⌊t * 2 ^ n⌋ * (2 ^ n)⁻¹ ≤ x := by + simp only [indexI._eq_1] + rw [← div_eq_mul_inv, div_le_iff, Int.le_floor] + positivity + +lemma iUnion_ge_I (n : ℕ) (t : ℝ) : + ⋃ (k) (_ : indexI n t ≤ k), I n k = Ici (⌊t * 2 ^ n⌋ * (2 ^ n)⁻¹ : ℝ) := by + ext x + simp [mem_I_iff_indexI, indexI_le_indexI_iff] + +lemma iInter_biUnion_I (x : ℝ) : ⋂ n, ⋃ (k) (_ : indexI n x ≤ k), I n k = Ici x := by + ext t + simp [iUnion_ge_I] + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · sorry + · intro n + refine le_trans ?_ h + rw [← div_eq_mul_inv, div_le_iff] + · exact Int.floor_le (x * 2 ^ n) + · positivity + +lemma iSup_ℱ : ⨆ n, ℱ n = borel ℝ := by + refine le_antisymm ?_ ?_ + · rw [iSup_le_iff] + exact ℱ.le + · conv_lhs => rw [borel_eq_generateFrom_Ici ℝ] + refine MeasurableSpace.generateFrom_le (fun s ⟨x, hx⟩ ↦ ?_) + rw [← hx, ← iInter_biUnion_I x] + refine MeasurableSet.iInter (fun n ↦ ?_) + refine MeasurableSet.biUnion ?_ (fun k hk ↦ ?_) + · sorry + · sorry + +end dissection_system + +variable {β : Type*} [MeasurableSpace β] + +-- issue with the following: joint measurability +--noncomputable +--def M' (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : ℝ≥0∞ := +-- (((κ a).restrict (univ ×ˢ s)).fst.trim (ℱ.le n)).rnDeriv (((kernel.fst κ a)).trim (ℱ.le n)) t + +noncomputable +def M (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : ℝ := + (κ a (I n (indexI n t) ×ˢ s) / kernel.fst κ a (I n (indexI n t))).toReal + +lemma m_def (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) : + M κ a s n = fun t ↦ (κ a (I n (indexI n t) ×ˢ s) / kernel.fst κ a (I n (indexI n t))).toReal := + rfl + +lemma measurable_m_aux (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : + Measurable (fun (p : α × ℝ) ↦ + κ p.1 (I n (indexI n p.2) ×ˢ s) / kernel.fst κ p.1 (I n (indexI n p.2))) := by + change Measurable ((fun (p : α × ℤ) ↦ + κ p.1 (I n p.2 ×ˢ s) / kernel.fst κ p.1 (I n p.2)) ∘ (fun (p : α × ℝ) ↦ (p.1, indexI n p.2))) + have h1 : Measurable (fun (p : α × ℤ) ↦ κ p.1 (I n p.2 ×ˢ s) / kernel.fst κ p.1 (I n p.2)) := by + refine Measurable.div ?_ ?_ + · have h_swap : Measurable fun (p : ℤ × α) ↦ κ p.2 (I n p.1 ×ˢ s) := by + refine measurable_uncurry_of_continuous_of_measurable + (u := fun k a ↦ κ a (I n k ×ˢ s)) ?_ ?_ + · exact fun _ ↦ continuous_bot + · exact fun _ ↦ kernel.measurable_coe _ ((measurableSet_I _ _).prod hs) + change Measurable ((fun (p : ℤ × α) ↦ κ p.2 (I n p.1 ×ˢ s)) ∘ Prod.swap) + exact h_swap.comp measurable_swap + · have h_swap : Measurable fun (p : ℤ × α) ↦ kernel.fst κ p.2 (I n p.1) := by + refine measurable_uncurry_of_continuous_of_measurable + (u := fun k a ↦ kernel.fst κ a (I n k)) ?_ ?_ + · exact fun _ ↦ continuous_bot + · exact fun _ ↦ kernel.measurable_coe _ (measurableSet_I _ _) + change Measurable ((fun (p : ℤ × α) ↦ kernel.fst κ p.2 (I n p.1)) ∘ Prod.swap) + exact h_swap.comp measurable_swap + refine h1.comp (measurable_fst.prod_mk ?_) + exact (Measurable.mono (measurable_indexI n) (ℱ.le n) le_rfl).comp measurable_snd + +lemma measurable_m (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : + Measurable (fun (p : α × ℝ) ↦ M κ p.1 s n p.2) := + (measurable_m_aux κ hs n).ennreal_toReal + +lemma measurable_m_left (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) (n : ℕ) (t : ℝ) : + Measurable (fun a ↦ M κ a s n t) := + (measurable_m κ hs n).comp (measurable_id.prod_mk measurable_const) + +lemma measurable_m_right (κ : kernel α (ℝ × β)) {s : Set β} (a : α) (hs : MeasurableSet s) (n : ℕ) : + Measurable (M κ a s n) := + (measurable_m κ hs n).comp (measurable_const.prod_mk measurable_id) + +lemma adapted_m (κ : kernel α (ℝ × β)) (a : α) (s : Set β) : Adapted ℱ (M κ a s) := by + intro n + rw [m_def] + refine Measurable.stronglyMeasurable ?_ + refine @Measurable.ennreal_toReal _ (ℱ n) _ ?_ + refine Measurable.div ?_ ?_ + · change Measurable[ℱ n] ((fun k ↦ κ a (I n k ×ˢ s)) ∘ (indexI n)) + refine Measurable.comp ?_ (measurable_indexI n) + exact measurable_of_countable _ + · change Measurable[ℱ n] ((fun k ↦ (kernel.fst κ) a (I n k)) ∘ (indexI n)) + refine Measurable.comp ?_ (measurable_indexI n) + exact measurable_of_countable _ + +lemma m_nonneg (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : + 0 ≤ M κ a s n t := + ENNReal.toReal_nonneg + +lemma m_le_one (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : + M κ a s n t ≤ 1 := by + rw [M] + refine ENNReal.toReal_le_of_le_ofReal zero_le_one ?_ + rw [ENNReal.ofReal_one] + refine ENNReal.div_le_of_le_mul ?_ + rw [one_mul, kernel.fst_apply' _ _ (measurableSet_I _ _)] + refine measure_mono (fun x ↦ ?_) + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + +lemma snorm_m_le_one (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] + (a : α) (s : Set β) (n : ℕ) : + snorm (M κ a s n) 1 (kernel.fst κ a) ≤ 1 := by + refine (snorm_mono_real (g := fun _ ↦ 1) ?_).trans ?_ + · intro x + simp only [Real.norm_eq_abs, abs_le] + constructor + · have h := m_nonneg κ a s n x + linarith + · exact m_le_one _ _ _ _ _ + · by_cases h0 : kernel.fst κ a = 0 + · simp [h0] + · rw [snorm_const _ one_ne_zero h0] + simp + +lemma integrable_m (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] + (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : + Integrable (M κ a s n) (kernel.fst κ a) := by + rw [← memℒp_one_iff_integrable] + refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ + · exact measurable_m_right κ a hs n + · exact (snorm_m_le_one κ a s n).trans_lt ENNReal.one_lt_top + +lemma set_integral_m_I (κ : kernel α (ℝ × β)) [IsFiniteKernel (kernel.fst κ)] + (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) (k : ℤ) : + ∫ t in I n k, M κ a s n t ∂(kernel.fst κ a) = (κ a (I n k ×ˢ s)).toReal := by + simp_rw [M] + rw [integral_toReal] + rotate_left + · refine Measurable.aemeasurable ?_ + have h := measurable_m_aux κ hs n + sorry + · sorry + congr + have : ∫⁻ t in I n k, κ a (I n (indexI n t) ×ˢ s) + / kernel.fst κ a (I n (indexI n t)) ∂(kernel.fst κ) a + = ∫⁻ t in I n k, κ a (I n k ×ˢ s) / kernel.fst κ a (I n k) ∂(kernel.fst κ) a := by + refine set_lintegral_congr_fun (measurableSet_I _ _) (ae_of_all _ (fun t ht ↦ ?_)) + rw [indexI_of_mem _ _ _ ht] + rw [this] + simp only [lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter] + by_cases h0 : kernel.fst κ a (I n k) = 0 + · simp only [h0, mul_zero] + rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0 + refine (measure_mono_null ?_ h0).symm + intro p + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + rw [div_eq_mul_inv, mul_assoc, ENNReal.inv_mul_cancel h0, mul_one] + exact measure_ne_top _ _ + +lemma integral_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] + (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : + ∫ t, M κ a s n t ∂(kernel.fst κ a) = (κ a (univ ×ˢ s)).toReal := by + rw [← integral_univ, ← iUnion_I n, iUnion_prod_const, measure_iUnion] + rotate_left + · intro i j hij + simp only [Set.disjoint_prod, disjoint_self, bot_eq_empty] + exact Or.inl (pairwise_disjoint_I n hij) + · exact fun k ↦ (measurableSet_I n k).prod hs + rw [integral_iUnion (measurableSet_I n) (pairwise_disjoint_I n) + (integrable_m κ a hs n).integrableOn] + rw [ENNReal.tsum_toReal_eq (fun _ ↦ measure_ne_top _ _)] + congr with k + rw [set_integral_m_I _ _ hs] + +lemma set_integral_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] + (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : + ∫ t in A, M κ a s n t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal := by + suffices MeasurableSet A ∧ ∫ t in A, M κ a s n t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal by + exact this.2 + refine MeasurableSpace.generateFrom_induction + (p := fun A' ↦ MeasurableSet A' + ∧ ∫ t in A', M κ a s n t ∂(kernel.fst κ) a = ENNReal.toReal (κ a (A' ×ˢ s))) + (C := {s | ∃ k, s = I n k}) ?_ ?_ ?_ ?_ hA + · rintro _ ⟨k, rfl⟩ + rw [set_integral_m_I _ _ hs] + exact ⟨measurableSet_I _ _, rfl⟩ + · simp only [MeasurableSet.empty, Measure.restrict_empty, integral_zero_measure, empty_prod, + OuterMeasure.empty', ENNReal.zero_toReal, and_self] + · intro A ⟨hA, hA_eq⟩ + have h := integral_add_compl hA (integrable_m κ a hs n) + refine ⟨hA.compl, ?_⟩ + rw [hA_eq, integral_m κ a hs] at h + have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by + rw [prod_diff_prod, compl_eq_univ_diff] + simp + rw [this, measure_diff (by intro x; simp) (hA.prod hs) (measure_ne_top (κ a) _), + ENNReal.toReal_sub_of_le (measure_mono (by intro x; simp)) (measure_ne_top _ _)] + rw [eq_tsub_iff_add_eq_of_le, add_comm] + · exact h + · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] + exact measure_mono (by intro x; simp) + · intro f hf + simp only at hf + -- todo: introduce disjointed sets, etc. + sorry + +lemma condexp_m (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] (a : α) (s : Set β) + {i j : ℕ} (hij : i ≤ j) : + (kernel.fst κ a)[M κ a s j | ℱ i] =ᵐ[kernel.fst κ a] M κ a s i := by + sorry + +lemma martingale_m (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] (a : α) (s : Set β) : + Martingale (M κ a s) ℱ (kernel.fst κ a) := + ⟨adapted_m κ a s, fun _ _ ↦ condexp_m κ a s⟩ + +lemma m_mono_set (κ : kernel α (ℝ × β)) (a : α) {s s' : Set β} (h : s ⊆ s') (n : ℕ) (t : ℝ) : + M κ a s n t ≤ M κ a s' n t := by + rw [M, M, ENNReal.toReal_le_toReal] + · gcongr + rw [prod_subset_prod_iff] + simp [subset_rfl, h] + · rw [ne_eq, ENNReal.div_eq_top] + push_neg + constructor + · sorry + · sorry + · rw [ne_eq, ENNReal.div_eq_top] + push_neg + constructor + · sorry + · sorry + +lemma m_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ) (t : ℝ) : + M κ a univ n t = if kernel.fst κ a (I n (indexI n t)) = 0 then 0 else 1 := by + rw [M] + by_cases h : kernel.fst κ a (I n (indexI n t)) = 0 + · simp [h] + by_cases h' : κ a (I n (indexI n t) ×ˢ univ) = 0 + · simp [h'] + · rw [ENNReal.div_zero h'] + simp + · simp only [h, ite_false] + rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] + have : I n (indexI n t) ×ˢ univ = {p : ℝ × β | p.1 ∈ I n (indexI n t)} := by + ext x + simp + rw [this, ENNReal.div_self] + · simp + · rwa [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h + · exact measure_ne_top _ _ + +lemma m_empty (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ) (t : ℝ) : + M κ a ∅ n t = 0 := by + rw [M] + simp + +lemma m_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ) : + ∀ᵐ t ∂(kernel.fst κ a), M κ a univ n t = 1 := by + rw [ae_iff] + have : {t | ¬M κ a univ n t = 1} ⊆ {t | kernel.fst κ a (I n (indexI n t)) = 0} := by + intro t ht + simp only [mem_setOf_eq] at ht ⊢ + rw [m_univ] at ht + simpa using ht + refine measure_mono_null this ?_ + have : {t | kernel.fst κ a (I n (indexI n t)) = 0} + ⊆ ⋃ (k) (_ : kernel.fst κ a (I n k) = 0), I n k := by + intro t + simp only [mem_setOf_eq, mem_iUnion, exists_prop] + intro ht + exact ⟨indexI n t, ht, mem_I_indexI _ _⟩ + refine measure_mono_null this ?_ + rw [measure_iUnion_null] + intro i + simp + +lemma tendsto_m_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] + (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) (t : ℝ) : + Tendsto (fun m ↦ M κ a (s m) n t) atTop (𝓝 (M κ a univ n t)) := by + simp_rw [M] + refine (ENNReal.tendsto_toReal ?_).comp ?_ + · rw [ne_eq, ENNReal.div_eq_top] + push_neg + sorry + by_cases h0 : kernel.fst κ a (I n (indexI n t)) = 0 + · simp only [h0] + sorry + refine ENNReal.Tendsto.div_const ?_ ?_ + · have h := tendsto_measure_iUnion (μ := κ a) (s := fun m ↦ I n (indexI n t) ×ˢ s m) ?_ + swap + · intro m m' hmm' + simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] + exact Or.inl <| hs hmm' + convert h + rw [← prod_iUnion, hs_iUnion] + · exact Or.inr h0 + +lemma tendsto_m_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] + (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) : + ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ M κ a (s m) n t) atTop (𝓝 1) := by + filter_upwards [m_univ_ae κ a n] with t ht + rw [← ht] + exact tendsto_m_atTop_univ_of_monotone κ a s hs hs_iUnion n t + +lemma tendsto_m_atTop_empty_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] + (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) + (hs_meas : ∀ n, MeasurableSet (s n)) (n : ℕ) (t : ℝ) : + Tendsto (fun m ↦ M κ a (s m) n t) atTop (𝓝 (M κ a ∅ n t)) := by + simp_rw [M] + refine (ENNReal.tendsto_toReal ?_).comp ?_ + · rw [ne_eq, ENNReal.div_eq_top] + push_neg + sorry + by_cases h0 : kernel.fst κ a (I n (indexI n t)) = 0 + · simp only [h0, prod_empty, OuterMeasure.empty', ENNReal.zero_div] + sorry + refine ENNReal.Tendsto.div_const ?_ ?_ + · have h := tendsto_measure_iInter (μ := κ a) (s := fun m ↦ I n (indexI n t) ×ˢ s m) ?_ ?_ ?_ + · convert h + rw [← prod_iInter, hs_iInter] + · exact fun n ↦ MeasurableSet.prod (measurableSet_I _ _) (hs_meas n) + · intro m m' hmm' + simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] + exact Or.inl <| hs hmm' + · exact ⟨0, measure_ne_top _ _⟩ + · simp only [prod_empty, OuterMeasure.empty', ne_eq, not_true_eq_false, false_or, h0, + not_false_iff] + +lemma tendsto_m_atTop_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] + (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) + (hs_meas : ∀ n, MeasurableSet (s n)) (n : ℕ) (t : ℝ) : + Tendsto (fun m ↦ M κ a (s m) n t) atTop (𝓝 0) := by + rw [← m_empty κ a n t] + exact tendsto_m_atTop_empty_of_antitone κ a s hs hs_iInter hs_meas n t + +lemma tendsto_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsMarkovKernel (kernel.fst κ)] + (s : Set β) : + ∀ᵐ t ∂(kernel.fst κ a), + Tendsto (fun n ↦ M κ a s n t) atTop (𝓝 (ℱ.limitProcess (M κ a s) (kernel.fst κ a) t)) := by + refine Submartingale.ae_tendsto_limitProcess (martingale_m κ a s).submartingale (R := 1) ?_ + intro n + rw [ENNReal.coe_one] + exact snorm_m_le_one κ a s n + +lemma limitProcess_mem_L1 (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] + (a : α) (s : Set β) : + Memℒp (ℱ.limitProcess (M κ a s) (kernel.fst κ a)) 1 (kernel.fst κ a) := by + refine Submartingale.memℒp_limitProcess (martingale_m κ a s).submartingale (R := 1) ?_ + intro n + rw [ENNReal.coe_one] + exact snorm_m_le_one κ a s n + +noncomputable +def MLimsup (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (t : ℝ) : ℝ := + limsup (fun n ↦ M κ a s n t) atTop + +lemma mLimsup_ae_eq_mLim (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] + (a : α) (s : Set β) : + MLimsup κ a s =ᵐ[kernel.fst κ a] ℱ.limitProcess (M κ a s) (kernel.fst κ a) := by + filter_upwards [tendsto_m_limitProcess κ a s] with t ht using ht.limsup_eq + +lemma tendsto_m_mLimsup (κ : kernel α (ℝ × β)) (a : α) [IsMarkovKernel (kernel.fst κ)] (s : Set β) : + ∀ᵐ t ∂(kernel.fst κ a), + Tendsto (fun n ↦ M κ a s n t) atTop (𝓝 (MLimsup κ a s t)) := by + filter_upwards [tendsto_m_limitProcess κ a s, mLimsup_ae_eq_mLim κ a s] with t h1 h2 + rw [h2] + exact h1 + +lemma measurable_mLimsup (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) : + Measurable (fun (p : α × ℝ) ↦ MLimsup κ p.1 s p.2) := + measurable_limsup (fun n ↦ measurable_m κ hs n) + +lemma measurable_mLimsup_left (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) (t : ℝ) : + Measurable (fun a ↦ MLimsup κ a s t) := by + sorry + +lemma measurable_mLimsup_right (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) (a : α) : + Measurable (MLimsup κ a s) := by + sorry + +lemma mLimsup_mono_set (κ : kernel α (ℝ × β)) (a : α) {s s' : Set β} (h : s ⊆ s') (t : ℝ) : + MLimsup κ a s t ≤ MLimsup κ a s' t := by + rw [MLimsup, MLimsup] + refine limsup_le_limsup ?_ ?_ ?_ + · exact Filter.eventually_of_forall (fun n ↦ m_mono_set κ a h n t) + · sorry + · sorry + +lemma mLimsup_nonneg (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (t : ℝ) : + 0 ≤ MLimsup κ a s t := by + refine le_limsup_of_frequently_le ?_ ?_ + · exact Filter.frequently_of_forall (fun n ↦ m_nonneg _ _ _ _ _) + · sorry + +lemma mLimsup_le_one (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (t : ℝ) : + MLimsup κ a s t ≤ 1 := by + refine limsup_le_of_le ?_ ?_ + · sorry + · exact Filter.eventually_of_forall (fun n ↦ m_le_one _ _ _ _ _) + +lemma mLimsup_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) : + ∀ᵐ t ∂(kernel.fst κ a), MLimsup κ a Set.univ t = 1 := by + have h := m_univ_ae κ a + rw [← ae_all_iff] at h + filter_upwards [h] with t ht + rw [MLimsup] + simp_rw [ht] + rw [limsup_const] -- should be simp + +lemma tendsto_mLimsup_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] + (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) : + ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 1) := by + sorry + +lemma tendsto_mLimsup_atTop_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] + (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) + (hs_meas : ∀ n, MeasurableSet (s n)) (n : ℕ) (t : ℝ) : + Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 0) := by + sorry + +section Iic_Q + +noncomputable +def todo1' (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : ℝ := MLimsup κ a (Set.Iic q) t + +lemma measurable_todo1' (κ : kernel α (ℝ × ℝ)) (q : ℚ) : + Measurable (fun p : α × ℝ ↦ todo1' κ p.1 p.2 q) := by + sorry + +lemma monotone_todo1' (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) : Monotone (todo1' κ a t) := by + intro q r hqr + rw [todo1', todo1'] + refine mLimsup_mono_set κ a ?_ t + refine Iic_subset_Iic.mpr ?_ + exact_mod_cast hqr + +lemma todo1'_nonneg (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : 0 ≤ todo1' κ a t q := + mLimsup_nonneg κ a _ t + +lemma todo1'_le_one (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : todo1' κ a t q ≤ 1 := + mLimsup_le_one κ a _ t + +lemma tendsto_atTop_todo1' (κ : kernel α (ℝ × ℝ)) (a : α) : + ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ todo1' κ a t q) atTop (𝓝 1) := by + sorry + +lemma tendsto_atBot_todo1' (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) : + Tendsto (fun q ↦ todo1' κ a t q) atBot (𝓝 0) := by + sorry + +lemma iInf_rat_gt_todo1'_eq (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : + ⨅ r : Ioi q, todo1' κ a t r = todo1' κ a t q := by + sorry + +end Iic_Q + +section Rat + +lemma measurableSet_tendstoAtTopSet (κ : kernel α (ℝ × ℝ)) : + MeasurableSet {p : α × ℝ | Tendsto (fun q ↦ todo1' κ p.1 p.2 q) atTop (𝓝 1)} := + measurableSet_tendsto_nhds (fun q ↦ measurable_todo1' κ q) 1 + +open Classical in +noncomputable +def todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : ℚ → ℝ := + if Tendsto (fun q ↦ todo1' κ p.1 p.2 q) atTop (𝓝 1) + then fun q ↦ todo1' κ p.1 p.2 q + else fun q ↦ if q < 0 then 0 else 1 + +lemma measurable_todo2' (κ : kernel α (ℝ × ℝ)) (q : ℚ) : + Measurable (fun p ↦ todo2' κ p q) := by + classical + simp only [todo2', ite_apply] + exact Measurable.ite (measurableSet_tendstoAtTopSet κ) (measurable_todo1' κ q) measurable_const + +lemma monotone_todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : + Monotone (todo2' κ p) := by + unfold todo2' + split_ifs with h + · exact monotone_todo1' κ p.1 p.2 + · intro x y hxy + dsimp only + split_ifs with h_1 h_2 h_2 + exacts [le_rfl, zero_le_one, absurd (hxy.trans_lt h_2) h_1, le_rfl] + +lemma todo2'_nonneg (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : + 0 ≤ todo2' κ p := by + unfold todo2' + split_ifs with h + · exact todo1'_nonneg κ p.1 p.2 + · intro q + simp only [Pi.one_apply] + split_ifs <;> simp + +lemma todo2'_le_one (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : + todo2' κ p ≤ 1 := by + unfold todo2' + split_ifs with h + · exact todo1'_le_one κ p.1 p.2 + · intro q + simp only [Pi.one_apply] + split_ifs <;> simp + +lemma tendsto_atTop_todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : + Tendsto (todo2' κ p) atTop (𝓝 1) := by + unfold todo2' + split_ifs with h + · exact h + · refine' (tendsto_congr' _).mp tendsto_const_nhds + rw [EventuallyEq, eventually_atTop] + exact ⟨0, fun q hq ↦ (if_neg (not_lt.mpr hq)).symm⟩ + +lemma tendsto_atBot_todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : + Tendsto (todo2' κ p) atBot (𝓝 0) := by + unfold todo2' + split_ifs with h + · exact tendsto_atBot_todo1' κ p.1 p.2 + · refine' (tendsto_congr' _).mp tendsto_const_nhds + rw [EventuallyEq, eventually_atBot] + refine' ⟨-1, fun q hq ↦ (if_pos (hq.trans_lt _)).symm⟩ + linarith + +theorem inf_gt_todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) (t : ℚ) : + ⨅ r : Ioi t, todo2' κ p r = todo2' κ p t := by + rw [todo2'] + split_ifs with hp + · simp_rw [iInf_rat_gt_todo1'_eq] + · simp only + have h_bdd : BddBelow (range fun r : ↥(Ioi t) ↦ ite ((r : ℚ) < 0) (0 : ℝ) 1) := by + refine' ⟨0, fun x hx ↦ _⟩ + obtain ⟨y, rfl⟩ := mem_range.mpr hx + dsimp only + split_ifs + exacts [le_rfl, zero_le_one] + split_ifs with h + · refine' le_antisymm _ (le_ciInf fun x ↦ _) + · obtain ⟨q, htq, hq_neg⟩ : ∃ q, t < q ∧ q < 0 := by + refine' ⟨t / 2, _, _⟩ + · linarith + · linarith + refine' (ciInf_le h_bdd ⟨q, htq⟩).trans _ + rw [if_pos] + rwa [Subtype.coe_mk] + · split_ifs + exacts [le_rfl, zero_le_one] + · refine' le_antisymm _ _ + · refine' (ciInf_le h_bdd ⟨t + 1, lt_add_one t⟩).trans _ + split_ifs + exacts [zero_le_one, le_rfl] + · refine' le_ciInf fun x ↦ _ + rw [if_neg] + rw [not_lt] at h ⊢ + exact h.trans (mem_Ioi.mp x.prop).le + +lemma isCDFLike_todo2' (κ : kernel α (ℝ × ℝ)) : IsCDFLike (todo2' κ) where + mono := monotone_todo2' κ + nonneg := todo2'_nonneg κ + le_one := todo2'_le_one κ + tendsto_atTop_one := tendsto_atTop_todo2' κ + tendsto_atBot_zero := tendsto_atBot_todo2' κ + iInf_rat_gt_eq := inf_gt_todo2' κ + measurable := measurable_todo2' κ + +end Rat + +noncomputable +def kernel.condexpReal (κ : kernel α (ℝ × ℝ)) : kernel (α × ℝ) ℝ := + (isCDFLike_todo2' κ).kernel + +end Real + +variable {Ω' : Type*} [MeasurableSpace Ω'] [StandardBorelSpace Ω'] [Nonempty Ω'] + +def kernel.condexp (κ : kernel α (Ω × Ω')) [IsMarkovKernel (kernel.fst κ)] : + kernel (α × Ω) Ω' := + sorry + +theorem kernel.eq_compProd (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : + κ = kernel.fst κ ⊗ₖ (kernel.condexp κ) := by + sorry + +end ProbabilityTheory From a9b011ca82efd21f118f493e435a11d140b24cfe Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 1 Feb 2024 10:23:13 +0100 Subject: [PATCH 002/129] refactor defaultRatCDF --- Mathlib/Probability/Kernel/CondCdf.lean | 60 +++------- Mathlib/Probability/Kernel/StieltjesReal.lean | 86 ++++++++++++++ Mathlib/Probability/Kernel/draft.lean | 106 ++++++------------ 3 files changed, 135 insertions(+), 117 deletions(-) diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index 475eded804785..a2e9244a7cb6b 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -553,17 +553,20 @@ theorem mem_condCDFSet_ae (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : end HasCondCDF +-- 10:05, 10:10, 10:17, 10:22 + open scoped Classical /-- Conditional cdf of the measure given the value on `α`, restricted to the rationals. It is defined to be `pre_cdf` if `a ∈ condCDFSet`, and a default cdf-like function otherwise. This is an auxiliary definition used to define `cond_cdf`. -/ noncomputable def condCDFRat (ρ : Measure (α × ℝ)) : α → ℚ → ℝ := fun a => - if a ∈ condCDFSet ρ then fun r => (preCDF ρ r a).toReal else fun r => if r < 0 then 0 else 1 + if a ∈ condCDFSet ρ then fun r => (preCDF ρ r a).toReal else defaultRatCDF #align probability_theory.cond_cdf_rat ProbabilityTheory.condCDFRat theorem condCDFRat_of_not_mem (ρ : Measure (α × ℝ)) (a : α) (h : a ∉ condCDFSet ρ) {r : ℚ} : - condCDFRat ρ a r = if r < 0 then 0 else 1 := by simp only [condCDFRat, h, if_false] + condCDFRat ρ a r = if r < 0 then 0 else 1 := by + simp only [condCDFRat, h, if_false, defaultRatCDF] #align probability_theory.cond_cdf_rat_of_not_mem ProbabilityTheory.condCDFRat_of_not_mem theorem condCDFRat_of_mem (ρ : Measure (α × ℝ)) (a : α) (h : a ∈ condCDFSet ρ) (r : ℚ) : @@ -580,10 +583,7 @@ theorem monotone_condCDFRat (ρ : Measure (α × ℝ)) (a : α) : Monotone (cond rw [ENNReal.toReal_le_toReal (h_ne_top _) (h_ne_top _)] exact h'.1 hrr' · simp only [condCDFRat, h, if_false] - intro x y hxy - dsimp only - split_ifs with h_1 h_2 h_2 - exacts [le_rfl, zero_le_one, absurd (hxy.trans_lt h_2) h_1, le_rfl] + exact monotone_defaultRatCDF #align probability_theory.monotone_cond_cdf_rat ProbabilityTheory.monotone_condCDFRat theorem measurable_condCDFRat (ρ : Measure (α × ℝ)) (q : ℚ) : @@ -595,13 +595,10 @@ theorem measurable_condCDFRat (ρ : Measure (α × ℝ)) (q : ℚ) : #align probability_theory.measurable_cond_cdf_rat ProbabilityTheory.measurable_condCDFRat theorem condCDFRat_nonneg (ρ : Measure (α × ℝ)) (a : α) (r : ℚ) : 0 ≤ condCDFRat ρ a r := by - -- Porting note: was - -- unfold condCDFRat; split_ifs; exacts [ENNReal.toReal_nonneg, le_rfl, zero_le_one] - unfold condCDFRat; split_ifs - · exact ENNReal.toReal_nonneg - dsimp only + unfold condCDFRat split_ifs - exacts [le_rfl, zero_le_one] + · exact ENNReal.toReal_nonneg + · exact defaultRatCDF_nonneg _ #align probability_theory.cond_cdf_rat_nonneg ProbabilityTheory.condCDFRat_nonneg theorem condCDFRat_le_one (ρ : Measure (α × ℝ)) (a : α) (r : ℚ) : condCDFRat ρ a r ≤ 1 := by @@ -610,9 +607,7 @@ theorem condCDFRat_le_one (ρ : Measure (α × ℝ)) (a : α) (r : ℚ) : condCD · refine' ENNReal.toReal_le_of_le_ofReal zero_le_one _ rw [ENNReal.ofReal_one] exact (hasCondCDF_of_mem_condCDFSet h).le_one r - -- Porting note: added - dsimp only; split_ifs - exacts [zero_le_one, le_rfl] + · exact defaultRatCDF_le_one _ #align probability_theory.cond_cdf_rat_le_one ProbabilityTheory.condCDFRat_le_one theorem tendsto_condCDFRat_atBot (ρ : Measure (α × ℝ)) (a : α) : @@ -624,10 +619,7 @@ theorem tendsto_condCDFRat_atBot (ρ : Measure (α × ℝ)) (a : α) : · have h' := hasCondCDF_of_mem_condCDFSet h exact fun r => ((h'.le_one r).trans_lt ENNReal.one_lt_top).ne · exact ENNReal.zero_ne_top - · refine' (tendsto_congr' _).mp tendsto_const_nhds - rw [EventuallyEq, eventually_atBot] - refine' ⟨-1, fun q hq => (if_pos (hq.trans_lt _)).symm⟩ - linarith + · exact tendsto_defaultRatCDF_atBot #align probability_theory.tendsto_cond_cdf_rat_at_bot ProbabilityTheory.tendsto_condCDFRat_atBot theorem tendsto_condCDFRat_atTop (ρ : Measure (α × ℝ)) (a : α) : @@ -639,9 +631,7 @@ theorem tendsto_condCDFRat_atTop (ρ : Measure (α × ℝ)) (a : α) : · exact h'.tendsto_atTop_one · exact fun r => ((h'.le_one r).trans_lt ENNReal.one_lt_top).ne · exact ENNReal.one_ne_top - · refine' (tendsto_congr' _).mp tendsto_const_nhds - rw [EventuallyEq, eventually_atTop] - exact ⟨0, fun q hq => (if_neg (not_lt.mpr hq)).symm⟩ + · exact tendsto_defaultRatCDF_atTop #align probability_theory.tendsto_cond_cdf_rat_at_top ProbabilityTheory.tendsto_condCDFRat_atTop theorem condCDFRat_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) : @@ -666,31 +656,7 @@ theorem inf_gt_condCDFRat (ρ : Measure (α × ℝ)) (a : α) (t : ℚ) : rw [← ha'.iInf_rat_gt_eq] · exact fun r => ((ha'.le_one r).trans_lt ENNReal.one_lt_top).ne · simp_rw [condCDFRat_of_not_mem ρ a ha] - have h_bdd : BddBelow (range fun r : ↥(Ioi t) => ite ((r : ℚ) < 0) (0 : ℝ) 1) := by - refine' ⟨0, fun x hx => _⟩ - obtain ⟨y, rfl⟩ := mem_range.mpr hx - dsimp only - split_ifs - exacts [le_rfl, zero_le_one] - split_ifs with h - · refine' le_antisymm _ (le_ciInf fun x => _) - · obtain ⟨q, htq, hq_neg⟩ : ∃ q, t < q ∧ q < 0 := by - refine' ⟨t / 2, _, _⟩ - · linarith - · linarith - refine' (ciInf_le h_bdd ⟨q, htq⟩).trans _ - rw [if_pos] - rwa [Subtype.coe_mk] - · split_ifs - exacts [le_rfl, zero_le_one] - · refine' le_antisymm _ _ - · refine' (ciInf_le h_bdd ⟨t + 1, lt_add_one t⟩).trans _ - split_ifs - exacts [zero_le_one, le_rfl] - · refine' le_ciInf fun x => _ - rw [if_neg] - rw [not_lt] at h ⊢ - exact h.trans (mem_Ioi.mp x.prop).le + exact inf_gt_rat_defaultRatCDF _ #align probability_theory.inf_gt_cond_cdf_rat ProbabilityTheory.inf_gt_condCDFRat lemma isCDFLike_condCDFRat (ρ : Measure (α × ℝ)) : IsCDFLike (condCDFRat ρ) where diff --git a/Mathlib/Probability/Kernel/StieltjesReal.lean b/Mathlib/Probability/Kernel/StieltjesReal.lean index a8e9742432c03..b82c09b155f40 100644 --- a/Mathlib/Probability/Kernel/StieltjesReal.lean +++ b/Mathlib/Probability/Kernel/StieltjesReal.lean @@ -32,8 +32,94 @@ structure IsCDFLike (f : α → ℚ → ℝ) : Prop where iInf_rat_gt_eq : ∀ a, ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t measurable : ∀ q, Measurable (fun a ↦ f a q) +lemma IsCDFLike.ite {s : Set α} (hs : MeasurableSet s) [DecidablePred (fun a ↦ a ∈ s)] + {f g : α → ℚ → ℝ} (hf : IsCDFLike f) (hg : IsCDFLike g) : + IsCDFLike (fun a q ↦ if a ∈ s then f a q else g a q) where + mono a := by split_ifs; exacts [hf.mono a, hg.mono a] + nonneg a := by split_ifs; exacts [hf.nonneg a, hg.nonneg a] + le_one a := by split_ifs; exacts [hf.le_one a, hg.le_one a] + tendsto_atTop_one a := by split_ifs; exacts [hf.tendsto_atTop_one a, hg.tendsto_atTop_one a] + tendsto_atBot_zero a := by split_ifs; exacts [hf.tendsto_atBot_zero a, hg.tendsto_atBot_zero a] + iInf_rat_gt_eq a := by split_ifs; exacts [hf.iInf_rat_gt_eq a, hg.iInf_rat_gt_eq a] + measurable q := Measurable.ite hs (hf.measurable q) (hg.measurable q) + end IsCDFLike +section DefaultRatCDF + +def defaultRatCDF (q : ℚ) := if q < 0 then (0 : ℝ) else 1 + +lemma monotone_defaultRatCDF : Monotone defaultRatCDF := by + unfold defaultRatCDF + intro x y hxy + dsimp only + split_ifs with h_1 h_2 h_2 + exacts [le_rfl, zero_le_one, absurd (hxy.trans_lt h_2) h_1, le_rfl] + +lemma defaultRatCDF_nonneg (q : ℚ) : 0 ≤ defaultRatCDF q := by + unfold defaultRatCDF + split_ifs + exacts [le_rfl, zero_le_one] + +lemma defaultRatCDF_le_one (q : ℚ) : defaultRatCDF q ≤ 1 := by + unfold defaultRatCDF + split_ifs <;> simp + +lemma tendsto_defaultRatCDF_atTop : Tendsto defaultRatCDF atTop (𝓝 1) := by + refine (tendsto_congr' ?_).mp tendsto_const_nhds + rw [EventuallyEq, eventually_atTop] + exact ⟨0, fun q hq => (if_neg (not_lt.mpr hq)).symm⟩ + +lemma tendsto_defaultRatCDF_atBot : Tendsto defaultRatCDF atBot (𝓝 0) := by + refine (tendsto_congr' ?_).mp tendsto_const_nhds + rw [EventuallyEq, eventually_atBot] + refine ⟨-1, fun q hq => (if_pos (hq.trans_lt ?_)).symm⟩ + linarith + +lemma inf_gt_rat_defaultRatCDF (t : ℚ) : + ⨅ r : Ioi t, defaultRatCDF r = defaultRatCDF t := by + simp only [defaultRatCDF] + have h_bdd : BddBelow (range fun r : ↥(Ioi t) ↦ ite ((r : ℚ) < 0) (0 : ℝ) 1) := by + refine' ⟨0, fun x hx ↦ _⟩ + obtain ⟨y, rfl⟩ := mem_range.mpr hx + dsimp only + split_ifs + exacts [le_rfl, zero_le_one] + split_ifs with h + · refine' le_antisymm _ (le_ciInf fun x ↦ _) + · obtain ⟨q, htq, hq_neg⟩ : ∃ q, t < q ∧ q < 0 := by + refine' ⟨t / 2, _, _⟩ + · linarith + · linarith + refine' (ciInf_le h_bdd ⟨q, htq⟩).trans _ + rw [if_pos] + rwa [Subtype.coe_mk] + · split_ifs + exacts [le_rfl, zero_le_one] + · refine' le_antisymm _ _ + · refine' (ciInf_le h_bdd ⟨t + 1, lt_add_one t⟩).trans _ + split_ifs + exacts [zero_le_one, le_rfl] + · refine' le_ciInf fun x ↦ _ + rw [if_neg] + rw [not_lt] at h ⊢ + exact h.trans (mem_Ioi.mp x.prop).le + +lemma measurable_defaultRatCDF (α : Type*) [MeasurableSpace α] (q : ℚ) : + Measurable (fun (_ : α) ↦ defaultRatCDF q) := measurable_const + +lemma isCDFLike_defaultRatCDF (α : Type*) [MeasurableSpace α] : + IsCDFLike (fun (_ : α) (q : ℚ) ↦ defaultRatCDF q) where + mono _ := monotone_defaultRatCDF + nonneg _ := defaultRatCDF_nonneg + le_one _ := defaultRatCDF_le_one + tendsto_atBot_zero _ := tendsto_defaultRatCDF_atBot + tendsto_atTop_one _ := tendsto_defaultRatCDF_atTop + iInf_rat_gt_eq _ := inf_gt_rat_defaultRatCDF + measurable := measurable_defaultRatCDF α + +end DefaultRatCDF + variable {f : α → ℚ → ℝ} (hf : IsCDFLike f) /-- Conditional cdf of the measure given the value on `α`, as a plain function. This is an auxiliary diff --git a/Mathlib/Probability/Kernel/draft.lean b/Mathlib/Probability/Kernel/draft.lean index cf82cfb80d778..7452f6b69691d 100644 --- a/Mathlib/Probability/Kernel/draft.lean +++ b/Mathlib/Probability/Kernel/draft.lean @@ -508,11 +508,13 @@ lemma measurable_mLimsup (κ : kernel α (ℝ × β)) {s : Set β} (hs : Measura lemma measurable_mLimsup_left (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) (t : ℝ) : Measurable (fun a ↦ MLimsup κ a s t) := by - sorry + change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ p.1 s p.2) ∘ (fun a ↦ (a, t))) + exact (measurable_mLimsup κ hs).comp (measurable_id.prod_mk measurable_const) lemma measurable_mLimsup_right (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) (a : α) : Measurable (MLimsup κ a s) := by - sorry + change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ p.1 s p.2) ∘ (fun t ↦ (a, t))) + exact (measurable_mLimsup κ hs).comp (measurable_const.prod_mk measurable_id) lemma mLimsup_mono_set (κ : kernel α (ℝ × β)) (a : α) {s s' : Set β} (h : s ⊆ s') (t : ℝ) : MLimsup κ a s t ≤ MLimsup κ a s' t := by @@ -557,35 +559,35 @@ lemma tendsto_mLimsup_atTop_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKe section Iic_Q noncomputable -def todo1' (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : ℝ := MLimsup κ a (Set.Iic q) t +def mLimsupIic (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : ℝ := MLimsup κ a (Set.Iic q) t -lemma measurable_todo1' (κ : kernel α (ℝ × ℝ)) (q : ℚ) : - Measurable (fun p : α × ℝ ↦ todo1' κ p.1 p.2 q) := by +lemma measurable_mLimsupIic (κ : kernel α (ℝ × ℝ)) (q : ℚ) : + Measurable (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2 q) := by sorry -lemma monotone_todo1' (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) : Monotone (todo1' κ a t) := by +lemma monotone_mLimsupIic (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) : Monotone (mLimsupIic κ a t) := by intro q r hqr - rw [todo1', todo1'] + rw [mLimsupIic, mLimsupIic] refine mLimsup_mono_set κ a ?_ t refine Iic_subset_Iic.mpr ?_ exact_mod_cast hqr -lemma todo1'_nonneg (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : 0 ≤ todo1' κ a t q := +lemma mLimsupIic_nonneg (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : 0 ≤ mLimsupIic κ a t q := mLimsup_nonneg κ a _ t -lemma todo1'_le_one (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : todo1' κ a t q ≤ 1 := +lemma mLimsupIic_le_one (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : mLimsupIic κ a t q ≤ 1 := mLimsup_le_one κ a _ t -lemma tendsto_atTop_todo1' (κ : kernel α (ℝ × ℝ)) (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ todo1' κ a t q) atTop (𝓝 1) := by +lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) (a : α) : + ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ a t q) atTop (𝓝 1) := by sorry -lemma tendsto_atBot_todo1' (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) : - Tendsto (fun q ↦ todo1' κ a t q) atBot (𝓝 0) := by +lemma tendsto_atBot_mLimsupIic (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) : + Tendsto (fun q ↦ mLimsupIic κ a t q) atBot (𝓝 0) := by sorry -lemma iInf_rat_gt_todo1'_eq (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : - ⨅ r : Ioi q, todo1' κ a t r = todo1' κ a t q := by +lemma iInf_rat_gt_mLimsupIic_eq (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : + ⨅ r : Ioi q, mLimsupIic κ a t r = mLimsupIic κ a t q := by sorry end Iic_Q @@ -593,100 +595,64 @@ end Iic_Q section Rat lemma measurableSet_tendstoAtTopSet (κ : kernel α (ℝ × ℝ)) : - MeasurableSet {p : α × ℝ | Tendsto (fun q ↦ todo1' κ p.1 p.2 q) atTop (𝓝 1)} := - measurableSet_tendsto_nhds (fun q ↦ measurable_todo1' κ q) 1 + MeasurableSet {p : α × ℝ | Tendsto (fun q ↦ mLimsupIic κ p.1 p.2 q) atTop (𝓝 1)} := + measurableSet_tendsto_nhds (fun q ↦ measurable_mLimsupIic κ q) 1 open Classical in noncomputable def todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : ℚ → ℝ := - if Tendsto (fun q ↦ todo1' κ p.1 p.2 q) atTop (𝓝 1) - then fun q ↦ todo1' κ p.1 p.2 q - else fun q ↦ if q < 0 then 0 else 1 + if Tendsto (fun q ↦ mLimsupIic κ p.1 p.2 q) atTop (𝓝 1) + then fun q ↦ mLimsupIic κ p.1 p.2 q + else defaultRatCDF lemma measurable_todo2' (κ : kernel α (ℝ × ℝ)) (q : ℚ) : Measurable (fun p ↦ todo2' κ p q) := by classical simp only [todo2', ite_apply] - exact Measurable.ite (measurableSet_tendstoAtTopSet κ) (measurable_todo1' κ q) measurable_const + exact Measurable.ite (measurableSet_tendstoAtTopSet κ) (measurable_mLimsupIic κ q) + measurable_const lemma monotone_todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : Monotone (todo2' κ p) := by unfold todo2' split_ifs with h - · exact monotone_todo1' κ p.1 p.2 - · intro x y hxy - dsimp only - split_ifs with h_1 h_2 h_2 - exacts [le_rfl, zero_le_one, absurd (hxy.trans_lt h_2) h_1, le_rfl] + · exact monotone_mLimsupIic κ p.1 p.2 + · exact monotone_defaultRatCDF lemma todo2'_nonneg (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : 0 ≤ todo2' κ p := by unfold todo2' split_ifs with h - · exact todo1'_nonneg κ p.1 p.2 - · intro q - simp only [Pi.one_apply] - split_ifs <;> simp + · exact mLimsupIic_nonneg κ p.1 p.2 + · exact defaultRatCDF_nonneg lemma todo2'_le_one (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : todo2' κ p ≤ 1 := by unfold todo2' split_ifs with h - · exact todo1'_le_one κ p.1 p.2 - · intro q - simp only [Pi.one_apply] - split_ifs <;> simp + · exact mLimsupIic_le_one κ p.1 p.2 + · exact defaultRatCDF_le_one lemma tendsto_atTop_todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : Tendsto (todo2' κ p) atTop (𝓝 1) := by unfold todo2' split_ifs with h · exact h - · refine' (tendsto_congr' _).mp tendsto_const_nhds - rw [EventuallyEq, eventually_atTop] - exact ⟨0, fun q hq ↦ (if_neg (not_lt.mpr hq)).symm⟩ + · exact tendsto_defaultRatCDF_atTop lemma tendsto_atBot_todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : Tendsto (todo2' κ p) atBot (𝓝 0) := by unfold todo2' split_ifs with h - · exact tendsto_atBot_todo1' κ p.1 p.2 - · refine' (tendsto_congr' _).mp tendsto_const_nhds - rw [EventuallyEq, eventually_atBot] - refine' ⟨-1, fun q hq ↦ (if_pos (hq.trans_lt _)).symm⟩ - linarith + · exact tendsto_atBot_mLimsupIic κ p.1 p.2 + · exact tendsto_defaultRatCDF_atBot theorem inf_gt_todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) (t : ℚ) : ⨅ r : Ioi t, todo2' κ p r = todo2' κ p t := by - rw [todo2'] + unfold todo2' split_ifs with hp - · simp_rw [iInf_rat_gt_todo1'_eq] - · simp only - have h_bdd : BddBelow (range fun r : ↥(Ioi t) ↦ ite ((r : ℚ) < 0) (0 : ℝ) 1) := by - refine' ⟨0, fun x hx ↦ _⟩ - obtain ⟨y, rfl⟩ := mem_range.mpr hx - dsimp only - split_ifs - exacts [le_rfl, zero_le_one] - split_ifs with h - · refine' le_antisymm _ (le_ciInf fun x ↦ _) - · obtain ⟨q, htq, hq_neg⟩ : ∃ q, t < q ∧ q < 0 := by - refine' ⟨t / 2, _, _⟩ - · linarith - · linarith - refine' (ciInf_le h_bdd ⟨q, htq⟩).trans _ - rw [if_pos] - rwa [Subtype.coe_mk] - · split_ifs - exacts [le_rfl, zero_le_one] - · refine' le_antisymm _ _ - · refine' (ciInf_le h_bdd ⟨t + 1, lt_add_one t⟩).trans _ - split_ifs - exacts [zero_le_one, le_rfl] - · refine' le_ciInf fun x ↦ _ - rw [if_neg] - rw [not_lt] at h ⊢ - exact h.trans (mem_Ioi.mp x.prop).le + · simp_rw [iInf_rat_gt_mLimsupIic_eq] + · exact inf_gt_rat_defaultRatCDF t lemma isCDFLike_todo2' (κ : kernel α (ℝ × ℝ)) : IsCDFLike (todo2' κ) where mono := monotone_todo2' κ From 8d1da47674825533fa47e0dc15af8f2c68ff533f Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 1 Feb 2024 10:23:31 +0100 Subject: [PATCH 003/129] remove comment --- Mathlib/Probability/Kernel/CondCdf.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index a2e9244a7cb6b..331fba61ed049 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -553,8 +553,6 @@ theorem mem_condCDFSet_ae (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : end HasCondCDF --- 10:05, 10:10, 10:17, 10:22 - open scoped Classical /-- Conditional cdf of the measure given the value on `α`, restricted to the rationals. From 099b859c1bd7c374cf146d35ef34e6d48ec09921 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 1 Feb 2024 10:57:33 +0100 Subject: [PATCH 004/129] minor --- Mathlib/Probability/Kernel/draft.lean | 8 ++++++-- 1 file changed, 6 insertions(+), 2 deletions(-) diff --git a/Mathlib/Probability/Kernel/draft.lean b/Mathlib/Probability/Kernel/draft.lean index 7452f6b69691d..bb82ff92cdc84 100644 --- a/Mathlib/Probability/Kernel/draft.lean +++ b/Mathlib/Probability/Kernel/draft.lean @@ -545,9 +545,13 @@ lemma mLimsup_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) : simp_rw [ht] rw [limsup_const] -- should be simp -lemma tendsto_mLimsup_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma tendsto_mLimsup_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 1) := by + have h1 := tendsto_m_atTop_ae_of_monotone κ a s hs hs_iUnion + have h2 := fun (n : ℕ) ↦ tendsto_m_mLimsup κ a (s n) + rw [← ae_all_iff] at h1 h2 + filter_upwards [h1, h2] with t h_tendsto_set h_tendsto_nat sorry lemma tendsto_mLimsup_atTop_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] @@ -578,7 +582,7 @@ lemma mLimsupIic_nonneg (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : lemma mLimsupIic_le_one (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : mLimsupIic κ a t q ≤ 1 := mLimsup_le_one κ a _ t -lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) (a : α) : +lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ a t q) atTop (𝓝 1) := by sorry From bf37c05f7cd9aa4d1a6efb7539c0a480ab5ff12a Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sat, 10 Feb 2024 18:05:54 +0100 Subject: [PATCH 005/129] more work --- Mathlib/Probability/Kernel/BuildKernel.lean | 14 +- Mathlib/Probability/Kernel/CondCdf.lean | 248 ++--- Mathlib/Probability/Kernel/StieltjesReal.lean | 188 +++- Mathlib/Probability/Kernel/draft.lean | 875 ++++++++++++++---- 4 files changed, 922 insertions(+), 403 deletions(-) diff --git a/Mathlib/Probability/Kernel/BuildKernel.lean b/Mathlib/Probability/Kernel/BuildKernel.lean index 2d3c36e85bc89..1f2c6fd17c371 100644 --- a/Mathlib/Probability/Kernel/BuildKernel.lean +++ b/Mathlib/Probability/Kernel/BuildKernel.lean @@ -18,11 +18,17 @@ open scoped NNReal ENNReal MeasureTheory Topology namespace ProbabilityTheory -variable {α : Type*} [MeasurableSpace α] +variable {α : Type*} [MeasurableSpace α] {f : α → ℚ → ℝ} {hf : ∀ q, Measurable fun a ↦ f a q} noncomputable -def IsCDFLike.kernel {f : α → ℚ → ℝ} (hf : IsCDFLike f) : kernel α ℝ where - val (a : α) := (todo2 hf a).measure - property := measurable_measure_todo2 hf +def cdfKernel (hf : ∀ q, Measurable fun a ↦ f a q) : kernel α ℝ where + val a := (todo3 f hf a).measure + property := measurable_measure_todo3 hf + +instance instIsMarkovKernel_cdfKernel : IsMarkovKernel (cdfKernel hf) := + ⟨fun _ ↦ instIsProbabilityMeasure_todo3 _ _⟩ + +lemma cdfKernel_Iic (a : α) (x : ℝ) : + cdfKernel hf a (Iic x) = ENNReal.ofReal (todo3 f hf a x) := measure_todo3_Iic hf a x end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index 331fba61ed049..33f3f81299378 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -508,164 +508,24 @@ theorem inf_gt_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : rw [set_lintegral_iInf_gt_preCDF ρ t hs, set_lintegral_preCDF_fst ρ t hs] #align probability_theory.inf_gt_pre_cdf ProbabilityTheory.inf_gt_preCDF -section HasCondCDF - -/-- A product measure on `α × ℝ` is said to have a conditional cdf at `a : α` if `preCDF` is -monotone with limit 0 at -∞ and 1 at +∞, and is right continuous. -This property holds almost everywhere (see `has_cond_cdf_ae`). -/ -structure HasCondCDF (ρ : Measure (α × ℝ)) (a : α) : Prop where - mono : Monotone fun r => preCDF ρ r a - le_one : ∀ r, preCDF ρ r a ≤ 1 - tendsto_atTop_one : Tendsto (fun r => preCDF ρ r a) atTop (𝓝 1) - tendsto_atBot_zero : Tendsto (fun r => preCDF ρ r a) atBot (𝓝 0) - iInf_rat_gt_eq : ∀ t : ℚ, ⨅ r : Ioi t, preCDF ρ r a = preCDF ρ t a -#align probability_theory.has_cond_cdf ProbabilityTheory.HasCondCDF - -theorem hasCondCDF_ae (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : ∀ᵐ a ∂ρ.fst, HasCondCDF ρ a := by - filter_upwards [monotone_preCDF ρ, preCDF_le_one ρ, tendsto_preCDF_atTop_one ρ, - tendsto_preCDF_atBot_zero ρ, inf_gt_preCDF ρ] with a h1 h2 h3 h4 h5 - exact ⟨h1, h2, h3, h4, h5⟩ -#align probability_theory.has_cond_cdf_ae ProbabilityTheory.hasCondCDF_ae - -/-- A measurable set of elements of `α` such that `ρ` has a conditional cdf at all -`a ∈ condCDFSet`. -/ -def condCDFSet (ρ : Measure (α × ℝ)) : Set α := - (toMeasurable ρ.fst {b | ¬HasCondCDF ρ b})ᶜ -#align probability_theory.cond_cdf_set ProbabilityTheory.condCDFSet - -theorem measurableSet_condCDFSet (ρ : Measure (α × ℝ)) : MeasurableSet (condCDFSet ρ) := - (measurableSet_toMeasurable _ _).compl -#align probability_theory.measurable_set_cond_cdf_set ProbabilityTheory.measurableSet_condCDFSet - -theorem hasCondCDF_of_mem_condCDFSet {ρ : Measure (α × ℝ)} {a : α} (h : a ∈ condCDFSet ρ) : - HasCondCDF ρ a := by - rw [condCDFSet, mem_compl_iff] at h - have h_ss := subset_toMeasurable ρ.fst {b | ¬HasCondCDF ρ b} - by_contra ha - exact h (h_ss ha) -#align probability_theory.has_cond_cdf_of_mem_cond_cdf_set ProbabilityTheory.hasCondCDF_of_mem_condCDFSet - -theorem mem_condCDFSet_ae (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - ∀ᵐ a ∂ρ.fst, a ∈ condCDFSet ρ := by - simp_rw [ae_iff, condCDFSet, not_mem_compl_iff, setOf_mem_eq, measure_toMeasurable] - exact hasCondCDF_ae ρ -#align probability_theory.mem_cond_cdf_set_ae ProbabilityTheory.mem_condCDFSet_ae - -end HasCondCDF - -open scoped Classical - -/-- Conditional cdf of the measure given the value on `α`, restricted to the rationals. -It is defined to be `pre_cdf` if `a ∈ condCDFSet`, and a default cdf-like function -otherwise. This is an auxiliary definition used to define `cond_cdf`. -/ -noncomputable def condCDFRat (ρ : Measure (α × ℝ)) : α → ℚ → ℝ := fun a => - if a ∈ condCDFSet ρ then fun r => (preCDF ρ r a).toReal else defaultRatCDF -#align probability_theory.cond_cdf_rat ProbabilityTheory.condCDFRat - -theorem condCDFRat_of_not_mem (ρ : Measure (α × ℝ)) (a : α) (h : a ∉ condCDFSet ρ) {r : ℚ} : - condCDFRat ρ a r = if r < 0 then 0 else 1 := by - simp only [condCDFRat, h, if_false, defaultRatCDF] -#align probability_theory.cond_cdf_rat_of_not_mem ProbabilityTheory.condCDFRat_of_not_mem - -theorem condCDFRat_of_mem (ρ : Measure (α × ℝ)) (a : α) (h : a ∈ condCDFSet ρ) (r : ℚ) : - condCDFRat ρ a r = (preCDF ρ r a).toReal := by simp only [condCDFRat, h, if_true] -#align probability_theory.cond_cdf_rat_of_mem ProbabilityTheory.condCDFRat_of_mem - -theorem monotone_condCDFRat (ρ : Measure (α × ℝ)) (a : α) : Monotone (condCDFRat ρ a) := by - by_cases h : a ∈ condCDFSet ρ - · simp only [condCDFRat, h, if_true, forall_const, and_self_iff] - intro r r' hrr' - have h' := hasCondCDF_of_mem_condCDFSet h - have h_ne_top : ∀ r, preCDF ρ r a ≠ ∞ := fun r => - ((h'.le_one r).trans_lt ENNReal.one_lt_top).ne - rw [ENNReal.toReal_le_toReal (h_ne_top _) (h_ne_top _)] - exact h'.1 hrr' - · simp only [condCDFRat, h, if_false] - exact monotone_defaultRatCDF -#align probability_theory.monotone_cond_cdf_rat ProbabilityTheory.monotone_condCDFRat - -theorem measurable_condCDFRat (ρ : Measure (α × ℝ)) (q : ℚ) : - Measurable fun a => condCDFRat ρ a q := by - simp_rw [condCDFRat, ite_apply] - exact - Measurable.ite (measurableSet_condCDFSet ρ) measurable_preCDF.ennreal_toReal - measurable_const -#align probability_theory.measurable_cond_cdf_rat ProbabilityTheory.measurable_condCDFRat - -theorem condCDFRat_nonneg (ρ : Measure (α × ℝ)) (a : α) (r : ℚ) : 0 ≤ condCDFRat ρ a r := by - unfold condCDFRat - split_ifs - · exact ENNReal.toReal_nonneg - · exact defaultRatCDF_nonneg _ -#align probability_theory.cond_cdf_rat_nonneg ProbabilityTheory.condCDFRat_nonneg - -theorem condCDFRat_le_one (ρ : Measure (α × ℝ)) (a : α) (r : ℚ) : condCDFRat ρ a r ≤ 1 := by - unfold condCDFRat - split_ifs with h - · refine' ENNReal.toReal_le_of_le_ofReal zero_le_one _ - rw [ENNReal.ofReal_one] - exact (hasCondCDF_of_mem_condCDFSet h).le_one r - · exact defaultRatCDF_le_one _ -#align probability_theory.cond_cdf_rat_le_one ProbabilityTheory.condCDFRat_le_one - -theorem tendsto_condCDFRat_atBot (ρ : Measure (α × ℝ)) (a : α) : - Tendsto (condCDFRat ρ a) atBot (𝓝 0) := by - unfold condCDFRat - split_ifs with h - · rw [← ENNReal.zero_toReal, ENNReal.tendsto_toReal_iff] - · exact (hasCondCDF_of_mem_condCDFSet h).tendsto_atBot_zero - · have h' := hasCondCDF_of_mem_condCDFSet h - exact fun r => ((h'.le_one r).trans_lt ENNReal.one_lt_top).ne - · exact ENNReal.zero_ne_top - · exact tendsto_defaultRatCDF_atBot -#align probability_theory.tendsto_cond_cdf_rat_at_bot ProbabilityTheory.tendsto_condCDFRat_atBot - -theorem tendsto_condCDFRat_atTop (ρ : Measure (α × ℝ)) (a : α) : - Tendsto (condCDFRat ρ a) atTop (𝓝 1) := by - unfold condCDFRat - split_ifs with h - · have h' := hasCondCDF_of_mem_condCDFSet h - rw [← ENNReal.one_toReal, ENNReal.tendsto_toReal_iff] - · exact h'.tendsto_atTop_one - · exact fun r => ((h'.le_one r).trans_lt ENNReal.one_lt_top).ne - · exact ENNReal.one_ne_top - · exact tendsto_defaultRatCDF_atTop -#align probability_theory.tendsto_cond_cdf_rat_at_top ProbabilityTheory.tendsto_condCDFRat_atTop - -theorem condCDFRat_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) : - (fun a => condCDFRat ρ a r) =ᵐ[ρ.fst] fun a => (preCDF ρ r a).toReal := by - filter_upwards [mem_condCDFSet_ae ρ] with a ha using condCDFRat_of_mem ρ a ha r -#align probability_theory.cond_cdf_rat_ae_eq ProbabilityTheory.condCDFRat_ae_eq - -theorem ofReal_condCDFRat_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) : - (fun a => ENNReal.ofReal (condCDFRat ρ a r)) =ᵐ[ρ.fst] preCDF ρ r := by - filter_upwards [condCDFRat_ae_eq ρ r, preCDF_le_one ρ] with a ha ha_le_one - rw [ha, ENNReal.ofReal_toReal] - exact ((ha_le_one r).trans_lt ENNReal.one_lt_top).ne -#align probability_theory.of_real_cond_cdf_rat_ae_eq ProbabilityTheory.ofReal_condCDFRat_ae_eq - -theorem inf_gt_condCDFRat (ρ : Measure (α × ℝ)) (a : α) (t : ℚ) : - ⨅ r : Ioi t, condCDFRat ρ a r = condCDFRat ρ a t := by - by_cases ha : a ∈ condCDFSet ρ - · simp_rw [condCDFRat_of_mem ρ a ha] - have ha' := hasCondCDF_of_mem_condCDFSet ha - rw [← ENNReal.toReal_iInf] - · suffices ⨅ i : ↥(Ioi t), preCDF ρ (↑i) a = preCDF ρ t a by rw [this] - rw [← ha'.iInf_rat_gt_eq] - · exact fun r => ((ha'.le_one r).trans_lt ENNReal.one_lt_top).ne - · simp_rw [condCDFRat_of_not_mem ρ a ha] - exact inf_gt_rat_defaultRatCDF _ -#align probability_theory.inf_gt_cond_cdf_rat ProbabilityTheory.inf_gt_condCDFRat - -lemma isCDFLike_condCDFRat (ρ : Measure (α × ℝ)) : IsCDFLike (condCDFRat ρ) where - mono := monotone_condCDFRat ρ - nonneg := condCDFRat_nonneg ρ - le_one := condCDFRat_le_one ρ - tendsto_atTop_one := tendsto_condCDFRat_atTop ρ - tendsto_atBot_zero := tendsto_condCDFRat_atBot ρ - iInf_rat_gt_eq := inf_gt_condCDFRat ρ - measurable := measurable_condCDFRat ρ - +#noalign probability_theory.has_cond_cdf +#noalign probability_theory.has_cond_cdf_ae +#noalign probability_theory.cond_cdf_set +#noalign probability_theory.measurable_set_cond_cdf_set +#noalign probability_theory.has_cond_cdf_of_mem_cond_cdf_set +#noalign probability_theory.mem_cond_cdf_set_ae +#noalign probability_theory.cond_cdf_rat +#noalign probability_theory.cond_cdf_rat_of_not_mem +#noalign probability_theory.cond_cdf_rat_of_mem +#noalign probability_theory.monotone_cond_cdf_rat +#noalign probability_theory.measurable_cond_cdf_rat +#noalign probability_theory.cond_cdf_rat_nonneg +#noalign probability_theory.cond_cdf_rat_le_one +#noalign probability_theory.tendsto_cond_cdf_rat_at_bot +#noalign probability_theory.tendsto_cond_cdf_rat_at_top +#noalign probability_theory.cond_cdf_rat_ae_eq +#noalign probability_theory.of_real_cond_cdf_rat_ae_eq +#noalign probability_theory.inf_gt_cond_cdf_rat #noalign probability_theory.cond_cdf' #noalign probability_theory.cond_cdf'_def #noalign probability_theory.cond_cdf'_eq_cond_cdf_rat @@ -676,43 +536,65 @@ lemma isCDFLike_condCDFRat (ρ : Measure (α × ℝ)) : IsCDFLike (condCDFRat ρ /-! ### Conditional cdf -/ +lemma isRatStieltjesPoint_ae (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : + ∀ᵐ a ∂ρ.fst, IsRatStieltjesPoint (fun a r ↦ (preCDF ρ r a).toReal) a := by + filter_upwards [monotone_preCDF ρ, preCDF_le_one ρ, tendsto_preCDF_atTop_one ρ, + tendsto_preCDF_atBot_zero ρ, inf_gt_preCDF ρ] with a h1 h2 h3 h4 h5 + constructor + · intro r r' hrr' + have h_ne_top : ∀ r, preCDF ρ r a ≠ ∞ := fun r ↦ + ((h2 r).trans_lt ENNReal.one_lt_top).ne + rw [ENNReal.toReal_le_toReal (h_ne_top _) (h_ne_top _)] + exact h1 hrr' + · exact fun _ ↦ ENNReal.toReal_nonneg + · refine fun r ↦ ENNReal.toReal_le_of_le_ofReal zero_le_one ?_ + rw [ENNReal.ofReal_one] + exact h2 r + · rw [← ENNReal.one_toReal, ENNReal.tendsto_toReal_iff] + · exact h3 + · exact fun r ↦ ((h2 r).trans_lt ENNReal.one_lt_top).ne + · exact ENNReal.one_ne_top + · rw [← ENNReal.zero_toReal, ENNReal.tendsto_toReal_iff] + · exact h4 + · exact fun r ↦ ((h2 r).trans_lt ENNReal.one_lt_top).ne + · exact ENNReal.zero_ne_top + · intro q + rw [← ENNReal.toReal_iInf] + · suffices ⨅ i : ↥(Ioi q), preCDF ρ (↑i) a = preCDF ρ q a by rw [this] + rw [← h5] + · exact fun r ↦ ((h2 r).trans_lt ENNReal.one_lt_top).ne /-- Conditional cdf of the measure given the value on `α`, as a Stieltjes function. -/ noncomputable def condCDF (ρ : Measure (α × ℝ)) (a : α) : StieltjesFunction := - todo2 (isCDFLike_condCDFRat ρ) a + todo3 (fun a r ↦ (preCDF ρ r a).toReal) (fun _ ↦ measurable_preCDF.ennreal_toReal) a #align probability_theory.cond_cdf ProbabilityTheory.condCDF -theorem condCDF_eq_condCDFRat (ρ : Measure (α × ℝ)) (a : α) (r : ℚ) : - condCDF ρ a r = condCDFRat ρ a r := - todo2_eq _ _ r -#align probability_theory.cond_cdf_eq_cond_cdf_rat ProbabilityTheory.condCDF_eq_condCDFRat +#noalign probability_theory.cond_cdf_eq_cond_cdf_rat /-- The conditional cdf is non-negative for all `a : α`. -/ theorem condCDF_nonneg (ρ : Measure (α × ℝ)) (a : α) (r : ℝ) : 0 ≤ condCDF ρ a r := - todo2_nonneg _ a r + todo3_nonneg _ a r #align probability_theory.cond_cdf_nonneg ProbabilityTheory.condCDF_nonneg /-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ theorem condCDF_le_one (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : condCDF ρ a x ≤ 1 := - todo2_le_one _ _ _ + todo3_le_one _ _ _ #align probability_theory.cond_cdf_le_one ProbabilityTheory.condCDF_le_one /-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ theorem tendsto_condCDF_atBot (ρ : Measure (α × ℝ)) (a : α) : - Tendsto (condCDF ρ a) atBot (𝓝 0) := - tendsto_todo2_atBot _ _ + Tendsto (condCDF ρ a) atBot (𝓝 0) := tendsto_todo3_atBot _ _ #align probability_theory.tendsto_cond_cdf_at_bot ProbabilityTheory.tendsto_condCDF_atBot /-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ theorem tendsto_condCDF_atTop (ρ : Measure (α × ℝ)) (a : α) : - Tendsto (condCDF ρ a) atTop (𝓝 1) := - tendsto_todo2_atTop _ _ + Tendsto (condCDF ρ a) atTop (𝓝 1) := tendsto_todo3_atTop _ _ #align probability_theory.tendsto_cond_cdf_at_top ProbabilityTheory.tendsto_condCDF_atTop theorem condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) : - (fun a => condCDF ρ a r) =ᵐ[ρ.fst] fun a => (preCDF ρ r a).toReal := by - filter_upwards [mem_condCDFSet_ae ρ] with a ha using - (condCDF_eq_condCDFRat ρ a r).trans (condCDFRat_of_mem ρ a ha r) + (fun a ↦ condCDF ρ a r) =ᵐ[ρ.fst] fun a ↦ (preCDF ρ r a).toReal := by + filter_upwards [isRatStieltjesPoint_ae ρ] with a ha + rw [condCDF, todo3_eq, toCDFLike_of_isRatStieltjesPoint ha] #align probability_theory.cond_cdf_ae_eq ProbabilityTheory.condCDF_ae_eq theorem ofReal_condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) : @@ -724,7 +606,7 @@ theorem ofReal_condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r /-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ theorem measurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : Measurable fun a => condCDF ρ a x := - measurable_todo2 _ _ + measurable_todo3 _ _ #align probability_theory.measurable_cond_cdf ProbabilityTheory.measurable_condCDF /-- Auxiliary lemma for `set_lintegral_cond_cdf`. -/ @@ -790,8 +672,7 @@ theorem lintegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : /-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ theorem stronglyMeasurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : - StronglyMeasurable fun a => condCDF ρ a x := - stronglyMeasurable_todo2 _ _ + StronglyMeasurable fun a => condCDF ρ a x := stronglyMeasurable_todo3 _ _ #align probability_theory.strongly_measurable_cond_cdf ProbabilityTheory.stronglyMeasurable_condCDF theorem integrable_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) : @@ -804,9 +685,9 @@ theorem integrable_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : -- Porting note: was exact_mod_cast condCDF_le_one _ _ _ simp only [ENNReal.coe_le_one_iff] exact condCDF_le_one _ _ _ - refine' - (set_lintegral_mono (measurable_condCDF _ _).ennnorm measurable_one fun y _ => this y).trans - _ + refine + (set_lintegral_mono (measurable_condCDF _ _).ennnorm measurable_one fun y _ ↦ this y).trans + ?_ simp only [Pi.one_apply, lintegral_one, Measure.restrict_apply, MeasurableSet.univ, univ_inter] exact measure_mono (subset_univ _) #align probability_theory.integrable_cond_cdf ProbabilityTheory.integrable_condCDF @@ -829,22 +710,19 @@ theorem integral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : section Measure theorem measure_condCDF_Iic (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : - (condCDF ρ a).measure (Iic x) = ENNReal.ofReal (condCDF ρ a x) := - measure_todo2_Iic _ _ _ + (condCDF ρ a).measure (Iic x) = ENNReal.ofReal (condCDF ρ a x) := measure_todo3_Iic _ _ _ #align probability_theory.measure_cond_cdf_Iic ProbabilityTheory.measure_condCDF_Iic theorem measure_condCDF_univ (ρ : Measure (α × ℝ)) (a : α) : (condCDF ρ a).measure univ = 1 := - measure_todo2_univ _ _ + measure_todo3_univ _ _ #align probability_theory.measure_cond_cdf_univ ProbabilityTheory.measure_condCDF_univ instance instIsProbabilityMeasure (ρ : Measure (α × ℝ)) (a : α) : - IsProbabilityMeasure (condCDF ρ a).measure := by - rw [condCDF]; infer_instance + IsProbabilityMeasure (condCDF ρ a).measure := by rw [condCDF]; infer_instance /-- The function `a ↦ (condCDF ρ a).measure` is measurable. -/ theorem measurable_measure_condCDF (ρ : Measure (α × ℝ)) : - Measurable fun a => (condCDF ρ a).measure := - measurable_measure_todo2 _ + Measurable fun a => (condCDF ρ a).measure := measurable_measure_todo3 _ #align probability_theory.measurable_measure_cond_cdf ProbabilityTheory.measurable_measure_condCDF end Measure diff --git a/Mathlib/Probability/Kernel/StieltjesReal.lean b/Mathlib/Probability/Kernel/StieltjesReal.lean index b82c09b155f40..9991f53c63ae2 100644 --- a/Mathlib/Probability/Kernel/StieltjesReal.lean +++ b/Mathlib/Probability/Kernel/StieltjesReal.lean @@ -6,6 +6,7 @@ Authors: Rémy Degenne import Mathlib.MeasureTheory.Measure.Stieltjes import Mathlib.MeasureTheory.Decomposition.RadonNikodym import Mathlib.MeasureTheory.Constructions.Prod.Basic +import Mathlib.MeasureTheory.Constructions.Polish /-! @@ -17,12 +18,69 @@ open MeasureTheory Set Filter TopologicalSpace open scoped NNReal ENNReal MeasureTheory Topology +lemma measurableSet_tendsto_nhds {β γ ι : Type*} [MeasurableSpace β] + [TopologicalSpace γ] [PolishSpace γ] [MeasurableSpace γ] + [hγ : OpensMeasurableSpace γ] [Countable ι] {l : Filter ι} + [l.IsCountablyGenerated] {f : ι → β → γ} (hf : ∀ i, Measurable (f i)) (c : γ) : + MeasurableSet { x | Tendsto (fun n ↦ f n x) l (𝓝 c) } := sorry + +lemma measurableSet_tendsto_fun {β γ ι : Type*} [MeasurableSpace β] + [TopologicalSpace γ] [PolishSpace γ] [MeasurableSpace γ] + [hγ : OpensMeasurableSpace γ] [Countable ι] {l : Filter ι} + [l.IsCountablyGenerated] {f : ι → β → γ} (hf : ∀ i, Measurable (f i)) {g : β → γ} + (hg : Measurable g) : + MeasurableSet { x | Tendsto (fun n ↦ f n x) l (𝓝 (g x)) } := by + letI := upgradePolishSpace γ + have : { x | Tendsto (fun n ↦ f n x) l (𝓝 (g x)) } + = { x | Tendsto (fun n ↦ dist (f n x) (g x)) l (𝓝 0) } := by + ext x + simp only [mem_setOf_eq] + rw [tendsto_iff_dist_tendsto_zero] + rw [this] + exact measurableSet_tendsto_nhds (fun n ↦ (hf n).dist hg) 0 + namespace ProbabilityTheory -variable {α β ι : Type*} [MeasurableSpace α] +variable {α β ι : Type*} [MeasurableSpace α] {f : α → ℚ → ℝ} section IsCDFLike +structure IsRatStieltjesPoint (f : α → ℚ → ℝ) (a : α) : Prop where + mono : Monotone (f a) + nonneg : ∀ q, 0 ≤ f a q + le_one : ∀ q, f a q ≤ 1 + tendsto_atTop_one : Tendsto (f a) atTop (𝓝 1) + tendsto_atBot_zero : Tendsto (f a) atBot (𝓝 0) + iInf_rat_gt_eq : ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t + +lemma measurableSet_isRatStieltjesPoint (hf : ∀ q, Measurable (fun a ↦ f a q)) : + MeasurableSet {a | IsRatStieltjesPoint f a} := by + have h1 : MeasurableSet {a | Monotone (f a)} := by + change MeasurableSet {a | ∀ q r (hqr : q ≤ r), f a q ≤ f a r} + simp_rw [Set.setOf_forall] + refine MeasurableSet.iInter (fun q ↦ ?_) + refine MeasurableSet.iInter (fun r ↦ ?_) + refine MeasurableSet.iInter (fun hqr ↦ ?_) + exact measurableSet_le (hf q) (hf r) + have h2 : MeasurableSet {a | ∀ q, 0 ≤ f a q} := by + simp_rw [Set.setOf_forall] + refine MeasurableSet.iInter (fun q ↦ ?_) + exact measurableSet_le measurable_const (hf q) + have h3 : MeasurableSet {a | ∀ q, f a q ≤ 1} := by + simp_rw [Set.setOf_forall] + refine MeasurableSet.iInter (fun q ↦ ?_) + exact measurableSet_le (hf q) measurable_const + have h4 : MeasurableSet {a | Tendsto (f a) atTop (𝓝 1)} := + measurableSet_tendsto_nhds (fun q ↦ hf q) 1 + have h5 : MeasurableSet {a | Tendsto (f a) atBot (𝓝 0)} := + measurableSet_tendsto_nhds (fun q ↦ hf q) 0 + have h6 : MeasurableSet {a | ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t} := by + rw [Set.setOf_forall] + refine MeasurableSet.iInter (fun q ↦ ?_) + exact measurableSet_eq_fun (measurable_iInf fun _ ↦ hf _) (hf _) + have h := ((((h1.inter h2).inter h3).inter h4).inter h5).inter h6 + sorry + structure IsCDFLike (f : α → ℚ → ℝ) : Prop where mono : ∀ a, Monotone (f a) nonneg : ∀ a q, 0 ≤ f a q @@ -120,6 +178,38 @@ lemma isCDFLike_defaultRatCDF (α : Type*) [MeasurableSpace α] : end DefaultRatCDF +section ToCDFLike + +open Classical in +noncomputable +def toCDFLike (f : α → ℚ → ℝ) : α → ℚ → ℝ := fun a q ↦ + if IsRatStieltjesPoint f a then f a q else defaultRatCDF q + +lemma toCDFLike_of_isRatStieltjesPoint {a : α} (h : IsRatStieltjesPoint f a) (q : ℚ) : + toCDFLike f a q = f a q := by + unfold toCDFLike; simp [h] + +lemma isCDFLike_toCDFLike (hf : ∀ q, Measurable fun a ↦ f a q) : + IsCDFLike (toCDFLike f) where + mono a := by + unfold toCDFLike; split_ifs with h; exacts [h.mono, monotone_defaultRatCDF] + nonneg a := by + unfold toCDFLike; split_ifs with h; exacts [h.nonneg, defaultRatCDF_nonneg] + le_one a := by + unfold toCDFLike; split_ifs with h; exacts [h.le_one, defaultRatCDF_le_one] + tendsto_atTop_one a := by + unfold toCDFLike; split_ifs with h; exacts [h.tendsto_atTop_one, tendsto_defaultRatCDF_atTop] + tendsto_atBot_zero a := by + unfold toCDFLike; split_ifs with h; exacts [h.tendsto_atBot_zero, tendsto_defaultRatCDF_atBot] + iInf_rat_gt_eq a := by + unfold toCDFLike; split_ifs with h; exacts [h.iInf_rat_gt_eq, inf_gt_rat_defaultRatCDF] + measurable q := + Measurable.ite (measurableSet_isRatStieltjesPoint hf) (hf q) (measurable_defaultRatCDF α q) + +end ToCDFLike + +section IsCDFLike.stieltjesFunction + variable {f : α → ℚ → ℝ} (hf : IsCDFLike f) /-- Conditional cdf of the measure given the value on `α`, as a plain function. This is an auxiliary @@ -184,18 +274,18 @@ theorem continuousWithinAt_todo1_Ici (a : α) (x : ℝ) : /-- Conditional cdf of the measure given the value on `α`, as a Stieltjes function. -/ -noncomputable def todo2 (a : α) : StieltjesFunction where +noncomputable def IsCDFLike.stieltjesFunction (a : α) : StieltjesFunction where toFun := todo1 f a mono' := monotone_todo1 hf a right_continuous' x := continuousWithinAt_todo1_Ici hf a x -theorem todo2_eq (a : α) (r : ℚ) : todo2 hf a r = f a r := todo1_eq hf a r +theorem todo2_eq (a : α) (r : ℚ) : hf.stieltjesFunction a r = f a r := todo1_eq hf a r /-- The conditional cdf is non-negative for all `a : α`. -/ -theorem todo2_nonneg (a : α) (r : ℝ) : 0 ≤ todo2 hf a r := todo1_nonneg hf a r +theorem todo2_nonneg (a : α) (r : ℝ) : 0 ≤ hf.stieltjesFunction a r := todo1_nonneg hf a r /-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ -theorem todo2_le_one (a : α) (x : ℝ) : todo2 hf a x ≤ 1 := by +theorem todo2_le_one (a : α) (x : ℝ) : hf.stieltjesFunction a x ≤ 1 := by obtain ⟨r, hrx⟩ := exists_rat_gt x rw [← StieltjesFunction.iInf_rat_gt_eq] simp_rw [todo2_eq] @@ -204,7 +294,7 @@ theorem todo2_le_one (a : α) (x : ℝ) : todo2 hf a x ≤ 1 := by /-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ theorem tendsto_todo2_atBot (a : α) : - Tendsto (todo2 hf a) atBot (𝓝 0) := by + Tendsto (hf.stieltjesFunction a) atBot (𝓝 0) := by have h_exists : ∀ x : ℝ, ∃ q : ℚ, x < q ∧ ↑q < x + 1 := fun x ↦ exists_rat_btwn (lt_add_one x) let qs : ℝ → ℚ := fun x ↦ (h_exists x).choose have hqs_tendsto : Tendsto qs atBot atBot := by @@ -217,11 +307,11 @@ theorem tendsto_todo2_atBot (a : α) : refine tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds ((hf.tendsto_atBot_zero a).comp hqs_tendsto) (todo2_nonneg hf a) fun x ↦ ?_ rw [Function.comp_apply, ← todo2_eq hf] - exact (todo2 hf a).mono (h_exists x).choose_spec.1.le + exact (hf.stieltjesFunction a).mono (h_exists x).choose_spec.1.le /-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ theorem tendsto_todo2_atTop (a : α) : - Tendsto (todo2 hf a) atTop (𝓝 1) := by + Tendsto (hf.stieltjesFunction a) atTop (𝓝 1) := by have h_exists : ∀ x : ℝ, ∃ q : ℚ, x - 1 < q ∧ ↑q < x := fun x ↦ exists_rat_btwn (sub_one_lt x) let qs : ℝ → ℚ := fun x ↦ (h_exists x).choose have hqs_tendsto : Tendsto qs atTop atTop := by @@ -234,11 +324,11 @@ theorem tendsto_todo2_atTop (a : α) : tendsto_const_nhds ?_ (todo2_le_one hf a) intro x rw [Function.comp_apply, ← todo2_eq hf] - exact (todo2 hf a).mono (le_of_lt (h_exists x).choose_spec.2) + exact (hf.stieltjesFunction a).mono (le_of_lt (h_exists x).choose_spec.2) /-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ -theorem measurable_todo2 (x : ℝ) : Measurable fun a ↦ todo2 hf a x := by - have : (fun a ↦ todo2 hf a x) = fun a ↦ ⨅ r : { r' : ℚ // x < r' }, f a ↑r := by +theorem measurable_todo2 (x : ℝ) : Measurable fun a ↦ hf.stieltjesFunction a x := by + have : (fun a ↦ hf.stieltjesFunction a x) = fun a ↦ ⨅ r : { r' : ℚ // x < r' }, f a ↑r := by ext1 a rw [← StieltjesFunction.iInf_rat_gt_eq] congr with q @@ -248,42 +338,43 @@ theorem measurable_todo2 (x : ℝ) : Measurable fun a ↦ todo2 hf a x := by /-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ theorem stronglyMeasurable_todo2 (x : ℝ) : - StronglyMeasurable fun a ↦ todo2 hf a x := + StronglyMeasurable fun a ↦ hf.stieltjesFunction a x := (measurable_todo2 hf x).stronglyMeasurable section Measure theorem measure_todo2_Iic (a : α) (x : ℝ) : - (todo2 hf a).measure (Iic x) = ENNReal.ofReal (todo2 hf a x) := by - rw [← sub_zero (todo2 hf a x)] - exact (todo2 hf a).measure_Iic (tendsto_todo2_atBot hf a) _ + (hf.stieltjesFunction a).measure (Iic x) = ENNReal.ofReal (hf.stieltjesFunction a x) := by + rw [← sub_zero (hf.stieltjesFunction a x)] + exact (hf.stieltjesFunction a).measure_Iic (tendsto_todo2_atBot hf a) _ -theorem measure_todo2_univ (a : α) : (todo2 hf a).measure univ = 1 := by +theorem measure_todo2_univ (a : α) : (hf.stieltjesFunction a).measure univ = 1 := by rw [← ENNReal.ofReal_one, ← sub_zero (1 : ℝ)] exact StieltjesFunction.measure_univ _ (tendsto_todo2_atBot hf a) (tendsto_todo2_atTop hf a) instance instIsProbabilityMeasure_todo2 (a : α) : - IsProbabilityMeasure (todo2 hf a).measure := + IsProbabilityMeasure (hf.stieltjesFunction a).measure := ⟨measure_todo2_univ hf a⟩ /-- The function `a ↦ (condCDF ρ a).measure` is measurable. -/ theorem measurable_measure_todo2 : - Measurable fun a ↦ (todo2 hf a).measure := by + Measurable fun a ↦ (hf.stieltjesFunction a).measure := by rw [Measure.measurable_measure] refine' fun s hs ↦ ?_ -- Porting note: supplied `C` refine' MeasurableSpace.induction_on_inter - (C := fun s ↦ Measurable fun b ↦ StieltjesFunction.measure (todo2 hf b) s) + (C := fun s ↦ Measurable fun b ↦ StieltjesFunction.measure (hf.stieltjesFunction b) s) (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic _ _ _ _ hs · simp only [measure_empty, measurable_const] · rintro S ⟨u, rfl⟩ simp_rw [measure_todo2_Iic hf _ u] exact (measurable_todo2 hf u).ennreal_ofReal · intro t ht ht_cd_meas - have : (fun a ↦ (todo2 hf a).measure tᶜ) = - (fun a ↦ (todo2 hf a).measure univ) - fun a ↦ (todo2 hf a).measure t := by + have : (fun a ↦ (hf.stieltjesFunction a).measure tᶜ) = + (fun a ↦ (hf.stieltjesFunction a).measure univ) + - fun a ↦ (hf.stieltjesFunction a).measure t := by ext1 a - rw [measure_compl ht (measure_ne_top (todo2 hf a).measure _), Pi.sub_apply] + rw [measure_compl ht (measure_ne_top (hf.stieltjesFunction a).measure _), Pi.sub_apply] simp_rw [this, measure_todo2_univ hf] exact Measurable.sub measurable_const ht_cd_meas · intro f hf_disj hf_meas hf_cd_meas @@ -292,4 +383,55 @@ theorem measurable_measure_todo2 : end Measure -end ProbabilityTheory +end IsCDFLike.stieltjesFunction + +variable {f : α → ℚ → ℝ} + +noncomputable +def todo3 (f : α → ℚ → ℝ) (hf : ∀ q, Measurable fun a ↦ f a q) : α → StieltjesFunction := + (isCDFLike_toCDFLike hf).stieltjesFunction + +theorem todo3_eq (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℚ) : + todo3 f hf a r = toCDFLike f a r := todo2_eq _ a r + +/-- The conditional cdf is non-negative for all `a : α`. -/ +theorem todo3_nonneg (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℝ) : + 0 ≤ todo3 f hf a r := todo2_nonneg _ a r + +/-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ +theorem todo3_le_one (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : + todo3 f hf a x ≤ 1 := todo2_le_one _ a x + +/-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ +theorem tendsto_todo3_atBot (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : + Tendsto (todo3 f hf a) atBot (𝓝 0) := tendsto_todo2_atBot _ a + +/-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ +theorem tendsto_todo3_atTop (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : + Tendsto (todo3 f hf a) atTop (𝓝 1) := tendsto_todo2_atTop _ a + +/-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ +theorem measurable_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : + Measurable fun a ↦ todo3 f hf a x := measurable_todo2 _ x + +/-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ +theorem stronglyMeasurable_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : + StronglyMeasurable fun a ↦ todo3 f hf a x := stronglyMeasurable_todo2 _ x + +section Measure + +theorem measure_todo3_Iic (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : + (todo3 f hf a).measure (Iic x) = ENNReal.ofReal (todo3 f hf a x) := + measure_todo2_Iic _ _ _ + +theorem measure_todo3_univ (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : + (todo3 f hf a).measure univ = 1 := measure_todo2_univ _ _ + +instance instIsProbabilityMeasure_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : + IsProbabilityMeasure (todo3 f hf a).measure := + instIsProbabilityMeasure_todo2 _ _ + +theorem measurable_measure_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) : + Measurable fun a ↦ (todo3 f hf a).measure := measurable_measure_todo2 _ + +end Measure diff --git a/Mathlib/Probability/Kernel/draft.lean b/Mathlib/Probability/Kernel/draft.lean index bb82ff92cdc84..e25f9d547cb5f 100644 --- a/Mathlib/Probability/Kernel/draft.lean +++ b/Mathlib/Probability/Kernel/draft.lean @@ -4,34 +4,13 @@ import Mathlib.Probability.Kernel.BuildKernel open MeasureTheory Set Filter -open scoped ENNReal MeasureTheory Topology ProbabilityTheory +open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory variable {α Ω : Type*} {mα : MeasurableSpace α} [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] -lemma measurableSet_tendsto_nhds {β γ ι : Type*} [MeasurableSpace β] - [TopologicalSpace γ] [PolishSpace γ] [MeasurableSpace γ] - [hγ : OpensMeasurableSpace γ] [Countable ι] {l : Filter ι} - [l.IsCountablyGenerated] {f : ι → β → γ} (hf : ∀ i, Measurable (f i)) (c : γ) : - MeasurableSet { x | Tendsto (fun n ↦ f n x) l (𝓝 c) } := sorry - -lemma measurableSet_tendsto_fun {β γ ι : Type*} [MeasurableSpace β] - [TopologicalSpace γ] [PolishSpace γ] [MeasurableSpace γ] - [hγ : OpensMeasurableSpace γ] [Countable ι] {l : Filter ι} - [l.IsCountablyGenerated] {f : ι → β → γ} (hf : ∀ i, Measurable (f i)) {g : β → γ} - (hg : Measurable g) : - MeasurableSet { x | Tendsto (fun n ↦ f n x) l (𝓝 (g x)) } := by - letI := upgradePolishSpace γ - have : { x | Tendsto (fun n ↦ f n x) l (𝓝 (g x)) } - = { x | Tendsto (fun n ↦ dist (f n x) (g x)) l (𝓝 0) } := by - ext x - simp only [mem_setOf_eq] - rw [tendsto_iff_dist_tendsto_zero] - rw [this] - exact measurableSet_tendsto_nhds (fun n ↦ (hf n).dist hg) 0 - section Real section dissection_system @@ -40,16 +19,37 @@ def I (n : ℕ) (k : ℤ) : Set ℝ := Set.Ico (k * (2⁻¹ : ℝ) ^ n) ((k + 1) lemma measurableSet_I (n : ℕ) (k : ℤ) : MeasurableSet (I n k) := measurableSet_Ico +lemma Measure.iInf_Iic_gt_prod {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] + {s : Set α} (hs : MeasurableSet s) (t : ℚ) : + ⨅ r : { r' : ℚ // t < r' }, ρ (s ×ˢ Iic (r : ℝ)) = ρ (s ×ˢ Iic (t : ℝ)) := by + have h := Measure.iInf_IicSnd_gt ρ t hs + simp_rw [Measure.IicSnd_apply ρ _ hs] at h + rw [← h] + lemma pairwise_disjoint_I (n : ℕ) : Pairwise (Disjoint on fun k ↦ I n k) := by + intro i j hij + rw [Function.onFun, Set.disjoint_iff] + intro x + simp only [I, inv_pow, mem_inter_iff, mem_Ico, mem_empty_iff_false, and_imp, imp_false, not_lt] + intro h1 h2 h3 sorry lemma I_succ_union (n : ℕ) (k : ℤ) : I (n+1) (2 * k) ∪ I (n+1) (2 * k + 1) = I n k := by ext x cases lt_or_le x ((2 * k + 1) * ((2 : ℝ) ^ (n+1))⁻¹) with - | inl h=> + | inl h => simp only [I, inv_pow, mem_Ico, Int.cast_mul, Int.int_cast_ofNat, Int.cast_add, Int.cast_one, mem_union, h, and_true, not_le.mpr h, false_and, or_false] - sorry + have : x < (k + 1) * (2 ^ n)⁻¹ := by + refine h.trans_le ?_ + rw [pow_add, pow_one, mul_inv, mul_comm _ 2⁻¹, ← mul_assoc] + gcongr + rw [add_mul, one_mul, mul_comm, ← mul_assoc, inv_mul_cancel two_ne_zero, one_mul] + gcongr + norm_num + simp only [this, and_true] + rw [pow_add, pow_one, mul_inv, mul_comm _ 2⁻¹, ← mul_assoc, mul_comm _ 2⁻¹, ← mul_assoc, + inv_mul_cancel two_ne_zero, one_mul] | inr h => simp only [I, inv_pow, mem_Ico, Int.cast_mul, Int.int_cast_ofNat, Int.cast_add, Int.cast_one, mem_union, not_lt.mpr h, and_false, h, true_and, false_or] @@ -157,9 +157,9 @@ lemma iSup_ℱ : ⨆ n, ℱ n = borel ℝ := by refine MeasurableSpace.generateFrom_le (fun s ⟨x, hx⟩ ↦ ?_) rw [← hx, ← iInter_biUnion_I x] refine MeasurableSet.iInter (fun n ↦ ?_) - refine MeasurableSet.biUnion ?_ (fun k hk ↦ ?_) - · sorry - · sorry + refine MeasurableSet.biUnion ?_ (fun k _ ↦ ?_) + · exact to_countable _ + · exact le_iSup ℱ n _ (measurableSet_ℱ_I n k) end dissection_system @@ -214,10 +214,9 @@ lemma measurable_m_right (κ : kernel α (ℝ × β)) {s : Set β} (a : α) (hs Measurable (M κ a s n) := (measurable_m κ hs n).comp (measurable_const.prod_mk measurable_id) -lemma adapted_m (κ : kernel α (ℝ × β)) (a : α) (s : Set β) : Adapted ℱ (M κ a s) := by - intro n +lemma measurable_ℱ_m (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) : + Measurable[ℱ n] (M κ a s n) := by rw [m_def] - refine Measurable.stronglyMeasurable ?_ refine @Measurable.ennreal_toReal _ (ℱ n) _ ?_ refine Measurable.div ?_ ?_ · change Measurable[ℱ n] ((fun k ↦ κ a (I n k ×ˢ s)) ∘ (indexI n)) @@ -227,6 +226,13 @@ lemma adapted_m (κ : kernel α (ℝ × β)) (a : α) (s : Set β) : Adapted ℱ refine Measurable.comp ?_ (measurable_indexI n) exact measurable_of_countable _ +lemma stronglyMeasurable_ℱ_m (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) : + StronglyMeasurable[ℱ n] (M κ a s n) := + (measurable_ℱ_m κ a s n).stronglyMeasurable + +lemma adapted_m (κ : kernel α (ℝ × β)) (a : α) (s : Set β) : Adapted ℱ (M κ a s) := + stronglyMeasurable_ℱ_m κ a s + lemma m_nonneg (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : 0 ≤ M κ a s n t := ENNReal.toReal_nonneg @@ -245,17 +251,9 @@ lemma m_le_one (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : lemma snorm_m_le_one (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] (a : α) (s : Set β) (n : ℕ) : snorm (M κ a s n) 1 (kernel.fst κ a) ≤ 1 := by - refine (snorm_mono_real (g := fun _ ↦ 1) ?_).trans ?_ - · intro x - simp only [Real.norm_eq_abs, abs_le] - constructor - · have h := m_nonneg κ a s n x - linarith - · exact m_le_one _ _ _ _ _ - · by_cases h0 : kernel.fst κ a = 0 - · simp [h0] - · rw [snorm_const _ one_ne_zero h0] - simp + refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun x ↦ ?_))).trans ?_ + · simp only [Real.norm_eq_abs, abs_of_nonneg (m_nonneg κ a s n x), m_le_one κ a s n x] + · simp lemma integrable_m (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : @@ -265,7 +263,7 @@ lemma integrable_m (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] · exact measurable_m_right κ a hs n · exact (snorm_m_le_one κ a s n).trans_lt ENNReal.one_lt_top -lemma set_integral_m_I (κ : kernel α (ℝ × β)) [IsFiniteKernel (kernel.fst κ)] +lemma set_integral_m_I (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) (k : ℤ) : ∫ t in I n k, M κ a s n t ∂(kernel.fst κ a) = (κ a (I n k ×ˢ s)).toReal := by simp_rw [M] @@ -273,12 +271,23 @@ lemma set_integral_m_I (κ : kernel α (ℝ × β)) [IsFiniteKernel (kernel.fst rotate_left · refine Measurable.aemeasurable ?_ have h := measurable_m_aux κ hs n - sorry - · sorry + change Measurable + ((fun (p : α × ℝ) ↦ κ p.1 (I n (indexI n p.2) ×ˢ s) / kernel.fst κ p.1 (I n (indexI n p.2))) + ∘ (fun t ↦ (a, t))) + exact h.comp measurable_prod_mk_left + · refine ae_of_all _ (fun t ↦ ?_) + by_cases h0 : kernel.fst κ a (I n (indexI n t)) = 0 + · suffices κ a (I n (indexI n t) ×ˢ s) = 0 by simp [h0, this] + rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0 + refine measure_mono_null (fun x ↦ ?_) h0 + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + · refine ENNReal.div_lt_top ?_ h0 + exact measure_ne_top _ _ congr have : ∫⁻ t in I n k, κ a (I n (indexI n t) ×ˢ s) / kernel.fst κ a (I n (indexI n t)) ∂(kernel.fst κ) a - = ∫⁻ t in I n k, κ a (I n k ×ˢ s) / kernel.fst κ a (I n k) ∂(kernel.fst κ) a := by + = ∫⁻ _ in I n k, κ a (I n k ×ˢ s) / kernel.fst κ a (I n k) ∂(kernel.fst κ) a := by refine set_lintegral_congr_fun (measurableSet_I _ _) (ae_of_all _ (fun t ht ↦ ?_)) rw [indexI_of_mem _ _ _ ht] rw [this] @@ -311,60 +320,86 @@ lemma integral_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] lemma set_integral_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : ∫ t in A, M κ a s n t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal := by - suffices MeasurableSet A ∧ ∫ t in A, M κ a s n t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal by - exact this.2 - refine MeasurableSpace.generateFrom_induction - (p := fun A' ↦ MeasurableSet A' - ∧ ∫ t in A', M κ a s n t ∂(kernel.fst κ) a = ENNReal.toReal (κ a (A' ×ˢ s))) - (C := {s | ∃ k, s = I n k}) ?_ ?_ ?_ ?_ hA + refine MeasurableSpace.induction_on_inter (m := ℱ n) (s := {s | ∃ k, s = I n k}) + (C := fun A ↦ ∫ t in A, M κ a s n t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal) rfl + ?_ ?_ ?_ ?_ ?_ hA + · rintro s ⟨i, rfl⟩ t ⟨j, rfl⟩ hst + refine ⟨i, ?_⟩ + suffices i = j by rw [this, inter_self] + by_contra h_ne + have h_disj := pairwise_disjoint_I n h_ne + rw [nonempty_iff_ne_empty] at hst + refine hst ?_ + rwa [Function.onFun, disjoint_iff_inter_eq_empty] at h_disj + · simp · rintro _ ⟨k, rfl⟩ rw [set_integral_m_I _ _ hs] - exact ⟨measurableSet_I _ _, rfl⟩ - · simp only [MeasurableSet.empty, Measure.restrict_empty, integral_zero_measure, empty_prod, - OuterMeasure.empty', ENNReal.zero_toReal, and_self] - · intro A ⟨hA, hA_eq⟩ - have h := integral_add_compl hA (integrable_m κ a hs n) - refine ⟨hA.compl, ?_⟩ + · intro A hA hA_eq + have hA' : MeasurableSet A := ℱ.le _ _ hA + have h := integral_add_compl hA' (integrable_m κ a hs n) rw [hA_eq, integral_m κ a hs] at h have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by rw [prod_diff_prod, compl_eq_univ_diff] simp - rw [this, measure_diff (by intro x; simp) (hA.prod hs) (measure_ne_top (κ a) _), + rw [this, measure_diff (by intro x; simp) (hA'.prod hs) (measure_ne_top (κ a) _), ENNReal.toReal_sub_of_le (measure_mono (by intro x; simp)) (measure_ne_top _ _)] rw [eq_tsub_iff_add_eq_of_le, add_comm] · exact h · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] exact measure_mono (by intro x; simp) - · intro f hf - simp only at hf - -- todo: introduce disjointed sets, etc. - sorry - -lemma condexp_m (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] (a : α) (s : Set β) - {i j : ℕ} (hij : i ≤ j) : + · intro f hf_disj hf h_eq + rw [integral_iUnion (fun i ↦ ℱ.le n _ (hf i)) hf_disj (integrable_m κ _ hs _).integrableOn] + simp_rw [h_eq] + rw [iUnion_prod_const, measure_iUnion _ (fun i ↦ (ℱ.le n _ (hf i)).prod hs)] + · rw [ENNReal.tsum_toReal_eq] + exact fun _ ↦ measure_ne_top _ _ + · intro i j hij + rw [Function.onFun, Set.disjoint_prod] + exact Or.inl (hf_disj hij) + +lemma set_integral_m_of_le (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] + (a : α) {s : Set β} (hs : MeasurableSet s) {n m : ℕ} (hnm : n ≤ m) + {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : + ∫ t in A, M κ a s m t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal := + set_integral_m κ a hs m (ℱ.mono hnm A hA) + +lemma condexp_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] (a : α) {s : Set β} + (hs : MeasurableSet s) {i j : ℕ} (hij : i ≤ j) : (kernel.fst κ a)[M κ a s j | ℱ i] =ᵐ[kernel.fst κ a] M κ a s i := by - sorry - -lemma martingale_m (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] (a : α) (s : Set β) : + symm + refine ae_eq_condexp_of_forall_set_integral_eq ?_ ?_ ?_ ?_ ?_ + · exact integrable_m κ a hs j + · refine fun t _ _ ↦ Integrable.integrableOn ?_ + exact integrable_m κ _ hs _ + · intro t ht _ + rw [set_integral_m κ a hs i ht, set_integral_m_of_le κ a hs hij ht] + · exact StronglyMeasurable.aeStronglyMeasurable' (stronglyMeasurable_ℱ_m κ a s i) + +lemma martingale_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] (a : α) {s : Set β} + (hs : MeasurableSet s) : Martingale (M κ a s) ℱ (kernel.fst κ a) := - ⟨adapted_m κ a s, fun _ _ ↦ condexp_m κ a s⟩ + ⟨adapted_m κ a s, fun _ _ ↦ condexp_m κ a hs⟩ lemma m_mono_set (κ : kernel α (ℝ × β)) (a : α) {s s' : Set β} (h : s ⊆ s') (n : ℕ) (t : ℝ) : M κ a s n t ≤ M κ a s' n t := by + have h_ne_top : ∀ s, κ a (I n (indexI n t) ×ˢ s) / kernel.fst κ a (I n (indexI n t)) ≠ ⊤ := by + intro s + rw [ne_eq, ENNReal.div_eq_top] + push_neg + simp_rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] + constructor + · refine fun h h0 ↦ h (measure_mono_null (fun x ↦ ?_) h0) + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + · refine fun h_top ↦ eq_top_mono (measure_mono (fun x ↦ ?_)) h_top + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h rw [M, M, ENNReal.toReal_le_toReal] · gcongr rw [prod_subset_prod_iff] simp [subset_rfl, h] - · rw [ne_eq, ENNReal.div_eq_top] - push_neg - constructor - · sorry - · sorry - · rw [ne_eq, ENNReal.div_eq_top] - push_neg - constructor - · sorry - · sorry + · exact h_ne_top s + · exact h_ne_top s' lemma m_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ) (t : ℝ) : M κ a univ n t = if kernel.fst κ a (I n (indexI n t)) = 0 then 0 else 1 := by @@ -385,7 +420,7 @@ lemma m_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ) · rwa [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h · exact measure_ne_top _ _ -lemma m_empty (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ) (t : ℝ) : +lemma m_empty (κ : kernel α (ℝ × β)) (a : α) (n : ℕ) (t : ℝ) : M κ a ∅ n t = 0 := by rw [M] simp @@ -393,7 +428,7 @@ lemma m_empty (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ lemma m_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ) : ∀ᵐ t ∂(kernel.fst κ a), M κ a univ n t = 1 := by rw [ae_iff] - have : {t | ¬M κ a univ n t = 1} ⊆ {t | kernel.fst κ a (I n (indexI n t)) = 0} := by + have : {t | ¬ M κ a univ n t = 1} ⊆ {t | kernel.fst κ a (I n (indexI n t)) = 0} := by intro t ht simp only [mem_setOf_eq] at ht ⊢ rw [m_univ] at ht @@ -410,17 +445,32 @@ lemma m_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : intro i simp -lemma tendsto_m_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma tendsto_m_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) (t : ℝ) : Tendsto (fun m ↦ M κ a (s m) n t) atTop (𝓝 (M κ a univ n t)) := by simp_rw [M] refine (ENNReal.tendsto_toReal ?_).comp ?_ · rw [ne_eq, ENNReal.div_eq_top] push_neg - sorry + simp_rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] + constructor + · refine fun h h0 ↦ h (measure_mono_null (fun x ↦ ?_) h0) + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + · refine fun h_top ↦ eq_top_mono (measure_mono (fun x ↦ ?_)) h_top + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h by_cases h0 : kernel.fst κ a (I n (indexI n t)) = 0 - · simp only [h0] - sorry + · rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0 ⊢ + suffices ∀ m, κ a (I n (indexI n t) ×ˢ s m) = 0 by + simp only [this, h0, ENNReal.zero_div, tendsto_const_nhds_iff] + suffices κ a (I n (indexI n t) ×ˢ univ) = 0 by simp only [this, ENNReal.zero_div] + convert h0 + ext x + simp only [mem_prod, mem_univ, and_true, mem_setOf_eq] + refine fun m ↦ measure_mono_null (fun x ↦ ?_) h0 + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h refine ENNReal.Tendsto.div_const ?_ ?_ · have h := tendsto_measure_iUnion (μ := κ a) (s := fun m ↦ I n (indexI n t) ×ˢ s m) ?_ swap @@ -446,10 +496,17 @@ lemma tendsto_m_atTop_empty_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKe refine (ENNReal.tendsto_toReal ?_).comp ?_ · rw [ne_eq, ENNReal.div_eq_top] push_neg - sorry + simp by_cases h0 : kernel.fst κ a (I n (indexI n t)) = 0 · simp only [h0, prod_empty, OuterMeasure.empty', ENNReal.zero_div] - sorry + have : ∀ m, (κ a) (I n (indexI n t) ×ˢ s m) = 0 := by + intro m + rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0 + refine measure_mono_null ?_ h0 + intro x hx + rw [mem_prod] at hx + exact hx.1 + simp [this] refine ENNReal.Tendsto.div_const ?_ ?_ · have h := tendsto_measure_iInter (μ := κ a) (s := fun m ↦ I n (indexI n t) ×ˢ s m) ?_ ?_ ?_ · convert h @@ -469,36 +526,78 @@ lemma tendsto_m_atTop_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKernel rw [← m_empty κ a n t] exact tendsto_m_atTop_empty_of_antitone κ a s hs hs_iInter hs_meas n t -lemma tendsto_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsMarkovKernel (kernel.fst κ)] - (s : Set β) : +lemma tendsto_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsMarkovKernel κ] + {s : Set β} (hs : MeasurableSet s) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun n ↦ M κ a s n t) atTop (𝓝 (ℱ.limitProcess (M κ a s) (kernel.fst κ a) t)) := by - refine Submartingale.ae_tendsto_limitProcess (martingale_m κ a s).submartingale (R := 1) ?_ + refine Submartingale.ae_tendsto_limitProcess (martingale_m κ a hs).submartingale (R := 1) ?_ intro n rw [ENNReal.coe_one] exact snorm_m_le_one κ a s n -lemma limitProcess_mem_L1 (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] - (a : α) (s : Set β) : +lemma limitProcess_mem_L1 (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] + (a : α) {s : Set β} (hs : MeasurableSet s) : Memℒp (ℱ.limitProcess (M κ a s) (kernel.fst κ a)) 1 (kernel.fst κ a) := by - refine Submartingale.memℒp_limitProcess (martingale_m κ a s).submartingale (R := 1) ?_ + refine Submartingale.memℒp_limitProcess (martingale_m κ a hs).submartingale (R := 1) ?_ intro n rw [ENNReal.coe_one] exact snorm_m_le_one κ a s n +lemma tendsto_snorm_one_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsMarkovKernel κ] + {s : Set β} (hs : MeasurableSet s) : + Tendsto + (fun n ↦ snorm (M κ a s n - ℱ.limitProcess (M κ a s) (kernel.fst κ a)) 1 (kernel.fst κ a)) + atTop (𝓝 0) := by + refine Submartingale.tendsto_snorm_one_limitProcess ?_ ?_ + · exact (martingale_m κ a hs).submartingale + · refine uniformIntegrable_of le_rfl ENNReal.one_ne_top ?_ ?_ + · exact fun n ↦ (measurable_m_right κ a hs n).aestronglyMeasurable + · intro ε _ + refine ⟨2, fun n ↦ ?_⟩ + refine le_of_eq_of_le ?_ (?_ : 0 ≤ ENNReal.ofReal ε) + · have : {x | 2 ≤ ‖M κ a s n x‖₊} = ∅ := by + ext x + simp only [mem_setOf_eq, mem_empty_iff_false, iff_false, not_le] + refine (?_ : _ ≤ (1 : ℝ≥0)).trans_lt one_lt_two + rw [Real.nnnorm_of_nonneg (m_nonneg _ _ _ _ _)] + exact mod_cast (m_le_one _ _ _ _ _) + rw [this] + simp + · simp + +lemma snorm_restrict_le [NormedAddCommGroup β] {p : ℝ≥0∞} {f : α → β} {μ : Measure α} (s : Set α) : + snorm f p (μ.restrict s) ≤ snorm f p μ := + snorm_mono_measure f Measure.restrict_le_self + +lemma tendsto_snorm_restrict_zero {α β ι : Type*} {mα : MeasurableSpace α} [NormedAddCommGroup β] + {p : ℝ≥0∞} {f : ι → α → β} {μ : Measure α} {l : Filter ι} + (h : Tendsto (fun n ↦ snorm (f n) p μ) l (𝓝 0)) (s : Set α) : + Tendsto (fun n ↦ snorm (f n) p (μ.restrict s)) l (𝓝 0) := by + refine tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds h ?_ ?_ + · exact fun _ ↦ zero_le _ + · exact fun _ ↦ snorm_restrict_le _ + +lemma tendsto_snorm_one_restrict_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsMarkovKernel κ] + {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : + Tendsto (fun n ↦ snorm (M κ a s n - ℱ.limitProcess (M κ a s) (kernel.fst κ a)) 1 + ((kernel.fst κ a).restrict A)) + atTop (𝓝 0) := + tendsto_snorm_restrict_zero (tendsto_snorm_one_m_limitProcess κ a hs) A + noncomputable def MLimsup (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (t : ℝ) : ℝ := limsup (fun n ↦ M κ a s n t) atTop -lemma mLimsup_ae_eq_mLim (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] - (a : α) (s : Set β) : +lemma mLimsup_ae_eq_limitProcess (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] + (a : α) {s : Set β} (hs : MeasurableSet s) : MLimsup κ a s =ᵐ[kernel.fst κ a] ℱ.limitProcess (M κ a s) (kernel.fst κ a) := by - filter_upwards [tendsto_m_limitProcess κ a s] with t ht using ht.limsup_eq + filter_upwards [tendsto_m_limitProcess κ a hs] with t ht using ht.limsup_eq -lemma tendsto_m_mLimsup (κ : kernel α (ℝ × β)) (a : α) [IsMarkovKernel (kernel.fst κ)] (s : Set β) : +lemma tendsto_m_mLimsup (κ : kernel α (ℝ × β)) (a : α) [IsMarkovKernel κ] + {s : Set β} (hs : MeasurableSet s) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun n ↦ M κ a s n t) atTop (𝓝 (MLimsup κ a s t)) := by - filter_upwards [tendsto_m_limitProcess κ a s, mLimsup_ae_eq_mLim κ a s] with t h1 h2 + filter_upwards [tendsto_m_limitProcess κ a hs, mLimsup_ae_eq_limitProcess κ a hs] with t h1 h2 rw [h2] exact h1 @@ -520,21 +619,37 @@ lemma mLimsup_mono_set (κ : kernel α (ℝ × β)) (a : α) {s s' : Set β} (h MLimsup κ a s t ≤ MLimsup κ a s' t := by rw [MLimsup, MLimsup] refine limsup_le_limsup ?_ ?_ ?_ - · exact Filter.eventually_of_forall (fun n ↦ m_mono_set κ a h n t) - · sorry - · sorry + · exact eventually_of_forall (fun n ↦ m_mono_set κ a h n t) + · -- todo: extract lemma (of find it) + refine ⟨0, ?_⟩ + simp only [eventually_map, eventually_atTop, ge_iff_le, forall_exists_index] + intro x n h + specialize h n le_rfl + exact (m_nonneg _ _ _ _ _).trans h + · -- todo: extract lemma (of find it) + refine ⟨1, ?_⟩ + simp only [eventually_map, eventually_atTop, ge_iff_le] + exact ⟨0, fun n _ ↦ m_le_one _ _ _ _ _⟩ lemma mLimsup_nonneg (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (t : ℝ) : 0 ≤ MLimsup κ a s t := by refine le_limsup_of_frequently_le ?_ ?_ - · exact Filter.frequently_of_forall (fun n ↦ m_nonneg _ _ _ _ _) - · sorry + · exact frequently_of_forall (fun n ↦ m_nonneg _ _ _ _ _) + · -- todo: extract lemma (of find it) + refine ⟨1, ?_⟩ + simp only [eventually_map, eventually_atTop, ge_iff_le] + exact ⟨0, fun n _ ↦ m_le_one _ _ _ _ _⟩ lemma mLimsup_le_one (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (t : ℝ) : MLimsup κ a s t ≤ 1 := by refine limsup_le_of_le ?_ ?_ - · sorry - · exact Filter.eventually_of_forall (fun n ↦ m_le_one _ _ _ _ _) + · -- todo: extract lemma (of find it) + refine ⟨0, ?_⟩ + simp only [eventually_map, eventually_atTop, ge_iff_le, forall_exists_index] + intro x n h + specialize h n le_rfl + exact (m_nonneg _ _ _ _ _).trans h + · exact eventually_of_forall (fun n ↦ m_le_one _ _ _ _ _) lemma mLimsup_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), MLimsup κ a Set.univ t = 1 := by @@ -545,19 +660,147 @@ lemma mLimsup_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) : simp_rw [ht] rw [limsup_const] -- should be simp +lemma snorm_mLimsup_le_one (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] + (a : α) (s : Set β) : + snorm (fun t ↦ MLimsup κ a s t) 1 (kernel.fst κ a) ≤ 1 := by + refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ + · simp only [Real.norm_eq_abs, abs_of_nonneg (mLimsup_nonneg κ a s t), + mLimsup_le_one κ a s t] + · simp + +lemma integrable_mLimsup (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] + (a : α) {s : Set β} (hs : MeasurableSet s) : + Integrable (fun t ↦ MLimsup κ a s t) (kernel.fst κ a) := by + rw [← memℒp_one_iff_integrable] + refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ + · exact measurable_mLimsup_right κ hs a + · exact (snorm_mLimsup_le_one κ a s).trans_lt ENNReal.one_lt_top + +lemma tendsto_integral_of_L1' {ι G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] + {μ : Measure α} + (f : α → G) (hfi : Integrable f μ) + {F : ι → α → G} {l : Filter ι} + (hFi : ∀ᶠ i in l, Integrable (F i) μ) + (hF : Tendsto (fun i ↦ snorm (F i - f) 1 μ) l (𝓝 0)) : + Tendsto (fun i ↦ ∫ x, F i x ∂μ) l (𝓝 (∫ x, f x ∂μ)) := by + refine tendsto_integral_of_L1 f hfi hFi ?_ + simp_rw [snorm_one_eq_lintegral_nnnorm, Pi.sub_apply] at hF + exact hF + +lemma tendsto_set_integral_of_L1 {ι G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] + {μ : Measure α} + (f : α → G) (hfi : Integrable f μ) + {F : ι → α → G} {l : Filter ι} + (hFi : ∀ᶠ i in l, Integrable (F i) μ) + (hF : Tendsto (fun i ↦ ∫⁻ x, ‖F i x - f x‖₊ ∂μ) l (𝓝 0)) (s : Set α) : + Tendsto (fun i ↦ ∫ x in s, F i x ∂μ) l (𝓝 (∫ x in s, f x ∂μ)) := by + refine tendsto_integral_of_L1 f hfi.restrict ?_ ?_ + · filter_upwards [hFi] with i hi using hi.restrict + · simp_rw [← snorm_one_eq_lintegral_nnnorm] at hF ⊢ + exact tendsto_snorm_restrict_zero hF s + +lemma tendsto_set_integral_of_L1' {ι G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] + {μ : Measure α} + (f : α → G) (hfi : Integrable f μ) + {F : ι → α → G} {l : Filter ι} + (hFi : ∀ᶠ i in l, Integrable (F i) μ) + (hF : Tendsto (fun i ↦ snorm (F i - f) 1 μ) l (𝓝 0)) (s : Set α) : + Tendsto (fun i ↦ ∫ x in s, F i x ∂μ) l (𝓝 (∫ x in s, f x ∂μ)) := by + refine tendsto_set_integral_of_L1 f hfi hFi ?_ s + simp_rw [snorm_one_eq_lintegral_nnnorm, Pi.sub_apply] at hF + exact hF + +lemma tendsto_set_integral_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] + (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : + Tendsto (fun i ↦ ∫ x in A, M κ a s i x ∂(kernel.fst κ) a) atTop + (𝓝 (∫ x in A, MLimsup κ a s x ∂(kernel.fst κ) a)) := by + refine tendsto_set_integral_of_L1' (μ := kernel.fst κ a) (MLimsup κ a s) + (integrable_mLimsup κ a hs) (F := fun i t ↦ M κ a s i t) (l := atTop) + (eventually_of_forall (fun n ↦ integrable_m _ _ hs _)) ?_ A + refine (tendsto_congr fun n ↦ ?_).mp (tendsto_snorm_one_m_limitProcess κ a hs) + refine snorm_congr_ae ?_ + refine EventuallyEq.sub EventuallyEq.rfl ?_ + exact (mLimsup_ae_eq_limitProcess κ a hs).symm + +lemma set_integral_mLimsup_of_measurableSet (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] + (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) + {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : + ∫ t in A, MLimsup κ a s t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal := by + suffices ∫ t in A, MLimsup κ a s t ∂(kernel.fst κ a) = ∫ t in A, M κ a s n t ∂(kernel.fst κ a) by + rw [this] + exact set_integral_m _ _ hs _ hA + suffices ∫ t in A, MLimsup κ a s t ∂(kernel.fst κ a) + = limsup (fun i ↦ ∫ t in A, M κ a s i t ∂(kernel.fst κ a)) atTop by + rw [this, ← limsup_const (α := ℕ) (f := atTop) (∫ t in A, M κ a s n t ∂(kernel.fst κ a)), + limsup_congr] + simp only [eventually_atTop, ge_iff_le] + refine ⟨n, fun m hnm ↦ ?_⟩ + rw [set_integral_m_of_le _ _ hs hnm hA, set_integral_m _ _ hs _ hA] + -- use L1 convergence + have h := tendsto_set_integral_m κ a hs A + rw [h.limsup_eq] + +lemma integral_mLimsup (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] + (a : α) {s : Set β} (hs : MeasurableSet s) : + ∫ t, MLimsup κ a s t ∂(kernel.fst κ a) = (κ a (univ ×ˢ s)).toReal := by + rw [← integral_univ, set_integral_mLimsup_of_measurableSet κ a hs 0 MeasurableSet.univ] + +lemma set_integral_mLimsup (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] + (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set ℝ} (hA : MeasurableSet A) : + ∫ t in A, MLimsup κ a s t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal := by + have hA' : MeasurableSet[⨆ n, ℱ n] A := by rwa [iSup_ℱ] + refine MeasurableSpace.induction_on_inter (m := ⨆ n, ℱ n) + (C := fun A ↦ ∫ t in A, MLimsup κ a s t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal) + (MeasurableSpace.measurableSpace_iSup_eq ℱ) ?_ ?_ ?_ ?_ ?_ hA' + · rintro s ⟨n, hs⟩ t ⟨m, ht⟩ _ + exact ⟨max n m, (ℱ.mono (le_max_left n m) _ hs).inter (ℱ.mono (le_max_right n m) _ ht)⟩ + · simp + · intro A ⟨n, hA⟩ + exact set_integral_mLimsup_of_measurableSet κ a hs n hA + · intro A hA hA_eq + rw [iSup_ℱ] at hA + have h := integral_add_compl hA (integrable_mLimsup κ a hs) + rw [hA_eq, integral_mLimsup κ a hs] at h + have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by + rw [prod_diff_prod, compl_eq_univ_diff] + simp + rw [this, measure_diff (by intro x; simp) (hA.prod hs) (measure_ne_top (κ a) _), + ENNReal.toReal_sub_of_le (measure_mono (by intro x; simp)) (measure_ne_top _ _)] + rw [eq_tsub_iff_add_eq_of_le, add_comm] + · exact h + · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] + exact measure_mono (by intro x; simp) + · intro f hf_disj hf h_eq + rw [integral_iUnion _ hf_disj (integrable_mLimsup _ _ hs).integrableOn] + · simp_rw [h_eq] + rw [← ENNReal.tsum_toReal_eq (fun _ ↦ measure_ne_top _ _)] + congr + rw [iUnion_prod_const, measure_iUnion] + · intro i j hij + rw [Function.onFun, Set.disjoint_prod] + exact Or.inl (hf_disj hij) + · rw [iSup_ℱ] at hf + exact fun i ↦ (hf i).prod hs + · rwa [iSup_ℱ] at hf + lemma tendsto_mLimsup_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] - (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) : + (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) + (hs_meas : ∀ n, MeasurableSet (s n)) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 1) := by have h1 := tendsto_m_atTop_ae_of_monotone κ a s hs hs_iUnion - have h2 := fun (n : ℕ) ↦ tendsto_m_mLimsup κ a (s n) + have h2 := fun (n : ℕ) ↦ tendsto_m_mLimsup κ a (hs_meas n) rw [← ae_all_iff] at h1 h2 filter_upwards [h1, h2] with t h_tendsto_set h_tendsto_nat sorry -lemma tendsto_mLimsup_atTop_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma tendsto_mLimsup_atTop_ae_of_antitone (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) - (hs_meas : ∀ n, MeasurableSet (s n)) (n : ℕ) (t : ℝ) : - Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 0) := by + (hs_meas : ∀ n, MeasurableSet (s n)) : + ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 0) := by + have h1 := tendsto_m_atTop_of_antitone κ a s hs hs_iInter hs_meas + have h2 := fun (n : ℕ) ↦ tendsto_m_mLimsup κ a (hs_meas n) + rw [← ae_all_iff] at h2 + filter_upwards [h2] with t h_tendsto_nat sorry section Iic_Q @@ -566,8 +809,12 @@ noncomputable def mLimsupIic (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : ℝ := MLimsup κ a (Set.Iic q) t lemma measurable_mLimsupIic (κ : kernel α (ℝ × ℝ)) (q : ℚ) : - Measurable (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2 q) := by - sorry + Measurable (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2 q) := + measurable_mLimsup κ measurableSet_Iic + +lemma measurable_mLimsupIic_right (κ : kernel α (ℝ × ℝ)) (a : α) (q : ℚ) : + Measurable (fun t ↦ mLimsupIic κ a t q) := + measurable_mLimsup_right _ measurableSet_Iic _ lemma monotone_mLimsupIic (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) : Monotone (mLimsupIic κ a t) := by intro q r hqr @@ -582,96 +829,342 @@ lemma mLimsupIic_nonneg (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : lemma mLimsupIic_le_one (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : mLimsupIic κ a t q ≤ 1 := mLimsup_le_one κ a _ t -lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : +lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ a t q) atTop (𝓝 1) := by - sorry - -lemma tendsto_atBot_mLimsupIic (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) : - Tendsto (fun q ↦ mLimsupIic κ a t q) atBot (𝓝 0) := by - sorry + suffices ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ a t n) atTop (𝓝 1) by + sorry + let s : ℕ → Set ℝ := fun n ↦ Iic n + have hs : Monotone s := fun i j hij ↦ Iic_subset_Iic.mpr (by exact mod_cast hij) + have hs_iUnion : ⋃ i, s i = univ := by + ext x + simp only [mem_iUnion, mem_Iic, mem_univ, iff_true] + exact ⟨Nat.ceil x, Nat.le_ceil x⟩ + have hs_meas : ∀ n, MeasurableSet (s n) := fun _ ↦ measurableSet_Iic + filter_upwards [tendsto_mLimsup_atTop_ae_of_monotone κ a s hs hs_iUnion hs_meas] + with x hx using hx + +lemma tendsto_atBot_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) : + ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ a t q) atBot (𝓝 0) := by + suffices ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ a t (-q)) atTop (𝓝 0) by + filter_upwards [this] with t ht + have h_eq_neg : (fun q ↦ mLimsupIic κ a t q) = fun q ↦ mLimsupIic κ a t (- -q) := by + simp_rw [neg_neg] + rw [h_eq_neg] + exact ht.comp tendsto_neg_atBot_atTop + suffices ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ a t (-n)) atTop (𝓝 0) by + sorry + let s : ℕ → Set ℝ := fun n ↦ Iic (-n) + have hs : Antitone s := fun i j hij ↦ Iic_subset_Iic.mpr (neg_le_neg (by exact mod_cast hij)) + have hs_iInter : ⋂ i, s i = ∅ := by + ext x + simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, neg_lt] + exact exists_nat_gt (-x) + have hs_meas : ∀ n, MeasurableSet (s n) := fun _ ↦ measurableSet_Iic + convert tendsto_mLimsup_atTop_ae_of_antitone κ a s hs hs_iInter hs_meas with x n + rw [mLimsupIic] + simp -lemma iInf_rat_gt_mLimsupIic_eq (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : - ⨅ r : Ioi q, mLimsupIic κ a t r = mLimsupIic κ a t q := by - sorry +lemma set_integral_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] + (a : α) (q : ℚ) {A : Set ℝ} (hA : MeasurableSet A) : + ∫ t in A, mLimsupIic κ a t q ∂(kernel.fst κ a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := + set_integral_mLimsup κ a measurableSet_Iic hA + +lemma integrable_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel (kernel.fst κ)] + (a : α) (q : ℚ) : + Integrable (fun t ↦ mLimsupIic κ a t q) (kernel.fst κ a) := + integrable_mLimsup _ _ measurableSet_Iic + +lemma bddBelow_range_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel (kernel.fst κ)] + (a : α) (t : ℝ) (q : ℚ) : + BddBelow (range fun (r : Ioi q) ↦ mLimsupIic κ a t r) := by + refine ⟨0, ?_⟩ + rw [mem_lowerBounds] + rintro x ⟨y, rfl⟩ + exact mLimsupIic_nonneg _ _ _ _ + +lemma integrable_iInf_rat_gt_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) (q : ℚ) : + Integrable (fun t ↦ ⨅ r : Ioi q, mLimsupIic κ a t r) (kernel.fst κ a) := by + rw [← memℒp_one_iff_integrable] + refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ + · exact measurable_iInf fun i ↦ measurable_mLimsupIic_right κ a i + refine (?_ : _ ≤ (1 : ℝ≥0∞)).trans_lt ENNReal.one_lt_top + refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ + · rw [Real.norm_eq_abs, abs_of_nonneg] + · refine ciInf_le_of_le ?_ ?_ ?_ + · exact bddBelow_range_mLimsupIic κ a t _ + · exact ⟨q + 1, by simp⟩ + · exact mLimsupIic_le_one _ _ _ _ + · exact le_ciInf fun r ↦ mLimsupIic_nonneg κ a t r + · simp + +lemma set_integral_iInf_rat_gt_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] + (a : α) (q : ℚ) {A : Set ℝ} (hA : MeasurableSet A) : + ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ a t r ∂(kernel.fst κ a) + = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by + refine le_antisymm ?_ ?_ + · have h : ∀ r : Ioi q, ∫ t in A, ⨅ r' : Ioi q, mLimsupIic κ a t r' ∂(kernel.fst κ a) + ≤ (κ a (A ×ˢ Iic (r : ℝ))).toReal := by + intro r + rw [← set_integral_mLimsupIic κ a r hA] + refine set_integral_mono ?_ ?_ ?_ + · exact (integrable_iInf_rat_gt_mLimsupIic _ _ _).integrableOn + · exact (integrable_mLimsupIic _ _ _).integrableOn + · exact fun t ↦ ciInf_le (bddBelow_range_mLimsupIic _ _ _ _) r + calc ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ a t r ∂(kernel.fst κ a) + ≤ ⨅ r : Ioi q, (κ a (A ×ˢ Iic (r : ℝ))).toReal := le_ciInf h + _ = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by + rw [← Measure.iInf_Iic_gt_prod hA q] + exact (ENNReal.toReal_iInf (fun r ↦ measure_ne_top _ _)).symm + · rw [← set_integral_mLimsupIic κ a q hA] + refine set_integral_mono ?_ ?_ ?_ + · exact (integrable_mLimsupIic _ _ _).integrableOn + · exact (integrable_iInf_rat_gt_mLimsupIic _ _ _).integrableOn + · exact fun t ↦ le_ciInf (fun r ↦ monotone_mLimsupIic _ _ _ (le_of_lt r.prop)) + +lemma iInf_rat_gt_mLimsupIic_eq (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) : + ∀ᵐ t ∂(kernel.fst κ a), ∀ q : ℚ, ⨅ r : Ioi q, mLimsupIic κ a t r = mLimsupIic κ a t q := by + rw [ae_all_iff] + refine fun q ↦ ae_eq_of_forall_set_integral_eq_of_sigmaFinite (μ := kernel.fst κ a) ?_ ?_ ?_ + · intro s _ _ + refine Integrable.integrableOn ?_ + exact integrable_iInf_rat_gt_mLimsupIic κ _ _ + · exact fun s _ _ ↦ (integrable_mLimsupIic κ a q).integrableOn + · intro s hs _ + rw [set_integral_mLimsupIic _ _ _ hs, set_integral_iInf_rat_gt_mLimsupIic _ _ _ hs] end Iic_Q section Rat -lemma measurableSet_tendstoAtTopSet (κ : kernel α (ℝ × ℝ)) : - MeasurableSet {p : α × ℝ | Tendsto (fun q ↦ mLimsupIic κ p.1 p.2 q) atTop (𝓝 1)} := - measurableSet_tendsto_nhds (fun q ↦ measurable_mLimsupIic κ q) 1 - -open Classical in -noncomputable -def todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : ℚ → ℝ := - if Tendsto (fun q ↦ mLimsupIic κ p.1 p.2 q) atTop (𝓝 1) - then fun q ↦ mLimsupIic κ p.1 p.2 q - else defaultRatCDF - -lemma measurable_todo2' (κ : kernel α (ℝ × ℝ)) (q : ℚ) : - Measurable (fun p ↦ todo2' κ p q) := by - classical - simp only [todo2', ite_apply] - exact Measurable.ite (measurableSet_tendstoAtTopSet κ) (measurable_mLimsupIic κ q) - measurable_const - -lemma monotone_todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : - Monotone (todo2' κ p) := by - unfold todo2' - split_ifs with h - · exact monotone_mLimsupIic κ p.1 p.2 - · exact monotone_defaultRatCDF - -lemma todo2'_nonneg (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : - 0 ≤ todo2' κ p := by - unfold todo2' - split_ifs with h - · exact mLimsupIic_nonneg κ p.1 p.2 - · exact defaultRatCDF_nonneg - -lemma todo2'_le_one (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : - todo2' κ p ≤ 1 := by - unfold todo2' - split_ifs with h - · exact mLimsupIic_le_one κ p.1 p.2 - · exact defaultRatCDF_le_one - -lemma tendsto_atTop_todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : - Tendsto (todo2' κ p) atTop (𝓝 1) := by - unfold todo2' - split_ifs with h - · exact h - · exact tendsto_defaultRatCDF_atTop - -lemma tendsto_atBot_todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) : - Tendsto (todo2' κ p) atBot (𝓝 0) := by - unfold todo2' - split_ifs with h - · exact tendsto_atBot_mLimsupIic κ p.1 p.2 - · exact tendsto_defaultRatCDF_atBot - -theorem inf_gt_todo2' (κ : kernel α (ℝ × ℝ)) (p : α × ℝ) (t : ℚ) : - ⨅ r : Ioi t, todo2' κ p r = todo2' κ p t := by - unfold todo2' - split_ifs with hp - · simp_rw [iInf_rat_gt_mLimsupIic_eq] - · exact inf_gt_rat_defaultRatCDF t - -lemma isCDFLike_todo2' (κ : kernel α (ℝ × ℝ)) : IsCDFLike (todo2' κ) where - mono := monotone_todo2' κ - nonneg := todo2'_nonneg κ - le_one := todo2'_le_one κ - tendsto_atTop_one := tendsto_atTop_todo2' κ - tendsto_atBot_zero := tendsto_atBot_todo2' κ - iInf_rat_gt_eq := inf_gt_todo2' κ - measurable := measurable_todo2' κ +lemma isRatStieltjesPoint_mLimsupIic_ae (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) : + ∀ᵐ t ∂(kernel.fst κ a), IsRatStieltjesPoint (fun p q ↦ mLimsupIic κ p.1 p.2 q) (a, t) := by + filter_upwards [tendsto_atTop_mLimsupIic κ a, tendsto_atBot_mLimsupIic κ a, + iInf_rat_gt_mLimsupIic_eq κ a] with t ht_top ht_bot ht_iInf + constructor + · exact monotone_mLimsupIic κ a t + · exact mLimsupIic_nonneg κ a t + · exact mLimsupIic_le_one κ a t + · exact ht_top + · exact ht_bot + · exact ht_iInf + +lemma todo3_mLimsupIic_ae_eq (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) (q : ℚ) : + (fun t ↦ todo3 _ (measurable_mLimsupIic κ) (a, t) q) + =ᵐ[kernel.fst κ a] fun t ↦ mLimsupIic κ a t q := by + filter_upwards [isRatStieltjesPoint_mLimsupIic_ae κ a] with a ha + rw [todo3_eq, toCDFLike_of_isRatStieltjesPoint ha] end Rat +-- todo: name? noncomputable def kernel.condexpReal (κ : kernel α (ℝ × ℝ)) : kernel (α × ℝ) ℝ := - (isCDFLike_todo2' κ).kernel + cdfKernel (measurable_mLimsupIic κ) + +instance (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : IsMarkovKernel (kernel.condexpReal κ) := by + unfold kernel.condexpReal; infer_instance + +lemma condexpReal_Iic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) (t x : ℝ) : + kernel.condexpReal κ (a, t) (Iic x) + = ENNReal.ofReal (todo3 _ (measurable_mLimsupIic κ) (a, t) x) := + cdfKernel_Iic _ _ + +lemma set_lintegral_condexpReal_Iic_rat (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) (q : ℚ) + {s : Set ℝ} (hs : MeasurableSet s) : + ∫⁻ t in s, kernel.condexpReal κ (a, t) (Iic q) ∂(kernel.fst κ a) = κ a (s ×ˢ Iic (q : ℝ)) := by + simp_rw [condexpReal_Iic] + rw [← ofReal_integral_eq_lintegral_ofReal] + · rw [set_integral_congr_ae (g := fun t ↦ mLimsupIic κ a t q) hs, + set_integral_mLimsupIic κ _ _ hs, ENNReal.ofReal_toReal] + · exact measure_ne_top _ _ + · filter_upwards [todo3_mLimsupIic_ae_eq κ a q] with t ht + exact fun _ ↦ ht + · refine Integrable.restrict ?_ + rw [integrable_congr (todo3_mLimsupIic_ae_eq κ a q)] + exact integrable_mLimsupIic _ _ _ + · exact ae_of_all _ (fun x ↦ todo3_nonneg _ _ _) + +lemma set_lintegral_condexpReal_Iic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) (x : ℝ) + {s : Set ℝ} (hs : MeasurableSet s) : + ∫⁻ t in s, kernel.condexpReal κ (a, t) (Iic x) ∂(kernel.fst κ a) = κ a (s ×ˢ Iic x) := by + -- We have the result for `x : ℚ` thanks to `set_lintegral_condexpReal_Iic_rat`. + -- We use the equality `condCDF ρ a x = ⨅ r : {r' : ℚ // x < r'}, condCDF ρ a r` and a monotone + -- convergence argument to extend it to the reals. + by_cases hρ_zero : (kernel.fst κ a).restrict s = 0 + · rw [hρ_zero, lintegral_zero_measure] + refine le_antisymm (zero_le _) ?_ + calc κ a (s ×ˢ Iic x) + ≤ κ a (Prod.fst ⁻¹' s) := measure_mono (prod_subset_preimage_fst s (Iic x)) + _ = kernel.fst κ a s := by rw [kernel.fst_apply' _ _ hs]; rfl + _ = (kernel.fst κ a).restrict s univ := by rw [Measure.restrict_apply_univ] + _ = 0 := by simp only [hρ_zero, Measure.coe_zero, Pi.zero_apply] + have h : ∫⁻ t in s, kernel.condexpReal κ (a, t) (Iic x) ∂(kernel.fst κ a) + = ∫⁻ t in s, ⨅ r : { r' : ℚ // x < r' }, + kernel.condexpReal κ (a, t) (Iic r) ∂(kernel.fst κ a) := by + congr with t : 1 + rw [← measure_iInter_eq_iInf] + · congr with y : 1 + simp only [mem_Iic, mem_iInter, Subtype.forall] + refine ⟨fun h a ha ↦ h.trans ?_, fun h ↦ ?_⟩ + · exact mod_cast ha.le + · refine le_of_forall_lt_rat_imp_le fun q hq ↦ h q ?_ + exact mod_cast hq + · exact fun _ ↦ measurableSet_Iic + · refine Monotone.directed_ge fun r r' hrr' ↦ ?_ + refine Iic_subset_Iic.mpr ?_ + exact mod_cast hrr' + · obtain ⟨q, hq⟩ := exists_rat_gt x + exact ⟨⟨q, hq⟩, measure_ne_top _ _⟩ + have h_nonempty : Nonempty { r' : ℚ // x < ↑r' } := by + obtain ⟨r, hrx⟩ := exists_rat_gt x + exact ⟨⟨r, hrx⟩⟩ + rw [h, lintegral_iInf_directed_of_measurable hρ_zero fun q : { r' : ℚ // x < ↑r' } ↦ ?_] + rotate_left + · intro b + rw [set_lintegral_condexpReal_Iic_rat _ _ _ hs] + exact measure_ne_top _ _ + · refine Monotone.directed_ge fun i j hij t ↦ measure_mono (Iic_subset_Iic.mpr ?_) + exact mod_cast hij + · exact (kernel.measurable_coe _ measurableSet_Iic).comp measurable_prod_mk_left + simp_rw [set_lintegral_condexpReal_Iic_rat κ _ _ hs] + rw [← measure_iInter_eq_iInf] + · rw [← prod_iInter] + congr with y + simp only [mem_iInter, mem_Iic, Subtype.forall, Subtype.coe_mk] + exact ⟨le_of_forall_lt_rat_imp_le, fun hyx q hq ↦ hyx.trans hq.le⟩ + · exact fun i ↦ hs.prod measurableSet_Iic + · refine Monotone.directed_ge fun i j hij ↦ ?_ + refine prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, Iic_subset_Iic.mpr ?_⟩) + exact mod_cast hij + · exact ⟨h_nonempty.some, measure_ne_top _ _⟩ + +lemma set_lintegral_condexpReal_univ (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) + {s : Set ℝ} (hs : MeasurableSet s) : + ∫⁻ t in s, kernel.condexpReal κ (a, t) univ ∂(kernel.fst κ a) = κ a (s ×ˢ univ) := by + simp only [measure_univ, lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter, + one_mul, kernel.fst_apply' _ _ hs] + congr with x + simp only [mem_prod, mem_univ, and_true, mem_setOf_eq] + +lemma lintegral_condexpReal_univ (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) : + ∫⁻ t, kernel.condexpReal κ (a, t) univ ∂(kernel.fst κ a) = κ a univ := by + rw [← set_lintegral_univ, set_lintegral_condexpReal_univ κ a MeasurableSet.univ, + univ_prod_univ] + +lemma set_lintegral_condexpReal_prod (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) + {s t : Set ℝ} (hs : MeasurableSet s) (ht : MeasurableSet t) : + ∫⁻ x in s, kernel.condexpReal κ (a, x) t ∂(kernel.fst κ a) = κ a (s ×ˢ t) := by + -- `set_lintegral_condKernelReal_Iic` gives the result for `t = Iic x`. These sets form a + -- π-system that generates the Borel σ-algebra, hence we can get the same equality for any + -- measurable set `t`. + apply MeasurableSpace.induction_on_inter (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic _ _ _ _ ht + · simp only [measure_empty, lintegral_const, zero_mul, prod_empty] + · rintro t ⟨q, rfl⟩ + exact set_lintegral_condexpReal_Iic κ a _ hs + · intro t ht ht_lintegral + calc ∫⁻ x in s, kernel.condexpReal κ (a, x) tᶜ ∂(kernel.fst κ a) + = ∫⁻ x in s, kernel.condexpReal κ (a, x) univ - kernel.condexpReal κ (a, x) t ∂(kernel.fst κ a) := by + congr with x; rw [measure_compl ht (measure_ne_top (kernel.condexpReal κ (a, x)) _)] + _ = ∫⁻ x in s, kernel.condexpReal κ (a, x) univ ∂(kernel.fst κ a) + - ∫⁻ x in s, kernel.condexpReal κ (a, x) t ∂(kernel.fst κ a) := by + rw [lintegral_sub] + · exact (kernel.measurable_coe (kernel.condexpReal κ) ht).comp measurable_prod_mk_left + · rw [ht_lintegral] + exact measure_ne_top _ _ + · exact eventually_of_forall fun a ↦ measure_mono (subset_univ _) + _ = κ a (s ×ˢ univ) - κ a (s ×ˢ t) := by + rw [set_lintegral_condexpReal_univ κ a hs, ht_lintegral] + _ = κ a (s ×ˢ tᶜ) := by + rw [← measure_diff _ (hs.prod ht) (measure_ne_top _ _)] + · rw [prod_diff_prod, compl_eq_univ_diff] + simp only [diff_self, empty_prod, union_empty] + · rw [prod_subset_prod_iff] + exact Or.inl ⟨subset_rfl, subset_univ t⟩ + · intro f hf_disj hf_meas hf_eq + simp_rw [measure_iUnion hf_disj hf_meas] + rw [lintegral_tsum, prod_iUnion, measure_iUnion] + · simp_rw [hf_eq] + · intro i j hij + rw [Function.onFun, Set.disjoint_prod] + exact Or.inr (hf_disj hij) + · exact fun i ↦ MeasurableSet.prod hs (hf_meas i) + · exact fun i ↦ + ((kernel.measurable_coe _ (hf_meas i)).comp measurable_prod_mk_left).aemeasurable.restrict + +lemma lintegral_condexpReal_mem (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) + {s : Set (ℝ × ℝ)} (hs : MeasurableSet s) : + ∫⁻ x, kernel.condexpReal κ (a, x) {y | (x, y) ∈ s} ∂(kernel.fst κ a) = κ a s := by + -- `set_lintegral_condKernelReal_prod` gives the result for sets of the form `t₁ × t₂`. These + -- sets form a π-system that generates the product σ-algebra, hence we can get the same equality + -- for any measurable set `s`. + apply MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod _ _ _ _ hs + · simp only [mem_empty_iff_false, setOf_false, measure_empty, lintegral_const, + zero_mul] + · rintro _ ⟨t₁, ht₁, t₂, ht₂, rfl⟩ + simp only [mem_setOf_eq] at ht₁ ht₂ + have h_prod_eq_snd : ∀ a ∈ t₁, {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} = t₂ := by + intro a ha + simp only [ha, prod_mk_mem_set_prod_eq, true_and_iff, setOf_mem_eq] + rw [← lintegral_add_compl _ ht₁] + have h_eq1 : ∫⁻ x in t₁, kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(kernel.fst κ a) + = ∫⁻ x in t₁, kernel.condexpReal κ (a, x) t₂ ∂(kernel.fst κ a) := by + refine' set_lintegral_congr_fun ht₁ (eventually_of_forall fun a ha ↦ _) + rw [h_prod_eq_snd a ha] + have h_eq2 : ∫⁻ x in t₁ᶜ, kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(kernel.fst κ a) = 0 := by + suffices h_eq_zero : ∀ x ∈ t₁ᶜ, kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} = 0 + · rw [set_lintegral_congr_fun ht₁.compl (eventually_of_forall h_eq_zero)] + simp only [lintegral_const, zero_mul] + intro a hat₁ + rw [mem_compl_iff] at hat₁ + simp only [hat₁, prod_mk_mem_set_prod_eq, false_and_iff, setOf_false, measure_empty] + rw [h_eq1, h_eq2, add_zero] + exact set_lintegral_condexpReal_prod κ a ht₁ ht₂ + · intro t ht ht_eq + calc ∫⁻ x, kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ tᶜ} ∂(kernel.fst κ a) + = ∫⁻ x, kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t}ᶜ ∂(kernel.fst κ a) := rfl + _ = ∫⁻ x, kernel.condexpReal κ (a, x) univ + - kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t} ∂(kernel.fst κ a) := by + congr with x : 1 + exact measure_compl (measurable_prod_mk_left ht) + (measure_ne_top (kernel.condexpReal κ (a, x)) _) + _ = ∫⁻ x, kernel.condexpReal κ (a, x) univ ∂(kernel.fst κ a) - + ∫⁻ x, kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t} ∂(kernel.fst κ a) := by + have h_le : (fun x ↦ kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t}) + ≤ᵐ[kernel.fst κ a] fun x ↦ kernel.condexpReal κ (a, x) univ := + eventually_of_forall fun _ ↦ measure_mono (subset_univ _) + rw [lintegral_sub _ _ h_le] + · exact kernel.measurable_kernel_prod_mk_left' ht a + refine ((lintegral_mono_ae h_le).trans_lt ?_).ne + rw [lintegral_condexpReal_univ] + exact measure_lt_top _ univ + _ = κ a univ - κ a t := by rw [ht_eq, lintegral_condexpReal_univ] + _ = κ a tᶜ := (measure_compl ht (measure_ne_top _ _)).symm + · intro f hf_disj hf_meas hf_eq + have h_eq : ∀ a, {x | (a, x) ∈ ⋃ i, f i} = ⋃ i, {x | (a, x) ∈ f i} := by + intro a; ext x; simp only [mem_iUnion, mem_setOf_eq] + simp_rw [h_eq] + have h_disj : ∀ a, Pairwise (Disjoint on fun i ↦ {x | (a, x) ∈ f i}) := by + intro a i j hij + have h_disj := hf_disj hij + rw [Function.onFun, disjoint_iff_inter_eq_empty] at h_disj ⊢ + ext1 x + simp only [mem_inter_iff, mem_setOf_eq, mem_empty_iff_false, iff_false_iff] + intro h_mem_both + suffices (a, x) ∈ ∅ by rwa [mem_empty_iff_false] at this + rwa [← h_disj, mem_inter_iff] + calc ∫⁻ x, kernel.condexpReal κ (a, x) (⋃ i, {y | (x, y) ∈ f i}) ∂(kernel.fst κ a) + = ∫⁻ x, ∑' i, kernel.condexpReal κ (a, x) {y | (x, y) ∈ f i} ∂(kernel.fst κ a) := by + congr with x : 1 + rw [measure_iUnion (h_disj x) fun i ↦ measurable_prod_mk_left (hf_meas i)] + _ = ∑' i, ∫⁻ x, kernel.condexpReal κ (a, x) {y | (x, y) ∈ f i} ∂(kernel.fst κ a) := + lintegral_tsum fun i ↦ (kernel.measurable_kernel_prod_mk_left' (hf_meas i) a).aemeasurable + _ = ∑' i, κ a (f i) := by simp_rw [hf_eq] + _ = κ a (iUnion f) := (measure_iUnion hf_disj hf_meas).symm + +lemma kernel.eq_compProd_condexpReal (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : + κ = kernel.fst κ ⊗ₖ kernel.condexpReal κ := by + ext a s hs + rw [kernel.compProd_apply _ _ _ hs, lintegral_condexpReal_mem κ a hs] end Real @@ -682,7 +1175,7 @@ def kernel.condexp (κ : kernel α (Ω × Ω')) [IsMarkovKernel (kernel.fst κ)] sorry theorem kernel.eq_compProd (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : - κ = kernel.fst κ ⊗ₖ (kernel.condexp κ) := by + κ = kernel.fst κ ⊗ₖ kernel.condexp κ := by sorry end ProbabilityTheory From 2c9e4493dcf66062e57610a51a09b82cda195252 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 12 Feb 2024 15:13:45 +0100 Subject: [PATCH 006/129] almost done --- Mathlib/Probability/Kernel/BuildKernel.lean | 293 +++++++- Mathlib/Probability/Kernel/CondCdf.lean | 101 ++- .../Probability/Kernel/Disintegration.lean | 141 +--- Mathlib/Probability/Kernel/StieltjesReal.lean | 74 +- Mathlib/Probability/Kernel/draft.lean | 708 ++++++++++++------ 5 files changed, 871 insertions(+), 446 deletions(-) diff --git a/Mathlib/Probability/Kernel/BuildKernel.lean b/Mathlib/Probability/Kernel/BuildKernel.lean index 1f2c6fd17c371..1c1714211b740 100644 --- a/Mathlib/Probability/Kernel/BuildKernel.lean +++ b/Mathlib/Probability/Kernel/BuildKernel.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.StieltjesReal -import Mathlib.Probability.Kernel.Basic +import Mathlib.Probability.Kernel.MeasureCompProd /-! @@ -14,21 +14,302 @@ import Mathlib.Probability.Kernel.Basic open MeasureTheory Set Filter TopologicalSpace -open scoped NNReal ENNReal MeasureTheory Topology +open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory -variable {α : Type*} [MeasurableSpace α] {f : α → ℚ → ℝ} {hf : ∀ q, Measurable fun a ↦ f a q} +variable {α β : Type*} [MeasurableSpace α] + +section kernel + +variable {f : α → ℚ → ℝ} {hf : ∀ q, Measurable fun a ↦ f a q} noncomputable -def cdfKernel (hf : ∀ q, Measurable fun a ↦ f a q) : kernel α ℝ where +def cdfKernel (f : α → ℚ → ℝ) (hf : ∀ q, Measurable fun a ↦ f a q) : kernel α ℝ where val a := (todo3 f hf a).measure property := measurable_measure_todo3 hf -instance instIsMarkovKernel_cdfKernel : IsMarkovKernel (cdfKernel hf) := +instance instIsMarkovKernel_cdfKernel : IsMarkovKernel (cdfKernel f hf) := ⟨fun _ ↦ instIsProbabilityMeasure_todo3 _ _⟩ lemma cdfKernel_Iic (a : α) (x : ℝ) : - cdfKernel hf a (Iic x) = ENNReal.ofReal (todo3 f hf a x) := measure_todo3_Iic hf a x + cdfKernel f hf a (Iic x) = ENNReal.ofReal (todo3 f hf a x) := measure_todo3_Iic hf a x + +lemma cdfKernel_unit_prod_Iic (a : α) (x : ℝ) : + cdfKernel (fun p : Unit × α ↦ f p.2) (fun q ↦ (hf q).comp measurable_snd) ((), a) (Iic x) + = cdfKernel f hf a (Iic x) := by + simp only [cdfKernel_Iic] + rw [todo3_unit_prod hf a] + +end kernel + +section + +variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} + {f : α × β → ℚ → ℝ} {μ : kernel α (β × ℝ)} {ν : kernel α β} + +structure IsRatKernelCDF (f : α × β → ℚ → ℝ) (μ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := + (measurable (q : ℚ) : Measurable fun p ↦ f p q) + (isRatStieltjesPoint_ae (a : α) : ∀ᵐ t ∂(ν a), IsRatStieltjesPoint f (a, t)) + (integrable (a : α) (q : ℚ) : Integrable (fun t ↦ f (a, t) q) (ν a)) + (isCDF (a : α) {s : Set β} (_hs : MeasurableSet s) (q : ℚ) : + ∫ t in s, f (a, t) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal) + +lemma todo3_ae_eq (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) : + (fun t ↦ todo3 f hf.measurable (a, t) q) =ᵐ[ν a] fun t ↦ f (a, t) q := by + filter_upwards [hf.isRatStieltjesPoint_ae a] with a ha + rw [todo3_eq, toCDFLike_of_isRatStieltjesPoint ha] + +lemma set_integral_todo3_rat (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) + {s : Set β} (hs : MeasurableSet s) : + ∫ t in s, todo3 f hf.measurable (a, t) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal := by + rw [set_integral_congr_ae hs (g := fun t ↦ f (a, t) q) ?_, hf.isCDF a hs] + filter_upwards [todo3_ae_eq hf a q] with b hb using fun _ ↦ hb + +lemma set_lintegral_todo3_rat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) + (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : + ∫⁻ t in s, ENNReal.ofReal (todo3 f hf.measurable (a, t) q) ∂(ν a) = μ a (s ×ˢ Iic (q : ℝ)) := by + rw [← ofReal_integral_eq_lintegral_ofReal] + · rw [set_integral_todo3_rat hf a q hs, ENNReal.ofReal_toReal] + exact measure_ne_top _ _ + · refine Integrable.restrict ?_ + rw [integrable_congr (todo3_ae_eq hf a q)] + exact hf.integrable a q + · exact ae_of_all _ (fun x ↦ todo3_nonneg _ _ _) + +lemma set_lintegral_cdfKernel_Iic_rat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) + (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : + ∫⁻ t in s, cdfKernel f hf.measurable (a, t) (Iic q) ∂(ν a) = μ a (s ×ˢ Iic (q : ℝ)) := by + simp_rw [cdfKernel_Iic, set_lintegral_todo3_rat hf a q hs] + +theorem lintegral_iInf_directed_of_measurable' {mα : MeasurableSpace α} [Countable β] + {f : β → α → ℝ≥0∞} {μ : Measure α} (hμ : μ ≠ 0) (hf : ∀ b, Measurable (f b)) + (hf_int : ∀ b, ∫⁻ a, f b a ∂μ ≠ ∞) (h_directed : Directed (· ≥ ·) f) : + ∫⁻ a, ⨅ b, f b a ∂μ = ⨅ b, ∫⁻ a, f b a ∂μ := sorry + +lemma set_lintegral_cdfKernel_Iic [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) + (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : + ∫⁻ t in s, cdfKernel f hf.measurable (a, t) (Iic x) ∂(ν a) = μ a (s ×ˢ Iic x) := by + -- We have the result for `x : ℚ` thanks to `set_lintegral_cdfKernel_Iic_rat`. + -- We use the equality `condCDF ρ a x = ⨅ r : {r' : ℚ // x < r'}, condCDF ρ a r` and a monotone + -- convergence argument to extend it to the reals. + by_cases hρ_zero : (ν a).restrict s = 0 + · rw [hρ_zero, lintegral_zero_measure] + have ⟨q, hq⟩ := exists_rat_gt x + suffices μ a (s ×ˢ Iic (q : ℝ)) = 0 by + symm + refine measure_mono_null (fun p ↦ ?_) this + simp only [mem_prod, mem_Iic, and_imp] + exact fun h1 h2 ↦ ⟨h1, h2.trans hq.le⟩ + suffices (μ a (s ×ˢ Iic (q : ℝ))).toReal = 0 by + rw [ENNReal.toReal_eq_zero_iff] at this + simpa [measure_ne_top] using this + rw [← hf.isCDF a hs q] + simp [hρ_zero] + have h : ∫⁻ t in s, cdfKernel f hf.measurable (a, t) (Iic x) ∂(ν a) + = ∫⁻ t in s, ⨅ r : { r' : ℚ // x < r' }, + cdfKernel f hf.measurable (a, t) (Iic r) ∂(ν a) := by + congr with t : 1 + rw [← measure_iInter_eq_iInf] + · congr with y : 1 + simp only [mem_Iic, mem_iInter, Subtype.forall] + refine ⟨fun h a ha ↦ h.trans ?_, fun h ↦ ?_⟩ + · exact mod_cast ha.le + · refine le_of_forall_lt_rat_imp_le fun q hq ↦ h q ?_ + exact mod_cast hq + · exact fun _ ↦ measurableSet_Iic + · refine Monotone.directed_ge fun r r' hrr' ↦ ?_ + refine Iic_subset_Iic.mpr ?_ + exact mod_cast hrr' + · obtain ⟨q, hq⟩ := exists_rat_gt x + exact ⟨⟨q, hq⟩, measure_ne_top _ _⟩ + have h_nonempty : Nonempty { r' : ℚ // x < ↑r' } := by + obtain ⟨r, hrx⟩ := exists_rat_gt x + exact ⟨⟨r, hrx⟩⟩ + rw [h, lintegral_iInf_directed_of_measurable' hρ_zero fun q : { r' : ℚ // x < ↑r' } ↦ ?_] + rotate_left + · intro b + rw [set_lintegral_cdfKernel_Iic_rat hf a _ hs] + exact measure_ne_top _ _ + · refine Monotone.directed_ge fun i j hij t ↦ measure_mono (Iic_subset_Iic.mpr ?_) + exact mod_cast hij + · exact (kernel.measurable_coe _ measurableSet_Iic).comp measurable_prod_mk_left + simp_rw [set_lintegral_cdfKernel_Iic_rat hf _ _ hs] + rw [← measure_iInter_eq_iInf] + · rw [← prod_iInter] + congr with y + simp only [mem_iInter, mem_Iic, Subtype.forall, Subtype.coe_mk] + exact ⟨le_of_forall_lt_rat_imp_le, fun hyx q hq ↦ hyx.trans hq.le⟩ + · exact fun i ↦ hs.prod measurableSet_Iic + · refine Monotone.directed_ge fun i j hij ↦ ?_ + refine prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, Iic_subset_Iic.mpr ?_⟩) + exact mod_cast hij + · exact ⟨h_nonempty.some, measure_ne_top _ _⟩ + +theorem Real.iUnion_Iic_rat' : ⋃ r : ℚ, Iic (r : ℝ) = univ := sorry + +lemma set_lintegral_cdfKernel_univ [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) + (a : α) {s : Set β} (hs : MeasurableSet s) : + ∫⁻ t in s, cdfKernel f hf.measurable (a, t) univ ∂(ν a) = μ a (s ×ˢ univ) := by + rw [← Real.iUnion_Iic_rat', prod_iUnion] + have h_dir : Directed (fun x y ↦ x ⊆ y) fun q : ℚ ↦ Iic (q : ℝ) := by + refine Monotone.directed_le fun r r' hrr' ↦ Iic_subset_Iic.mpr ?_ + exact mod_cast hrr' + have h_dir_prod : Directed (fun x y ↦ x ⊆ y) fun q : ℚ ↦ s ×ˢ Iic (q : ℝ) := by + refine Monotone.directed_le fun i j hij ↦ ?_ + refine prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, Iic_subset_Iic.mpr ?_⟩) + exact mod_cast hij + simp_rw [measure_iUnion_eq_iSup h_dir, measure_iUnion_eq_iSup h_dir_prod] + rw [lintegral_iSup_directed] + · simp_rw [set_lintegral_cdfKernel_Iic hf _ _ hs] + · refine fun q ↦ Measurable.aemeasurable ?_ + exact (kernel.measurable_coe _ measurableSet_Iic).comp measurable_prod_mk_left + · refine Monotone.directed_le fun i j hij t ↦ measure_mono (Iic_subset_Iic.mpr ?_) + exact mod_cast hij + +lemma lintegral_cdfKernel_univ [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) : + ∫⁻ t, cdfKernel f hf.measurable (a, t) univ ∂(ν a) = μ a univ := by + rw [← set_lintegral_univ, set_lintegral_cdfKernel_univ hf a MeasurableSet.univ, univ_prod_univ] + +lemma set_lintegral_cdfKernel_prod [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) + (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set ℝ} (ht : MeasurableSet t) : + ∫⁻ x in s, cdfKernel f hf.measurable (a, x) t ∂(ν a) = μ a (s ×ˢ t) := by + -- `set_lintegral_cdfKernel_Iic` gives the result for `t = Iic x`. These sets form a + -- π-system that generates the Borel σ-algebra, hence we can get the same equality for any + -- measurable set `t`. + apply MeasurableSpace.induction_on_inter (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic _ _ _ _ ht + · simp only [measure_empty, lintegral_const, zero_mul, prod_empty] + · rintro t ⟨q, rfl⟩ + exact set_lintegral_cdfKernel_Iic hf a _ hs + · intro t ht ht_lintegral + calc ∫⁻ x in s, cdfKernel f hf.measurable (a, x) tᶜ ∂(ν a) + = ∫⁻ x in s, cdfKernel f hf.measurable (a, x) univ - cdfKernel f hf.measurable (a, x) t ∂(ν a) := by + congr with x; rw [measure_compl ht (measure_ne_top (cdfKernel f hf.measurable (a, x)) _)] + _ = ∫⁻ x in s, cdfKernel f hf.measurable (a, x) univ ∂(ν a) + - ∫⁻ x in s, cdfKernel f hf.measurable (a, x) t ∂(ν a) := by + rw [lintegral_sub] + · exact (kernel.measurable_coe (cdfKernel f hf.measurable) ht).comp measurable_prod_mk_left + · rw [ht_lintegral] + exact measure_ne_top _ _ + · exact eventually_of_forall fun a ↦ measure_mono (subset_univ _) + _ = μ a (s ×ˢ univ) - μ a (s ×ˢ t) := by + rw [set_lintegral_cdfKernel_univ hf a hs, ht_lintegral] + _ = μ a (s ×ˢ tᶜ) := by + rw [← measure_diff _ (hs.prod ht) (measure_ne_top _ _)] + · rw [prod_diff_prod, compl_eq_univ_diff] + simp only [diff_self, empty_prod, union_empty] + · rw [prod_subset_prod_iff] + exact Or.inl ⟨subset_rfl, subset_univ t⟩ + · intro f hf_disj hf_meas hf_eq + simp_rw [measure_iUnion hf_disj hf_meas] + rw [lintegral_tsum, prod_iUnion, measure_iUnion] + · simp_rw [hf_eq] + · intro i j hij + rw [Function.onFun, Set.disjoint_prod] + exact Or.inr (hf_disj hij) + · exact fun i ↦ MeasurableSet.prod hs (hf_meas i) + · exact fun i ↦ + ((kernel.measurable_coe _ (hf_meas i)).comp measurable_prod_mk_left).aemeasurable.restrict + +lemma lintegral_cdfKernel_mem [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) + (a : α) {s : Set (β × ℝ)} (hs : MeasurableSet s) : + ∫⁻ x, cdfKernel f hf.measurable (a, x) {y | (x, y) ∈ s} ∂(ν a) = μ a s := by + -- `set_lintegral_condKernelReal_prod` gives the result for sets of the form `t₁ × t₂`. These + -- sets form a π-system that generates the product σ-algebra, hence we can get the same equality + -- for any measurable set `s`. + apply MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod _ _ _ _ hs + · simp only [mem_empty_iff_false, setOf_false, measure_empty, lintegral_const, + zero_mul] + · rintro _ ⟨t₁, ht₁, t₂, ht₂, rfl⟩ + simp only [mem_setOf_eq] at ht₁ ht₂ + have h_prod_eq_snd : ∀ a ∈ t₁, {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} = t₂ := by + intro a ha + simp only [ha, prod_mk_mem_set_prod_eq, true_and_iff, setOf_mem_eq] + rw [← lintegral_add_compl _ ht₁] + have h_eq1 : ∫⁻ x in t₁, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) + = ∫⁻ x in t₁, cdfKernel f hf.measurable (a, x) t₂ ∂(ν a) := by + refine' set_lintegral_congr_fun ht₁ (eventually_of_forall fun a ha ↦ _) + rw [h_prod_eq_snd a ha] + have h_eq2 : ∫⁻ x in t₁ᶜ, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) = 0 := by + suffices h_eq_zero : ∀ x ∈ t₁ᶜ, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} = 0 + · rw [set_lintegral_congr_fun ht₁.compl (eventually_of_forall h_eq_zero)] + simp only [lintegral_const, zero_mul] + intro a hat₁ + rw [mem_compl_iff] at hat₁ + simp only [hat₁, prod_mk_mem_set_prod_eq, false_and_iff, setOf_false, measure_empty] + rw [h_eq1, h_eq2, add_zero] + exact set_lintegral_cdfKernel_prod hf a ht₁ ht₂ + · intro t ht ht_eq + calc ∫⁻ x, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ tᶜ} ∂(ν a) + = ∫⁻ x, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t}ᶜ ∂(ν a) := rfl + _ = ∫⁻ x, cdfKernel f hf.measurable (a, x) univ + - cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t} ∂(ν a) := by + congr with x : 1 + exact measure_compl (measurable_prod_mk_left ht) + (measure_ne_top (cdfKernel f hf.measurable (a, x)) _) + _ = ∫⁻ x, cdfKernel f hf.measurable (a, x) univ ∂(ν a) - + ∫⁻ x, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t} ∂(ν a) := by + have h_le : (fun x ↦ cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t}) + ≤ᵐ[ν a] fun x ↦ cdfKernel f hf.measurable (a, x) univ := + eventually_of_forall fun _ ↦ measure_mono (subset_univ _) + rw [lintegral_sub _ _ h_le] + · exact kernel.measurable_kernel_prod_mk_left' ht a + refine ((lintegral_mono_ae h_le).trans_lt ?_).ne + rw [lintegral_cdfKernel_univ hf] + exact measure_lt_top _ univ + _ = μ a univ - μ a t := by rw [ht_eq, lintegral_cdfKernel_univ hf] + _ = μ a tᶜ := (measure_compl ht (measure_ne_top _ _)).symm + · intro f' hf_disj hf_meas hf_eq + have h_eq : ∀ a, {x | (a, x) ∈ ⋃ i, f' i} = ⋃ i, {x | (a, x) ∈ f' i} := by + intro a; ext x; simp only [mem_iUnion, mem_setOf_eq] + simp_rw [h_eq] + have h_disj : ∀ a, Pairwise (Disjoint on fun i ↦ {x | (a, x) ∈ f' i}) := by + intro a i j hij + have h_disj := hf_disj hij + rw [Function.onFun, disjoint_iff_inter_eq_empty] at h_disj ⊢ + ext1 x + simp only [mem_inter_iff, mem_setOf_eq, mem_empty_iff_false, iff_false_iff] + intro h_mem_both + suffices (a, x) ∈ ∅ by rwa [mem_empty_iff_false] at this + rwa [← h_disj, mem_inter_iff] + calc ∫⁻ x, cdfKernel f hf.measurable (a, x) (⋃ i, {y | (x, y) ∈ f' i}) ∂(ν a) + = ∫⁻ x, ∑' i, cdfKernel f hf.measurable (a, x) {y | (x, y) ∈ f' i} ∂(ν a) := by + congr with x : 1 + rw [measure_iUnion (h_disj x) fun i ↦ measurable_prod_mk_left (hf_meas i)] + _ = ∑' i, ∫⁻ x, cdfKernel f hf.measurable (a, x) {y | (x, y) ∈ f' i} ∂(ν a) := + lintegral_tsum fun i ↦ (kernel.measurable_kernel_prod_mk_left' (hf_meas i) a).aemeasurable + _ = ∑' i, μ a (f' i) := by simp_rw [hf_eq] + _ = μ a (iUnion f') := (measure_iUnion hf_disj hf_meas).symm + +lemma kernel.eq_compProd_cdfKernel [IsFiniteKernel μ] [IsSFiniteKernel ν] + (hf : IsRatKernelCDF f μ ν) : + μ = ν ⊗ₖ cdfKernel f hf.measurable := by + ext a s hs + rw [kernel.compProd_apply _ _ _ hs, lintegral_cdfKernel_mem hf a hs] + +lemma ae_cdfKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsRatKernelCDF f μ ν) (a : α) + {s : Set ℝ} (hs : MeasurableSet s) (hμs : μ a {x | x.snd ∈ sᶜ} = 0) : + ∀ᵐ t ∂(ν a), cdfKernel f hf.measurable (a, t) s = 1 := by + have h_eq : μ = ν ⊗ₖ cdfKernel f hf.measurable := kernel.eq_compProd_cdfKernel hf + have h : μ a {x | x.snd ∈ sᶜ} = (ν ⊗ₖ cdfKernel f hf.measurable) a {x | x.snd ∈ sᶜ} := by + rw [← h_eq] + rw [hμs, kernel.compProd_apply] at h + swap; · exact measurable_snd hs.compl + rw [eq_comm, lintegral_eq_zero_iff] at h + swap + · simp only [mem_compl_iff, mem_setOf_eq] + change Measurable ((fun p ↦ cdfKernel f _ p {c | c ∉ s}) ∘ (fun b : β ↦ (a, b))) + exact (kernel.measurable_coe _ hs.compl).comp measurable_prod_mk_left + simp only [mem_compl_iff, mem_setOf_eq, kernel.prodMkLeft_apply'] at h + filter_upwards [h] with t ht + change cdfKernel f hf.measurable (a, t) sᶜ = 0 at ht + rwa [prob_compl_eq_zero_iff hs] at ht + +lemma measurableSet_eq_one (hf : IsRatKernelCDF f μ ν) {s : Set ℝ} (hs : MeasurableSet s) : + MeasurableSet {p | cdfKernel f hf.measurable p s = 1} := + (kernel.measurable_coe _ hs) (measurableSet_singleton 1) + + +end end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index 33f3f81299378..a898d8bab75c9 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.Data.Set.Lattice -import Mathlib.Probability.Kernel.StieltjesReal +import Mathlib.Probability.Kernel.BuildKernel #align_import probability.kernel.cond_cdf from "leanprover-community/mathlib"@"3b88f4005dc2e28d42f974cc1ce838f0dafb39b8" @@ -348,6 +348,16 @@ theorem preCDF_le_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : exact Measure.IicSnd_le_fst ρ r s hs #align probability_theory.pre_cdf_le_one ProbabilityTheory.preCDF_le_one +theorem set_integral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) {s : Set α} (hs : MeasurableSet s) + [IsFiniteMeasure ρ] : + ∫ x in s, (preCDF ρ r x).toReal ∂ρ.fst = (ρ.IicSnd r s).toReal := by + rw [integral_toReal] + · rw [set_lintegral_preCDF_fst _ _ hs] + · exact measurable_preCDF.aemeasurable + · refine ae_restrict_of_ae ?_ + filter_upwards [preCDF_le_one ρ] with a ha + exact (ha r).trans_lt ENNReal.one_lt_top + theorem tendsto_lintegral_preCDF_atTop (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : Tendsto (fun r => ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (ρ univ)) := by convert ρ.tendsto_IicSnd_atTop MeasurableSet.univ @@ -564,11 +574,38 @@ lemma isRatStieltjesPoint_ae (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : rw [← h5] · exact fun r ↦ ((h2 r).trans_lt ENNReal.one_lt_top).ne +theorem integrable_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℚ) : + Integrable (fun a ↦ (preCDF ρ x a).toReal) ρ.fst := by + refine' integrable_of_forall_fin_meas_le _ (measure_lt_top ρ.fst univ) _ fun t _ _ ↦ _ + · exact measurable_preCDF.ennreal_toReal.aestronglyMeasurable + · simp_rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_of_nonneg ENNReal.toReal_nonneg] + rw [← lintegral_one] + refine (set_lintegral_le_lintegral _ _).trans (lintegral_mono_ae ?_) + filter_upwards [preCDF_le_one ρ] with a ha using ENNReal.ofReal_toReal_le.trans (ha _) + +lemma isRatKernelCDF_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : + IsRatKernelCDF (fun p r ↦ (preCDF ρ r p.2).toReal) + (kernel.const Unit ρ) (kernel.const Unit ρ.fst) where + measurable q := measurable_preCDF.ennreal_toReal.comp measurable_snd + isRatStieltjesPoint_ae a := by + filter_upwards [isRatStieltjesPoint_ae ρ] with a ha + exact ⟨ha.mono, ha.nonneg, ha.le_one, ha.tendsto_atTop_one, ha.tendsto_atBot_zero, + ha.iInf_rat_gt_eq⟩ + integrable _ q := integrable_preCDF ρ q + isCDF a s hs q := by rw [kernel.const_apply, kernel.const_apply, set_integral_preCDF_fst _ _ hs, + Measure.IicSnd_apply _ _ hs] + /-- Conditional cdf of the measure given the value on `α`, as a Stieltjes function. -/ noncomputable def condCDF (ρ : Measure (α × ℝ)) (a : α) : StieltjesFunction := todo3 (fun a r ↦ (preCDF ρ r a).toReal) (fun _ ↦ measurable_preCDF.ennreal_toReal) a #align probability_theory.cond_cdf ProbabilityTheory.condCDF +lemma condCDF_eq_todo3_unit_prod (ρ : Measure (α × ℝ)) (a : α) : + condCDF ρ a = todo3 (fun (p : Unit × α) r ↦ (preCDF ρ r p.2).toReal) + (fun _ ↦ measurable_preCDF.ennreal_toReal.comp measurable_snd) ((), a) := by + ext x + rw [condCDF, ← todo3_unit_prod] + #noalign probability_theory.cond_cdf_eq_cond_cdf_rat /-- The conditional cdf is non-negative for all `a : α`. -/ @@ -609,60 +646,18 @@ theorem measurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : Measurable fun measurable_todo3 _ _ #align probability_theory.measurable_cond_cdf ProbabilityTheory.measurable_condCDF -/-- Auxiliary lemma for `set_lintegral_cond_cdf`. -/ -theorem set_lintegral_condCDF_rat (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) {s : Set α} - (hs : MeasurableSet s) : - ∫⁻ a in s, ENNReal.ofReal (condCDF ρ a r) ∂ρ.fst = ρ (s ×ˢ Iic (r : ℝ)) := by - have : ∀ᵐ a ∂ρ.fst, a ∈ s → ENNReal.ofReal (condCDF ρ a r) = preCDF ρ r a := by - filter_upwards [ofReal_condCDF_ae_eq ρ r] with a ha using fun _ => ha - rw [set_lintegral_congr_fun hs this, set_lintegral_preCDF_fst ρ r hs] - exact ρ.IicSnd_apply r hs -#align probability_theory.set_lintegral_cond_cdf_rat ProbabilityTheory.set_lintegral_condCDF_rat +#noalign probability_theory.set_lintegral_cond_cdf_rat theorem set_lintegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) {s : Set α} - (hs : MeasurableSet s) : ∫⁻ a in s, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x) := by - -- We have the result for `x : ℚ` thanks to `set_lintegral_condCDF_rat`. We use the equality - -- `condCDF ρ a x = ⨅ r : {r' : ℚ // x < r'}, condCDF ρ a r` and a monotone convergence - -- argument to extend it to the reals. - by_cases hρ_zero : ρ.fst.restrict s = 0 - · rw [hρ_zero, lintegral_zero_measure] - refine' le_antisymm (zero_le _) _ - calc - ρ (s ×ˢ Iic x) ≤ ρ (Prod.fst ⁻¹' s) := measure_mono (prod_subset_preimage_fst s (Iic x)) - _ = ρ.fst s := by rw [Measure.fst_apply hs] - _ = ρ.fst.restrict s univ := by rw [Measure.restrict_apply_univ] - _ = 0 := by simp only [hρ_zero, Measure.coe_zero, Pi.zero_apply] - have h : - ∫⁻ a in s, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = - ∫⁻ a in s, ENNReal.ofReal (⨅ r : { r' : ℚ // x < r' }, condCDF ρ a r) ∂ρ.fst := by - congr with a : 1 - rw [← (condCDF ρ a).iInf_rat_gt_eq x] - have h_nonempty : Nonempty { r' : ℚ // x < ↑r' } := by - obtain ⟨r, hrx⟩ := exists_rat_gt x - exact ⟨⟨r, hrx⟩⟩ - rw [h] - simp_rw [ENNReal.ofReal_cinfi] - have h_coe : ∀ b : { r' : ℚ // x < ↑r' }, (b : ℝ) = ((b : ℚ) : ℝ) := fun _ => by congr - rw [lintegral_iInf_directed_of_measurable hρ_zero fun q : { r' : ℚ // x < ↑r' } => - (measurable_condCDF ρ q).ennreal_ofReal] - rotate_left - · intro b - rw [set_lintegral_condCDF_rat ρ _ hs] - exact measure_ne_top ρ _ - · refine' Monotone.directed_ge fun i j hij a => ENNReal.ofReal_le_ofReal ((condCDF ρ a).mono _) - rw [h_coe, h_coe] - exact mod_cast hij - simp_rw [set_lintegral_condCDF_rat ρ _ hs] - rw [← measure_iInter_eq_iInf] - · rw [← prod_iInter] - congr with y - simp only [mem_iInter, mem_Iic, Subtype.forall, Subtype.coe_mk] - exact ⟨le_of_forall_lt_rat_imp_le, fun hyx q hq => hyx.trans hq.le⟩ - · exact fun i => hs.prod measurableSet_Iic - · refine' Monotone.directed_ge fun i j hij => _ - refine' prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, Iic_subset_Iic.mpr _⟩) - exact mod_cast hij - · exact ⟨h_nonempty.some, measure_ne_top _ _⟩ + (hs : MeasurableSet s) : + ∫⁻ a in s, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x) := by + have h := set_lintegral_cdfKernel_Iic (isRatKernelCDF_preCDF ρ) () x hs + simp only [kernel.const_apply] at h + rw [← h] + simp_rw [cdfKernel_Iic] + congr with a + congr + exact condCDF_eq_todo3_unit_prod _ _ #align probability_theory.set_lintegral_cond_cdf ProbabilityTheory.set_lintegral_condCDF theorem lintegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) : diff --git a/Mathlib/Probability/Kernel/Disintegration.lean b/Mathlib/Probability/Kernel/Disintegration.lean index b59c70d867087..7c7c2faa68141 100644 --- a/Mathlib/Probability/Kernel/Disintegration.lean +++ b/Mathlib/Probability/Kernel/Disintegration.lean @@ -60,23 +60,24 @@ section Real /-- Conditional measure on the second space of the product given the value on the first, as a kernel. Use the more general `condKernel`. -/ -noncomputable def condKernelReal (ρ : Measure (α × ℝ)) : kernel α ℝ where - val a := (condCDF ρ a).measure - property := measurable_measure_condCDF ρ +noncomputable def condKernelReal (ρ : Measure (α × ℝ)) : kernel α ℝ := + kernel.comap (cdfKernel (fun (p : Unit × α) r ↦ (preCDF ρ r p.2).toReal) + (fun _ ↦ measurable_preCDF.ennreal_toReal.comp measurable_snd)) (fun a ↦ ((), a)) + measurable_prod_mk_left #align probability_theory.cond_kernel_real ProbabilityTheory.condKernelReal -instance (ρ : Measure (α × ℝ)) : IsMarkovKernel (condKernelReal ρ) := - ⟨fun a => by rw [condKernelReal]; exact instIsProbabilityMeasure ρ a⟩ +instance (ρ : Measure (α × ℝ)) : IsMarkovKernel (condKernelReal ρ) := by + rw [condKernelReal]; infer_instance theorem condKernelReal_Iic (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : - condKernelReal ρ a (Iic x) = ENNReal.ofReal (condCDF ρ a x) := - measure_condCDF_Iic ρ a x + condKernelReal ρ a (Iic x) = ENNReal.ofReal (condCDF ρ a x) := by + rw [condKernelReal, kernel.comap_apply, cdfKernel_Iic, condCDF_eq_todo3_unit_prod] #align probability_theory.cond_kernel_real_Iic ProbabilityTheory.condKernelReal_Iic theorem set_lintegral_condKernelReal_Iic (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) {s : Set α} (hs : MeasurableSet s) : - ∫⁻ a in s, condKernelReal ρ a (Iic x) ∂ρ.fst = ρ (s ×ˢ Iic x) := by - simp_rw [condKernelReal_Iic]; exact set_lintegral_condCDF ρ x hs + ∫⁻ a in s, condKernelReal ρ a (Iic x) ∂ρ.fst = ρ (s ×ˢ Iic x) := + set_lintegral_cdfKernel_Iic (isRatKernelCDF_preCDF ρ) () x hs #align probability_theory.set_lintegral_cond_kernel_real_Iic ProbabilityTheory.set_lintegral_condKernelReal_Iic theorem set_lintegral_condKernelReal_univ (ρ : Measure (α × ℝ)) {s : Set α} (hs : MeasurableSet s) : @@ -94,111 +95,14 @@ theorem lintegral_condKernelReal_univ (ρ : Measure (α × ℝ)) : variable (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] theorem set_lintegral_condKernelReal_prod {s : Set α} (hs : MeasurableSet s) {t : Set ℝ} - (ht : MeasurableSet t) : ∫⁻ a in s, condKernelReal ρ a t ∂ρ.fst = ρ (s ×ˢ t) := by - -- `set_lintegral_condKernelReal_Iic` gives the result for `t = Iic x`. These sets form a - -- π-system that generates the Borel σ-algebra, hence we can get the same equality for any - -- measurable set `t`. - apply MeasurableSpace.induction_on_inter (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic _ _ _ _ ht - · simp only [measure_empty, lintegral_const, zero_mul, prod_empty] - · rintro t ⟨q, rfl⟩ - exact set_lintegral_condKernelReal_Iic ρ q hs - · intro t ht ht_lintegral - calc - ∫⁻ a in s, condKernelReal ρ a tᶜ ∂ρ.fst = - ∫⁻ a in s, condKernelReal ρ a univ - condKernelReal ρ a t ∂ρ.fst := by - congr with a; rw [measure_compl ht (measure_ne_top (condKernelReal ρ a) _)] - _ = ∫⁻ a in s, condKernelReal ρ a univ ∂ρ.fst - ∫⁻ a in s, condKernelReal ρ a t ∂ρ.fst := by - rw [lintegral_sub (kernel.measurable_coe (condKernelReal ρ) ht)] - · rw [ht_lintegral] - exact measure_ne_top ρ _ - · exact eventually_of_forall fun a => measure_mono (subset_univ _) - _ = ρ (s ×ˢ univ) - ρ (s ×ˢ t) := by - rw [set_lintegral_condKernelReal_univ ρ hs, ht_lintegral] - _ = ρ (s ×ˢ tᶜ) := by - rw [← measure_diff _ (hs.prod ht) (measure_ne_top ρ _)] - · rw [prod_diff_prod, compl_eq_univ_diff] - simp only [diff_self, empty_prod, union_empty] - · rw [prod_subset_prod_iff] - exact Or.inl ⟨subset_rfl, subset_univ t⟩ - · intro f hf_disj hf_meas hf_eq - simp_rw [measure_iUnion hf_disj hf_meas] - rw [lintegral_tsum fun i => (kernel.measurable_coe _ (hf_meas i)).aemeasurable.restrict, - prod_iUnion, measure_iUnion] - · simp_rw [hf_eq] - · intro i j hij - rw [Function.onFun, Set.disjoint_prod] - exact Or.inr (hf_disj hij) - · exact fun i => MeasurableSet.prod hs (hf_meas i) + (ht : MeasurableSet t) : + ∫⁻ a in s, condKernelReal ρ a t ∂ρ.fst = ρ (s ×ˢ t) := + set_lintegral_cdfKernel_prod (isRatKernelCDF_preCDF ρ) () hs ht #align probability_theory.set_lintegral_cond_kernel_real_prod ProbabilityTheory.set_lintegral_condKernelReal_prod theorem lintegral_condKernelReal_mem {s : Set (α × ℝ)} (hs : MeasurableSet s) : - ∫⁻ a, condKernelReal ρ a {x | (a, x) ∈ s} ∂ρ.fst = ρ s := by - -- `set_lintegral_condKernelReal_prod` gives the result for sets of the form `t₁ × t₂`. These - -- sets form a π-system that generates the product σ-algebra, hence we can get the same equality - -- for any measurable set `s`. - apply MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod _ _ _ _ hs - · simp only [mem_empty_iff_false, setOf_false, measure_empty, lintegral_const, - zero_mul] - · rintro _ ⟨t₁, ht₁, t₂, ht₂, rfl⟩ - have h_prod_eq_snd : ∀ a ∈ t₁, {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} = t₂ := by - intro a ha - simp only [ha, prod_mk_mem_set_prod_eq, true_and_iff, setOf_mem_eq] - rw [← lintegral_add_compl _ ht₁] - have h_eq1 : ∫⁻ a in t₁, condKernelReal ρ a {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} ∂ρ.fst = - ∫⁻ a in t₁, condKernelReal ρ a t₂ ∂ρ.fst := by - refine' set_lintegral_congr_fun ht₁ (eventually_of_forall fun a ha => _) - rw [h_prod_eq_snd a ha] - have h_eq2 : ∫⁻ a in t₁ᶜ, condKernelReal ρ a {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} ∂ρ.fst = 0 := by - suffices h_eq_zero : ∀ a ∈ t₁ᶜ, condKernelReal ρ a {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} = 0 - · rw [set_lintegral_congr_fun ht₁.compl (eventually_of_forall h_eq_zero)] - simp only [lintegral_const, zero_mul] - intro a hat₁ - rw [mem_compl_iff] at hat₁ - simp only [hat₁, prod_mk_mem_set_prod_eq, false_and_iff, setOf_false, measure_empty] - rw [h_eq1, h_eq2, add_zero] - exact set_lintegral_condKernelReal_prod ρ ht₁ ht₂ - · intro t ht ht_eq - calc - ∫⁻ a, condKernelReal ρ a {x : ℝ | (a, x) ∈ tᶜ} ∂ρ.fst = - ∫⁻ a, condKernelReal ρ a {x : ℝ | (a, x) ∈ t}ᶜ ∂ρ.fst := rfl - _ = ∫⁻ a, condKernelReal ρ a univ - condKernelReal ρ a {x : ℝ | (a, x) ∈ t} ∂ρ.fst := by - congr with a : 1 - exact measure_compl (measurable_prod_mk_left ht) (measure_ne_top (condKernelReal ρ a) _) - _ = ∫⁻ a, condKernelReal ρ a univ ∂ρ.fst - - ∫⁻ a, condKernelReal ρ a {x : ℝ | (a, x) ∈ t} ∂ρ.fst := by - have h_le : (fun a => condKernelReal ρ a {x : ℝ | (a, x) ∈ t}) ≤ᵐ[ρ.fst] fun a => - condKernelReal ρ a univ := eventually_of_forall fun a => measure_mono (subset_univ _) - rw [lintegral_sub _ _ h_le] - · exact kernel.measurable_kernel_prod_mk_left ht - refine' ((lintegral_mono_ae h_le).trans_lt _).ne - rw [lintegral_condKernelReal_univ] - exact measure_lt_top ρ univ - _ = ρ univ - ρ t := by rw [ht_eq, lintegral_condKernelReal_univ] - _ = ρ tᶜ := (measure_compl ht (measure_ne_top _ _)).symm - · intro f hf_disj hf_meas hf_eq - have h_eq : ∀ a, {x | (a, x) ∈ ⋃ i, f i} = ⋃ i, {x | (a, x) ∈ f i} := by - intro a - ext1 x - simp only [mem_iUnion, mem_setOf_eq] - simp_rw [h_eq] - have h_disj : ∀ a, Pairwise (Disjoint on fun i => {x | (a, x) ∈ f i}) := by - intro a i j hij - have h_disj := hf_disj hij - rw [Function.onFun, disjoint_iff_inter_eq_empty] at h_disj ⊢ - ext1 x - simp only [mem_inter_iff, mem_setOf_eq, mem_empty_iff_false, iff_false_iff] - intro h_mem_both - suffices (a, x) ∈ ∅ by rwa [mem_empty_iff_false] at this - rwa [← h_disj, mem_inter_iff] - calc - ∫⁻ a, condKernelReal ρ a (⋃ i, {x | (a, x) ∈ f i}) ∂ρ.fst = - ∫⁻ a, ∑' i, condKernelReal ρ a {x | (a, x) ∈ f i} ∂ρ.fst := by - congr with a : 1 - rw [measure_iUnion (h_disj a) fun i => measurable_prod_mk_left (hf_meas i)] - _ = ∑' i, ∫⁻ a, condKernelReal ρ a {x | (a, x) ∈ f i} ∂ρ.fst := - (lintegral_tsum fun i => (kernel.measurable_kernel_prod_mk_left (hf_meas i)).aemeasurable) - _ = ∑' i, ρ (f i) := by simp_rw [hf_eq] - _ = ρ (iUnion f) := (measure_iUnion hf_disj hf_meas).symm + ∫⁻ a, condKernelReal ρ a {x | (a, x) ∈ s} ∂ρ.fst = ρ s := + lintegral_cdfKernel_mem (isRatKernelCDF_preCDF ρ) () hs #align probability_theory.lintegral_cond_kernel_real_mem ProbabilityTheory.lintegral_condKernelReal_mem theorem kernel.const_eq_compProd_real (γ : Type*) [MeasurableSpace γ] (ρ : Measure (α × ℝ)) @@ -222,18 +126,9 @@ theorem lintegral_condKernelReal {f : α × ℝ → ℝ≥0∞} (hf : Measurable theorem ae_condKernelReal_eq_one {s : Set ℝ} (hs : MeasurableSet s) (hρ : ρ {x | x.snd ∈ sᶜ} = 0) : ∀ᵐ a ∂ρ.fst, condKernelReal ρ a s = 1 := by - have h : ρ {x | x.snd ∈ sᶜ} = (ρ.fst ⊗ₘ condKernelReal ρ) {x | x.snd ∈ sᶜ} := by - rw [← measure_eq_compProd_real] - rw [hρ, Measure.compProd_apply] at h - swap; · exact measurable_snd hs.compl - rw [eq_comm, lintegral_eq_zero_iff] at h - swap - · simp only [mem_compl_iff, mem_setOf_eq] - exact kernel.measurable_coe _ hs.compl - simp only [mem_compl_iff, mem_setOf_eq, kernel.prodMkLeft_apply'] at h - filter_upwards [h] with a ha - change condKernelReal ρ a sᶜ = 0 at ha - rwa [prob_compl_eq_zero_iff hs] at ha + have h := ae_cdfKernel_eq_one (isRatKernelCDF_preCDF ρ) () hs + simp only [kernel.const_apply] at h + exact h hρ #align probability_theory.ae_cond_kernel_real_eq_one ProbabilityTheory.ae_condKernelReal_eq_one end Real diff --git a/Mathlib/Probability/Kernel/StieltjesReal.lean b/Mathlib/Probability/Kernel/StieltjesReal.lean index 9991f53c63ae2..11033196e7f98 100644 --- a/Mathlib/Probability/Kernel/StieltjesReal.lean +++ b/Mathlib/Probability/Kernel/StieltjesReal.lean @@ -41,10 +41,12 @@ lemma measurableSet_tendsto_fun {β γ ι : Type*} [MeasurableSpace β] namespace ProbabilityTheory -variable {α β ι : Type*} [MeasurableSpace α] {f : α → ℚ → ℝ} +variable {α β ι : Type*} [MeasurableSpace α] section IsCDFLike +variable {f : α → ℚ → ℝ} + structure IsRatStieltjesPoint (f : α → ℚ → ℝ) (a : α) : Prop where mono : Monotone (f a) nonneg : ∀ q, 0 ≤ f a q @@ -53,23 +55,30 @@ structure IsRatStieltjesPoint (f : α → ℚ → ℝ) (a : α) : Prop where tendsto_atBot_zero : Tendsto (f a) atBot (𝓝 0) iInf_rat_gt_eq : ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t +lemma isRatStieltjesPoint_unit_prod_iff (f : α → ℚ → ℝ) (a : α) : + IsRatStieltjesPoint (fun p : Unit × α ↦ f p.2) ((), a) + ↔ IsRatStieltjesPoint f a := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · exact ⟨h.mono, h.nonneg, h.le_one, h.tendsto_atTop_one, h.tendsto_atBot_zero, + h.iInf_rat_gt_eq⟩ + · exact ⟨h.mono, h.nonneg, h.le_one, h.tendsto_atTop_one, h.tendsto_atBot_zero, + h.iInf_rat_gt_eq⟩ + lemma measurableSet_isRatStieltjesPoint (hf : ∀ q, Measurable (fun a ↦ f a q)) : MeasurableSet {a | IsRatStieltjesPoint f a} := by have h1 : MeasurableSet {a | Monotone (f a)} := by - change MeasurableSet {a | ∀ q r (hqr : q ≤ r), f a q ≤ f a r} + change MeasurableSet {a | ∀ q r (_ : q ≤ r), f a q ≤ f a r} simp_rw [Set.setOf_forall] refine MeasurableSet.iInter (fun q ↦ ?_) refine MeasurableSet.iInter (fun r ↦ ?_) - refine MeasurableSet.iInter (fun hqr ↦ ?_) + refine MeasurableSet.iInter (fun _ ↦ ?_) exact measurableSet_le (hf q) (hf r) have h2 : MeasurableSet {a | ∀ q, 0 ≤ f a q} := by simp_rw [Set.setOf_forall] - refine MeasurableSet.iInter (fun q ↦ ?_) - exact measurableSet_le measurable_const (hf q) + exact MeasurableSet.iInter (fun q ↦ measurableSet_le measurable_const (hf q)) have h3 : MeasurableSet {a | ∀ q, f a q ≤ 1} := by simp_rw [Set.setOf_forall] - refine MeasurableSet.iInter (fun q ↦ ?_) - exact measurableSet_le (hf q) measurable_const + exact MeasurableSet.iInter (fun q ↦ measurableSet_le (hf q) measurable_const) have h4 : MeasurableSet {a | Tendsto (f a) atTop (𝓝 1)} := measurableSet_tendsto_nhds (fun q ↦ hf q) 1 have h5 : MeasurableSet {a | Tendsto (f a) atBot (𝓝 0)} := @@ -78,8 +87,18 @@ lemma measurableSet_isRatStieltjesPoint (hf : ∀ q, Measurable (fun a ↦ f a q rw [Set.setOf_forall] refine MeasurableSet.iInter (fun q ↦ ?_) exact measurableSet_eq_fun (measurable_iInf fun _ ↦ hf _) (hf _) - have h := ((((h1.inter h2).inter h3).inter h4).inter h5).inter h6 - sorry + suffices {a | IsRatStieltjesPoint f a} + = ({a | Monotone (f a)} ∩ {a | ∀ (q : ℚ), 0 ≤ f a q} ∩ {a | ∀ (q : ℚ), f a q ≤ 1} + ∩ {a | Tendsto (f a) atTop (𝓝 1)} ∩ {a | Tendsto (f a) atBot (𝓝 0)} ∩ + {a | ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t}) by + rw [this] + exact ((((h1.inter h2).inter h3).inter h4).inter h5).inter h6 + ext a + simp only [mem_setOf_eq, mem_inter_iff] + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · exact ⟨⟨⟨⟨⟨h.mono, h.nonneg⟩, h.le_one⟩, h.tendsto_atTop_one⟩, h.tendsto_atBot_zero⟩, + h.iInf_rat_gt_eq⟩ + · exact ⟨h.1.1.1.1.1, h.1.1.1.1.2, h.1.1.1.2, h.1.1.2, h.1.2, h.2⟩ structure IsCDFLike (f : α → ℚ → ℝ) : Prop where mono : ∀ a, Monotone (f a) @@ -90,17 +109,6 @@ structure IsCDFLike (f : α → ℚ → ℝ) : Prop where iInf_rat_gt_eq : ∀ a, ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t measurable : ∀ q, Measurable (fun a ↦ f a q) -lemma IsCDFLike.ite {s : Set α} (hs : MeasurableSet s) [DecidablePred (fun a ↦ a ∈ s)] - {f g : α → ℚ → ℝ} (hf : IsCDFLike f) (hg : IsCDFLike g) : - IsCDFLike (fun a q ↦ if a ∈ s then f a q else g a q) where - mono a := by split_ifs; exacts [hf.mono a, hg.mono a] - nonneg a := by split_ifs; exacts [hf.nonneg a, hg.nonneg a] - le_one a := by split_ifs; exacts [hf.le_one a, hg.le_one a] - tendsto_atTop_one a := by split_ifs; exacts [hf.tendsto_atTop_one a, hg.tendsto_atTop_one a] - tendsto_atBot_zero a := by split_ifs; exacts [hf.tendsto_atBot_zero a, hg.tendsto_atBot_zero a] - iInf_rat_gt_eq a := by split_ifs; exacts [hf.iInf_rat_gt_eq a, hg.iInf_rat_gt_eq a] - measurable q := Measurable.ite hs (hf.measurable q) (hg.measurable q) - end IsCDFLike section DefaultRatCDF @@ -180,6 +188,8 @@ end DefaultRatCDF section ToCDFLike +variable {f : α → ℚ → ℝ} + open Classical in noncomputable def toCDFLike (f : α → ℚ → ℝ) : α → ℚ → ℝ := fun a q ↦ @@ -206,6 +216,11 @@ lemma isCDFLike_toCDFLike (hf : ∀ q, Measurable fun a ↦ f a q) : measurable q := Measurable.ite (measurableSet_isRatStieltjesPoint hf) (hf q) (measurable_defaultRatCDF α q) +lemma toCDFLike_unit_prod (a : α) : + toCDFLike (fun (p : Unit × α) ↦ f p.2) ((), a) = toCDFLike f a := by + unfold toCDFLike + rw [isRatStieltjesPoint_unit_prod_iff] + end ToCDFLike section IsCDFLike.stieltjesFunction @@ -217,6 +232,9 @@ definition used to define `cond_cdf`. -/ noncomputable irreducible_def todo1 (f : α → ℚ → ℝ) : α → ℝ → ℝ := fun a t ↦ ⨅ r : { r' : ℚ // t < r' }, f a r +lemma todo1_def' (f : α → ℚ → ℝ) (a : α) : + todo1 f a = fun (t : ℝ) ↦ ⨅ r : { r' : ℚ // t < r' }, f a r := by ext t; exact todo1_def f a t + lemma todo1_eq (a : α) (r : ℚ) : todo1 f a r = f a r := by rw [← hf.iInf_rat_gt_eq a r, todo1] @@ -229,6 +247,9 @@ lemma todo1_eq (a : α) (r : ℚ) : · intro t simp only [Equiv.coe_fn_mk, Subtype.coe_mk] +lemma todo1_unit_prod (a : α) : + todo1 (fun (p : Unit × α) ↦ f p.2) ((), a) = todo1 f a := by simp_rw [todo1_def'] + theorem todo1_nonneg (a : α) (r : ℝ) : 0 ≤ todo1 f a r := by have : Nonempty { r' : ℚ // r < ↑r' } := by obtain ⟨r, hrx⟩ := exists_rat_gt r @@ -385,6 +406,8 @@ end Measure end IsCDFLike.stieltjesFunction +section todo3 + variable {f : α → ℚ → ℝ} noncomputable @@ -394,6 +417,15 @@ def todo3 (f : α → ℚ → ℝ) (hf : ∀ q, Measurable fun a ↦ f a q) : α theorem todo3_eq (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℚ) : todo3 f hf a r = toCDFLike f a r := todo2_eq _ a r +lemma todo3_unit_prod (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : + todo3 (fun (p : Unit × α) ↦ f p.2) (fun q ↦ (hf q).comp measurable_snd) ((), a) + = todo3 f hf a := by + simp_rw [todo3,IsCDFLike.stieltjesFunction, ← todo1_unit_prod a] + congr with x + congr 1 with p : 1 + cases p with + | mk _ b => rw [← toCDFLike_unit_prod b] + /-- The conditional cdf is non-negative for all `a : α`. -/ theorem todo3_nonneg (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℝ) : 0 ≤ todo3 f hf a r := todo2_nonneg _ a r @@ -435,3 +467,5 @@ theorem measurable_measure_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) : Measurable fun a ↦ (todo3 f hf a).measure := measurable_measure_todo2 _ end Measure + +end todo3 diff --git a/Mathlib/Probability/Kernel/draft.lean b/Mathlib/Probability/Kernel/draft.lean index e25f9d547cb5f..064d19569b55b 100644 --- a/Mathlib/Probability/Kernel/draft.lean +++ b/Mathlib/Probability/Kernel/draft.lean @@ -17,6 +17,15 @@ section dissection_system def I (n : ℕ) (k : ℤ) : Set ℝ := Set.Ico (k * (2⁻¹ : ℝ) ^ n) ((k + 1) * ((2 : ℝ) ^ n)⁻¹) +lemma mem_I_iff_mul {n : ℕ} {k : ℤ} (x : ℝ) : x ∈ I n k ↔ k ≤ x * 2 ^ n ∧ x * 2 ^ n < k + 1 := by + simp only [I, inv_pow, mem_Ico] + rw [← div_eq_mul_inv, div_le_iff, ← div_eq_mul_inv, lt_div_iff] + · positivity + · positivity + +lemma mem_I_iff_floor {n : ℕ} {k : ℤ} (x : ℝ) : x ∈ I n k ↔ ⌊x * 2 ^ n⌋ = k := by + simp [mem_I_iff_mul, Int.floor_eq_iff] + lemma measurableSet_I (n : ℕ) (k : ℤ) : MeasurableSet (I n k) := measurableSet_Ico lemma Measure.iInf_Iic_gt_prod {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] @@ -30,9 +39,10 @@ lemma pairwise_disjoint_I (n : ℕ) : Pairwise (Disjoint on fun k ↦ I n k) := intro i j hij rw [Function.onFun, Set.disjoint_iff] intro x - simp only [I, inv_pow, mem_inter_iff, mem_Ico, mem_empty_iff_false, and_imp, imp_false, not_lt] - intro h1 h2 h3 - sorry + simp only [mem_inter_iff, mem_I_iff_floor, mem_empty_iff_false, and_imp, imp_false] + intro hi hj + rw [hi] at hj + exact hij hj lemma I_succ_union (n : ℕ) (k : ℤ) : I (n+1) (2 * k) ∪ I (n+1) (2 * k + 1) = I n k := by ext x @@ -53,7 +63,17 @@ lemma I_succ_union (n : ℕ) (k : ℤ) : I (n+1) (2 * k) ∪ I (n+1) (2 * k + 1) | inr h => simp only [I, inv_pow, mem_Ico, Int.cast_mul, Int.int_cast_ofNat, Int.cast_add, Int.cast_one, mem_union, not_lt.mpr h, and_false, h, true_and, false_or] - sorry + have : k * (2 ^ n)⁻¹ ≤ x := by + refine le_trans ?_ h + rw [pow_add, pow_one, mul_inv, mul_comm _ 2⁻¹, ← mul_assoc, mul_comm _ 2⁻¹, mul_add, + ← mul_assoc, inv_mul_cancel two_ne_zero, mul_one, one_mul, add_mul] + simp only [le_add_iff_nonneg_right, gt_iff_lt, inv_pos, zero_lt_two, + mul_nonneg_iff_of_pos_left, inv_nonneg] + positivity + simp only [this, true_and] + rw [pow_add, pow_one, mul_inv, mul_comm _ 2⁻¹, ← mul_assoc, mul_comm _ 2⁻¹, add_assoc] + norm_num + rw [one_div, mul_add, ← mul_assoc, inv_mul_cancel two_ne_zero, one_mul] -- todo : `Filtration` should be renamed to `filtration` def ℱ : Filtration ℕ (borel ℝ) where @@ -104,21 +124,7 @@ lemma measurable_indexI (n : ℕ) : Measurable[ℱ n] (indexI n) := by refine @measurable_to_countable' ℤ ℝ _ _ (ℱ n) _ (fun k ↦ ?_) have : (fun t ↦ ⌊t * (2 : ℝ) ^ n⌋) ⁻¹' {k} = I n k := by ext t - simp only [mem_preimage, mem_singleton_iff, I, inv_pow, mem_Ico] - rw [Int.floor_eq_iff] - refine ⟨fun ⟨h1, h2⟩ ↦ ⟨?_, ?_⟩, fun ⟨h1, h2⟩ ↦ ⟨?_, ?_⟩⟩ - · rw [mul_inv_le_iff, mul_comm] - · exact h1 - · positivity - · rw [← div_eq_mul_inv, lt_div_iff] - · exact h2 - · positivity - · rw [mul_inv_le_iff, mul_comm] at h1 - · exact h1 - · positivity - · rw [← div_eq_mul_inv, lt_div_iff] at h2 - · exact h2 - · positivity + simp only [mem_I_iff_floor, mem_preimage, mem_singleton_iff] rw [this] exact measurableSet_ℱ_I n k @@ -142,7 +148,11 @@ lemma iInter_biUnion_I (x : ℝ) : ⋂ n, ⋃ (k) (_ : indexI n x ≤ k), I n k ext t simp [iUnion_ge_I] refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ - · sorry + · by_contra h_lt + push_neg at h_lt + --have h_pos : ∀ i, 0 < (2 : ℝ) ^ i := fun i ↦ by positivity + --simp_rw [← div_eq_mul_inv, div_le_iff (h_pos _)] at h + sorry · intro n refine le_trans ?_ h rw [← div_eq_mul_inv, div_le_iff] @@ -783,25 +793,258 @@ lemma set_integral_mLimsup (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] exact fun i ↦ (hf i).prod hs · rwa [iSup_ℱ] at hf +lemma tendsto_integral_mLimsup_of_monotone (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] + (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) + (hs_meas : ∀ n, MeasurableSet (s n)) : + Tendsto (fun m ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop (𝓝 1) := by + simp_rw [integral_mLimsup κ a (hs_meas _)] + rw [← ENNReal.one_toReal] + have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := 1) ?_ + swap + · rw [mem_nhds_iff] + refine ⟨Iio 2, fun x hx ↦ ne_top_of_lt (?_ : x < 2), isOpen_Iio, ENNReal.one_lt_two⟩ + simpa using hx + refine h_cont.tendsto.comp ?_ + have h := tendsto_measure_iUnion (s := fun n ↦ univ ×ˢ s n) (μ := κ a) ?_ + swap; · intro n m hnm x; simp only [mem_prod, mem_univ, true_and]; exact fun h ↦ hs hnm h + convert h + rw [← prod_iUnion, hs_iUnion] + simp only [univ_prod_univ, measure_univ] + +lemma tendsto_integral_mLimsup_of_antitone (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] + (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) + (hs_meas : ∀ n, MeasurableSet (s n)) : + Tendsto (fun m ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop (𝓝 0) := by + simp_rw [integral_mLimsup κ a (hs_meas _)] + rw [← ENNReal.zero_toReal] + have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := 0) ?_ + swap + · rw [mem_nhds_iff] + refine ⟨Iio 1, fun x hx ↦ ne_top_of_lt (?_ : x < 1), isOpen_Iio, ?_⟩ + · simpa using hx + · simp + refine h_cont.tendsto.comp ?_ + have h := tendsto_measure_iInter (s := fun n ↦ univ ×ˢ s n) (μ := κ a) + (fun n ↦ MeasurableSet.univ.prod (hs_meas n)) ?_ ?_ + rotate_left + · intro n m hnm x; simp only [mem_prod, mem_univ, true_and]; exact fun h ↦ hs hnm h + · refine ⟨0, measure_ne_top _ _⟩ + convert h + rw [← prod_iInter, hs_iInter] + simp only [ne_eq, prod_empty, OuterMeasure.empty', forall_exists_index] + +lemma ae_eq_of_integral_eq_of_ae_le {μ : Measure α} {f g : α → ℝ} (hf : Integrable f μ) + (hg : Integrable g μ) (h_le : f ≤ᵐ[μ] g) (h_eq : ∫ a, f a ∂μ = ∫ a, g a ∂μ) : + f =ᵐ[μ] g := by + suffices g - f =ᵐ[μ] 0 by + filter_upwards [this] with a ha + symm + simpa only [Pi.sub_apply, Pi.zero_apply, sub_eq_zero] using ha + have h_eq' : ∫ a, (g - f) a ∂μ = 0 := by + simp_rw [Pi.sub_apply] + rwa [integral_sub hg hf, sub_eq_zero, eq_comm] + rwa [integral_eq_zero_iff_of_nonneg_ae _ (hg.sub hf)] at h_eq' + filter_upwards [h_le] with a ha + simpa + +lemma integral_tendsto_of_tendsto_of_monotone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} + (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Monotone fun n ↦ f n x) + (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : + Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by + let f' := fun n x ↦ f n x - f 0 x + have hf'_nonneg : ∀ᵐ x ∂μ, ∀ n, 0 ≤ f' n x := by + filter_upwards [h_mono] with a ha n + simp [ha (zero_le n)] + have hf'_meas : ∀ n, Integrable (f' n) μ := fun n ↦ (hf n).sub (hf 0) + suffices Tendsto (fun n ↦ ∫ x, f' n x ∂μ) atTop (𝓝 (∫ x, (F - f 0) x ∂μ)) by + rw [integral_sub' hF (hf 0)] at this + have h_sub : ∀ n, ∫ x, f' n x ∂μ = ∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ := by + intro n + simp only + rw [integral_sub (hf n) (hf 0)] + simp_rw [h_sub] at this + have h1 : (fun n ↦ ∫ x, f n x ∂μ) + = fun n ↦ (∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by ext n; abel + have h2 : ∫ x, F x ∂μ = (∫ x, F x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by abel + rw [h1, h2] + exact this.add tendsto_const_nhds + have hF_ge : 0 ≤ᵐ[μ] fun x ↦ (F - f 0) x := by + filter_upwards [h_tendsto, h_mono] with x hx_tendsto hx_mono + simp only [Pi.zero_apply, Pi.sub_apply, sub_nonneg] + exact ge_of_tendsto' hx_tendsto (fun n ↦ hx_mono (zero_le _)) + rw [ae_all_iff] at hf'_nonneg + simp_rw [integral_eq_lintegral_of_nonneg_ae (hf'_nonneg _) (hf'_meas _).1] + rw [integral_eq_lintegral_of_nonneg_ae hF_ge (hF.1.sub (hf 0).1)] + have h_cont := ENNReal.continuousOn_toReal.continuousAt + (x := ∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ) ?_ + swap + · rw [mem_nhds_iff] + refine ⟨Iio (∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ + 1), ?_, isOpen_Iio, ?_⟩ + · intro x + simp only [Pi.sub_apply, mem_Iio, ne_eq, mem_setOf_eq] + exact ne_top_of_lt + · simp only [Pi.sub_apply, mem_Iio] + refine ENNReal.lt_add_right ?_ one_ne_zero + rw [← ofReal_integral_eq_lintegral_ofReal] + · exact ENNReal.ofReal_ne_top + · exact hF.sub (hf 0) + · exact hF_ge + refine h_cont.tendsto.comp ?_ + refine lintegral_tendsto_of_tendsto_of_monotone ?_ ?_ ?_ + · exact fun n ↦ ((hf n).sub (hf 0)).aemeasurable.ennreal_ofReal + · filter_upwards [h_mono] with x hx + intro n m hnm + refine ENNReal.ofReal_le_ofReal ?_ + simp only [tsub_le_iff_right, sub_add_cancel] + exact hx hnm + · filter_upwards [h_tendsto] with x hx + refine (ENNReal.continuous_ofReal.tendsto _).comp ?_ + simp only [Pi.sub_apply] + exact Tendsto.sub hx tendsto_const_nhds + +lemma integral_tendsto_of_tendsto_of_antitone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} + (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Antitone fun n ↦ f n x) + (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : + Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by + suffices Tendsto (fun n ↦ ∫ x, -f n x ∂μ) atTop (𝓝 (∫ x, -F x ∂μ)) by + suffices Tendsto (fun n ↦ ∫ x, - -f n x ∂μ) atTop (𝓝 (∫ x, - -F x ∂μ)) by + simp_rw [neg_neg] at this + exact this + convert this.neg <;> rw [integral_neg] + refine integral_tendsto_of_tendsto_of_monotone (fun n ↦ (hf n).neg) hF.neg ?_ ?_ + · filter_upwards [h_mono] with x hx + intro n m hnm + simp only [neg_le_neg_iff] + exact hx hnm + · filter_upwards [h_tendsto] with x hx + exact hx.neg + +theorem tendsto_atTop_atBot_iff_of_antitone {α β : Type*} + [Nonempty α] [SemilatticeSup α] [Preorder β] {f : α → β} + (hf : Antitone f) : + Tendsto f atTop atBot ↔ ∀ b : β, ∃ a : α, f a ≤ b := + @tendsto_atTop_atTop_iff_of_monotone _ βᵒᵈ _ _ _ _ hf + lemma tendsto_mLimsup_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 1) := by - have h1 := tendsto_m_atTop_ae_of_monotone κ a s hs hs_iUnion - have h2 := fun (n : ℕ) ↦ tendsto_m_mLimsup κ a (hs_meas n) - rw [← ae_all_iff] at h1 h2 - filter_upwards [h1, h2] with t h_tendsto_set h_tendsto_nat - sorry + have h_mono : ∀ t, Monotone (fun m ↦ MLimsup κ a (s m) t) := + fun t n m hnm ↦ mLimsup_mono_set κ a (hs hnm) t + have h_le_one : ∀ m t, MLimsup κ a (s m) t ≤ 1 := fun m t ↦ mLimsup_le_one κ a (s m) t + -- for all `t`, `fun m ↦ MLimsup κ a (s m) t` has a limit + have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 l) := by + intro t + have h_tendsto : Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop atTop ∨ + ∃ l, Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 l) := + tendsto_of_monotone (h_mono t) + cases' h_tendsto with h_absurd h_tendsto + · rw [tendsto_atTop_atTop_iff_of_monotone (h_mono t)] at h_absurd + obtain ⟨r, hr⟩ := h_absurd 2 + exact absurd (hr.trans (h_le_one r t)) one_lt_two.not_le + · exact h_tendsto + -- let `F` be the pointwise limit of `fun m ↦ MLimsup κ a (s m) t` for all `t` + let F : ℝ → ℝ := fun t ↦ (h_exists t).choose + have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 (F t)) := + fun t ↦ (h_exists t).choose_spec + have hF_nonneg : ∀ t, 0 ≤ F t := + fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ mLimsup_nonneg κ a (s m) t) + have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) + have hF_int : Integrable F (kernel.fst κ a) := by + rw [← memℒp_one_iff_integrable] + refine ⟨?_, ?_⟩ + · refine aestronglyMeasurable_of_tendsto_ae atTop (fun n ↦ ?_) (ae_of_all _ hF_tendsto) + exact (measurable_mLimsup_right κ (hs_meas _) a).aestronglyMeasurable + · rw [snorm_one_eq_lintegral_nnnorm] + calc ∫⁻ x, ‖F x‖₊ ∂(kernel.fst κ a) ≤ ∫⁻ _, 1 ∂(kernel.fst κ a) := by + refine lintegral_mono (fun x ↦ ?_) + rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, ENNReal.ofReal_le_one, + abs_of_nonneg (hF_nonneg _)] + exact hF_le_one _ + _ < ⊤ := by simp only [lintegral_const, measure_univ, mul_one, ENNReal.one_lt_top] + -- it suffices to show that the limit `F` is 1 a.e. + suffices ∀ᵐ t ∂(kernel.fst κ a), F t = 1 by + filter_upwards [this] with t ht_eq + rw [← ht_eq] + exact hF_tendsto t + -- since `F` is at most 1, proving that its integral is the same as the integral of 1 will tell + -- us that `F` is 1 a.e. + refine ae_eq_of_integral_eq_of_ae_le hF_int (integrable_const _) (ae_of_all _ hF_le_one) ?_ + have h_integral : + Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop + (𝓝 (∫ t, F t ∂(kernel.fst κ a))) := by + refine integral_tendsto_of_tendsto_of_monotone ?_ hF_int ?_ ?_ + · exact fun n ↦ integrable_mLimsup _ _ (hs_meas n) + · exact ae_of_all _ h_mono + · exact ae_of_all _ hF_tendsto + have h_integral' : + Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop + (𝓝 (∫ _, 1 ∂(kernel.fst κ a))) := by + rw [integral_const, measure_univ] + simp only [ENNReal.one_toReal, smul_eq_mul, mul_one] + exact tendsto_integral_mLimsup_of_monotone κ a s hs hs_iUnion hs_meas + exact tendsto_nhds_unique h_integral h_integral' lemma tendsto_mLimsup_atTop_ae_of_antitone (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 0) := by - have h1 := tendsto_m_atTop_of_antitone κ a s hs hs_iInter hs_meas - have h2 := fun (n : ℕ) ↦ tendsto_m_mLimsup κ a (hs_meas n) - rw [← ae_all_iff] at h2 - filter_upwards [h2] with t h_tendsto_nat - sorry + have h_anti : ∀ t, Antitone (fun m ↦ MLimsup κ a (s m) t) := + fun t n m hnm ↦ mLimsup_mono_set κ a (hs hnm) t + have h_le_one : ∀ m t, MLimsup κ a (s m) t ≤ 1 := fun m t ↦ mLimsup_le_one κ a (s m) t + -- for all `t`, `fun m ↦ MLimsup κ a (s m) t` has a limit + have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 l) := by + intro t + have h_tendsto : Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop atBot ∨ + ∃ l, Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 l) := + tendsto_of_antitone (h_anti t) + cases' h_tendsto with h_absurd h_tendsto + · rw [tendsto_atTop_atBot_iff_of_antitone (h_anti t)] at h_absurd + obtain ⟨r, hr⟩ := h_absurd (-1) + have h_nonneg := mLimsup_nonneg κ a (s r) t + linarith + · exact h_tendsto + -- let `F` be the pointwise limit of `fun m ↦ MLimsup κ a (s m) t` for all `t` + let F : ℝ → ℝ := fun t ↦ (h_exists t).choose + have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 (F t)) := + fun t ↦ (h_exists t).choose_spec + have hF_nonneg : ∀ t, 0 ≤ F t := + fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ mLimsup_nonneg κ a (s m) t) + have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) + have hF_int : Integrable F (kernel.fst κ a) := by + rw [← memℒp_one_iff_integrable] + refine ⟨?_, ?_⟩ + · refine aestronglyMeasurable_of_tendsto_ae atTop (fun n ↦ ?_) (ae_of_all _ hF_tendsto) + exact (measurable_mLimsup_right κ (hs_meas _) a).aestronglyMeasurable + · rw [snorm_one_eq_lintegral_nnnorm] + calc ∫⁻ x, ‖F x‖₊ ∂(kernel.fst κ a) ≤ ∫⁻ _, 1 ∂(kernel.fst κ a) := by + refine lintegral_mono (fun x ↦ ?_) + rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, ENNReal.ofReal_le_one, + abs_of_nonneg (hF_nonneg _)] + exact hF_le_one _ + _ < ⊤ := by simp only [lintegral_const, measure_univ, mul_one, ENNReal.one_lt_top] + -- it suffices to show that the limit `F` is 0 a.e. + suffices ∀ᵐ t ∂(kernel.fst κ a), F t = 0 by + filter_upwards [this] with t ht_eq + rw [← ht_eq] + exact hF_tendsto t + -- since `F` is nonnegative, proving that its integral is 0 is sufficient to get that + -- `F` is 0 a.e. + suffices ∀ᵐ (t : ℝ) ∂(kernel.fst κ) a, 0 = F t by filter_upwards [this] with a ha; simp [ha] + refine ae_eq_of_integral_eq_of_ae_le (integrable_const _) hF_int (ae_of_all _ hF_nonneg) ?_ + have h_integral : + Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop + (𝓝 (∫ t, F t ∂(kernel.fst κ a))) := by + refine integral_tendsto_of_tendsto_of_antitone ?_ hF_int ?_ ?_ + · exact fun n ↦ integrable_mLimsup _ _ (hs_meas n) + · exact ae_of_all _ h_anti + · exact ae_of_all _ hF_tendsto + have h_integral' : + Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop + (𝓝 (∫ _, 0 ∂(kernel.fst κ a))) := by + simp only [integral_zero] + exact tendsto_integral_mLimsup_of_antitone κ a s hs hs_iInter hs_meas + exact (tendsto_nhds_unique h_integral h_integral').symm section Iic_Q @@ -829,10 +1072,24 @@ lemma mLimsupIic_nonneg (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : lemma mLimsupIic_le_one (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : mLimsupIic κ a t q ≤ 1 := mLimsup_le_one κ a _ t +theorem tendsto_nat_ceil_atTop {α : Type*} [LinearOrderedSemiring α] [FloorSemiring α] : + Tendsto (fun x : α ↦ ⌈x⌉₊) atTop atTop := by + refine Nat.ceil_mono.tendsto_atTop_atTop (fun x ↦ ⟨x, ?_⟩) + simp only [Nat.ceil_natCast, le_refl] + lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ a t q) atTop (𝓝 1) := by suffices ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ a t n) atTop (𝓝 1) by - sorry + filter_upwards [this] with t ht + let f := fun q : ℚ ↦ mLimsupIic κ a t ⌊q⌋₊ + let g := fun q : ℚ ↦ mLimsupIic κ a t ⌈q⌉₊ + have hf_le : ∀ᶠ q in atTop, f q ≤ mLimsupIic κ a t q := by + simp only [eventually_atTop, ge_iff_le] + exact ⟨0, fun q hq ↦ monotone_mLimsupIic κ a t (Nat.floor_le hq)⟩ + have hg_le : ∀ q, mLimsupIic κ a t q ≤ g q := fun q ↦ monotone_mLimsupIic κ a t (Nat.le_ceil _) + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ hf_le (eventually_of_forall hg_le) + · exact ht.comp tendsto_nat_floor_atTop + · exact ht.comp tendsto_nat_ceil_atTop let s : ℕ → Set ℝ := fun n ↦ Iic n have hs : Monotone s := fun i j hij ↦ Iic_subset_Iic.mpr (by exact mod_cast hij) have hs_iUnion : ⋃ i, s i = univ := by @@ -852,7 +1109,21 @@ lemma tendsto_atBot_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] rw [h_eq_neg] exact ht.comp tendsto_neg_atBot_atTop suffices ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ a t (-n)) atTop (𝓝 0) by - sorry + filter_upwards [this] with t ht + let f := fun q : ℚ ↦ mLimsupIic κ a t (-⌊q⌋₊) + let g := fun q : ℚ ↦ mLimsupIic κ a t (-⌈q⌉₊) + have hf_le : ∀ᶠ q in atTop, mLimsupIic κ a t (-q) ≤ f q := by + simp only [eventually_atTop, ge_iff_le] + refine ⟨0, fun q hq ↦ monotone_mLimsupIic κ a t ?_⟩ + rw [neg_le_neg_iff] + exact Nat.floor_le hq + have hg_le : ∀ q, g q ≤ mLimsupIic κ a t (-q) := by + refine fun q ↦ monotone_mLimsupIic κ a t ?_ + rw [neg_le_neg_iff] + exact Nat.le_ceil _ + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ (eventually_of_forall hg_le) hf_le + · exact ht.comp tendsto_nat_ceil_atTop + · exact ht.comp tendsto_nat_floor_atTop let s : ℕ → Set ℝ := fun n ↦ Iic (-n) have hs : Antitone s := fun i j hij ↦ Iic_subset_Iic.mpr (neg_le_neg (by exact mod_cast hij)) have hs_iInter : ⋂ i, s i = ∅ := by @@ -948,234 +1219,183 @@ lemma isRatStieltjesPoint_mLimsupIic_ae (κ : kernel α (ℝ × ℝ)) [IsMarkovK · exact ht_bot · exact ht_iInf -lemma todo3_mLimsupIic_ae_eq (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) (q : ℚ) : - (fun t ↦ todo3 _ (measurable_mLimsupIic κ) (a, t) q) - =ᵐ[kernel.fst κ a] fun t ↦ mLimsupIic κ a t q := by - filter_upwards [isRatStieltjesPoint_mLimsupIic_ae κ a] with a ha - rw [todo3_eq, toCDFLike_of_isRatStieltjesPoint ha] +lemma isRatKernelCDF_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : + IsRatKernelCDF (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2) κ (kernel.fst κ) where + measurable := measurable_mLimsupIic κ + isRatStieltjesPoint_ae := isRatStieltjesPoint_mLimsupIic_ae κ + integrable := integrable_mLimsupIic κ + isCDF := fun _ _ hs _ ↦ set_integral_mLimsupIic _ _ _ hs end Rat -- todo: name? noncomputable def kernel.condexpReal (κ : kernel α (ℝ × ℝ)) : kernel (α × ℝ) ℝ := - cdfKernel (measurable_mLimsupIic κ) + cdfKernel _ (measurable_mLimsupIic κ) instance (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : IsMarkovKernel (kernel.condexpReal κ) := by unfold kernel.condexpReal; infer_instance -lemma condexpReal_Iic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) (t x : ℝ) : - kernel.condexpReal κ (a, t) (Iic x) - = ENNReal.ofReal (todo3 _ (measurable_mLimsupIic κ) (a, t) x) := - cdfKernel_Iic _ _ - -lemma set_lintegral_condexpReal_Iic_rat (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) (q : ℚ) - {s : Set ℝ} (hs : MeasurableSet s) : - ∫⁻ t in s, kernel.condexpReal κ (a, t) (Iic q) ∂(kernel.fst κ a) = κ a (s ×ˢ Iic (q : ℝ)) := by - simp_rw [condexpReal_Iic] - rw [← ofReal_integral_eq_lintegral_ofReal] - · rw [set_integral_congr_ae (g := fun t ↦ mLimsupIic κ a t q) hs, - set_integral_mLimsupIic κ _ _ hs, ENNReal.ofReal_toReal] - · exact measure_ne_top _ _ - · filter_upwards [todo3_mLimsupIic_ae_eq κ a q] with t ht - exact fun _ ↦ ht - · refine Integrable.restrict ?_ - rw [integrable_congr (todo3_mLimsupIic_ae_eq κ a q)] - exact integrable_mLimsupIic _ _ _ - · exact ae_of_all _ (fun x ↦ todo3_nonneg _ _ _) - -lemma set_lintegral_condexpReal_Iic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) (x : ℝ) - {s : Set ℝ} (hs : MeasurableSet s) : - ∫⁻ t in s, kernel.condexpReal κ (a, t) (Iic x) ∂(kernel.fst κ a) = κ a (s ×ˢ Iic x) := by - -- We have the result for `x : ℚ` thanks to `set_lintegral_condexpReal_Iic_rat`. - -- We use the equality `condCDF ρ a x = ⨅ r : {r' : ℚ // x < r'}, condCDF ρ a r` and a monotone - -- convergence argument to extend it to the reals. - by_cases hρ_zero : (kernel.fst κ a).restrict s = 0 - · rw [hρ_zero, lintegral_zero_measure] - refine le_antisymm (zero_le _) ?_ - calc κ a (s ×ˢ Iic x) - ≤ κ a (Prod.fst ⁻¹' s) := measure_mono (prod_subset_preimage_fst s (Iic x)) - _ = kernel.fst κ a s := by rw [kernel.fst_apply' _ _ hs]; rfl - _ = (kernel.fst κ a).restrict s univ := by rw [Measure.restrict_apply_univ] - _ = 0 := by simp only [hρ_zero, Measure.coe_zero, Pi.zero_apply] - have h : ∫⁻ t in s, kernel.condexpReal κ (a, t) (Iic x) ∂(kernel.fst κ a) - = ∫⁻ t in s, ⨅ r : { r' : ℚ // x < r' }, - kernel.condexpReal κ (a, t) (Iic r) ∂(kernel.fst κ a) := by - congr with t : 1 - rw [← measure_iInter_eq_iInf] - · congr with y : 1 - simp only [mem_Iic, mem_iInter, Subtype.forall] - refine ⟨fun h a ha ↦ h.trans ?_, fun h ↦ ?_⟩ - · exact mod_cast ha.le - · refine le_of_forall_lt_rat_imp_le fun q hq ↦ h q ?_ - exact mod_cast hq - · exact fun _ ↦ measurableSet_Iic - · refine Monotone.directed_ge fun r r' hrr' ↦ ?_ - refine Iic_subset_Iic.mpr ?_ - exact mod_cast hrr' - · obtain ⟨q, hq⟩ := exists_rat_gt x - exact ⟨⟨q, hq⟩, measure_ne_top _ _⟩ - have h_nonempty : Nonempty { r' : ℚ // x < ↑r' } := by - obtain ⟨r, hrx⟩ := exists_rat_gt x - exact ⟨⟨r, hrx⟩⟩ - rw [h, lintegral_iInf_directed_of_measurable hρ_zero fun q : { r' : ℚ // x < ↑r' } ↦ ?_] - rotate_left - · intro b - rw [set_lintegral_condexpReal_Iic_rat _ _ _ hs] - exact measure_ne_top _ _ - · refine Monotone.directed_ge fun i j hij t ↦ measure_mono (Iic_subset_Iic.mpr ?_) - exact mod_cast hij - · exact (kernel.measurable_coe _ measurableSet_Iic).comp measurable_prod_mk_left - simp_rw [set_lintegral_condexpReal_Iic_rat κ _ _ hs] - rw [← measure_iInter_eq_iInf] - · rw [← prod_iInter] - congr with y - simp only [mem_iInter, mem_Iic, Subtype.forall, Subtype.coe_mk] - exact ⟨le_of_forall_lt_rat_imp_le, fun hyx q hq ↦ hyx.trans hq.le⟩ - · exact fun i ↦ hs.prod measurableSet_Iic - · refine Monotone.directed_ge fun i j hij ↦ ?_ - refine prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, Iic_subset_Iic.mpr ?_⟩) - exact mod_cast hij - · exact ⟨h_nonempty.some, measure_ne_top _ _⟩ - -lemma set_lintegral_condexpReal_univ (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) - {s : Set ℝ} (hs : MeasurableSet s) : - ∫⁻ t in s, kernel.condexpReal κ (a, t) univ ∂(kernel.fst κ a) = κ a (s ×ˢ univ) := by - simp only [measure_univ, lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter, - one_mul, kernel.fst_apply' _ _ hs] - congr with x - simp only [mem_prod, mem_univ, and_true, mem_setOf_eq] - -lemma lintegral_condexpReal_univ (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) : - ∫⁻ t, kernel.condexpReal κ (a, t) univ ∂(kernel.fst κ a) = κ a univ := by - rw [← set_lintegral_univ, set_lintegral_condexpReal_univ κ a MeasurableSet.univ, - univ_prod_univ] - -lemma set_lintegral_condexpReal_prod (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) - {s t : Set ℝ} (hs : MeasurableSet s) (ht : MeasurableSet t) : - ∫⁻ x in s, kernel.condexpReal κ (a, x) t ∂(kernel.fst κ a) = κ a (s ×ˢ t) := by - -- `set_lintegral_condKernelReal_Iic` gives the result for `t = Iic x`. These sets form a - -- π-system that generates the Borel σ-algebra, hence we can get the same equality for any - -- measurable set `t`. - apply MeasurableSpace.induction_on_inter (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic _ _ _ _ ht - · simp only [measure_empty, lintegral_const, zero_mul, prod_empty] - · rintro t ⟨q, rfl⟩ - exact set_lintegral_condexpReal_Iic κ a _ hs - · intro t ht ht_lintegral - calc ∫⁻ x in s, kernel.condexpReal κ (a, x) tᶜ ∂(kernel.fst κ a) - = ∫⁻ x in s, kernel.condexpReal κ (a, x) univ - kernel.condexpReal κ (a, x) t ∂(kernel.fst κ a) := by - congr with x; rw [measure_compl ht (measure_ne_top (kernel.condexpReal κ (a, x)) _)] - _ = ∫⁻ x in s, kernel.condexpReal κ (a, x) univ ∂(kernel.fst κ a) - - ∫⁻ x in s, kernel.condexpReal κ (a, x) t ∂(kernel.fst κ a) := by - rw [lintegral_sub] - · exact (kernel.measurable_coe (kernel.condexpReal κ) ht).comp measurable_prod_mk_left - · rw [ht_lintegral] - exact measure_ne_top _ _ - · exact eventually_of_forall fun a ↦ measure_mono (subset_univ _) - _ = κ a (s ×ˢ univ) - κ a (s ×ˢ t) := by - rw [set_lintegral_condexpReal_univ κ a hs, ht_lintegral] - _ = κ a (s ×ˢ tᶜ) := by - rw [← measure_diff _ (hs.prod ht) (measure_ne_top _ _)] - · rw [prod_diff_prod, compl_eq_univ_diff] - simp only [diff_self, empty_prod, union_empty] - · rw [prod_subset_prod_iff] - exact Or.inl ⟨subset_rfl, subset_univ t⟩ - · intro f hf_disj hf_meas hf_eq - simp_rw [measure_iUnion hf_disj hf_meas] - rw [lintegral_tsum, prod_iUnion, measure_iUnion] - · simp_rw [hf_eq] - · intro i j hij - rw [Function.onFun, Set.disjoint_prod] - exact Or.inr (hf_disj hij) - · exact fun i ↦ MeasurableSet.prod hs (hf_meas i) - · exact fun i ↦ - ((kernel.measurable_coe _ (hf_meas i)).comp measurable_prod_mk_left).aemeasurable.restrict - -lemma lintegral_condexpReal_mem (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) - {s : Set (ℝ × ℝ)} (hs : MeasurableSet s) : - ∫⁻ x, kernel.condexpReal κ (a, x) {y | (x, y) ∈ s} ∂(kernel.fst κ a) = κ a s := by - -- `set_lintegral_condKernelReal_prod` gives the result for sets of the form `t₁ × t₂`. These - -- sets form a π-system that generates the product σ-algebra, hence we can get the same equality - -- for any measurable set `s`. - apply MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod _ _ _ _ hs - · simp only [mem_empty_iff_false, setOf_false, measure_empty, lintegral_const, - zero_mul] - · rintro _ ⟨t₁, ht₁, t₂, ht₂, rfl⟩ - simp only [mem_setOf_eq] at ht₁ ht₂ - have h_prod_eq_snd : ∀ a ∈ t₁, {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} = t₂ := by - intro a ha - simp only [ha, prod_mk_mem_set_prod_eq, true_and_iff, setOf_mem_eq] - rw [← lintegral_add_compl _ ht₁] - have h_eq1 : ∫⁻ x in t₁, kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(kernel.fst κ a) - = ∫⁻ x in t₁, kernel.condexpReal κ (a, x) t₂ ∂(kernel.fst κ a) := by - refine' set_lintegral_congr_fun ht₁ (eventually_of_forall fun a ha ↦ _) - rw [h_prod_eq_snd a ha] - have h_eq2 : ∫⁻ x in t₁ᶜ, kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(kernel.fst κ a) = 0 := by - suffices h_eq_zero : ∀ x ∈ t₁ᶜ, kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} = 0 - · rw [set_lintegral_congr_fun ht₁.compl (eventually_of_forall h_eq_zero)] - simp only [lintegral_const, zero_mul] - intro a hat₁ - rw [mem_compl_iff] at hat₁ - simp only [hat₁, prod_mk_mem_set_prod_eq, false_and_iff, setOf_false, measure_empty] - rw [h_eq1, h_eq2, add_zero] - exact set_lintegral_condexpReal_prod κ a ht₁ ht₂ - · intro t ht ht_eq - calc ∫⁻ x, kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ tᶜ} ∂(kernel.fst κ a) - = ∫⁻ x, kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t}ᶜ ∂(kernel.fst κ a) := rfl - _ = ∫⁻ x, kernel.condexpReal κ (a, x) univ - - kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t} ∂(kernel.fst κ a) := by - congr with x : 1 - exact measure_compl (measurable_prod_mk_left ht) - (measure_ne_top (kernel.condexpReal κ (a, x)) _) - _ = ∫⁻ x, kernel.condexpReal κ (a, x) univ ∂(kernel.fst κ a) - - ∫⁻ x, kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t} ∂(kernel.fst κ a) := by - have h_le : (fun x ↦ kernel.condexpReal κ (a, x) {y : ℝ | (x, y) ∈ t}) - ≤ᵐ[kernel.fst κ a] fun x ↦ kernel.condexpReal κ (a, x) univ := - eventually_of_forall fun _ ↦ measure_mono (subset_univ _) - rw [lintegral_sub _ _ h_le] - · exact kernel.measurable_kernel_prod_mk_left' ht a - refine ((lintegral_mono_ae h_le).trans_lt ?_).ne - rw [lintegral_condexpReal_univ] - exact measure_lt_top _ univ - _ = κ a univ - κ a t := by rw [ht_eq, lintegral_condexpReal_univ] - _ = κ a tᶜ := (measure_compl ht (measure_ne_top _ _)).symm - · intro f hf_disj hf_meas hf_eq - have h_eq : ∀ a, {x | (a, x) ∈ ⋃ i, f i} = ⋃ i, {x | (a, x) ∈ f i} := by - intro a; ext x; simp only [mem_iUnion, mem_setOf_eq] - simp_rw [h_eq] - have h_disj : ∀ a, Pairwise (Disjoint on fun i ↦ {x | (a, x) ∈ f i}) := by - intro a i j hij - have h_disj := hf_disj hij - rw [Function.onFun, disjoint_iff_inter_eq_empty] at h_disj ⊢ - ext1 x - simp only [mem_inter_iff, mem_setOf_eq, mem_empty_iff_false, iff_false_iff] - intro h_mem_both - suffices (a, x) ∈ ∅ by rwa [mem_empty_iff_false] at this - rwa [← h_disj, mem_inter_iff] - calc ∫⁻ x, kernel.condexpReal κ (a, x) (⋃ i, {y | (x, y) ∈ f i}) ∂(kernel.fst κ a) - = ∫⁻ x, ∑' i, kernel.condexpReal κ (a, x) {y | (x, y) ∈ f i} ∂(kernel.fst κ a) := by - congr with x : 1 - rw [measure_iUnion (h_disj x) fun i ↦ measurable_prod_mk_left (hf_meas i)] - _ = ∑' i, ∫⁻ x, kernel.condexpReal κ (a, x) {y | (x, y) ∈ f i} ∂(kernel.fst κ a) := - lintegral_tsum fun i ↦ (kernel.measurable_kernel_prod_mk_left' (hf_meas i) a).aemeasurable - _ = ∑' i, κ a (f i) := by simp_rw [hf_eq] - _ = κ a (iUnion f) := (measure_iUnion hf_disj hf_meas).symm - lemma kernel.eq_compProd_condexpReal (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : - κ = kernel.fst κ ⊗ₖ kernel.condexpReal κ := by - ext a s hs - rw [kernel.compProd_apply _ _ _ hs, lintegral_condexpReal_mem κ a hs] + κ = kernel.fst κ ⊗ₖ kernel.condexpReal κ := + kernel.eq_compProd_cdfKernel (isRatKernelCDF_mLimsupIic κ) end Real variable {Ω' : Type*} [MeasurableSpace Ω'] [StandardBorelSpace Ω'] [Nonempty Ω'] -def kernel.condexp (κ : kernel α (Ω × Ω')) [IsMarkovKernel (kernel.fst κ)] : - kernel (α × Ω) Ω' := - sorry +noncomputable +def measurableEmbedding_real (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : Ω → ℝ := + (exists_measurableEmbedding_real Ω).choose + +lemma measurableEmbedding_measurableEmbedding_real + (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : + MeasurableEmbedding (measurableEmbedding_real Ω) := + (exists_measurableEmbedding_real Ω).choose_spec + +noncomputable +def condKernelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel κ] : kernel (α × ℝ) Ω' := + let f := measurableEmbedding_real Ω' + let hf := measurableEmbedding_measurableEmbedding_real Ω' + let κ' := kernel.map κ (Prod.map (id : ℝ → ℝ) f) (measurable_id.prod_map hf.measurable) + let x₀ := (range_nonempty f).choose + kernel.comapRight + (kernel.piecewise (measurableSet_eq_one (isRatKernelCDF_mLimsupIic κ') hf.measurableSet_range) + (kernel.condexpReal κ') (kernel.deterministic (fun _ ↦ x₀) measurable_const)) + hf + +instance instIsMarkovKernel_condKernelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel κ] : + IsMarkovKernel (condKernelAux κ) := by + rw [condKernelAux] + refine kernel.IsMarkovKernel.comapRight _ _ fun a ↦ ?_ + rw [kernel.piecewise_apply'] + split_ifs with h_mem + · exact h_mem + · classical + rw [kernel.deterministic_apply' _ _ + (measurableEmbedding_measurableEmbedding_real Ω').measurableSet_range, + Set.indicator_apply, if_pos] + exact (range_nonempty (measurableEmbedding_real Ω')).choose_spec + +lemma compProd_fst_condKernelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel κ] : + kernel.fst κ ⊗ₖ condKernelAux κ = κ := by + let f := measurableEmbedding_real Ω' + let hf := measurableEmbedding_measurableEmbedding_real Ω' + let κ' := kernel.map κ (Prod.map (id : ℝ → ℝ) f) (measurable_id.prod_map hf.measurable) + have h_prod_embed : MeasurableEmbedding (Prod.map (id : ℝ → ℝ) f) := + MeasurableEmbedding.id.prod_mk hf + have h_fst : kernel.fst κ' = kernel.fst κ := by + ext a u hu + unfold_let κ' + rw [kernel.fst_apply' _ _ hu, kernel.fst_apply' _ _ hu, + kernel.map_apply' κ h_prod_embed.measurable] + · rfl + · exact measurable_fst hu + have : κ = kernel.comapRight κ' h_prod_embed := by + ext c t ht : 2 + unfold_let κ' + rw [kernel.comapRight_apply' _ _ _ ht, kernel.map_apply' κ h_prod_embed.measurable + _ (h_prod_embed.measurableSet_image.mpr ht)] + congr with x : 1 + rw [← @Prod.mk.eta _ _ x] + simp only [id.def, mem_preimage, Prod.map_mk, mem_image, Prod.mk.inj_iff, Prod.exists] + refine' ⟨fun h => ⟨x.1, x.2, h, rfl, rfl⟩, _⟩ + rintro ⟨a, b, h_mem, rfl, hf_eq⟩ + rwa [hf.injective hf_eq] at h_mem + conv_rhs => rw [this, kernel.eq_compProd_condexpReal κ'] + ext c t ht : 2 + rw [kernel.comapRight_apply' _ _ _ ht, + kernel.compProd_apply _ _ _ (h_prod_embed.measurableSet_image.mpr ht), + h_fst, kernel.compProd_apply _ _ _ ht] + refine lintegral_congr_ae ?_ + let ρ_set := {p : α × ℝ | kernel.condexpReal κ' p (range f) = 1} + have h_ae : ∀ a, ∀ᵐ t ∂(kernel.fst κ a), (a, t) ∈ ρ_set := by + intro a + rw [← h_fst] + refine ae_cdfKernel_eq_one (isRatKernelCDF_mLimsupIic κ') a hf.measurableSet_range ?_ + simp only [mem_compl_iff, mem_range, not_exists] + rw [kernel.map_apply'] + · have h_empty : {a : ℝ × Ω' | ∀ (x : Ω'), ¬f x = f a.2} = ∅ := by + ext x + simp only [mem_setOf_eq, mem_empty_iff_false, iff_false, not_forall, not_not, + exists_apply_eq_apply] + simp [h_empty] + · have : {x : ℝ × ℝ | ∀ (y : Ω'), ¬ f y = x.2} = univ ×ˢ (range f)ᶜ := by + ext x + simp only [mem_setOf_eq, mem_prod, mem_univ, mem_compl_iff, mem_range, not_exists, true_and] + rw [this] + exact MeasurableSet.univ.prod hf.measurableSet_range.compl + filter_upwards [h_ae c] with a ha + rw [condKernelAux, kernel.comapRight_apply'] + swap; · exact measurable_prod_mk_left ht + have h1 : {c : ℝ | (a, c) ∈ Prod.map id f '' t} = f '' {c : Ω' | (a, c) ∈ t} := by + ext1 x + simp only [Prod_map, id.def, mem_image, Prod.mk.inj_iff, Prod.exists, mem_setOf_eq] + constructor + · rintro ⟨a', b, h_mem, rfl, hf_eq⟩ + exact ⟨b, h_mem, hf_eq⟩ + · rintro ⟨b, h_mem, hf_eq⟩ + exact ⟨a, b, h_mem, rfl, hf_eq⟩ + rw [h1, kernel.piecewise_apply, if_pos] + exact ha -theorem kernel.eq_compProd (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : - κ = kernel.fst κ ⊗ₖ kernel.condexp κ := by - sorry +noncomputable +def condKernel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : kernel (α × Ω) Ω' := + let f := measurableEmbedding_real Ω + let hf := measurableEmbedding_measurableEmbedding_real Ω + let κ' := kernel.map κ (Prod.map f (id : Ω' → Ω')) (hf.measurable.prod_map measurable_id) + kernel.comap (condKernelAux κ') (Prod.map (id : α → α) f) (measurable_id.prod_map hf.measurable) + +instance instIsMarkovKernel_condKernel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : + IsMarkovKernel (condKernel κ) := by rw [condKernel]; infer_instance + +lemma compProd_fst_condKernel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : + kernel.fst κ ⊗ₖ condKernel κ = κ := by + let f := measurableEmbedding_real Ω + let hf := measurableEmbedding_measurableEmbedding_real Ω + let κ' : kernel α (ℝ × Ω') := + kernel.map κ (Prod.map f (id : Ω' → Ω')) (hf.measurable.prod_map measurable_id) + have h_condKernel : condKernel κ + = kernel.comap (condKernelAux κ') (Prod.map (id : α → α) f) + (measurable_id.prod_map hf.measurable) := rfl + have h_compProd := compProd_fst_condKernelAux κ' + have h_prod_embed : MeasurableEmbedding (Prod.map f (id : Ω' → Ω')) := + hf.prod_mk MeasurableEmbedding.id + have : κ = kernel.comapRight κ' h_prod_embed := by + ext c t ht : 2 + unfold_let κ' + rw [kernel.comapRight_apply' _ _ _ ht, kernel.map_apply' κ h_prod_embed.measurable + _ (h_prod_embed.measurableSet_image.mpr ht)] + congr with x : 1 + rw [← @Prod.mk.eta _ _ x] + simp only [Prod.mk.eta, Prod_map, id_eq, mem_preimage, mem_image, Prod.mk.injEq, Prod.exists, + exists_eq_right_right] + refine ⟨fun h ↦ ⟨x.1, h, rfl⟩, ?_⟩ + rintro ⟨ω, h_mem, h⟩ + rwa [hf.injective h] at h_mem + have h_fst : kernel.fst κ' = kernel.map (kernel.fst κ) f hf.measurable := by + ext a s hs + unfold_let κ' + rw [kernel.map_apply' _ _ _ hs, kernel.fst_apply', kernel.fst_apply', kernel.map_apply'] + · congr + · exact measurable_fst hs + · exact hf.measurable hs + · exact hs + conv_rhs => rw [this, ← h_compProd] + ext a s hs + rw [h_condKernel, h_fst] + rw [kernel.comapRight_apply' _ _ _ hs, kernel.compProd_apply _ _ _ hs, kernel.compProd_apply] + swap; · exact h_prod_embed.measurableSet_image.mpr hs + rw [kernel.map_apply, lintegral_map] + congr with ω + rw [kernel.comap_apply'] + · congr with ω' + simp only [mem_setOf_eq, Prod_map, id_eq, mem_image, Prod.mk.injEq, Prod.exists, + exists_eq_right_right] + refine ⟨fun h ↦ ⟨ω, h, rfl⟩, ?_⟩ + rintro ⟨a, h_mem, h⟩ + rwa [hf.injective h] at h_mem + · exact kernel.measurable_kernel_prod_mk_left' (h_prod_embed.measurableSet_image.mpr hs) _ + · exact hf.measurable end ProbabilityTheory From d28bd4c172f175efd86bb607f029b49c484bd0b3 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 12 Feb 2024 16:11:02 +0100 Subject: [PATCH 007/129] move files --- Mathlib/Probability/Kernel/CondCdf.lean | 2 +- .../{ => Disintegration}/BuildKernel.lean | 135 ++++++++++++++---- .../{ => Disintegration}/StieltjesReal.lean | 0 .../Kernel/{ => Disintegration}/draft.lean | 2 +- 4 files changed, 106 insertions(+), 33 deletions(-) rename Mathlib/Probability/Kernel/{ => Disintegration}/BuildKernel.lean (81%) rename Mathlib/Probability/Kernel/{ => Disintegration}/StieltjesReal.lean (100%) rename Mathlib/Probability/Kernel/{ => Disintegration}/draft.lean (99%) diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index fc1b2b5ba21f5..4d7b661260422 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -5,7 +5,7 @@ Authors: Rémy Degenne -/ import Mathlib.Logic.Encodable.Basic import Mathlib.Data.Set.Lattice -import Mathlib.Probability.Kernel.BuildKernel +import Mathlib.Probability.Kernel.Disintegration.BuildKernel #align_import probability.kernel.cond_cdf from "leanprover-community/mathlib"@"3b88f4005dc2e28d42f974cc1ce838f0dafb39b8" diff --git a/Mathlib/Probability/Kernel/BuildKernel.lean b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean similarity index 81% rename from Mathlib/Probability/Kernel/BuildKernel.lean rename to Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean index 1c1714211b740..4969ea5081065 100644 --- a/Mathlib/Probability/Kernel/BuildKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.StieltjesReal +import Mathlib.Probability.Kernel.Disintegration.StieltjesReal import Mathlib.Probability.Kernel.MeasureCompProd /-! @@ -20,30 +20,7 @@ namespace ProbabilityTheory variable {α β : Type*} [MeasurableSpace α] -section kernel - -variable {f : α → ℚ → ℝ} {hf : ∀ q, Measurable fun a ↦ f a q} - -noncomputable -def cdfKernel (f : α → ℚ → ℝ) (hf : ∀ q, Measurable fun a ↦ f a q) : kernel α ℝ where - val a := (todo3 f hf a).measure - property := measurable_measure_todo3 hf - -instance instIsMarkovKernel_cdfKernel : IsMarkovKernel (cdfKernel f hf) := - ⟨fun _ ↦ instIsProbabilityMeasure_todo3 _ _⟩ - -lemma cdfKernel_Iic (a : α) (x : ℝ) : - cdfKernel f hf a (Iic x) = ENNReal.ofReal (todo3 f hf a x) := measure_todo3_Iic hf a x - -lemma cdfKernel_unit_prod_Iic (a : α) (x : ℝ) : - cdfKernel (fun p : Unit × α ↦ f p.2) (fun q ↦ (hf q).comp measurable_snd) ((), a) (Iic x) - = cdfKernel f hf a (Iic x) := by - simp only [cdfKernel_Iic] - rw [todo3_unit_prod hf a] - -end kernel - -section +section todo3 variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} {f : α × β → ℚ → ℝ} {μ : kernel α (β × ℝ)} {ν : kernel α β} @@ -55,6 +32,13 @@ structure IsRatKernelCDF (f : α × β → ℚ → ℝ) (μ : kernel α (β × (isCDF (a : α) {s : Set β} (_hs : MeasurableSet s) (q : ℚ) : ∫ t in s, f (a, t) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal) +structure IsKernelCDF (f : α × β → StieltjesFunction) (μ : kernel α (β × ℝ)) (ν : kernel α β) : + Prop := + (measurable (x : ℝ) : Measurable fun p ↦ f p x) + (integrable (a : α) (x : ℝ) : Integrable (fun t ↦ f (a, t) x) (ν a)) + (isCDF (a : α) {s : Set β} (_hs : MeasurableSet s) (x : ℝ) : + ∫ t in s, f (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal) + lemma todo3_ae_eq (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) : (fun t ↦ todo3 f hf.measurable (a, t) q) =ᵐ[ν a] fun t ↦ f (a, t) q := by filter_upwards [hf.isRatStieltjesPoint_ae a] with a ha @@ -77,16 +61,105 @@ lemma set_lintegral_todo3_rat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) exact hf.integrable a q · exact ae_of_all _ (fun x ↦ todo3_nonneg _ _ _) +lemma set_lintegral_todo3_Iic [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) + (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : + ∫⁻ t in s, ENNReal.ofReal (todo3 f hf.measurable (a, t) x) ∂(ν a) = μ a (s ×ˢ Iic x) := by + -- We have the result for `x : ℚ` thanks to `set_lintegral_todo3_rat`. + -- We use the equality `condCDF ρ a x = ⨅ r : {r' : ℚ // x < r'}, condCDF ρ a r` and a monotone + -- convergence argument to extend it to the reals. + by_cases hρ_zero : (ν a).restrict s = 0 + · rw [hρ_zero, lintegral_zero_measure] + have ⟨q, hq⟩ := exists_rat_gt x + suffices μ a (s ×ˢ Iic (q : ℝ)) = 0 by + symm + refine measure_mono_null (fun p ↦ ?_) this + simp only [mem_prod, mem_Iic, and_imp] + exact fun h1 h2 ↦ ⟨h1, h2.trans hq.le⟩ + suffices (μ a (s ×ˢ Iic (q : ℝ))).toReal = 0 by + rw [ENNReal.toReal_eq_zero_iff] at this + simpa [measure_ne_top] using this + rw [← hf.isCDF a hs q] + simp [hρ_zero] + have h : ∫⁻ t in s, ENNReal.ofReal (todo3 f hf.measurable (a, t) x) ∂(ν a) + = ∫⁻ t in s, ⨅ r : { r' : ℚ // x < r' }, + ENNReal.ofReal (todo3 f hf.measurable (a, t) r) ∂(ν a) := by + congr with t : 1 + simp_rw [← measure_todo3_Iic] + rw [← measure_iInter_eq_iInf] + · congr with y : 1 + simp only [mem_Iic, mem_iInter, Subtype.forall] + refine ⟨fun h a ha ↦ h.trans ?_, fun h ↦ ?_⟩ + · exact mod_cast ha.le + · refine le_of_forall_lt_rat_imp_le fun q hq ↦ h q ?_ + exact mod_cast hq + · exact fun _ ↦ measurableSet_Iic + · refine Monotone.directed_ge fun r r' hrr' ↦ ?_ + refine Iic_subset_Iic.mpr ?_ + exact mod_cast hrr' + · obtain ⟨q, hq⟩ := exists_rat_gt x + exact ⟨⟨q, hq⟩, measure_ne_top _ _⟩ + have h_nonempty : Nonempty { r' : ℚ // x < ↑r' } := by + obtain ⟨r, hrx⟩ := exists_rat_gt x + exact ⟨⟨r, hrx⟩⟩ + rw [h, lintegral_iInf_directed_of_measurable hρ_zero fun q : { r' : ℚ // x < ↑r' } ↦ ?_] + rotate_left + · intro b + rw [set_lintegral_todo3_rat hf a _ hs] + exact measure_ne_top _ _ + · refine Monotone.directed_ge fun i j hij t ↦ ?_ + simp_rw [← measure_todo3_Iic] + refine measure_mono (Iic_subset_Iic.mpr ?_) + exact mod_cast hij + · refine Measurable.ennreal_ofReal ?_ + exact (measurable_todo3 hf.measurable _).comp measurable_prod_mk_left + simp_rw [set_lintegral_todo3_rat hf _ _ hs] + rw [← measure_iInter_eq_iInf] + · rw [← prod_iInter] + congr with y + simp only [mem_iInter, mem_Iic, Subtype.forall, Subtype.coe_mk] + exact ⟨le_of_forall_lt_rat_imp_le, fun hyx q hq ↦ hyx.trans hq.le⟩ + · exact fun i ↦ hs.prod measurableSet_Iic + · refine Monotone.directed_ge fun i j hij ↦ ?_ + refine prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, Iic_subset_Iic.mpr ?_⟩) + exact mod_cast hij + · exact ⟨h_nonempty.some, measure_ne_top _ _⟩ + + +end todo3 + +section kernel + +variable {f : α → ℚ → ℝ} {hf : ∀ q, Measurable fun a ↦ f a q} + +noncomputable +def cdfKernel (f : α → ℚ → ℝ) (hf : ∀ q, Measurable fun a ↦ f a q) : kernel α ℝ where + val a := (todo3 f hf a).measure + property := measurable_measure_todo3 hf + +instance instIsMarkovKernel_cdfKernel : IsMarkovKernel (cdfKernel f hf) := + ⟨fun _ ↦ instIsProbabilityMeasure_todo3 _ _⟩ + +lemma cdfKernel_Iic (a : α) (x : ℝ) : + cdfKernel f hf a (Iic x) = ENNReal.ofReal (todo3 f hf a x) := measure_todo3_Iic hf a x + +lemma cdfKernel_unit_prod_Iic (a : α) (x : ℝ) : + cdfKernel (fun p : Unit × α ↦ f p.2) (fun q ↦ (hf q).comp measurable_snd) ((), a) (Iic x) + = cdfKernel f hf a (Iic x) := by + simp only [cdfKernel_Iic] + rw [todo3_unit_prod hf a] + +end kernel + +section + +variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} + {f : α × β → ℚ → ℝ} {μ : kernel α (β × ℝ)} {ν : kernel α β} + lemma set_lintegral_cdfKernel_Iic_rat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : ∫⁻ t in s, cdfKernel f hf.measurable (a, t) (Iic q) ∂(ν a) = μ a (s ×ˢ Iic (q : ℝ)) := by simp_rw [cdfKernel_Iic, set_lintegral_todo3_rat hf a q hs] -theorem lintegral_iInf_directed_of_measurable' {mα : MeasurableSpace α} [Countable β] - {f : β → α → ℝ≥0∞} {μ : Measure α} (hμ : μ ≠ 0) (hf : ∀ b, Measurable (f b)) - (hf_int : ∀ b, ∫⁻ a, f b a ∂μ ≠ ∞) (h_directed : Directed (· ≥ ·) f) : - ∫⁻ a, ⨅ b, f b a ∂μ = ⨅ b, ∫⁻ a, f b a ∂μ := sorry - lemma set_lintegral_cdfKernel_Iic [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : ∫⁻ t in s, cdfKernel f hf.measurable (a, t) (Iic x) ∂(ν a) = μ a (s ×ˢ Iic x) := by @@ -126,7 +199,7 @@ lemma set_lintegral_cdfKernel_Iic [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ have h_nonempty : Nonempty { r' : ℚ // x < ↑r' } := by obtain ⟨r, hrx⟩ := exists_rat_gt x exact ⟨⟨r, hrx⟩⟩ - rw [h, lintegral_iInf_directed_of_measurable' hρ_zero fun q : { r' : ℚ // x < ↑r' } ↦ ?_] + rw [h, lintegral_iInf_directed_of_measurable hρ_zero fun q : { r' : ℚ // x < ↑r' } ↦ ?_] rotate_left · intro b rw [set_lintegral_cdfKernel_Iic_rat hf a _ hs] diff --git a/Mathlib/Probability/Kernel/StieltjesReal.lean b/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean similarity index 100% rename from Mathlib/Probability/Kernel/StieltjesReal.lean rename to Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean diff --git a/Mathlib/Probability/Kernel/draft.lean b/Mathlib/Probability/Kernel/Disintegration/draft.lean similarity index 99% rename from Mathlib/Probability/Kernel/draft.lean rename to Mathlib/Probability/Kernel/Disintegration/draft.lean index 064d19569b55b..a5e0e21e1caca 100644 --- a/Mathlib/Probability/Kernel/draft.lean +++ b/Mathlib/Probability/Kernel/Disintegration/draft.lean @@ -1,6 +1,6 @@ import Mathlib.Probability.Kernel.Disintegration import Mathlib.Probability.Martingale.Convergence -import Mathlib.Probability.Kernel.BuildKernel +import Mathlib.Probability.Kernel.Disintegration.BuildKernel open MeasureTheory Set Filter From 48afecb780c25b0febc70ce9bb6b5dca4dd44a85 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Tue, 13 Feb 2024 16:06:52 +0100 Subject: [PATCH 008/129] save --- .../Kernel/Disintegration/BuildKernel.lean | 133 ++++++++---------- 1 file changed, 60 insertions(+), 73 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean index 4969ea5081065..2c5470bb64e39 100644 --- a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean @@ -32,13 +32,6 @@ structure IsRatKernelCDF (f : α × β → ℚ → ℝ) (μ : kernel α (β × (isCDF (a : α) {s : Set β} (_hs : MeasurableSet s) (q : ℚ) : ∫ t in s, f (a, t) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal) -structure IsKernelCDF (f : α × β → StieltjesFunction) (μ : kernel α (β × ℝ)) (ν : kernel α β) : - Prop := - (measurable (x : ℝ) : Measurable fun p ↦ f p x) - (integrable (a : α) (x : ℝ) : Integrable (fun t ↦ f (a, t) x) (ν a)) - (isCDF (a : α) {s : Set β} (_hs : MeasurableSet s) (x : ℝ) : - ∫ t in s, f (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal) - lemma todo3_ae_eq (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) : (fun t ↦ todo3 f hf.measurable (a, t) q) =ᵐ[ν a] fun t ↦ f (a, t) q := by filter_upwards [hf.isRatStieltjesPoint_ae a] with a ha @@ -61,7 +54,7 @@ lemma set_lintegral_todo3_rat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) exact hf.integrable a q · exact ae_of_all _ (fun x ↦ todo3_nonneg _ _ _) -lemma set_lintegral_todo3_Iic [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma set_lintegral_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : ∫⁻ t in s, ENNReal.ofReal (todo3 f hf.measurable (a, t) x) ∂(ν a) = μ a (s ×ˢ Iic x) := by -- We have the result for `x : ℚ` thanks to `set_lintegral_todo3_rat`. @@ -124,6 +117,57 @@ lemma set_lintegral_todo3_Iic [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) exact mod_cast hij · exact ⟨h_nonempty.some, measure_ne_top _ _⟩ +lemma lintegral_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) (x : ℝ) : + ∫⁻ t, ENNReal.ofReal (todo3 f hf.measurable (a, t) x) ∂(ν a) = μ a (univ ×ˢ Iic x) := by + rw [← set_lintegral_univ, set_lintegral_todo3 hf _ _ MeasurableSet.univ] + +lemma integrable_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) (x : ℝ) : + Integrable (fun t ↦ todo3 f hf.measurable (a, t) x) (ν a) := by + have : (fun t ↦ todo3 f hf.measurable (a, t) x) + = fun t ↦ (ENNReal.ofReal (todo3 f hf.measurable (a, t) x)).toReal := by + ext t + rw [ENNReal.toReal_ofReal] + exact todo3_nonneg _ _ _ + rw [this] + refine integrable_toReal_of_lintegral_ne_top ?_ ?_ + · refine (Measurable.ennreal_ofReal ?_).aemeasurable + exact (measurable_todo3 hf.measurable x).comp measurable_prod_mk_left + · rw [lintegral_todo3 hf] + exact measure_ne_top _ _ + +lemma set_integral_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) + (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : + ∫ t in s, todo3 f hf.measurable (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal := by + rw [← ENNReal.ofReal_eq_ofReal_iff, ENNReal.ofReal_toReal] + rotate_left + · exact measure_ne_top _ _ + · exact set_integral_nonneg hs (fun _ _ ↦ todo3_nonneg _ _ _) + · exact ENNReal.toReal_nonneg + rw [ofReal_integral_eq_lintegral_ofReal, set_lintegral_todo3 hf _ _ hs] + · exact (integrable_todo3 hf _ _).restrict + · exact ae_of_all _ (fun _ ↦ todo3_nonneg _ _ _) + +lemma integral_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) (x : ℝ) : + ∫ t, todo3 f hf.measurable (a, t) x ∂(ν a) = (μ a (univ ×ˢ Iic x)).toReal := by + rw [← integral_univ, set_integral_todo3 hf _ _ MeasurableSet.univ] + +structure IsKernelCDF (f : α × β → StieltjesFunction) (μ : kernel α (β × ℝ)) (ν : kernel α β) : + Prop := + (measurable (x : ℝ) : Measurable fun p ↦ f p x) + (integrable (a : α) (x : ℝ) : Integrable (fun t ↦ f (a, t) x) (ν a)) + (tendsto_atTop_one (p : α × β) : Tendsto (f p) atTop (𝓝 1)) + (tendsto_atBot_zero (p : α × β) : Tendsto (f p) atBot (𝓝 0)) + (isCDF (a : α) {s : Set β} (_hs : MeasurableSet s) (x : ℝ) : + ∫ t in s, f (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal) +-- todo: nonneg and le_one are consequences of tendsto_atTop_one and tendsto_atBot_zero + +lemma isKernelCDF_todo3 (hf : IsRatKernelCDF f μ ν) [IsFiniteKernel μ] : + IsKernelCDF (todo3 f hf.measurable) μ ν where + measurable := measurable_todo3 hf.measurable + integrable := integrable_todo3 hf + tendsto_atTop_one := tendsto_todo3_atTop hf.measurable + tendsto_atBot_zero := tendsto_todo3_atBot hf.measurable + isCDF a _ hs x := set_integral_todo3 hf a x hs end todo3 @@ -155,69 +199,11 @@ section variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} {f : α × β → ℚ → ℝ} {μ : kernel α (β × ℝ)} {ν : kernel α β} -lemma set_lintegral_cdfKernel_Iic_rat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) - (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ t in s, cdfKernel f hf.measurable (a, t) (Iic q) ∂(ν a) = μ a (s ×ˢ Iic (q : ℝ)) := by - simp_rw [cdfKernel_Iic, set_lintegral_todo3_rat hf a q hs] - lemma set_lintegral_cdfKernel_Iic [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : ∫⁻ t in s, cdfKernel f hf.measurable (a, t) (Iic x) ∂(ν a) = μ a (s ×ˢ Iic x) := by - -- We have the result for `x : ℚ` thanks to `set_lintegral_cdfKernel_Iic_rat`. - -- We use the equality `condCDF ρ a x = ⨅ r : {r' : ℚ // x < r'}, condCDF ρ a r` and a monotone - -- convergence argument to extend it to the reals. - by_cases hρ_zero : (ν a).restrict s = 0 - · rw [hρ_zero, lintegral_zero_measure] - have ⟨q, hq⟩ := exists_rat_gt x - suffices μ a (s ×ˢ Iic (q : ℝ)) = 0 by - symm - refine measure_mono_null (fun p ↦ ?_) this - simp only [mem_prod, mem_Iic, and_imp] - exact fun h1 h2 ↦ ⟨h1, h2.trans hq.le⟩ - suffices (μ a (s ×ˢ Iic (q : ℝ))).toReal = 0 by - rw [ENNReal.toReal_eq_zero_iff] at this - simpa [measure_ne_top] using this - rw [← hf.isCDF a hs q] - simp [hρ_zero] - have h : ∫⁻ t in s, cdfKernel f hf.measurable (a, t) (Iic x) ∂(ν a) - = ∫⁻ t in s, ⨅ r : { r' : ℚ // x < r' }, - cdfKernel f hf.measurable (a, t) (Iic r) ∂(ν a) := by - congr with t : 1 - rw [← measure_iInter_eq_iInf] - · congr with y : 1 - simp only [mem_Iic, mem_iInter, Subtype.forall] - refine ⟨fun h a ha ↦ h.trans ?_, fun h ↦ ?_⟩ - · exact mod_cast ha.le - · refine le_of_forall_lt_rat_imp_le fun q hq ↦ h q ?_ - exact mod_cast hq - · exact fun _ ↦ measurableSet_Iic - · refine Monotone.directed_ge fun r r' hrr' ↦ ?_ - refine Iic_subset_Iic.mpr ?_ - exact mod_cast hrr' - · obtain ⟨q, hq⟩ := exists_rat_gt x - exact ⟨⟨q, hq⟩, measure_ne_top _ _⟩ - have h_nonempty : Nonempty { r' : ℚ // x < ↑r' } := by - obtain ⟨r, hrx⟩ := exists_rat_gt x - exact ⟨⟨r, hrx⟩⟩ - rw [h, lintegral_iInf_directed_of_measurable hρ_zero fun q : { r' : ℚ // x < ↑r' } ↦ ?_] - rotate_left - · intro b - rw [set_lintegral_cdfKernel_Iic_rat hf a _ hs] - exact measure_ne_top _ _ - · refine Monotone.directed_ge fun i j hij t ↦ measure_mono (Iic_subset_Iic.mpr ?_) - exact mod_cast hij - · exact (kernel.measurable_coe _ measurableSet_Iic).comp measurable_prod_mk_left - simp_rw [set_lintegral_cdfKernel_Iic_rat hf _ _ hs] - rw [← measure_iInter_eq_iInf] - · rw [← prod_iInter] - congr with y - simp only [mem_iInter, mem_Iic, Subtype.forall, Subtype.coe_mk] - exact ⟨le_of_forall_lt_rat_imp_le, fun hyx q hq ↦ hyx.trans hq.le⟩ - · exact fun i ↦ hs.prod measurableSet_Iic - · refine Monotone.directed_ge fun i j hij ↦ ?_ - refine prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, Iic_subset_Iic.mpr ?_⟩) - exact mod_cast hij - · exact ⟨h_nonempty.some, measure_ne_top _ _⟩ + simp_rw [cdfKernel_Iic] + rw [set_lintegral_todo3 hf _ _ hs] theorem Real.iUnion_Iic_rat' : ⋃ r : ℚ, Iic (r : ℝ) = univ := sorry @@ -287,7 +273,7 @@ lemma set_lintegral_cdfKernel_prod [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ lemma lintegral_cdfKernel_mem [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) {s : Set (β × ℝ)} (hs : MeasurableSet s) : ∫⁻ x, cdfKernel f hf.measurable (a, x) {y | (x, y) ∈ s} ∂(ν a) = μ a s := by - -- `set_lintegral_condKernelReal_prod` gives the result for sets of the form `t₁ × t₂`. These + -- `set_lintegral_cdfKernel_prod` gives the result for sets of the form `t₁ × t₂`. These -- sets form a π-system that generates the product σ-algebra, hence we can get the same equality -- for any measurable set `s`. apply MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod _ _ _ _ hs @@ -303,9 +289,11 @@ lemma lintegral_cdfKernel_mem [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) = ∫⁻ x in t₁, cdfKernel f hf.measurable (a, x) t₂ ∂(ν a) := by refine' set_lintegral_congr_fun ht₁ (eventually_of_forall fun a ha ↦ _) rw [h_prod_eq_snd a ha] - have h_eq2 : ∫⁻ x in t₁ᶜ, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) = 0 := by - suffices h_eq_zero : ∀ x ∈ t₁ᶜ, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} = 0 - · rw [set_lintegral_congr_fun ht₁.compl (eventually_of_forall h_eq_zero)] + have h_eq2 : + ∫⁻ x in t₁ᶜ, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) = 0 := by + suffices h_eq_zero : + ∀ x ∈ t₁ᶜ, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} = 0 by + rw [set_lintegral_congr_fun ht₁.compl (eventually_of_forall h_eq_zero)] simp only [lintegral_const, zero_mul] intro a hat₁ rw [mem_compl_iff] at hat₁ @@ -382,7 +370,6 @@ lemma measurableSet_eq_one (hf : IsRatKernelCDF f μ ν) {s : Set ℝ} (hs : Mea MeasurableSet {p | cdfKernel f hf.measurable p s = 1} := (kernel.measurable_coe _ hs) (measurableSet_singleton 1) - end end ProbabilityTheory From 71e4115af0184a12366cb2bd6ccbc7e95e5c422a Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Tue, 13 Feb 2024 23:01:30 +0100 Subject: [PATCH 009/129] add IsKernelCDF --- Mathlib/Probability/Kernel/CondCdf.lean | 14 +- .../Probability/Kernel/Disintegration.lean | 24 +-- .../Kernel/Disintegration/BuildKernel.lean | 183 ++++++++++++------ .../Kernel/Disintegration/draft.lean | 22 ++- 4 files changed, 164 insertions(+), 79 deletions(-) diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index 4d7b661260422..05bfed4ba3479 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -576,10 +576,9 @@ theorem measurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : Measurable fun theorem set_lintegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) {s : Set α} (hs : MeasurableSet s) : ∫⁻ a in s, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x) := by - have h := set_lintegral_cdfKernel_Iic (isRatKernelCDF_preCDF ρ) () x hs + have h := set_lintegral_todo3 (isRatKernelCDF_preCDF ρ) () x hs simp only [kernel.const_apply] at h rw [← h] - simp_rw [cdfKernel_Iic] congr with a congr exact condCDF_eq_todo3_unit_prod _ _ @@ -596,7 +595,7 @@ theorem stronglyMeasurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : #align probability_theory.strongly_measurable_cond_cdf ProbabilityTheory.stronglyMeasurable_condCDF theorem integrable_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) : - Integrable (fun a => condCDF ρ a x) ρ.fst := by + Integrable (fun a ↦ condCDF ρ a x) ρ.fst := by refine' integrable_of_forall_fin_meas_le _ (measure_lt_top ρ.fst univ) _ fun t _ _ => _ · exact (stronglyMeasurable_condCDF ρ _).aestronglyMeasurable · have : ∀ y, (‖condCDF ρ y x‖₊ : ℝ≥0∞) ≤ 1 := by @@ -627,6 +626,15 @@ theorem integral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : rw [← set_integral_condCDF ρ _ MeasurableSet.univ, Measure.restrict_univ] #align probability_theory.integral_cond_cdf ProbabilityTheory.integral_condCDF +lemma isKernelCDF_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : + IsKernelCDF (fun p : Unit × α ↦ condCDF ρ p.2) (kernel.const Unit ρ) + (kernel.const Unit ρ.fst) where + measurable x := (measurable_condCDF ρ x).comp measurable_snd + integrable _ x := integrable_condCDF ρ x + tendsto_atTop_one p := tendsto_condCDF_atTop ρ p.2 + tendsto_atBot_zero p := tendsto_condCDF_atBot ρ p.2 + set_integral _ _ hs x := set_integral_condCDF ρ x hs + section Measure theorem measure_condCDF_Iic (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : diff --git a/Mathlib/Probability/Kernel/Disintegration.lean b/Mathlib/Probability/Kernel/Disintegration.lean index 7c7c2faa68141..ec0ffc4848fa5 100644 --- a/Mathlib/Probability/Kernel/Disintegration.lean +++ b/Mathlib/Probability/Kernel/Disintegration.lean @@ -60,16 +60,15 @@ section Real /-- Conditional measure on the second space of the product given the value on the first, as a kernel. Use the more general `condKernel`. -/ -noncomputable def condKernelReal (ρ : Measure (α × ℝ)) : kernel α ℝ := - kernel.comap (cdfKernel (fun (p : Unit × α) r ↦ (preCDF ρ r p.2).toReal) - (fun _ ↦ measurable_preCDF.ennreal_toReal.comp measurable_snd)) (fun a ↦ ((), a)) - measurable_prod_mk_left +noncomputable def condKernelReal (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : kernel α ℝ := + kernel.comap (cdfKernel (fun (p : Unit × α) ↦ condCDF ρ p.2) (isKernelCDF_condCDF ρ)) + (fun a ↦ ((), a)) measurable_prod_mk_left #align probability_theory.cond_kernel_real ProbabilityTheory.condKernelReal -instance (ρ : Measure (α × ℝ)) : IsMarkovKernel (condKernelReal ρ) := by +instance (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : IsMarkovKernel (condKernelReal ρ) := by rw [condKernelReal]; infer_instance -theorem condKernelReal_Iic (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : +theorem condKernelReal_Iic (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (a : α) (x : ℝ) : condKernelReal ρ a (Iic x) = ENNReal.ofReal (condCDF ρ a x) := by rw [condKernelReal, kernel.comap_apply, cdfKernel_Iic, condCDF_eq_todo3_unit_prod] #align probability_theory.cond_kernel_real_Iic ProbabilityTheory.condKernelReal_Iic @@ -77,16 +76,17 @@ theorem condKernelReal_Iic (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : theorem set_lintegral_condKernelReal_Iic (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) {s : Set α} (hs : MeasurableSet s) : ∫⁻ a in s, condKernelReal ρ a (Iic x) ∂ρ.fst = ρ (s ×ˢ Iic x) := - set_lintegral_cdfKernel_Iic (isRatKernelCDF_preCDF ρ) () x hs + set_lintegral_cdfKernel_Iic (isKernelCDF_condCDF ρ) () x hs #align probability_theory.set_lintegral_cond_kernel_real_Iic ProbabilityTheory.set_lintegral_condKernelReal_Iic -theorem set_lintegral_condKernelReal_univ (ρ : Measure (α × ℝ)) {s : Set α} (hs : MeasurableSet s) : +theorem set_lintegral_condKernelReal_univ (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] + {s : Set α} (hs : MeasurableSet s) : ∫⁻ a in s, condKernelReal ρ a univ ∂ρ.fst = ρ (s ×ˢ univ) := by simp only [measure_univ, lintegral_const, Measure.restrict_apply, MeasurableSet.univ, univ_inter, one_mul, Measure.fst_apply hs, ← prod_univ] #align probability_theory.set_lintegral_cond_kernel_real_univ ProbabilityTheory.set_lintegral_condKernelReal_univ -theorem lintegral_condKernelReal_univ (ρ : Measure (α × ℝ)) : +theorem lintegral_condKernelReal_univ (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : ∫⁻ a, condKernelReal ρ a univ ∂ρ.fst = ρ univ := by rw [← set_lintegral_univ, set_lintegral_condKernelReal_univ ρ MeasurableSet.univ, univ_prod_univ] @@ -97,12 +97,12 @@ variable (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] theorem set_lintegral_condKernelReal_prod {s : Set α} (hs : MeasurableSet s) {t : Set ℝ} (ht : MeasurableSet t) : ∫⁻ a in s, condKernelReal ρ a t ∂ρ.fst = ρ (s ×ˢ t) := - set_lintegral_cdfKernel_prod (isRatKernelCDF_preCDF ρ) () hs ht + set_lintegral_cdfKernel_prod (isKernelCDF_condCDF ρ) () hs ht #align probability_theory.set_lintegral_cond_kernel_real_prod ProbabilityTheory.set_lintegral_condKernelReal_prod theorem lintegral_condKernelReal_mem {s : Set (α × ℝ)} (hs : MeasurableSet s) : ∫⁻ a, condKernelReal ρ a {x | (a, x) ∈ s} ∂ρ.fst = ρ s := - lintegral_cdfKernel_mem (isRatKernelCDF_preCDF ρ) () hs + lintegral_cdfKernel_mem (isKernelCDF_condCDF ρ) () hs #align probability_theory.lintegral_cond_kernel_real_mem ProbabilityTheory.lintegral_condKernelReal_mem theorem kernel.const_eq_compProd_real (γ : Type*) [MeasurableSpace γ] (ρ : Measure (α × ℝ)) @@ -126,7 +126,7 @@ theorem lintegral_condKernelReal {f : α × ℝ → ℝ≥0∞} (hf : Measurable theorem ae_condKernelReal_eq_one {s : Set ℝ} (hs : MeasurableSet s) (hρ : ρ {x | x.snd ∈ sᶜ} = 0) : ∀ᵐ a ∂ρ.fst, condKernelReal ρ a s = 1 := by - have h := ae_cdfKernel_eq_one (isRatKernelCDF_preCDF ρ) () hs + have h := ae_cdfKernel_eq_one (isKernelCDF_condCDF ρ) () hs simp only [kernel.const_apply] at h exact h hρ #align probability_theory.ae_cond_kernel_real_eq_one ProbabilityTheory.ae_condKernelReal_eq_one diff --git a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean index 2c5470bb64e39..dfe7c056eaf44 100644 --- a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean @@ -151,65 +151,130 @@ lemma integral_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) ∫ t, todo3 f hf.measurable (a, t) x ∂(ν a) = (μ a (univ ×ˢ Iic x)).toReal := by rw [← integral_univ, set_integral_todo3 hf _ _ MeasurableSet.univ] +end todo3 + +section IsKernelCDF + +variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} + {f : α × β → StieltjesFunction} {μ : kernel α (β × ℝ)} {ν : kernel α β} + structure IsKernelCDF (f : α × β → StieltjesFunction) (μ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := (measurable (x : ℝ) : Measurable fun p ↦ f p x) (integrable (a : α) (x : ℝ) : Integrable (fun t ↦ f (a, t) x) (ν a)) (tendsto_atTop_one (p : α × β) : Tendsto (f p) atTop (𝓝 1)) (tendsto_atBot_zero (p : α × β) : Tendsto (f p) atBot (𝓝 0)) - (isCDF (a : α) {s : Set β} (_hs : MeasurableSet s) (x : ℝ) : + (set_integral (a : α) {s : Set β} (_hs : MeasurableSet s) (x : ℝ) : ∫ t in s, f (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal) -- todo: nonneg and le_one are consequences of tendsto_atTop_one and tendsto_atBot_zero -lemma isKernelCDF_todo3 (hf : IsRatKernelCDF f μ ν) [IsFiniteKernel μ] : +lemma IsKernelCDF.nonneg (hf : IsKernelCDF f μ ν) (p : α × β) (x : ℝ) : 0 ≤ f p x := + Monotone.le_of_tendsto (f p).mono (hf.tendsto_atBot_zero p) x + +lemma IsKernelCDF.le_one (hf : IsKernelCDF f μ ν) (p : α × β) (x : ℝ) : f p x ≤ 1 := + Monotone.ge_of_tendsto (f p).mono (hf.tendsto_atTop_one p) x + +lemma IsKernelCDF.set_lintegral [IsFiniteKernel μ] + {f : α × β → StieltjesFunction} (hf : IsKernelCDF f μ ν) + (a : α) {s : Set β} (hs : MeasurableSet s) (x : ℝ) : + ∫⁻ t in s, ENNReal.ofReal (f (a, t) x) ∂(ν a) = μ a (s ×ˢ Iic x) := by + rw [← ofReal_integral_eq_lintegral_ofReal (hf.integrable a x).restrict + (ae_of_all _ (fun _ ↦ hf.nonneg _ _)), hf.set_integral a hs x, ENNReal.ofReal_toReal] + exact measure_ne_top _ _ + +lemma isKernelCDF_todo3 {f : α × β → ℚ → ℝ} (hf : IsRatKernelCDF f μ ν) [IsFiniteKernel μ] : IsKernelCDF (todo3 f hf.measurable) μ ν where measurable := measurable_todo3 hf.measurable integrable := integrable_todo3 hf tendsto_atTop_one := tendsto_todo3_atTop hf.measurable tendsto_atBot_zero := tendsto_todo3_atBot hf.measurable - isCDF a _ hs x := set_integral_todo3 hf a x hs + set_integral a _ hs x := set_integral_todo3 hf a x hs -end todo3 +end IsKernelCDF section kernel -variable {f : α → ℚ → ℝ} {hf : ∀ q, Measurable fun a ↦ f a q} +variable {_ : MeasurableSpace β} {f : α × β → StieltjesFunction} + {μ : kernel α (β × ℝ)} {ν : kernel α β} {hf : IsKernelCDF f μ ν} + +lemma isProbabilityMeasure_stieltjesFunction {f : StieltjesFunction} + (hf_bot : Tendsto f atBot (𝓝 0)) (hf_top : Tendsto f atTop (𝓝 1)) : + IsProbabilityMeasure f.measure := + ⟨by simp [StieltjesFunction.measure_univ _ hf_bot hf_top]⟩ + +lemma StieltjesFunction.measurable_measure {f : α → StieltjesFunction} + (hf : ∀ q, Measurable fun a ↦ f a q) + (hf_bot : ∀ a, Tendsto (f a) atBot (𝓝 0)) + (hf_top : ∀ a, Tendsto (f a) atTop (𝓝 1)) : + Measurable fun a ↦ (f a).measure := by + rw [Measure.measurable_measure] + have : ∀ a, IsProbabilityMeasure (f a).measure := + fun a ↦ isProbabilityMeasure_stieltjesFunction (hf_bot _) (hf_top _) + refine fun s hs ↦ ?_ + -- Porting note: supplied `C` + refine MeasurableSpace.induction_on_inter + (C := fun s ↦ Measurable fun b ↦ StieltjesFunction.measure (f b) s) + (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic ?_ ?_ ?_ ?_ hs + · simp only [measure_empty, measurable_const] + · rintro S ⟨u, rfl⟩ + simp_rw [StieltjesFunction.measure_Iic (f _) (hf_bot _)] + simp only [sub_zero] + exact (hf _).ennreal_ofReal + · intro t ht ht_cd_meas + have : (fun a ↦ (f a).measure tᶜ) = + (fun a ↦ (f a).measure univ) + - fun a ↦ (f a).measure t := by + ext1 a + rw [measure_compl ht, Pi.sub_apply] + exact measure_ne_top _ _ + simp_rw [this, measure_univ] + exact Measurable.sub measurable_const ht_cd_meas + · intro f hf_disj hf_meas hf_cd_meas + simp_rw [measure_iUnion hf_disj hf_meas] + exact Measurable.ennreal_tsum hf_cd_meas noncomputable -def cdfKernel (f : α → ℚ → ℝ) (hf : ∀ q, Measurable fun a ↦ f a q) : kernel α ℝ where - val a := (todo3 f hf a).measure - property := measurable_measure_todo3 hf +def cdfKernel (f : α × β → StieltjesFunction) (hf : IsKernelCDF f μ ν) : kernel (α × β) ℝ where + val p := (f p).measure + property := StieltjesFunction.measurable_measure hf.measurable + hf.tendsto_atBot_zero hf.tendsto_atTop_one + +lemma cdfKernel_apply (p : α × β) : cdfKernel f hf p = (f p).measure := rfl instance instIsMarkovKernel_cdfKernel : IsMarkovKernel (cdfKernel f hf) := - ⟨fun _ ↦ instIsProbabilityMeasure_todo3 _ _⟩ + ⟨fun _ ↦ isProbabilityMeasure_stieltjesFunction + (hf.tendsto_atBot_zero _) (hf.tendsto_atTop_one _)⟩ -lemma cdfKernel_Iic (a : α) (x : ℝ) : - cdfKernel f hf a (Iic x) = ENNReal.ofReal (todo3 f hf a x) := measure_todo3_Iic hf a x +lemma cdfKernel_Iic (p : α × β) (x : ℝ) : + cdfKernel f hf p (Iic x) = ENNReal.ofReal (f p x) := by + rw [cdfKernel_apply p, (f p).measure_Iic (hf.tendsto_atBot_zero p)] + simp -lemma cdfKernel_unit_prod_Iic (a : α) (x : ℝ) : - cdfKernel (fun p : Unit × α ↦ f p.2) (fun q ↦ (hf q).comp measurable_snd) ((), a) (Iic x) - = cdfKernel f hf a (Iic x) := by - simp only [cdfKernel_Iic] - rw [todo3_unit_prod hf a] +--lemma cdfKernel_unit_prod_Iic (a : α) (x : ℝ) : +-- cdfKernel (fun p : Unit × α ↦ f p.2) (fun q ↦ (hf q).comp measurable_snd) ((), a) (Iic x) +-- = cdfKernel f hf a (Iic x) := by +-- simp only [cdfKernel_Iic] +-- rw [todo3_unit_prod hf a] end kernel section variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} - {f : α × β → ℚ → ℝ} {μ : kernel α (β × ℝ)} {ν : kernel α β} + {f : α × β → StieltjesFunction} {μ : kernel α (β × ℝ)} {ν : kernel α β} + {hf : IsKernelCDF f μ ν} -lemma set_lintegral_cdfKernel_Iic [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma set_lintegral_cdfKernel_Iic [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ t in s, cdfKernel f hf.measurable (a, t) (Iic x) ∂(ν a) = μ a (s ×ˢ Iic x) := by + ∫⁻ t in s, cdfKernel f hf (a, t) (Iic x) ∂(ν a) = μ a (s ×ˢ Iic x) := by simp_rw [cdfKernel_Iic] - rw [set_lintegral_todo3 hf _ _ hs] + exact hf.set_lintegral _ hs _ theorem Real.iUnion_Iic_rat' : ⋃ r : ℚ, Iic (r : ℝ) = univ := sorry -lemma set_lintegral_cdfKernel_univ [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma set_lintegral_cdfKernel_univ [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ t in s, cdfKernel f hf.measurable (a, t) univ ∂(ν a) = μ a (s ×ˢ univ) := by + ∫⁻ t in s, cdfKernel f hf (a, t) univ ∂(ν a) = μ a (s ×ˢ univ) := by rw [← Real.iUnion_Iic_rat', prod_iUnion] have h_dir : Directed (fun x y ↦ x ⊆ y) fun q : ℚ ↦ Iic (q : ℝ) := by refine Monotone.directed_le fun r r' hrr' ↦ Iic_subset_Iic.mpr ?_ @@ -226,13 +291,13 @@ lemma set_lintegral_cdfKernel_univ [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ · refine Monotone.directed_le fun i j hij t ↦ measure_mono (Iic_subset_Iic.mpr ?_) exact mod_cast hij -lemma lintegral_cdfKernel_univ [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) : - ∫⁻ t, cdfKernel f hf.measurable (a, t) univ ∂(ν a) = μ a univ := by +lemma lintegral_cdfKernel_univ [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) : + ∫⁻ t, cdfKernel f hf (a, t) univ ∂(ν a) = μ a univ := by rw [← set_lintegral_univ, set_lintegral_cdfKernel_univ hf a MeasurableSet.univ, univ_prod_univ] -lemma set_lintegral_cdfKernel_prod [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma set_lintegral_cdfKernel_prod [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set ℝ} (ht : MeasurableSet t) : - ∫⁻ x in s, cdfKernel f hf.measurable (a, x) t ∂(ν a) = μ a (s ×ˢ t) := by + ∫⁻ x in s, cdfKernel f hf (a, x) t ∂(ν a) = μ a (s ×ˢ t) := by -- `set_lintegral_cdfKernel_Iic` gives the result for `t = Iic x`. These sets form a -- π-system that generates the Borel σ-algebra, hence we can get the same equality for any -- measurable set `t`. @@ -241,13 +306,13 @@ lemma set_lintegral_cdfKernel_prod [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ · rintro t ⟨q, rfl⟩ exact set_lintegral_cdfKernel_Iic hf a _ hs · intro t ht ht_lintegral - calc ∫⁻ x in s, cdfKernel f hf.measurable (a, x) tᶜ ∂(ν a) - = ∫⁻ x in s, cdfKernel f hf.measurable (a, x) univ - cdfKernel f hf.measurable (a, x) t ∂(ν a) := by - congr with x; rw [measure_compl ht (measure_ne_top (cdfKernel f hf.measurable (a, x)) _)] - _ = ∫⁻ x in s, cdfKernel f hf.measurable (a, x) univ ∂(ν a) - - ∫⁻ x in s, cdfKernel f hf.measurable (a, x) t ∂(ν a) := by + calc ∫⁻ x in s, cdfKernel f hf (a, x) tᶜ ∂(ν a) + = ∫⁻ x in s, cdfKernel f hf (a, x) univ - cdfKernel f hf (a, x) t ∂(ν a) := by + congr with x; rw [measure_compl ht (measure_ne_top (cdfKernel f hf (a, x)) _)] + _ = ∫⁻ x in s, cdfKernel f hf (a, x) univ ∂(ν a) + - ∫⁻ x in s, cdfKernel f hf (a, x) t ∂(ν a) := by rw [lintegral_sub] - · exact (kernel.measurable_coe (cdfKernel f hf.measurable) ht).comp measurable_prod_mk_left + · exact (kernel.measurable_coe (cdfKernel f hf) ht).comp measurable_prod_mk_left · rw [ht_lintegral] exact measure_ne_top _ _ · exact eventually_of_forall fun a ↦ measure_mono (subset_univ _) @@ -270,9 +335,9 @@ lemma set_lintegral_cdfKernel_prod [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ · exact fun i ↦ ((kernel.measurable_coe _ (hf_meas i)).comp measurable_prod_mk_left).aemeasurable.restrict -lemma lintegral_cdfKernel_mem [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma lintegral_cdfKernel_mem [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) {s : Set (β × ℝ)} (hs : MeasurableSet s) : - ∫⁻ x, cdfKernel f hf.measurable (a, x) {y | (x, y) ∈ s} ∂(ν a) = μ a s := by + ∫⁻ x, cdfKernel f hf (a, x) {y | (x, y) ∈ s} ∂(ν a) = μ a s := by -- `set_lintegral_cdfKernel_prod` gives the result for sets of the form `t₁ × t₂`. These -- sets form a π-system that generates the product σ-algebra, hence we can get the same equality -- for any measurable set `s`. @@ -285,14 +350,14 @@ lemma lintegral_cdfKernel_mem [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) intro a ha simp only [ha, prod_mk_mem_set_prod_eq, true_and_iff, setOf_mem_eq] rw [← lintegral_add_compl _ ht₁] - have h_eq1 : ∫⁻ x in t₁, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) - = ∫⁻ x in t₁, cdfKernel f hf.measurable (a, x) t₂ ∂(ν a) := by + have h_eq1 : ∫⁻ x in t₁, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) + = ∫⁻ x in t₁, cdfKernel f hf (a, x) t₂ ∂(ν a) := by refine' set_lintegral_congr_fun ht₁ (eventually_of_forall fun a ha ↦ _) rw [h_prod_eq_snd a ha] have h_eq2 : - ∫⁻ x in t₁ᶜ, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) = 0 := by + ∫⁻ x in t₁ᶜ, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) = 0 := by suffices h_eq_zero : - ∀ x ∈ t₁ᶜ, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} = 0 by + ∀ x ∈ t₁ᶜ, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} = 0 by rw [set_lintegral_congr_fun ht₁.compl (eventually_of_forall h_eq_zero)] simp only [lintegral_const, zero_mul] intro a hat₁ @@ -301,17 +366,17 @@ lemma lintegral_cdfKernel_mem [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) rw [h_eq1, h_eq2, add_zero] exact set_lintegral_cdfKernel_prod hf a ht₁ ht₂ · intro t ht ht_eq - calc ∫⁻ x, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ tᶜ} ∂(ν a) - = ∫⁻ x, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t}ᶜ ∂(ν a) := rfl - _ = ∫⁻ x, cdfKernel f hf.measurable (a, x) univ - - cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t} ∂(ν a) := by + calc ∫⁻ x, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ tᶜ} ∂(ν a) + = ∫⁻ x, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t}ᶜ ∂(ν a) := rfl + _ = ∫⁻ x, cdfKernel f hf (a, x) univ + - cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t} ∂(ν a) := by congr with x : 1 exact measure_compl (measurable_prod_mk_left ht) - (measure_ne_top (cdfKernel f hf.measurable (a, x)) _) - _ = ∫⁻ x, cdfKernel f hf.measurable (a, x) univ ∂(ν a) - - ∫⁻ x, cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t} ∂(ν a) := by - have h_le : (fun x ↦ cdfKernel f hf.measurable (a, x) {y : ℝ | (x, y) ∈ t}) - ≤ᵐ[ν a] fun x ↦ cdfKernel f hf.measurable (a, x) univ := + (measure_ne_top (cdfKernel f hf (a, x)) _) + _ = ∫⁻ x, cdfKernel f hf (a, x) univ ∂(ν a) - + ∫⁻ x, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t} ∂(ν a) := by + have h_le : (fun x ↦ cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t}) + ≤ᵐ[ν a] fun x ↦ cdfKernel f hf (a, x) univ := eventually_of_forall fun _ ↦ measure_mono (subset_univ _) rw [lintegral_sub _ _ h_le] · exact kernel.measurable_kernel_prod_mk_left' ht a @@ -333,26 +398,26 @@ lemma lintegral_cdfKernel_mem [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) intro h_mem_both suffices (a, x) ∈ ∅ by rwa [mem_empty_iff_false] at this rwa [← h_disj, mem_inter_iff] - calc ∫⁻ x, cdfKernel f hf.measurable (a, x) (⋃ i, {y | (x, y) ∈ f' i}) ∂(ν a) - = ∫⁻ x, ∑' i, cdfKernel f hf.measurable (a, x) {y | (x, y) ∈ f' i} ∂(ν a) := by + calc ∫⁻ x, cdfKernel f hf (a, x) (⋃ i, {y | (x, y) ∈ f' i}) ∂(ν a) + = ∫⁻ x, ∑' i, cdfKernel f hf (a, x) {y | (x, y) ∈ f' i} ∂(ν a) := by congr with x : 1 rw [measure_iUnion (h_disj x) fun i ↦ measurable_prod_mk_left (hf_meas i)] - _ = ∑' i, ∫⁻ x, cdfKernel f hf.measurable (a, x) {y | (x, y) ∈ f' i} ∂(ν a) := + _ = ∑' i, ∫⁻ x, cdfKernel f hf (a, x) {y | (x, y) ∈ f' i} ∂(ν a) := lintegral_tsum fun i ↦ (kernel.measurable_kernel_prod_mk_left' (hf_meas i) a).aemeasurable _ = ∑' i, μ a (f' i) := by simp_rw [hf_eq] _ = μ a (iUnion f') := (measure_iUnion hf_disj hf_meas).symm lemma kernel.eq_compProd_cdfKernel [IsFiniteKernel μ] [IsSFiniteKernel ν] - (hf : IsRatKernelCDF f μ ν) : - μ = ν ⊗ₖ cdfKernel f hf.measurable := by + (hf : IsKernelCDF f μ ν) : + μ = ν ⊗ₖ cdfKernel f hf := by ext a s hs rw [kernel.compProd_apply _ _ _ hs, lintegral_cdfKernel_mem hf a hs] -lemma ae_cdfKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsRatKernelCDF f μ ν) (a : α) +lemma ae_cdfKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsKernelCDF f μ ν) (a : α) {s : Set ℝ} (hs : MeasurableSet s) (hμs : μ a {x | x.snd ∈ sᶜ} = 0) : - ∀ᵐ t ∂(ν a), cdfKernel f hf.measurable (a, t) s = 1 := by - have h_eq : μ = ν ⊗ₖ cdfKernel f hf.measurable := kernel.eq_compProd_cdfKernel hf - have h : μ a {x | x.snd ∈ sᶜ} = (ν ⊗ₖ cdfKernel f hf.measurable) a {x | x.snd ∈ sᶜ} := by + ∀ᵐ t ∂(ν a), cdfKernel f hf (a, t) s = 1 := by + have h_eq : μ = ν ⊗ₖ cdfKernel f hf := kernel.eq_compProd_cdfKernel hf + have h : μ a {x | x.snd ∈ sᶜ} = (ν ⊗ₖ cdfKernel f hf) a {x | x.snd ∈ sᶜ} := by rw [← h_eq] rw [hμs, kernel.compProd_apply] at h swap; · exact measurable_snd hs.compl @@ -363,11 +428,11 @@ lemma ae_cdfKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsRatKe exact (kernel.measurable_coe _ hs.compl).comp measurable_prod_mk_left simp only [mem_compl_iff, mem_setOf_eq, kernel.prodMkLeft_apply'] at h filter_upwards [h] with t ht - change cdfKernel f hf.measurable (a, t) sᶜ = 0 at ht + change cdfKernel f hf (a, t) sᶜ = 0 at ht rwa [prob_compl_eq_zero_iff hs] at ht -lemma measurableSet_eq_one (hf : IsRatKernelCDF f μ ν) {s : Set ℝ} (hs : MeasurableSet s) : - MeasurableSet {p | cdfKernel f hf.measurable p s = 1} := +lemma measurableSet_eq_one (hf : IsKernelCDF f μ ν) {s : Set ℝ} (hs : MeasurableSet s) : + MeasurableSet {p | cdfKernel f hf p s = 1} := (kernel.measurable_coe _ hs) (measurableSet_singleton 1) end diff --git a/Mathlib/Probability/Kernel/Disintegration/draft.lean b/Mathlib/Probability/Kernel/Disintegration/draft.lean index a5e0e21e1caca..16fce343d7c8e 100644 --- a/Mathlib/Probability/Kernel/Disintegration/draft.lean +++ b/Mathlib/Probability/Kernel/Disintegration/draft.lean @@ -1228,17 +1228,29 @@ lemma isRatKernelCDF_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ end Rat +section KernelCDF + +noncomputable +def mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : α × ℝ → StieltjesFunction := + todo3 (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2) (isRatKernelCDF_mLimsupIic κ).measurable + +lemma isKernelCDF_mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : + IsKernelCDF (mLimsupCDF κ) κ (kernel.fst κ) := + isKernelCDF_todo3 (isRatKernelCDF_mLimsupIic κ) + +end KernelCDF + -- todo: name? noncomputable -def kernel.condexpReal (κ : kernel α (ℝ × ℝ)) : kernel (α × ℝ) ℝ := - cdfKernel _ (measurable_mLimsupIic κ) +def kernel.condexpReal (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : kernel (α × ℝ) ℝ := + cdfKernel _ (isKernelCDF_mLimsupCDF κ) instance (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : IsMarkovKernel (kernel.condexpReal κ) := by unfold kernel.condexpReal; infer_instance lemma kernel.eq_compProd_condexpReal (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : κ = kernel.fst κ ⊗ₖ kernel.condexpReal κ := - kernel.eq_compProd_cdfKernel (isRatKernelCDF_mLimsupIic κ) + kernel.eq_compProd_cdfKernel (isKernelCDF_mLimsupCDF κ) end Real @@ -1260,7 +1272,7 @@ def condKernelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel κ] : kernel (α let κ' := kernel.map κ (Prod.map (id : ℝ → ℝ) f) (measurable_id.prod_map hf.measurable) let x₀ := (range_nonempty f).choose kernel.comapRight - (kernel.piecewise (measurableSet_eq_one (isRatKernelCDF_mLimsupIic κ') hf.measurableSet_range) + (kernel.piecewise (measurableSet_eq_one (isKernelCDF_mLimsupCDF κ') hf.measurableSet_range) (kernel.condexpReal κ') (kernel.deterministic (fun _ ↦ x₀) measurable_const)) hf @@ -1312,7 +1324,7 @@ lemma compProd_fst_condKernelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel have h_ae : ∀ a, ∀ᵐ t ∂(kernel.fst κ a), (a, t) ∈ ρ_set := by intro a rw [← h_fst] - refine ae_cdfKernel_eq_one (isRatKernelCDF_mLimsupIic κ') a hf.measurableSet_range ?_ + refine ae_cdfKernel_eq_one (isKernelCDF_mLimsupCDF κ') a hf.measurableSet_range ?_ simp only [mem_compl_iff, mem_range, not_exists] rw [kernel.map_apply'] · have h_empty : {a : ℝ × Ω' | ∀ (x : Ω'), ¬f x = f a.2} = ∅ := by From 242d5f519d72d0c6a7a4305d1ab8330f757630f5 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 14 Feb 2024 09:25:58 +0100 Subject: [PATCH 010/129] rename todo2 lemmas --- Mathlib/Probability/Kernel/CondCdf.lean | 15 ---- .../Probability/Kernel/Disintegration.lean | 2 - .../Kernel/Disintegration/BuildKernel.lean | 22 +++++- .../Kernel/Disintegration/StieltjesReal.lean | 70 +++++++++---------- .../Kernel/Disintegration/draft.lean | 17 +++-- 5 files changed, 67 insertions(+), 59 deletions(-) diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index 05bfed4ba3479..5d69e0fc0f366 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -3,8 +3,6 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Logic.Encodable.Basic -import Mathlib.Data.Set.Lattice import Mathlib.Probability.Kernel.Disintegration.BuildKernel #align_import probability.kernel.cond_cdf from "leanprover-community/mathlib"@"3b88f4005dc2e28d42f974cc1ce838f0dafb39b8" @@ -51,19 +49,6 @@ section AuxLemmasToBeMoved variable {α β ι : Type*} -theorem Real.iUnion_Iic_rat : ⋃ r : ℚ, Iic (r : ℝ) = univ := by - ext1 x - simp only [mem_iUnion, mem_Iic, mem_univ, iff_true_iff] - obtain ⟨r, hr⟩ := exists_rat_gt x - exact ⟨r, hr.le⟩ -#align real.Union_Iic_rat Real.iUnion_Iic_rat - -theorem Real.iInter_Iic_rat : ⋂ r : ℚ, Iic (r : ℝ) = ∅ := by - ext1 x - simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false_iff, not_forall, not_le] - exact exists_rat_lt x -#align real.Inter_Iic_rat Real.iInter_Iic_rat - -- todo after the port: move to order/filter/at_top_bot theorem atBot_le_nhds_bot {α : Type*} [TopologicalSpace α] [LinearOrder α] [OrderBot α] [OrderTopology α] : (atBot : Filter α) ≤ 𝓝 ⊥ := by diff --git a/Mathlib/Probability/Kernel/Disintegration.lean b/Mathlib/Probability/Kernel/Disintegration.lean index ec0ffc4848fa5..c0f1d09cb1146 100644 --- a/Mathlib/Probability/Kernel/Disintegration.lean +++ b/Mathlib/Probability/Kernel/Disintegration.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne, Kexing Ying -/ import Mathlib.Probability.Kernel.CondCdf -import Mathlib.MeasureTheory.Constructions.Polish -import Mathlib.Probability.Kernel.MeasureCompProd #align_import probability.kernel.disintegration from "leanprover-community/mathlib"@"6315581f5650ffa2fbdbbbedc41243c8d7070981" diff --git a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean index dfe7c056eaf44..af7579db1db91 100644 --- a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean @@ -15,6 +15,24 @@ import Mathlib.Probability.Kernel.MeasureCompProd open MeasureTheory Set Filter TopologicalSpace open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory +section AuxLemmasToBeMoved + +variable {α β ι : Type*} + +theorem Real.iUnion_Iic_rat : ⋃ r : ℚ, Iic (r : ℝ) = univ := by + ext1 x + simp only [mem_iUnion, mem_Iic, mem_univ, iff_true_iff] + obtain ⟨r, hr⟩ := exists_rat_gt x + exact ⟨r, hr.le⟩ +#align real.Union_Iic_rat Real.iUnion_Iic_rat + +theorem Real.iInter_Iic_rat : ⋂ r : ℚ, Iic (r : ℝ) = ∅ := by + ext1 x + simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false_iff, not_forall, not_le] + exact exists_rat_lt x +#align real.Inter_Iic_rat Real.iInter_Iic_rat + +end AuxLemmasToBeMoved namespace ProbabilityTheory @@ -270,12 +288,10 @@ lemma set_lintegral_cdfKernel_Iic [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) simp_rw [cdfKernel_Iic] exact hf.set_lintegral _ hs _ -theorem Real.iUnion_Iic_rat' : ⋃ r : ℚ, Iic (r : ℝ) = univ := sorry - lemma set_lintegral_cdfKernel_univ [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) {s : Set β} (hs : MeasurableSet s) : ∫⁻ t in s, cdfKernel f hf (a, t) univ ∂(ν a) = μ a (s ×ˢ univ) := by - rw [← Real.iUnion_Iic_rat', prod_iUnion] + rw [← Real.iUnion_Iic_rat, prod_iUnion] have h_dir : Directed (fun x y ↦ x ⊆ y) fun q : ℚ ↦ Iic (q : ℝ) := by refine Monotone.directed_le fun r r' hrr' ↦ Iic_subset_Iic.mpr ?_ exact mod_cast hrr' diff --git a/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean b/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean index 11033196e7f98..1eaad20e3420a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean +++ b/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean @@ -300,21 +300,21 @@ noncomputable def IsCDFLike.stieltjesFunction (a : α) : StieltjesFunction where mono' := monotone_todo1 hf a right_continuous' x := continuousWithinAt_todo1_Ici hf a x -theorem todo2_eq (a : α) (r : ℚ) : hf.stieltjesFunction a r = f a r := todo1_eq hf a r +theorem IsCDFLike.stieltjesFunction_eq (a : α) (r : ℚ) : hf.stieltjesFunction a r = f a r := todo1_eq hf a r /-- The conditional cdf is non-negative for all `a : α`. -/ -theorem todo2_nonneg (a : α) (r : ℝ) : 0 ≤ hf.stieltjesFunction a r := todo1_nonneg hf a r +theorem IsCDFLike.stieltjesFunction_nonneg (a : α) (r : ℝ) : 0 ≤ hf.stieltjesFunction a r := todo1_nonneg hf a r /-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ -theorem todo2_le_one (a : α) (x : ℝ) : hf.stieltjesFunction a x ≤ 1 := by +theorem IsCDFLike.stieltjesFunction_le_one (a : α) (x : ℝ) : hf.stieltjesFunction a x ≤ 1 := by obtain ⟨r, hrx⟩ := exists_rat_gt x rw [← StieltjesFunction.iInf_rat_gt_eq] - simp_rw [todo2_eq] + simp_rw [IsCDFLike.stieltjesFunction_eq] refine ciInf_le_of_le (bddBelow_range_gt hf a x) ?_ (hf.le_one _ _) exact ⟨r, hrx⟩ /-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ -theorem tendsto_todo2_atBot (a : α) : +theorem IsCDFLike.tendsto_stieltjesFunction_atBot (a : α) : Tendsto (hf.stieltjesFunction a) atBot (𝓝 0) := by have h_exists : ∀ x : ℝ, ∃ q : ℚ, x < q ∧ ↑q < x + 1 := fun x ↦ exists_rat_btwn (lt_add_one x) let qs : ℝ → ℚ := fun x ↦ (h_exists x).choose @@ -326,12 +326,12 @@ theorem tendsto_todo2_atBot (a : α) : rw [sub_add_cancel] at h_le exact mod_cast h_le refine tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds - ((hf.tendsto_atBot_zero a).comp hqs_tendsto) (todo2_nonneg hf a) fun x ↦ ?_ - rw [Function.comp_apply, ← todo2_eq hf] + ((hf.tendsto_atBot_zero a).comp hqs_tendsto) (stieltjesFunction_nonneg hf a) fun x ↦ ?_ + rw [Function.comp_apply, ← stieltjesFunction_eq hf] exact (hf.stieltjesFunction a).mono (h_exists x).choose_spec.1.le /-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ -theorem tendsto_todo2_atTop (a : α) : +theorem IsCDFLike.tendsto_stieltjesFunction_atTop (a : α) : Tendsto (hf.stieltjesFunction a) atTop (𝓝 1) := by have h_exists : ∀ x : ℝ, ∃ q : ℚ, x - 1 < q ∧ ↑q < x := fun x ↦ exists_rat_btwn (sub_one_lt x) let qs : ℝ → ℚ := fun x ↦ (h_exists x).choose @@ -342,43 +342,43 @@ theorem tendsto_todo2_atTop (a : α) : rw [sub_le_iff_le_add] at h_le exact_mod_cast le_of_add_le_add_right (hy.trans h_le) refine tendsto_of_tendsto_of_tendsto_of_le_of_le ((hf.tendsto_atTop_one a).comp hqs_tendsto) - tendsto_const_nhds ?_ (todo2_le_one hf a) + tendsto_const_nhds ?_ (stieltjesFunction_le_one hf a) intro x - rw [Function.comp_apply, ← todo2_eq hf] + rw [Function.comp_apply, ← stieltjesFunction_eq hf] exact (hf.stieltjesFunction a).mono (le_of_lt (h_exists x).choose_spec.2) /-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ -theorem measurable_todo2 (x : ℝ) : Measurable fun a ↦ hf.stieltjesFunction a x := by +theorem IsCDFLike.measurable_stieltjesFunction (x : ℝ) : Measurable fun a ↦ hf.stieltjesFunction a x := by have : (fun a ↦ hf.stieltjesFunction a x) = fun a ↦ ⨅ r : { r' : ℚ // x < r' }, f a ↑r := by ext1 a rw [← StieltjesFunction.iInf_rat_gt_eq] congr with q - rw [todo2_eq] + rw [stieltjesFunction_eq] rw [this] exact measurable_iInf (fun q ↦ hf.measurable q) /-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ -theorem stronglyMeasurable_todo2 (x : ℝ) : +theorem IsCDFLike.stronglyMeasurable_stieltjesFunction (x : ℝ) : StronglyMeasurable fun a ↦ hf.stieltjesFunction a x := - (measurable_todo2 hf x).stronglyMeasurable + (measurable_stieltjesFunction hf x).stronglyMeasurable section Measure -theorem measure_todo2_Iic (a : α) (x : ℝ) : +theorem IsCDFLike.measure_stieltjesFunction_Iic (a : α) (x : ℝ) : (hf.stieltjesFunction a).measure (Iic x) = ENNReal.ofReal (hf.stieltjesFunction a x) := by rw [← sub_zero (hf.stieltjesFunction a x)] - exact (hf.stieltjesFunction a).measure_Iic (tendsto_todo2_atBot hf a) _ + exact (hf.stieltjesFunction a).measure_Iic (tendsto_stieltjesFunction_atBot hf a) _ -theorem measure_todo2_univ (a : α) : (hf.stieltjesFunction a).measure univ = 1 := by +theorem IsCDFLike.measure_stieltjesFunction_univ (a : α) : (hf.stieltjesFunction a).measure univ = 1 := by rw [← ENNReal.ofReal_one, ← sub_zero (1 : ℝ)] - exact StieltjesFunction.measure_univ _ (tendsto_todo2_atBot hf a) (tendsto_todo2_atTop hf a) + exact StieltjesFunction.measure_univ _ (tendsto_stieltjesFunction_atBot hf a) (tendsto_stieltjesFunction_atTop hf a) -instance instIsProbabilityMeasure_todo2 (a : α) : +instance IsCDFLike.instIsProbabilityMeasure_stieltjesFunction (a : α) : IsProbabilityMeasure (hf.stieltjesFunction a).measure := - ⟨measure_todo2_univ hf a⟩ + ⟨measure_stieltjesFunction_univ hf a⟩ /-- The function `a ↦ (condCDF ρ a).measure` is measurable. -/ -theorem measurable_measure_todo2 : +theorem IsCDFLike.measurable_measure_stieltjesFunction : Measurable fun a ↦ (hf.stieltjesFunction a).measure := by rw [Measure.measurable_measure] refine' fun s hs ↦ ?_ @@ -388,15 +388,15 @@ theorem measurable_measure_todo2 : (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic _ _ _ _ hs · simp only [measure_empty, measurable_const] · rintro S ⟨u, rfl⟩ - simp_rw [measure_todo2_Iic hf _ u] - exact (measurable_todo2 hf u).ennreal_ofReal + simp_rw [measure_stieltjesFunction_Iic hf _ u] + exact (measurable_stieltjesFunction hf u).ennreal_ofReal · intro t ht ht_cd_meas have : (fun a ↦ (hf.stieltjesFunction a).measure tᶜ) = (fun a ↦ (hf.stieltjesFunction a).measure univ) - fun a ↦ (hf.stieltjesFunction a).measure t := by ext1 a rw [measure_compl ht (measure_ne_top (hf.stieltjesFunction a).measure _), Pi.sub_apply] - simp_rw [this, measure_todo2_univ hf] + simp_rw [this, measure_stieltjesFunction_univ hf] exact Measurable.sub measurable_const ht_cd_meas · intro f hf_disj hf_meas hf_cd_meas simp_rw [measure_iUnion hf_disj hf_meas] @@ -415,7 +415,7 @@ def todo3 (f : α → ℚ → ℝ) (hf : ∀ q, Measurable fun a ↦ f a q) : α (isCDFLike_toCDFLike hf).stieltjesFunction theorem todo3_eq (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℚ) : - todo3 f hf a r = toCDFLike f a r := todo2_eq _ a r + todo3 f hf a r = toCDFLike f a r := IsCDFLike.stieltjesFunction_eq _ a r lemma todo3_unit_prod (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : todo3 (fun (p : Unit × α) ↦ f p.2) (fun q ↦ (hf q).comp measurable_snd) ((), a) @@ -428,43 +428,43 @@ lemma todo3_unit_prod (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : /-- The conditional cdf is non-negative for all `a : α`. -/ theorem todo3_nonneg (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℝ) : - 0 ≤ todo3 f hf a r := todo2_nonneg _ a r + 0 ≤ todo3 f hf a r := IsCDFLike.stieltjesFunction_nonneg _ a r /-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ theorem todo3_le_one (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : - todo3 f hf a x ≤ 1 := todo2_le_one _ a x + todo3 f hf a x ≤ 1 := IsCDFLike.stieltjesFunction_le_one _ a x /-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ theorem tendsto_todo3_atBot (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : - Tendsto (todo3 f hf a) atBot (𝓝 0) := tendsto_todo2_atBot _ a + Tendsto (todo3 f hf a) atBot (𝓝 0) := IsCDFLike.tendsto_stieltjesFunction_atBot _ a /-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ theorem tendsto_todo3_atTop (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : - Tendsto (todo3 f hf a) atTop (𝓝 1) := tendsto_todo2_atTop _ a + Tendsto (todo3 f hf a) atTop (𝓝 1) := IsCDFLike.tendsto_stieltjesFunction_atTop _ a /-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ theorem measurable_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : - Measurable fun a ↦ todo3 f hf a x := measurable_todo2 _ x + Measurable fun a ↦ todo3 f hf a x := IsCDFLike.measurable_stieltjesFunction _ x /-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ theorem stronglyMeasurable_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : - StronglyMeasurable fun a ↦ todo3 f hf a x := stronglyMeasurable_todo2 _ x + StronglyMeasurable fun a ↦ todo3 f hf a x := IsCDFLike.stronglyMeasurable_stieltjesFunction _ x section Measure theorem measure_todo3_Iic (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : (todo3 f hf a).measure (Iic x) = ENNReal.ofReal (todo3 f hf a x) := - measure_todo2_Iic _ _ _ + IsCDFLike.measure_stieltjesFunction_Iic _ _ _ theorem measure_todo3_univ (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : - (todo3 f hf a).measure univ = 1 := measure_todo2_univ _ _ + (todo3 f hf a).measure univ = 1 := IsCDFLike.measure_stieltjesFunction_univ _ _ instance instIsProbabilityMeasure_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : IsProbabilityMeasure (todo3 f hf a).measure := - instIsProbabilityMeasure_todo2 _ _ + IsCDFLike.instIsProbabilityMeasure_stieltjesFunction _ _ theorem measurable_measure_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) : - Measurable fun a ↦ (todo3 f hf a).measure := measurable_measure_todo2 _ + Measurable fun a ↦ (todo3 f hf a).measure := IsCDFLike.measurable_measure_stieltjesFunction _ end Measure diff --git a/Mathlib/Probability/Kernel/Disintegration/draft.lean b/Mathlib/Probability/Kernel/Disintegration/draft.lean index 16fce343d7c8e..ded81f98f2c76 100644 --- a/Mathlib/Probability/Kernel/Disintegration/draft.lean +++ b/Mathlib/Probability/Kernel/Disintegration/draft.lean @@ -1,4 +1,3 @@ -import Mathlib.Probability.Kernel.Disintegration import Mathlib.Probability.Martingale.Convergence import Mathlib.Probability.Kernel.Disintegration.BuildKernel @@ -31,9 +30,19 @@ lemma measurableSet_I (n : ℕ) (k : ℤ) : MeasurableSet (I n k) := measurableS lemma Measure.iInf_Iic_gt_prod {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] {s : Set α} (hs : MeasurableSet s) (t : ℚ) : ⨅ r : { r' : ℚ // t < r' }, ρ (s ×ˢ Iic (r : ℝ)) = ρ (s ×ˢ Iic (t : ℝ)) := by - have h := Measure.iInf_IicSnd_gt ρ t hs - simp_rw [Measure.IicSnd_apply ρ _ hs] at h - rw [← h] + rw [← measure_iInter_eq_iInf] + · rw [← prod_iInter] + congr with x : 1 + simp only [mem_iInter, mem_Iic, Subtype.forall, Subtype.coe_mk] + refine' ⟨fun h => _, fun h a hta => h.trans _⟩ + · refine' le_of_forall_lt_rat_imp_le fun q htq => h q _ + exact mod_cast htq + · exact mod_cast hta.le + · exact fun _ => hs.prod measurableSet_Iic + · refine' Monotone.directed_ge fun r r' hrr' => prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, _⟩) + refine' Iic_subset_Iic.mpr _ + exact mod_cast hrr' + · exact ⟨⟨t + 1, lt_add_one _⟩, measure_ne_top ρ _⟩ lemma pairwise_disjoint_I (n : ℕ) : Pairwise (Disjoint on fun k ↦ I n k) := by intro i j hij From 6ea94c411f6d0cc9f53fab53b62d9a14ea13b01d Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 14 Feb 2024 10:49:22 +0100 Subject: [PATCH 011/129] reorganize. WIP --- .../MeasureTheory/Constructions/Polish.lean | 10 + Mathlib/Probability/Kernel/CondCdf.lean | 23 +- .../Probability/Kernel/Disintegration.lean | 2 +- .../Kernel/Disintegration/Basic.lean | 237 ++++++++++++++++++ .../Kernel/Disintegration/BuildKernel.lean | 2 +- .../{draft.lean => KernelCDFBorel.lean} | 168 ------------- .../Kernel/Disintegration/StieltjesReal.lean | 105 ++++---- 7 files changed, 305 insertions(+), 242 deletions(-) create mode 100644 Mathlib/Probability/Kernel/Disintegration/Basic.lean rename Mathlib/Probability/Kernel/Disintegration/{draft.lean => KernelCDFBorel.lean} (88%) diff --git a/Mathlib/MeasureTheory/Constructions/Polish.lean b/Mathlib/MeasureTheory/Constructions/Polish.lean index ef4908a785719..291c04e5b8ed9 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish.lean @@ -1074,4 +1074,14 @@ theorem exists_measurableEmbedding_real : ∃ f : α → ℝ, MeasurableEmbeddin exact ⟨(↑) ∘ e, (MeasurableEmbedding.subtype_coe hs).comp e.measurableEmbedding⟩ #align measure_theory.exists_measurable_embedding_real MeasureTheory.exists_measurableEmbedding_real +/-- A measurable embedding of a standard Borel space into `ℝ`. -/ +noncomputable +def measurableEmbedding_real (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : Ω → ℝ := + (exists_measurableEmbedding_real Ω).choose + +lemma measurableEmbedding_measurableEmbedding_real + (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : + MeasurableEmbedding (measurableEmbedding_real Ω) := + (exists_measurableEmbedding_real Ω).choose_spec + end MeasureTheory diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index 5d69e0fc0f366..f4dd16d9cf0ed 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ +import Mathlib.MeasureTheory.Decomposition.RadonNikodym import Mathlib.Probability.Kernel.Disintegration.BuildKernel #align_import probability.kernel.cond_cdf from "leanprover-community/mathlib"@"3b88f4005dc2e28d42f974cc1ce838f0dafb39b8" @@ -620,24 +621,8 @@ lemma isKernelCDF_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : tendsto_atBot_zero p := tendsto_condCDF_atBot ρ p.2 set_integral _ _ hs x := set_integral_condCDF ρ x hs -section Measure - -theorem measure_condCDF_Iic (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : - (condCDF ρ a).measure (Iic x) = ENNReal.ofReal (condCDF ρ a x) := measure_todo3_Iic _ _ _ -#align probability_theory.measure_cond_cdf_Iic ProbabilityTheory.measure_condCDF_Iic - -theorem measure_condCDF_univ (ρ : Measure (α × ℝ)) (a : α) : (condCDF ρ a).measure univ = 1 := - measure_todo3_univ _ _ -#align probability_theory.measure_cond_cdf_univ ProbabilityTheory.measure_condCDF_univ - -instance instIsProbabilityMeasure (ρ : Measure (α × ℝ)) (a : α) : - IsProbabilityMeasure (condCDF ρ a).measure := by rw [condCDF]; infer_instance - -/-- The function `a ↦ (condCDF ρ a).measure` is measurable. -/ -theorem measurable_measure_condCDF (ρ : Measure (α × ℝ)) : - Measurable fun a => (condCDF ρ a).measure := measurable_measure_todo3 _ -#align probability_theory.measurable_measure_cond_cdf ProbabilityTheory.measurable_measure_condCDF - -end Measure +#noalign probability_theory.measure_cond_cdf_Iic +#noalign probability_theory.measure_cond_cdf_univ +#noalign probability_theory.measurable_measure_cond_cdf end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Disintegration.lean b/Mathlib/Probability/Kernel/Disintegration.lean index c0f1d09cb1146..aa30bf92fa818 100644 --- a/Mathlib/Probability/Kernel/Disintegration.lean +++ b/Mathlib/Probability/Kernel/Disintegration.lean @@ -63,7 +63,7 @@ noncomputable def condKernelReal (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (fun a ↦ ((), a)) measurable_prod_mk_left #align probability_theory.cond_kernel_real ProbabilityTheory.condKernelReal -instance (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : IsMarkovKernel (condKernelReal ρ) := by +instance (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : IsMarkovKernel (condKernelReal ρ) := by rw [condKernelReal]; infer_instance theorem condKernelReal_Iic (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (a : α) (x : ℝ) : diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean new file mode 100644 index 0000000000000..4be74195e3530 --- /dev/null +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -0,0 +1,237 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne, Kexing Ying +-/ +import Mathlib.Probability.Kernel.CondCdf +import Mathlib.Probability.Kernel.Disintegration.KernelCDFBorel + +/-! +# Disintegration of kernels and measures + +-/ + +open MeasureTheory Set Filter + +open scoped ENNReal MeasureTheory Topology ProbabilityTheory + +namespace ProbabilityTheory + +variable {α β Ω Ω': Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] + [MeasurableSpace Ω'] [StandardBorelSpace Ω'] [Nonempty Ω'] + +section BorelSnd + +noncomputable +def condKernelBorelSnd (κ : kernel α (β × Ω')) [IsMarkovKernel κ] + {f : α × β → StieltjesFunction} + (hf : IsKernelCDF f + (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) + (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω').measurable)) + (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) + (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω').measurable)))) : + kernel (α × β) Ω' := + let e := measurableEmbedding_real Ω' + let he := measurableEmbedding_measurableEmbedding_real Ω' + let x₀ := (range_nonempty e).choose + kernel.comapRight + (kernel.piecewise (measurableSet_eq_one hf he.measurableSet_range) + (cdfKernel f hf) (kernel.deterministic (fun _ ↦ x₀) measurable_const)) + he + +instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω')) [IsMarkovKernel κ] + {f : α × β → StieltjesFunction} + (hf : IsKernelCDF f + (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) + (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω').measurable)) + (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) + (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω').measurable)))) : + IsMarkovKernel (condKernelBorelSnd κ hf) := by + rw [condKernelBorelSnd] + refine kernel.IsMarkovKernel.comapRight _ _ fun a ↦ ?_ + rw [kernel.piecewise_apply'] + split_ifs with h_mem + · exact h_mem + · classical + rw [kernel.deterministic_apply' _ _ + (measurableEmbedding_measurableEmbedding_real Ω').measurableSet_range, + Set.indicator_apply, if_pos] + exact (range_nonempty (measurableEmbedding_real Ω')).choose_spec + +lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω')) [IsMarkovKernel κ] + {f : α × β → StieltjesFunction} + (hf : IsKernelCDF f + (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) + (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω').measurable)) + (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) + (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω').measurable)))) : + kernel.fst κ ⊗ₖ condKernelBorelSnd κ hf = κ := by + let e := measurableEmbedding_real Ω' + let he := measurableEmbedding_measurableEmbedding_real Ω' + let κ' := kernel.map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable) + have h_prod_embed : MeasurableEmbedding (Prod.map (id : β → β) e) := + MeasurableEmbedding.id.prod_mk he + have h_fst : kernel.fst κ' = kernel.fst κ := by + ext a u hu + unfold_let κ' + rw [kernel.fst_apply' _ _ hu, kernel.fst_apply' _ _ hu, + kernel.map_apply' κ h_prod_embed.measurable] + · rfl + · exact measurable_fst hu + have : κ = kernel.comapRight κ' h_prod_embed := by + ext c t ht : 2 + unfold_let κ' + rw [kernel.comapRight_apply' _ _ _ ht, kernel.map_apply' κ h_prod_embed.measurable + _ (h_prod_embed.measurableSet_image.mpr ht)] + congr with x : 1 + rw [← @Prod.mk.eta _ _ x] + simp only [id.def, mem_preimage, Prod.map_mk, mem_image, Prod.mk.inj_iff, Prod.exists] + refine' ⟨fun h => ⟨x.1, x.2, h, rfl, rfl⟩, _⟩ + rintro ⟨a, b, h_mem, rfl, he_eq⟩ + rwa [he.injective he_eq] at h_mem + conv_rhs => rw [this] + unfold_let κ' + conv_rhs => rw [kernel.eq_compProd_cdfKernel hf] + change kernel.fst κ ⊗ₖ condKernelBorelSnd κ hf + = kernel.comapRight (kernel.fst κ' ⊗ₖ cdfKernel f hf) h_prod_embed + ext c t ht : 2 + rw [kernel.comapRight_apply' _ _ _ ht, + kernel.compProd_apply _ _ _ (h_prod_embed.measurableSet_image.mpr ht)] + simp_rw [h_fst, kernel.compProd_apply _ _ _ ht] + refine lintegral_congr_ae ?_ + let ρ_set := {p : α × β | cdfKernel f hf p (range e) = 1} + have h_ae : ∀ a, ∀ᵐ t ∂(kernel.fst κ a), (a, t) ∈ ρ_set := by + intro a + rw [← h_fst] + refine ae_cdfKernel_eq_one hf a he.measurableSet_range ?_ + simp only [mem_compl_iff, mem_range, not_exists] + rw [kernel.map_apply'] + · have h_empty : {a : β × Ω' | ∀ (x : Ω'), ¬e x = e a.2} = ∅ := by + ext x + simp only [mem_setOf_eq, mem_empty_iff_false, iff_false, not_forall, not_not, + exists_apply_eq_apply] + simp [h_empty] + · have : {x : β × ℝ | ∀ (y : Ω'), ¬ e y = x.2} = univ ×ˢ (range e)ᶜ := by + ext x + simp only [mem_setOf_eq, mem_prod, mem_univ, mem_compl_iff, mem_range, not_exists, true_and] + rw [this] + exact MeasurableSet.univ.prod he.measurableSet_range.compl + filter_upwards [h_ae c] with a ha + rw [condKernelBorelSnd, kernel.comapRight_apply'] + swap; · exact measurable_prod_mk_left ht + have h1 : {c : ℝ | (a, c) ∈ Prod.map id e '' t} = e '' {c : Ω' | (a, c) ∈ t} := by + ext1 x + simp only [Prod_map, id.def, mem_image, Prod.mk.inj_iff, Prod.exists, mem_setOf_eq] + constructor + · rintro ⟨a', b, h_mem, rfl, hf_eq⟩ + exact ⟨b, h_mem, hf_eq⟩ + · rintro ⟨b, h_mem, hf_eq⟩ + exact ⟨a, b, h_mem, rfl, hf_eq⟩ + rw [h1, kernel.piecewise_apply, if_pos ha] + rfl + +end BorelSnd + +section StandardBorel + +noncomputable +def kernel.condKernelReal (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : kernel (α × ℝ) ℝ := + cdfKernel _ (isKernelCDF_mLimsupCDF κ) + +instance (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : IsMarkovKernel (kernel.condKernelReal κ) := by + unfold kernel.condKernelReal; infer_instance + +lemma kernel.eq_compProd_condKernelReal (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : + κ = kernel.fst κ ⊗ₖ kernel.condKernelReal κ := + kernel.eq_compProd_cdfKernel (isKernelCDF_mLimsupCDF κ) + +noncomputable +def condKernelBorelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel κ] : kernel (α × ℝ) Ω' := + let f := measurableEmbedding_real Ω' + let hf := measurableEmbedding_measurableEmbedding_real Ω' + let κ' := kernel.map κ (Prod.map (id : ℝ → ℝ) f) (measurable_id.prod_map hf.measurable) + condKernelBorelSnd κ (isKernelCDF_mLimsupCDF κ') + +instance instIsMarkovKernel_condKernelBorelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel κ] : + IsMarkovKernel (condKernelBorelAux κ) := by + rw [condKernelBorelAux] + infer_instance + +lemma compProd_fst_condKernelBorelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel κ] : + kernel.fst κ ⊗ₖ condKernelBorelAux κ = κ := by + rw [condKernelBorelAux, compProd_fst_condKernelBorelSnd] + +noncomputable +def condKernelBorel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : kernel (α × Ω) Ω' := + let f := measurableEmbedding_real Ω + let hf := measurableEmbedding_measurableEmbedding_real Ω + let κ' := kernel.map κ (Prod.map f (id : Ω' → Ω')) (hf.measurable.prod_map measurable_id) + kernel.comap (condKernelBorelAux κ') (Prod.map (id : α → α) f) + (measurable_id.prod_map hf.measurable) + +instance instIsMarkovKernel_condKernelBorel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : + IsMarkovKernel (condKernelBorel κ) := by rw [condKernelBorel]; infer_instance + +lemma compProd_fst_condKernel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : + kernel.fst κ ⊗ₖ condKernelBorel κ = κ := by + let f := measurableEmbedding_real Ω + let hf := measurableEmbedding_measurableEmbedding_real Ω + let κ' : kernel α (ℝ × Ω') := + kernel.map κ (Prod.map f (id : Ω' → Ω')) (hf.measurable.prod_map measurable_id) + have h_condKernel : condKernelBorel κ + = kernel.comap (condKernelBorelAux κ') (Prod.map (id : α → α) f) + (measurable_id.prod_map hf.measurable) := rfl + have h_compProd := compProd_fst_condKernelBorelAux κ' + have h_prod_embed : MeasurableEmbedding (Prod.map f (id : Ω' → Ω')) := + hf.prod_mk MeasurableEmbedding.id + have : κ = kernel.comapRight κ' h_prod_embed := by + ext c t ht : 2 + unfold_let κ' + rw [kernel.comapRight_apply' _ _ _ ht, kernel.map_apply' κ h_prod_embed.measurable + _ (h_prod_embed.measurableSet_image.mpr ht)] + congr with x : 1 + rw [← @Prod.mk.eta _ _ x] + simp only [Prod.mk.eta, Prod_map, id_eq, mem_preimage, mem_image, Prod.mk.injEq, Prod.exists, + exists_eq_right_right] + refine ⟨fun h ↦ ⟨x.1, h, rfl⟩, ?_⟩ + rintro ⟨ω, h_mem, h⟩ + rwa [hf.injective h] at h_mem + have h_fst : kernel.fst κ' = kernel.map (kernel.fst κ) f hf.measurable := by + ext a s hs + unfold_let κ' + rw [kernel.map_apply' _ _ _ hs, kernel.fst_apply', kernel.fst_apply', kernel.map_apply'] + · congr + · exact measurable_fst hs + · exact hf.measurable hs + · exact hs + conv_rhs => rw [this, ← h_compProd] + ext a s hs + rw [h_condKernel, h_fst] + rw [kernel.comapRight_apply' _ _ _ hs, kernel.compProd_apply _ _ _ hs, kernel.compProd_apply] + swap; · exact h_prod_embed.measurableSet_image.mpr hs + rw [kernel.map_apply, lintegral_map] + congr with ω + rw [kernel.comap_apply'] + · congr with ω' + simp only [mem_setOf_eq, Prod_map, id_eq, mem_image, Prod.mk.injEq, Prod.exists, + exists_eq_right_right] + refine ⟨fun h ↦ ⟨ω, h, rfl⟩, ?_⟩ + rintro ⟨a, h_mem, h⟩ + rwa [hf.injective h] at h_mem + · exact kernel.measurable_kernel_prod_mk_left' (h_prod_embed.measurableSet_image.mpr hs) _ + · exact hf.measurable + +end StandardBorel + +section Countable + +section Unit + + + +end Unit + +end Countable + +end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean index af7579db1db91..bec5c1621a242 100644 --- a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean @@ -368,7 +368,7 @@ lemma lintegral_cdfKernel_mem [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) rw [← lintegral_add_compl _ ht₁] have h_eq1 : ∫⁻ x in t₁, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) = ∫⁻ x in t₁, cdfKernel f hf (a, x) t₂ ∂(ν a) := by - refine' set_lintegral_congr_fun ht₁ (eventually_of_forall fun a ha ↦ _) + refine set_lintegral_congr_fun ht₁ (eventually_of_forall fun a ha ↦ ?_) rw [h_prod_eq_snd a ha] have h_eq2 : ∫⁻ x in t₁ᶜ, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) = 0 := by diff --git a/Mathlib/Probability/Kernel/Disintegration/draft.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean similarity index 88% rename from Mathlib/Probability/Kernel/Disintegration/draft.lean rename to Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean index ded81f98f2c76..711fab705fee4 100644 --- a/Mathlib/Probability/Kernel/Disintegration/draft.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean @@ -1249,174 +1249,6 @@ lemma isKernelCDF_mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : end KernelCDF --- todo: name? -noncomputable -def kernel.condexpReal (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : kernel (α × ℝ) ℝ := - cdfKernel _ (isKernelCDF_mLimsupCDF κ) - -instance (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : IsMarkovKernel (kernel.condexpReal κ) := by - unfold kernel.condexpReal; infer_instance - -lemma kernel.eq_compProd_condexpReal (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : - κ = kernel.fst κ ⊗ₖ kernel.condexpReal κ := - kernel.eq_compProd_cdfKernel (isKernelCDF_mLimsupCDF κ) - end Real -variable {Ω' : Type*} [MeasurableSpace Ω'] [StandardBorelSpace Ω'] [Nonempty Ω'] - -noncomputable -def measurableEmbedding_real (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : Ω → ℝ := - (exists_measurableEmbedding_real Ω).choose - -lemma measurableEmbedding_measurableEmbedding_real - (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : - MeasurableEmbedding (measurableEmbedding_real Ω) := - (exists_measurableEmbedding_real Ω).choose_spec - -noncomputable -def condKernelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel κ] : kernel (α × ℝ) Ω' := - let f := measurableEmbedding_real Ω' - let hf := measurableEmbedding_measurableEmbedding_real Ω' - let κ' := kernel.map κ (Prod.map (id : ℝ → ℝ) f) (measurable_id.prod_map hf.measurable) - let x₀ := (range_nonempty f).choose - kernel.comapRight - (kernel.piecewise (measurableSet_eq_one (isKernelCDF_mLimsupCDF κ') hf.measurableSet_range) - (kernel.condexpReal κ') (kernel.deterministic (fun _ ↦ x₀) measurable_const)) - hf - -instance instIsMarkovKernel_condKernelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel κ] : - IsMarkovKernel (condKernelAux κ) := by - rw [condKernelAux] - refine kernel.IsMarkovKernel.comapRight _ _ fun a ↦ ?_ - rw [kernel.piecewise_apply'] - split_ifs with h_mem - · exact h_mem - · classical - rw [kernel.deterministic_apply' _ _ - (measurableEmbedding_measurableEmbedding_real Ω').measurableSet_range, - Set.indicator_apply, if_pos] - exact (range_nonempty (measurableEmbedding_real Ω')).choose_spec - -lemma compProd_fst_condKernelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel κ] : - kernel.fst κ ⊗ₖ condKernelAux κ = κ := by - let f := measurableEmbedding_real Ω' - let hf := measurableEmbedding_measurableEmbedding_real Ω' - let κ' := kernel.map κ (Prod.map (id : ℝ → ℝ) f) (measurable_id.prod_map hf.measurable) - have h_prod_embed : MeasurableEmbedding (Prod.map (id : ℝ → ℝ) f) := - MeasurableEmbedding.id.prod_mk hf - have h_fst : kernel.fst κ' = kernel.fst κ := by - ext a u hu - unfold_let κ' - rw [kernel.fst_apply' _ _ hu, kernel.fst_apply' _ _ hu, - kernel.map_apply' κ h_prod_embed.measurable] - · rfl - · exact measurable_fst hu - have : κ = kernel.comapRight κ' h_prod_embed := by - ext c t ht : 2 - unfold_let κ' - rw [kernel.comapRight_apply' _ _ _ ht, kernel.map_apply' κ h_prod_embed.measurable - _ (h_prod_embed.measurableSet_image.mpr ht)] - congr with x : 1 - rw [← @Prod.mk.eta _ _ x] - simp only [id.def, mem_preimage, Prod.map_mk, mem_image, Prod.mk.inj_iff, Prod.exists] - refine' ⟨fun h => ⟨x.1, x.2, h, rfl, rfl⟩, _⟩ - rintro ⟨a, b, h_mem, rfl, hf_eq⟩ - rwa [hf.injective hf_eq] at h_mem - conv_rhs => rw [this, kernel.eq_compProd_condexpReal κ'] - ext c t ht : 2 - rw [kernel.comapRight_apply' _ _ _ ht, - kernel.compProd_apply _ _ _ (h_prod_embed.measurableSet_image.mpr ht), - h_fst, kernel.compProd_apply _ _ _ ht] - refine lintegral_congr_ae ?_ - let ρ_set := {p : α × ℝ | kernel.condexpReal κ' p (range f) = 1} - have h_ae : ∀ a, ∀ᵐ t ∂(kernel.fst κ a), (a, t) ∈ ρ_set := by - intro a - rw [← h_fst] - refine ae_cdfKernel_eq_one (isKernelCDF_mLimsupCDF κ') a hf.measurableSet_range ?_ - simp only [mem_compl_iff, mem_range, not_exists] - rw [kernel.map_apply'] - · have h_empty : {a : ℝ × Ω' | ∀ (x : Ω'), ¬f x = f a.2} = ∅ := by - ext x - simp only [mem_setOf_eq, mem_empty_iff_false, iff_false, not_forall, not_not, - exists_apply_eq_apply] - simp [h_empty] - · have : {x : ℝ × ℝ | ∀ (y : Ω'), ¬ f y = x.2} = univ ×ˢ (range f)ᶜ := by - ext x - simp only [mem_setOf_eq, mem_prod, mem_univ, mem_compl_iff, mem_range, not_exists, true_and] - rw [this] - exact MeasurableSet.univ.prod hf.measurableSet_range.compl - filter_upwards [h_ae c] with a ha - rw [condKernelAux, kernel.comapRight_apply'] - swap; · exact measurable_prod_mk_left ht - have h1 : {c : ℝ | (a, c) ∈ Prod.map id f '' t} = f '' {c : Ω' | (a, c) ∈ t} := by - ext1 x - simp only [Prod_map, id.def, mem_image, Prod.mk.inj_iff, Prod.exists, mem_setOf_eq] - constructor - · rintro ⟨a', b, h_mem, rfl, hf_eq⟩ - exact ⟨b, h_mem, hf_eq⟩ - · rintro ⟨b, h_mem, hf_eq⟩ - exact ⟨a, b, h_mem, rfl, hf_eq⟩ - rw [h1, kernel.piecewise_apply, if_pos] - exact ha - -noncomputable -def condKernel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : kernel (α × Ω) Ω' := - let f := measurableEmbedding_real Ω - let hf := measurableEmbedding_measurableEmbedding_real Ω - let κ' := kernel.map κ (Prod.map f (id : Ω' → Ω')) (hf.measurable.prod_map measurable_id) - kernel.comap (condKernelAux κ') (Prod.map (id : α → α) f) (measurable_id.prod_map hf.measurable) - -instance instIsMarkovKernel_condKernel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : - IsMarkovKernel (condKernel κ) := by rw [condKernel]; infer_instance - -lemma compProd_fst_condKernel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : - kernel.fst κ ⊗ₖ condKernel κ = κ := by - let f := measurableEmbedding_real Ω - let hf := measurableEmbedding_measurableEmbedding_real Ω - let κ' : kernel α (ℝ × Ω') := - kernel.map κ (Prod.map f (id : Ω' → Ω')) (hf.measurable.prod_map measurable_id) - have h_condKernel : condKernel κ - = kernel.comap (condKernelAux κ') (Prod.map (id : α → α) f) - (measurable_id.prod_map hf.measurable) := rfl - have h_compProd := compProd_fst_condKernelAux κ' - have h_prod_embed : MeasurableEmbedding (Prod.map f (id : Ω' → Ω')) := - hf.prod_mk MeasurableEmbedding.id - have : κ = kernel.comapRight κ' h_prod_embed := by - ext c t ht : 2 - unfold_let κ' - rw [kernel.comapRight_apply' _ _ _ ht, kernel.map_apply' κ h_prod_embed.measurable - _ (h_prod_embed.measurableSet_image.mpr ht)] - congr with x : 1 - rw [← @Prod.mk.eta _ _ x] - simp only [Prod.mk.eta, Prod_map, id_eq, mem_preimage, mem_image, Prod.mk.injEq, Prod.exists, - exists_eq_right_right] - refine ⟨fun h ↦ ⟨x.1, h, rfl⟩, ?_⟩ - rintro ⟨ω, h_mem, h⟩ - rwa [hf.injective h] at h_mem - have h_fst : kernel.fst κ' = kernel.map (kernel.fst κ) f hf.measurable := by - ext a s hs - unfold_let κ' - rw [kernel.map_apply' _ _ _ hs, kernel.fst_apply', kernel.fst_apply', kernel.map_apply'] - · congr - · exact measurable_fst hs - · exact hf.measurable hs - · exact hs - conv_rhs => rw [this, ← h_compProd] - ext a s hs - rw [h_condKernel, h_fst] - rw [kernel.comapRight_apply' _ _ _ hs, kernel.compProd_apply _ _ _ hs, kernel.compProd_apply] - swap; · exact h_prod_embed.measurableSet_image.mpr hs - rw [kernel.map_apply, lintegral_map] - congr with ω - rw [kernel.comap_apply'] - · congr with ω' - simp only [mem_setOf_eq, Prod_map, id_eq, mem_image, Prod.mk.injEq, Prod.exists, - exists_eq_right_right] - refine ⟨fun h ↦ ⟨ω, h, rfl⟩, ?_⟩ - rintro ⟨a, h_mem, h⟩ - rwa [hf.injective h] at h_mem - · exact kernel.measurable_kernel_prod_mk_left' (h_prod_embed.measurableSet_image.mpr hs) _ - · exact hf.measurable - end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean b/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean index 1eaad20e3420a..9120601ff26a8 100644 --- a/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean +++ b/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean @@ -3,10 +3,12 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.MeasureTheory.Measure.Stieltjes -import Mathlib.MeasureTheory.Decomposition.RadonNikodym -import Mathlib.MeasureTheory.Constructions.Prod.Basic +import Mathlib.Data.Complex.Abs import Mathlib.MeasureTheory.Constructions.Polish +import Mathlib.MeasureTheory.Measure.GiryMonad +import Mathlib.MeasureTheory.Measure.Stieltjes +import Mathlib.Analysis.Normed.Order.Lattice +import Mathlib.MeasureTheory.Function.StronglyMeasurable.Basic /-! @@ -146,27 +148,24 @@ lemma inf_gt_rat_defaultRatCDF (t : ℚ) : ⨅ r : Ioi t, defaultRatCDF r = defaultRatCDF t := by simp only [defaultRatCDF] have h_bdd : BddBelow (range fun r : ↥(Ioi t) ↦ ite ((r : ℚ) < 0) (0 : ℝ) 1) := by - refine' ⟨0, fun x hx ↦ _⟩ + refine ⟨0, fun x hx ↦ ?_⟩ obtain ⟨y, rfl⟩ := mem_range.mpr hx dsimp only split_ifs exacts [le_rfl, zero_le_one] split_ifs with h - · refine' le_antisymm _ (le_ciInf fun x ↦ _) - · obtain ⟨q, htq, hq_neg⟩ : ∃ q, t < q ∧ q < 0 := by - refine' ⟨t / 2, _, _⟩ - · linarith - · linarith - refine' (ciInf_le h_bdd ⟨q, htq⟩).trans _ + · refine le_antisymm ?_ (le_ciInf fun x ↦ ?_) + · obtain ⟨q, htq, hq_neg⟩ : ∃ q, t < q ∧ q < 0 := ⟨t / 2, by linarith, by linarith⟩ + refine (ciInf_le h_bdd ⟨q, htq⟩).trans ?_ rw [if_pos] rwa [Subtype.coe_mk] · split_ifs exacts [le_rfl, zero_le_one] - · refine' le_antisymm _ _ - · refine' (ciInf_le h_bdd ⟨t + 1, lt_add_one t⟩).trans _ + · refine le_antisymm ?_ ?_ + · refine (ciInf_le h_bdd ⟨t + 1, lt_add_one t⟩).trans ?_ split_ifs exacts [zero_le_one, le_rfl] - · refine' le_ciInf fun x ↦ _ + · refine le_ciInf fun x ↦ ?_ rw [if_neg] rw [not_lt] at h ⊢ exact h.trans (mem_Ioi.mp x.prop).le @@ -201,12 +200,9 @@ lemma toCDFLike_of_isRatStieltjesPoint {a : α} (h : IsRatStieltjesPoint f a) (q lemma isCDFLike_toCDFLike (hf : ∀ q, Measurable fun a ↦ f a q) : IsCDFLike (toCDFLike f) where - mono a := by - unfold toCDFLike; split_ifs with h; exacts [h.mono, monotone_defaultRatCDF] - nonneg a := by - unfold toCDFLike; split_ifs with h; exacts [h.nonneg, defaultRatCDF_nonneg] - le_one a := by - unfold toCDFLike; split_ifs with h; exacts [h.le_one, defaultRatCDF_le_one] + mono a := by unfold toCDFLike; split_ifs with h; exacts [h.mono, monotone_defaultRatCDF] + nonneg a := by unfold toCDFLike; split_ifs with h; exacts [h.nonneg, defaultRatCDF_nonneg] + le_one a := by unfold toCDFLike; split_ifs with h; exacts [h.le_one, defaultRatCDF_le_one] tendsto_atTop_one a := by unfold toCDFLike; split_ifs with h; exacts [h.tendsto_atTop_one, tendsto_defaultRatCDF_atTop] tendsto_atBot_zero a := by @@ -238,7 +234,7 @@ lemma todo1_def' (f : α → ℚ → ℝ) (a : α) : lemma todo1_eq (a : α) (r : ℚ) : todo1 f a r = f a r := by rw [← hf.iInf_rat_gt_eq a r, todo1] - refine' Equiv.iInf_congr _ _ + refine Equiv.iInf_congr ?_ ?_ · exact { toFun := fun t ↦ ⟨t.1, mod_cast t.2⟩ invFun := fun t ↦ ⟨t.1, mod_cast t.2⟩ @@ -250,36 +246,36 @@ lemma todo1_eq (a : α) (r : ℚ) : lemma todo1_unit_prod (a : α) : todo1 (fun (p : Unit × α) ↦ f p.2) ((), a) = todo1 f a := by simp_rw [todo1_def'] -theorem todo1_nonneg (a : α) (r : ℝ) : 0 ≤ todo1 f a r := by +lemma todo1_nonneg (a : α) (r : ℝ) : 0 ≤ todo1 f a r := by have : Nonempty { r' : ℚ // r < ↑r' } := by obtain ⟨r, hrx⟩ := exists_rat_gt r exact ⟨⟨r, hrx⟩⟩ rw [todo1_def] exact le_ciInf fun r' ↦ hf.nonneg a _ -theorem bddBelow_range_gt (a : α) (x : ℝ) : +lemma bddBelow_range_gt (a : α) (x : ℝ) : BddBelow (range fun r : { r' : ℚ // x < ↑r' } ↦ f a r) := by - refine' ⟨0, fun z ↦ _⟩; rintro ⟨u, rfl⟩; exact hf.nonneg a _ + refine ⟨0, fun z ↦ ?_⟩; rintro ⟨u, rfl⟩; exact hf.nonneg a _ -theorem monotone_todo1 (a : α) : Monotone (todo1 f a) := by +lemma monotone_todo1 (a : α) : Monotone (todo1 f a) := by intro x y hxy have : Nonempty { r' : ℚ // y < ↑r' } := by obtain ⟨r, hrx⟩ := exists_rat_gt y exact ⟨⟨r, hrx⟩⟩ simp_rw [todo1_def] - refine' le_ciInf fun r ↦ (ciInf_le _ _).trans_eq _ + refine le_ciInf fun r ↦ (ciInf_le ?_ ?_).trans_eq ?_ · exact bddBelow_range_gt hf a x · exact ⟨r.1, hxy.trans_lt r.prop⟩ · rfl -theorem continuousWithinAt_todo1_Ici (a : α) (x : ℝ) : +lemma continuousWithinAt_todo1_Ici (a : α) (x : ℝ) : ContinuousWithinAt (todo1 f a) (Ici x) x := by rw [← continuousWithinAt_Ioi_iff_Ici] convert Monotone.tendsto_nhdsWithin_Ioi (monotone_todo1 hf a) x rw [sInf_image'] have h' : ⨅ r : Ioi x, todo1 f a r = ⨅ r : { r' : ℚ // x < r' }, todo1 f a r := by - refine' Real.iInf_Ioi_eq_iInf_rat_gt x _ (monotone_todo1 hf a) - refine' ⟨0, fun z ↦ _⟩ + refine Real.iInf_Ioi_eq_iInf_rat_gt x ?_ (monotone_todo1 hf a) + refine ⟨0, fun z ↦ ?_⟩ rintro ⟨u, -, rfl⟩ exact todo1_nonneg hf a u have h'' : @@ -300,13 +296,15 @@ noncomputable def IsCDFLike.stieltjesFunction (a : α) : StieltjesFunction where mono' := monotone_todo1 hf a right_continuous' x := continuousWithinAt_todo1_Ici hf a x -theorem IsCDFLike.stieltjesFunction_eq (a : α) (r : ℚ) : hf.stieltjesFunction a r = f a r := todo1_eq hf a r +lemma IsCDFLike.stieltjesFunction_eq (a : α) (r : ℚ) : hf.stieltjesFunction a r = f a r := + todo1_eq hf a r /-- The conditional cdf is non-negative for all `a : α`. -/ -theorem IsCDFLike.stieltjesFunction_nonneg (a : α) (r : ℝ) : 0 ≤ hf.stieltjesFunction a r := todo1_nonneg hf a r +lemma IsCDFLike.stieltjesFunction_nonneg (a : α) (r : ℝ) : 0 ≤ hf.stieltjesFunction a r := + todo1_nonneg hf a r /-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ -theorem IsCDFLike.stieltjesFunction_le_one (a : α) (x : ℝ) : hf.stieltjesFunction a x ≤ 1 := by +lemma IsCDFLike.stieltjesFunction_le_one (a : α) (x : ℝ) : hf.stieltjesFunction a x ≤ 1 := by obtain ⟨r, hrx⟩ := exists_rat_gt x rw [← StieltjesFunction.iInf_rat_gt_eq] simp_rw [IsCDFLike.stieltjesFunction_eq] @@ -314,7 +312,7 @@ theorem IsCDFLike.stieltjesFunction_le_one (a : α) (x : ℝ) : hf.stieltjesFunc exact ⟨r, hrx⟩ /-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ -theorem IsCDFLike.tendsto_stieltjesFunction_atBot (a : α) : +lemma IsCDFLike.tendsto_stieltjesFunction_atBot (a : α) : Tendsto (hf.stieltjesFunction a) atBot (𝓝 0) := by have h_exists : ∀ x : ℝ, ∃ q : ℚ, x < q ∧ ↑q < x + 1 := fun x ↦ exists_rat_btwn (lt_add_one x) let qs : ℝ → ℚ := fun x ↦ (h_exists x).choose @@ -331,7 +329,7 @@ theorem IsCDFLike.tendsto_stieltjesFunction_atBot (a : α) : exact (hf.stieltjesFunction a).mono (h_exists x).choose_spec.1.le /-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ -theorem IsCDFLike.tendsto_stieltjesFunction_atTop (a : α) : +lemma IsCDFLike.tendsto_stieltjesFunction_atTop (a : α) : Tendsto (hf.stieltjesFunction a) atTop (𝓝 1) := by have h_exists : ∀ x : ℝ, ∃ q : ℚ, x - 1 < q ∧ ↑q < x := fun x ↦ exists_rat_btwn (sub_one_lt x) let qs : ℝ → ℚ := fun x ↦ (h_exists x).choose @@ -348,7 +346,8 @@ theorem IsCDFLike.tendsto_stieltjesFunction_atTop (a : α) : exact (hf.stieltjesFunction a).mono (le_of_lt (h_exists x).choose_spec.2) /-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ -theorem IsCDFLike.measurable_stieltjesFunction (x : ℝ) : Measurable fun a ↦ hf.stieltjesFunction a x := by +lemma IsCDFLike.measurable_stieltjesFunction (x : ℝ) : + Measurable fun a ↦ hf.stieltjesFunction a x := by have : (fun a ↦ hf.stieltjesFunction a x) = fun a ↦ ⨅ r : { r' : ℚ // x < r' }, f a ↑r := by ext1 a rw [← StieltjesFunction.iInf_rat_gt_eq] @@ -358,34 +357,34 @@ theorem IsCDFLike.measurable_stieltjesFunction (x : ℝ) : Measurable fun a ↦ exact measurable_iInf (fun q ↦ hf.measurable q) /-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ -theorem IsCDFLike.stronglyMeasurable_stieltjesFunction (x : ℝ) : +lemma IsCDFLike.stronglyMeasurable_stieltjesFunction (x : ℝ) : StronglyMeasurable fun a ↦ hf.stieltjesFunction a x := (measurable_stieltjesFunction hf x).stronglyMeasurable section Measure -theorem IsCDFLike.measure_stieltjesFunction_Iic (a : α) (x : ℝ) : +lemma IsCDFLike.measure_stieltjesFunction_Iic (a : α) (x : ℝ) : (hf.stieltjesFunction a).measure (Iic x) = ENNReal.ofReal (hf.stieltjesFunction a x) := by rw [← sub_zero (hf.stieltjesFunction a x)] exact (hf.stieltjesFunction a).measure_Iic (tendsto_stieltjesFunction_atBot hf a) _ -theorem IsCDFLike.measure_stieltjesFunction_univ (a : α) : (hf.stieltjesFunction a).measure univ = 1 := by +lemma IsCDFLike.measure_stieltjesFunction_univ (a : α) : + (hf.stieltjesFunction a).measure univ = 1 := by rw [← ENNReal.ofReal_one, ← sub_zero (1 : ℝ)] - exact StieltjesFunction.measure_univ _ (tendsto_stieltjesFunction_atBot hf a) (tendsto_stieltjesFunction_atTop hf a) + exact StieltjesFunction.measure_univ _ (tendsto_stieltjesFunction_atBot hf a) + (tendsto_stieltjesFunction_atTop hf a) instance IsCDFLike.instIsProbabilityMeasure_stieltjesFunction (a : α) : IsProbabilityMeasure (hf.stieltjesFunction a).measure := ⟨measure_stieltjesFunction_univ hf a⟩ /-- The function `a ↦ (condCDF ρ a).measure` is measurable. -/ -theorem IsCDFLike.measurable_measure_stieltjesFunction : +lemma IsCDFLike.measurable_measure_stieltjesFunction : Measurable fun a ↦ (hf.stieltjesFunction a).measure := by rw [Measure.measurable_measure] - refine' fun s hs ↦ ?_ - -- Porting note: supplied `C` - refine' MeasurableSpace.induction_on_inter + refine fun s hs ↦ MeasurableSpace.induction_on_inter (C := fun s ↦ Measurable fun b ↦ StieltjesFunction.measure (hf.stieltjesFunction b) s) - (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic _ _ _ _ hs + (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic ?_ ?_ ?_ ?_ hs · simp only [measure_empty, measurable_const] · rintro S ⟨u, rfl⟩ simp_rw [measure_stieltjesFunction_Iic hf _ u] @@ -414,7 +413,7 @@ noncomputable def todo3 (f : α → ℚ → ℝ) (hf : ∀ q, Measurable fun a ↦ f a q) : α → StieltjesFunction := (isCDFLike_toCDFLike hf).stieltjesFunction -theorem todo3_eq (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℚ) : +lemma todo3_eq (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℚ) : todo3 f hf a r = toCDFLike f a r := IsCDFLike.stieltjesFunction_eq _ a r lemma todo3_unit_prod (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : @@ -427,43 +426,43 @@ lemma todo3_unit_prod (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : | mk _ b => rw [← toCDFLike_unit_prod b] /-- The conditional cdf is non-negative for all `a : α`. -/ -theorem todo3_nonneg (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℝ) : +lemma todo3_nonneg (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℝ) : 0 ≤ todo3 f hf a r := IsCDFLike.stieltjesFunction_nonneg _ a r /-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ -theorem todo3_le_one (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : +lemma todo3_le_one (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : todo3 f hf a x ≤ 1 := IsCDFLike.stieltjesFunction_le_one _ a x /-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ -theorem tendsto_todo3_atBot (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : +lemma tendsto_todo3_atBot (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : Tendsto (todo3 f hf a) atBot (𝓝 0) := IsCDFLike.tendsto_stieltjesFunction_atBot _ a /-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ -theorem tendsto_todo3_atTop (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : +lemma tendsto_todo3_atTop (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : Tendsto (todo3 f hf a) atTop (𝓝 1) := IsCDFLike.tendsto_stieltjesFunction_atTop _ a /-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ -theorem measurable_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : +lemma measurable_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : Measurable fun a ↦ todo3 f hf a x := IsCDFLike.measurable_stieltjesFunction _ x /-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ -theorem stronglyMeasurable_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : +lemma stronglyMeasurable_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : StronglyMeasurable fun a ↦ todo3 f hf a x := IsCDFLike.stronglyMeasurable_stieltjesFunction _ x section Measure -theorem measure_todo3_Iic (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : +lemma measure_todo3_Iic (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : (todo3 f hf a).measure (Iic x) = ENNReal.ofReal (todo3 f hf a x) := IsCDFLike.measure_stieltjesFunction_Iic _ _ _ -theorem measure_todo3_univ (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : +lemma measure_todo3_univ (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : (todo3 f hf a).measure univ = 1 := IsCDFLike.measure_stieltjesFunction_univ _ _ instance instIsProbabilityMeasure_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : IsProbabilityMeasure (todo3 f hf a).measure := IsCDFLike.instIsProbabilityMeasure_stieltjesFunction _ _ -theorem measurable_measure_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) : +lemma measurable_measure_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) : Measurable fun a ↦ (todo3 f hf a).measure := IsCDFLike.measurable_measure_stieltjesFunction _ end Measure From 8012c6db140a51523d9365c926152593971f10f2 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 14 Feb 2024 15:00:17 +0100 Subject: [PATCH 012/129] split files --- .../Probability/Kernel/Disintegration.lean | 314 +----------------- .../Kernel/Disintegration/Basic.lean | 153 ++++++++- .../Kernel/Disintegration/Unique.lean | 146 ++++++++ 3 files changed, 308 insertions(+), 305 deletions(-) create mode 100644 Mathlib/Probability/Kernel/Disintegration/Unique.lean diff --git a/Mathlib/Probability/Kernel/Disintegration.lean b/Mathlib/Probability/Kernel/Disintegration.lean index aa30bf92fa818..1af8f1824d81b 100644 --- a/Mathlib/Probability/Kernel/Disintegration.lean +++ b/Mathlib/Probability/Kernel/Disintegration.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne, Kexing Ying -/ -import Mathlib.Probability.Kernel.CondCdf +import Mathlib.Probability.Kernel.Disintegration.Basic #align_import probability.kernel.disintegration from "leanprover-community/mathlib"@"6315581f5650ffa2fbdbbbedc41243c8d7070981" @@ -51,85 +51,17 @@ namespace ProbabilityTheory variable {α : Type*} {mα : MeasurableSpace α} -section Real - -/-! ### Disintegration of measures on `α × ℝ` -/ - - -/-- Conditional measure on the second space of the product given the value on the first, as a -kernel. Use the more general `condKernel`. -/ -noncomputable def condKernelReal (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : kernel α ℝ := - kernel.comap (cdfKernel (fun (p : Unit × α) ↦ condCDF ρ p.2) (isKernelCDF_condCDF ρ)) - (fun a ↦ ((), a)) measurable_prod_mk_left -#align probability_theory.cond_kernel_real ProbabilityTheory.condKernelReal - -instance (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : IsMarkovKernel (condKernelReal ρ) := by - rw [condKernelReal]; infer_instance - -theorem condKernelReal_Iic (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (a : α) (x : ℝ) : - condKernelReal ρ a (Iic x) = ENNReal.ofReal (condCDF ρ a x) := by - rw [condKernelReal, kernel.comap_apply, cdfKernel_Iic, condCDF_eq_todo3_unit_prod] -#align probability_theory.cond_kernel_real_Iic ProbabilityTheory.condKernelReal_Iic - -theorem set_lintegral_condKernelReal_Iic (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) - {s : Set α} (hs : MeasurableSet s) : - ∫⁻ a in s, condKernelReal ρ a (Iic x) ∂ρ.fst = ρ (s ×ˢ Iic x) := - set_lintegral_cdfKernel_Iic (isKernelCDF_condCDF ρ) () x hs -#align probability_theory.set_lintegral_cond_kernel_real_Iic ProbabilityTheory.set_lintegral_condKernelReal_Iic - -theorem set_lintegral_condKernelReal_univ (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] - {s : Set α} (hs : MeasurableSet s) : - ∫⁻ a in s, condKernelReal ρ a univ ∂ρ.fst = ρ (s ×ˢ univ) := by - simp only [measure_univ, lintegral_const, Measure.restrict_apply, MeasurableSet.univ, univ_inter, - one_mul, Measure.fst_apply hs, ← prod_univ] -#align probability_theory.set_lintegral_cond_kernel_real_univ ProbabilityTheory.set_lintegral_condKernelReal_univ - -theorem lintegral_condKernelReal_univ (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - ∫⁻ a, condKernelReal ρ a univ ∂ρ.fst = ρ univ := by - rw [← set_lintegral_univ, set_lintegral_condKernelReal_univ ρ MeasurableSet.univ, - univ_prod_univ] -#align probability_theory.lintegral_cond_kernel_real_univ ProbabilityTheory.lintegral_condKernelReal_univ - -variable (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] - -theorem set_lintegral_condKernelReal_prod {s : Set α} (hs : MeasurableSet s) {t : Set ℝ} - (ht : MeasurableSet t) : - ∫⁻ a in s, condKernelReal ρ a t ∂ρ.fst = ρ (s ×ˢ t) := - set_lintegral_cdfKernel_prod (isKernelCDF_condCDF ρ) () hs ht -#align probability_theory.set_lintegral_cond_kernel_real_prod ProbabilityTheory.set_lintegral_condKernelReal_prod - -theorem lintegral_condKernelReal_mem {s : Set (α × ℝ)} (hs : MeasurableSet s) : - ∫⁻ a, condKernelReal ρ a {x | (a, x) ∈ s} ∂ρ.fst = ρ s := - lintegral_cdfKernel_mem (isKernelCDF_condCDF ρ) () hs -#align probability_theory.lintegral_cond_kernel_real_mem ProbabilityTheory.lintegral_condKernelReal_mem - -theorem kernel.const_eq_compProd_real (γ : Type*) [MeasurableSpace γ] (ρ : Measure (α × ℝ)) - [IsFiniteMeasure ρ] : - kernel.const γ ρ = kernel.const γ ρ.fst ⊗ₖ kernel.prodMkLeft γ (condKernelReal ρ) := by - ext a s hs : 2 - rw [kernel.compProd_apply _ _ _ hs, kernel.const_apply, kernel.const_apply] - simp_rw [kernel.prodMkLeft_apply] - rw [lintegral_condKernelReal_mem ρ hs] -#align probability_theory.kernel.const_eq_comp_prod_real ProbabilityTheory.kernel.const_eq_compProd_real - -theorem measure_eq_compProd_real : ρ = ρ.fst ⊗ₘ condKernelReal ρ := by - rw [Measure.compProd, ← kernel.const_eq_compProd_real Unit ρ, kernel.const_apply] -#align probability_theory.measure_eq_comp_prod_real ProbabilityTheory.measure_eq_compProd_real - -theorem lintegral_condKernelReal {f : α × ℝ → ℝ≥0∞} (hf : Measurable f) : - ∫⁻ a, ∫⁻ y, f (a, y) ∂condKernelReal ρ a ∂ρ.fst = ∫⁻ x, f x ∂ρ := by - nth_rw 3 [measure_eq_compProd_real ρ] - rw [Measure.lintegral_compProd hf] -#align probability_theory.lintegral_cond_kernel_real ProbabilityTheory.lintegral_condKernelReal - -theorem ae_condKernelReal_eq_one {s : Set ℝ} (hs : MeasurableSet s) (hρ : ρ {x | x.snd ∈ sᶜ} = 0) : - ∀ᵐ a ∂ρ.fst, condKernelReal ρ a s = 1 := by - have h := ae_cdfKernel_eq_one (isKernelCDF_condCDF ρ) () hs - simp only [kernel.const_apply] at h - exact h hρ -#align probability_theory.ae_cond_kernel_real_eq_one ProbabilityTheory.ae_condKernelReal_eq_one - -end Real +#noalign probability_theory.cond_kernel_real +#noalign probability_theory.cond_kernel_real_Iic +#noalign probability_theory.set_lintegral_cond_kernel_real_Iic +#noalign probability_theory.set_lintegral_cond_kernel_real_univ +#noalign probability_theory.lintegral_cond_kernel_real_univ +#noalign probability_theory.set_lintegral_cond_kernel_real_prod +#noalign probability_theory.lintegral_cond_kernel_real_mem +#noalign probability_theory.kernel.const_eq_comp_prod_real +#noalign probability_theory.measure_eq_comp_prod_real +#noalign probability_theory.lintegral_cond_kernel_real +#noalign probability_theory.ae_cond_kernel_real_eq_one section StandardBorel @@ -142,125 +74,19 @@ property on `ℝ` to all these spaces. -/ variable {Ω : Type*} [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] -/-- Existence of a conditional kernel. Use the definition `condKernel` to get that kernel. -/ -theorem exists_cond_kernel (γ : Type*) [MeasurableSpace γ] : - ∃ (η : kernel α Ω) (_h : IsMarkovKernel η), kernel.const γ ρ = - (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ η) := by - obtain ⟨f, hf⟩ := exists_measurableEmbedding_real Ω - let ρ' : Measure (α × ℝ) := ρ.map (Prod.map id f) - -- The general idea is to define `η = kernel.comapRight (condKernelReal ρ') hf`. There is - -- however an issue: that `η` may not be a Markov kernel since its value is only a - -- probability distribution almost everywhere with respect to `ρ.fst`, not everywhere. - -- We modify it to obtain a Markov kernel which is almost everywhere equal. - let ρ_set := (toMeasurable ρ.fst {a | condKernelReal ρ' a (range f) = 1}ᶜ)ᶜ - have hm : MeasurableSet ρ_set := (measurableSet_toMeasurable _ _).compl - have h_eq_one_of_mem : ∀ a ∈ ρ_set, condKernelReal ρ' a (range f) = 1 := by - intro a ha - rw [mem_compl_iff] at ha - have h_ss := subset_toMeasurable ρ.fst {a : α | condKernelReal ρ' a (range f) = 1}ᶜ - suffices ha' : a ∉ {a : α | condKernelReal ρ' a (range f) = 1}ᶜ - · rwa [not_mem_compl_iff] at ha' - exact not_mem_subset h_ss ha - have h_prod_embed : MeasurableEmbedding (Prod.map (id : α → α) f) := - MeasurableEmbedding.id.prod_mk hf - have h_fst : ρ'.fst = ρ.fst := by - ext1 u hu - rw [Measure.fst_apply hu, Measure.fst_apply hu, - Measure.map_apply h_prod_embed.measurable (measurable_fst hu)] - rfl - have h_ae : ∀ᵐ a ∂ρ.fst, a ∈ ρ_set := by - rw [ae_iff] - simp only [not_mem_compl_iff, setOf_mem_eq, measure_toMeasurable] - change ρ.fst {a : α | a ∉ {a' : α | condKernelReal ρ' a' (range f) = 1}} = 0 - rw [← ae_iff, ← h_fst] - refine' ae_condKernelReal_eq_one ρ' hf.measurableSet_range _ - rw [Measure.map_apply h_prod_embed.measurable] - swap; · exact measurable_snd hf.measurableSet_range.compl - convert measure_empty (α := α × Ω) - ext1 x - simp only [mem_compl_iff, mem_range, preimage_setOf_eq, Prod_map, mem_setOf_eq, - mem_empty_iff_false, iff_false_iff, Classical.not_not, exists_apply_eq_apply] - classical - obtain ⟨x₀, hx₀⟩ : ∃ x, x ∈ range f := range_nonempty _ - let η' := - kernel.piecewise hm (condKernelReal ρ') (kernel.deterministic (fun _ => x₀) measurable_const) - -- We show that `kernel.comapRight η' hf` is a suitable Markov kernel. - refine' ⟨kernel.comapRight η' hf, _, _⟩ - · refine' kernel.IsMarkovKernel.comapRight _ _ fun a => _ - rw [kernel.piecewise_apply'] - split_ifs with h_mem - · exact h_eq_one_of_mem _ h_mem - · rw [kernel.deterministic_apply' _ _ hf.measurableSet_range, Set.indicator_apply, if_pos hx₀] - have : kernel.const γ ρ = kernel.comapRight (kernel.const γ ρ') h_prod_embed := by - ext c t ht : 2 - rw [kernel.const_apply, kernel.comapRight_apply' _ _ _ ht, kernel.const_apply, - Measure.map_apply h_prod_embed.measurable (h_prod_embed.measurableSet_image.mpr ht)] - congr with x : 1 - rw [← @Prod.mk.eta _ _ x] - simp only [id.def, mem_preimage, Prod.map_mk, mem_image, Prod.mk.inj_iff, Prod.exists] - refine' ⟨fun h => ⟨x.1, x.2, h, rfl, rfl⟩, _⟩ - rintro ⟨a, b, h_mem, rfl, hf_eq⟩ - rwa [hf.injective hf_eq] at h_mem - rw [this, kernel.const_eq_compProd_real _ ρ'] - ext c t ht : 2 - rw [kernel.comapRight_apply' _ _ _ ht, - kernel.compProd_apply _ _ _ (h_prod_embed.measurableSet_image.mpr ht), kernel.const_apply, - h_fst, kernel.compProd_apply _ _ _ ht, kernel.const_apply] - refine' lintegral_congr_ae _ - filter_upwards [h_ae] with a ha - rw [kernel.prodMkLeft_apply', kernel.prodMkLeft_apply', kernel.comapRight_apply'] - swap - · exact measurable_prod_mk_left ht - have h1 : {c : ℝ | (a, c) ∈ Prod.map id f '' t} = f '' {c : Ω | (a, c) ∈ t} := by - ext1 x - simp only [Prod_map, id.def, mem_image, Prod.mk.inj_iff, Prod.exists, mem_setOf_eq] - constructor - · rintro ⟨a', b, h_mem, rfl, hf_eq⟩ - exact ⟨b, h_mem, hf_eq⟩ - · rintro ⟨b, h_mem, hf_eq⟩ - exact ⟨a, b, h_mem, rfl, hf_eq⟩ - have h2 : condKernelReal ρ' (c, a).snd = η' (c, a).snd := by - rw [kernel.piecewise_apply, if_pos ha] - rw [h1, h2] -#align probability_theory.exists_cond_kernel ProbabilityTheory.exists_cond_kernel - -/-- Conditional kernel of a measure on a product space: a Markov kernel such that -`ρ = ρ.fst ⊗ₘ ρ.condKernel` (see `ProbabilityTheory.measure_eq_compProd`). -/ -noncomputable irreducible_def _root_.MeasureTheory.Measure.condKernel : kernel α Ω := - (exists_cond_kernel ρ Unit).choose -#align measure_theory.measure.cond_kernel MeasureTheory.Measure.condKernel - -theorem condKernel_def : ρ.condKernel = (exists_cond_kernel ρ Unit).choose := by - rw [MeasureTheory.Measure.condKernel] -#align probability_theory.cond_kernel_def ProbabilityTheory.condKernel_def - -instance : IsMarkovKernel ρ.condKernel := by - rw [condKernel_def]; exact (exists_cond_kernel ρ Unit).choose_spec.choose - -theorem kernel.const_unit_eq_compProd : - kernel.const Unit ρ = kernel.const Unit ρ.fst ⊗ₖ kernel.prodMkLeft Unit ρ.condKernel := by - simp_rw [condKernel_def]; exact (exists_cond_kernel ρ Unit).choose_spec.choose_spec -#align probability_theory.kernel.const_unit_eq_comp_prod ProbabilityTheory.kernel.const_unit_eq_compProd +#noalign probability_theory.exists_cond_kernel +#noalign probability_theory.cond_kernel_def +#noalign probability_theory.kernel.const_unit_eq_comp_prod /-- **Disintegration** of finite product measures on `α × Ω`, where `Ω` is standard Borel. Such a measure can be written as the composition-product of the constant kernel with value `ρ.fst` (marginal measure over `α`) and a Markov kernel from `α` to `Ω`. We call that Markov kernel `ProbabilityTheory.condKernel ρ`. -/ theorem measure_eq_compProd : ρ = ρ.fst ⊗ₘ ρ.condKernel := by - rw [Measure.compProd, ← kernel.const_unit_eq_compProd, kernel.const_apply] + rw [Measure.compProd_fst_condKernel] #align probability_theory.measure_eq_comp_prod ProbabilityTheory.measure_eq_compProd -/-- **Disintegration** of constant kernels. A constant kernel on a product space `α × Ω`, where `Ω` -is standard Borel, can be written as the composition-product of the constant kernel with -value `ρ.fst` (marginal measure over `α`) and a Markov kernel from `α` to `Ω`. We call that -Markov kernel `ProbabilityTheory.condKernel ρ`. -/ -theorem kernel.const_eq_compProd (γ : Type*) [MeasurableSpace γ] (ρ : Measure (α × Ω)) - [IsFiniteMeasure ρ] : - kernel.const γ ρ = kernel.const γ ρ.fst ⊗ₖ kernel.prodMkLeft γ ρ.condKernel := by - ext a s hs : 2 - simpa only [kernel.const_apply, kernel.compProd_apply _ _ _ hs, kernel.prodMkLeft_apply'] using - kernel.ext_iff'.mp (kernel.const_unit_eq_compProd ρ) () s hs -#align probability_theory.kernel.const_eq_comp_prod ProbabilityTheory.kernel.const_eq_compProd +#noalign probability_theory.kernel.const_eq_comp_prod /-- Auxiliary lemma for `condKernel_apply_of_ne_zero`. -/ lemma condKernel_apply_of_ne_zero_of_measurableSet [MeasurableSingletonClass α] @@ -387,112 +213,6 @@ theorem set_integral_condKernel_univ_left {ρ : Measure (α × Ω)} [IsFiniteMea end IntegralCondKernel -section Unique - -/-! ### Uniqueness - -This section will prove that the conditional kernel is unique almost everywhere. -/ - -/-- A s-finite kernel which satisfy the disintegration property of the given measure `ρ` is almost -everywhere equal to the disintegration kernel of `ρ` when evaluated on a measurable set. - -This theorem in the case of finite kernels is weaker than `eq_condKernel_of_measure_eq_compProd` -which asserts that the kernels are equal almost everywhere and not just on a given measurable -set. -/ -theorem eq_condKernel_of_measure_eq_compProd' (κ : kernel α Ω) [IsSFiniteKernel κ] - (hκ : ρ = ρ.fst ⊗ₘ κ) {s : Set Ω} (hs : MeasurableSet s) : - ∀ᵐ x ∂ρ.fst, κ x s = ρ.condKernel x s := by - refine' ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite - (kernel.measurable_coe κ hs) (kernel.measurable_coe ρ.condKernel hs) _ - intros t ht _ - conv_rhs => rw [set_lintegral_condKernel_eq_measure_prod _ ht hs, hκ] - simp only [Measure.compProd_apply (ht.prod hs), Set.mem_prod, ← lintegral_indicator _ ht] - congr; ext x - by_cases hx : x ∈ t - all_goals simp [hx] - --- This lemma establishes uniqueness of the disintegration kernel on ℝ -lemma eq_condKernel_of_measure_eq_compProd_real (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] - (κ : kernel α ℝ) [IsFiniteKernel κ] (hκ : ρ = ρ.fst ⊗ₘ κ) : - ∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by - have huniv : ∀ᵐ x ∂ρ.fst, κ x Set.univ = ρ.condKernel x Set.univ := - eq_condKernel_of_measure_eq_compProd' ρ κ hκ MeasurableSet.univ - suffices : ∀ᵐ x ∂ρ.fst, ∀ ⦃t⦄, MeasurableSet t → κ x t = ρ.condKernel x t - · filter_upwards [this] with x hx - ext t ht; exact hx ht - apply MeasurableSpace.ae_induction_on_inter Real.borel_eq_generateFrom_Iic_rat - Real.isPiSystem_Iic_rat - · simp only [OuterMeasure.empty', Filter.eventually_true] - · simp only [iUnion_singleton_eq_range, mem_range, forall_exists_index, forall_apply_eq_imp_iff] - exact ae_all_iff.2 fun q => eq_condKernel_of_measure_eq_compProd' ρ κ hκ measurableSet_Iic - · filter_upwards [huniv] with x hxuniv t ht heq - rw [measure_compl ht <| measure_ne_top _ _, heq, hxuniv, measure_compl ht <| measure_ne_top _ _] - · refine' ae_of_all _ (fun x f hdisj hf heq => _) - rw [measure_iUnion hdisj hf, measure_iUnion hdisj hf] - exact tsum_congr heq - -/-- A finite kernel which satisfies the disintegration property is almost everywhere equal to the -disintegration kernel. -/ -theorem eq_condKernel_of_measure_eq_compProd (κ : kernel α Ω) [IsFiniteKernel κ] - (hκ : ρ = ρ.fst ⊗ₘ κ) : - ∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by --- The idea is to transporting the question to `ℝ` from `Ω` using `exists_measurableEmbedding_real` --- and then constructing a measure on `α × ℝ` using the obtained measurable embedding - obtain ⟨f, hf⟩ := exists_measurableEmbedding_real Ω - set ρ' : Measure (α × ℝ) := ρ.map (Prod.map id f) with hρ'def - have hρ' : ρ'.fst = ρ.fst - · ext s hs - rw [hρ'def, Measure.fst_apply, Measure.fst_apply, Measure.map_apply] - exacts [rfl, Measurable.prod measurable_fst <| hf.measurable.comp measurable_snd, - measurable_fst hs, hs, hs] - have hρ'' : ∀ᵐ x ∂ρ'.fst, kernel.map κ f hf.measurable x = ρ'.condKernel x - · refine' eq_condKernel_of_measure_eq_compProd_real ρ' (kernel.map κ f hf.measurable) _ - ext s hs - simp only [Measure.map_apply (measurable_id.prod_map hf.measurable) hs] - conv_lhs => congr; rw [hκ] - rw [Measure.compProd_apply hs, Measure.compProd_apply - (measurable_id.prod_map hf.measurable hs), (_ : (ρ.map (Prod.map id f)).fst = ρ.fst)] - · congr - ext x - simp only [Set.mem_preimage, Prod_map, id_eq, kernel.prodMkLeft_apply, kernel.map_apply] - rw [Measure.map_apply hf.measurable] - · rfl - · exact measurable_prod_mk_left hs - · rw [Measure.fst_map_prod_mk] - · simp only [Prod_map, id_eq]; rfl - · exact (hf.measurable.comp measurable_snd) - rw [hρ'] at hρ'' - suffices : ∀ᵐ x ∂ρ.fst, ∀ s, MeasurableSet s → - ((ρ.map (Prod.map id f)).condKernel x) s = (ρ.condKernel x) (f ⁻¹' s) - · filter_upwards [hρ'', this] with x hx h - rw [kernel.map_apply] at hx - ext s hs - rw [← Set.preimage_image_eq s hf.injective, - ← Measure.map_apply hf.measurable <| hf.measurableSet_image.2 hs, hx, - h _ <| hf.measurableSet_image.2 hs] - have hprod : (ρ.map (Prod.map id f)).fst = ρ.fst - · ext s hs - rw [Measure.fst_apply hs, - Measure.map_apply (measurable_id.prod_map hf.measurable) (measurable_fst hs), - ← Set.preimage_comp, Measure.fst_apply hs] - rfl - suffices : ρ.map (Prod.map id f) = - ((ρ.map (Prod.map id f)).fst ⊗ₘ (kernel.map (Measure.condKernel ρ) f hf.measurable)) - · have heq := eq_condKernel_of_measure_eq_compProd_real _ _ this - rw [hprod] at heq - filter_upwards [heq] with x hx s hs - rw [← hx, kernel.map_apply, Measure.map_apply hf.measurable hs] - ext s hs - simp only [hprod, Measure.compProd_apply hs, kernel.map_apply _ hf.measurable] - rw [Measure.map_apply (measurable_id.prod_map hf.measurable) hs, ← lintegral_condKernel_mem] - · simp only [mem_preimage, Prod_map, id_eq] - congr with a - rw [Measure.map_apply hf.measurable (measurable_prod_mk_left hs)] - rfl - · exact measurable_id.prod_map hf.measurable hs - -end Unique - end StandardBorel end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 4be74195e3530..494526d8908b8 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Rémy Degenne, Kexing Ying +Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.CondCdf import Mathlib.Probability.Kernel.Disintegration.KernelCDFBorel @@ -24,8 +24,7 @@ variable {α β Ω Ω': Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace section BorelSnd noncomputable -def condKernelBorelSnd (κ : kernel α (β × Ω')) [IsMarkovKernel κ] - {f : α × β → StieltjesFunction} +def condKernelBorelSnd (κ : kernel α (β × Ω')) {f : α × β → StieltjesFunction} (hf : IsKernelCDF f (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω').measurable)) @@ -40,7 +39,7 @@ def condKernelBorelSnd (κ : kernel α (β × Ω')) [IsMarkovKernel κ] (cdfKernel f hf) (kernel.deterministic (fun _ ↦ x₀) measurable_const)) he -instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω')) [IsMarkovKernel κ] +instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω')) {f : α × β → StieltjesFunction} (hf : IsKernelCDF f (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) @@ -59,7 +58,7 @@ instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω')) [IsM Set.indicator_apply, if_pos] exact (range_nonempty (measurableEmbedding_real Ω')).choose_spec -lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω')) [IsMarkovKernel κ] +lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω')) [IsFiniteKernel κ] {f : α × β → StieltjesFunction} (hf : IsKernelCDF f (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) @@ -173,7 +172,7 @@ def condKernelBorel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : kernel ( instance instIsMarkovKernel_condKernelBorel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : IsMarkovKernel (condKernelBorel κ) := by rw [condKernelBorel]; infer_instance -lemma compProd_fst_condKernel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : +lemma compProd_fst_condKernelBorel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : kernel.fst κ ⊗ₖ condKernelBorel κ = κ := by let f := measurableEmbedding_real Ω let hf := measurableEmbedding_measurableEmbedding_real Ω @@ -224,14 +223,152 @@ lemma compProd_fst_condKernel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : end StandardBorel -section Countable - section Unit +section Real + +noncomputable def condKernelUnitReal (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : + kernel (Unit × α) ℝ := + cdfKernel (fun (p : Unit × α) ↦ condCDF (ρ ()) p.2) (isKernelCDF_condCDF (ρ ())) + +instance (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : IsMarkovKernel (condKernelUnitReal ρ) := by + rw [condKernelUnitReal]; infer_instance + +lemma fst_compProd_condKernelUnitReal (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : + kernel.fst ρ ⊗ₖ condKernelUnitReal ρ = ρ := by + have : ρ = kernel.const Unit (ρ ()) := by ext; simp + conv_rhs => rw [this, kernel.eq_compProd_cdfKernel (isKernelCDF_condCDF (ρ ()))] + +end Real + +section BorelSnd + +noncomputable +def condKernelUnitBorel (κ : kernel Unit (α × Ω')) [IsFiniteKernel κ] : kernel (Unit × α) Ω' := + let f := measurableEmbedding_real Ω' + let hf := measurableEmbedding_measurableEmbedding_real Ω' + let κ' := kernel.map κ (Prod.map (id : α → α) f) (measurable_id.prod_map hf.measurable) + condKernelBorelSnd κ (isKernelCDF_condCDF (κ' ())) + +instance instIsMarkovKernel_condKernelUnitBorel (κ : kernel Unit (α × Ω')) [IsFiniteKernel κ] : + IsMarkovKernel (condKernelUnitBorel κ) := by + rw [condKernelUnitBorel] + infer_instance + +lemma compProd_fst_condKernelUnitBorel (κ : kernel Unit (α × Ω')) [IsFiniteKernel κ] : + kernel.fst κ ⊗ₖ condKernelUnitBorel κ = κ := by + rw [condKernelUnitBorel, compProd_fst_condKernelBorelSnd] +end BorelSnd end Unit +section Measure + +/-- Conditional kernel of a measure on a product space: a Markov kernel such that +`ρ = ρ.fst ⊗ₘ ρ.condKernel` (see `ProbabilityTheory.measure_eq_compProd`). -/ +noncomputable +def _root_.MeasureTheory.Measure.condKernel (ρ : Measure (α × Ω')) [IsFiniteMeasure ρ] : + kernel α Ω' where + val a := condKernelUnitBorel (kernel.const Unit ρ) ((), a) + property := (condKernelUnitBorel (kernel.const Unit ρ)).property.comp measurable_prod_mk_left +#align measure_theory.measure.cond_kernel MeasureTheory.Measure.condKernel + +lemma _root_.MeasureTheory.Measure.condKernel_apply (ρ : Measure (α × Ω')) [IsFiniteMeasure ρ] + (a : α) : + ρ.condKernel a = condKernelUnitBorel (kernel.const Unit ρ) ((), a) := rfl + +instance _root_.MeasureTheory.Measure.instIsMarkovKernel_condKernel + (ρ : Measure (α × Ω')) [IsFiniteMeasure ρ] : + IsMarkovKernel ρ.condKernel := by + constructor + intro a + change IsProbabilityMeasure (condKernelUnitBorel (kernel.const Unit ρ) ((), a)) + infer_instance + +lemma _root_.MeasureTheory.Measure.compProd_fst_condKernel + (ρ : Measure (α × Ω')) [IsFiniteMeasure ρ] : + ρ.fst ⊗ₘ ρ.condKernel = ρ := by + have h1 : kernel.const Unit (Measure.fst ρ) = kernel.fst (kernel.const Unit ρ) := by + ext + simp only [kernel.fst_apply, Measure.fst, kernel.const_apply] + have h2 : kernel.prodMkLeft Unit (Measure.condKernel ρ) + = condKernelUnitBorel (kernel.const Unit ρ) := by + ext + simp only [kernel.prodMkLeft_apply, Measure.condKernel_apply] + rw [Measure.compProd, h1, h2, compProd_fst_condKernelUnitBorel] + simp + +end Measure + +section Countable + +variable [MeasurableSingletonClass α] [Countable α] + +noncomputable +def condKernelCountable (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : kernel (α × β) Ω' where + val p := (κ p.1).condKernel p.2 + property := by + change Measurable ((fun q : β × α ↦ (κ q.2).condKernel q.1) ∘ Prod.swap) + refine (measurable_from_prod_countable (fun a ↦ ?_)).comp measurable_swap + exact kernel.measurable (κ a).condKernel + +lemma condKernelCountable_apply (κ : kernel α (β × Ω')) [IsFiniteKernel κ] (p : α × β) : + condKernelCountable κ p = (κ p.1).condKernel p.2 := rfl + +instance instIsMarkovKernel_condKernelCountable (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : + IsMarkovKernel (condKernelCountable κ) := + ⟨fun p ↦ (Measure.instIsMarkovKernel_condKernel (κ p.1)).isProbabilityMeasure p.2⟩ + +lemma compProd_fst_condKernelCountable (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : + kernel.fst κ ⊗ₖ condKernelCountable κ = κ := by + ext a s hs + have h := (κ a).compProd_fst_condKernel + conv_rhs => rw [← h] + simp_rw [kernel.compProd_apply _ _ _ hs, condKernelCountable_apply, Measure.compProd_apply hs] + congr + end Countable +section CountableOrStandardBorel + +class CountableOrStandardBorel (α β : Type*) [MeasurableSpace α] [MeasurableSpace β] : Prop := + (countableOrStandardBorel : (Countable α ∧ MeasurableSingletonClass α) ∨ StandardBorelSpace β) + +instance instCountableOrStandardBorel_of_countable + [h1 : Countable α] [h2 : MeasurableSingletonClass α] : + CountableOrStandardBorel α β := ⟨Or.inl ⟨h1, h2⟩⟩ + +instance instCountableOrStandardBorel_of_standardBorelSpace [h : StandardBorelSpace β] : + CountableOrStandardBorel α β := ⟨Or.inr h⟩ + +open Classical in +noncomputable +def kernel.condKernel [h : CountableOrStandardBorel α β] + (κ : kernel α (β × Ω')) [IsMarkovKernel κ] : + kernel (α × β) Ω' := + if hα : Countable α ∧ MeasurableSingletonClass α then + letI := hα.1; letI := hα.2; condKernelCountable κ + else + letI := h.countableOrStandardBorel.resolve_left hα; condKernelBorel κ + +instance instIsMarkovKernel_condKernel [CountableOrStandardBorel α β] + (κ : kernel α (β × Ω')) [IsMarkovKernel κ] : + IsMarkovKernel (kernel.condKernel κ) := by + unfold kernel.condKernel + split_ifs <;> infer_instance + +lemma compProd_fst_condKernel [hαβ : CountableOrStandardBorel α β] + (κ : kernel α (β × Ω')) [IsMarkovKernel κ] : + kernel.fst κ ⊗ₖ kernel.condKernel κ = κ := by + unfold kernel.condKernel + split_ifs with h + · have := h.1 + have := h.2 + exact compProd_fst_condKernelCountable κ + · have := hαβ.countableOrStandardBorel.resolve_left h + exact compProd_fst_condKernelBorel κ + +end CountableOrStandardBorel + end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean new file mode 100644 index 0000000000000..9c398cf773354 --- /dev/null +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -0,0 +1,146 @@ +/- +Copyright (c) 2023 Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kexing Ying +-/ +import Mathlib.Probability.Kernel.Disintegration + +/-! +# Disintegration of measures on product spaces + +Let `ρ` be a finite measure on `α × Ω`, where `Ω` is a standard Borel space. In mathlib terms, `Ω` +verifies `[Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω]`. +Then there exists a kernel `ρ.condKernel : kernel α Ω` such that for any measurable set +`s : Set (α × Ω)`, `ρ s = ∫⁻ a, ρ.condKernel a {x | (a, x) ∈ s} ∂ρ.fst`. + +In terms of kernels, `ρ.condKernel` is such that for any measurable space `γ`, we +have a disintegration of the constant kernel from `γ` with value `ρ`: +`kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ (condKernel ρ))`, +where `ρ.fst` is the marginal measure of `ρ` on `α`. In particular, `ρ = ρ.fst ⊗ₘ ρ.condKernel`. + +In order to obtain a disintegration for any standard Borel space, we use that these spaces embed +measurably into `ℝ`: it then suffices to define a suitable kernel for `Ω = ℝ`. In the real case, +we define a conditional kernel by taking for each `a : α` the measure associated to the Stieltjes +function `condCDF ρ a` (the conditional cumulative distribution function). + +## Main statements + +* `ProbabilityTheory.eq_condKernel_of_measure_eq_compProd`: a.e. uniqueness of `Measure.condKernel` + +-/ + + +open MeasureTheory Set Filter + +open scoped ENNReal MeasureTheory Topology ProbabilityTheory + +namespace ProbabilityTheory + +variable {α : Type*} {mα : MeasurableSpace α} + +variable {Ω : Type*} [MeasurableSpace Ω] [StandardBorelSpace Ω] + [Nonempty Ω] (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] + +/-! ### Uniqueness + +This section will prove that the conditional kernel is unique almost everywhere. -/ + +/-- A s-finite kernel which satisfy the disintegration property of the given measure `ρ` is almost +everywhere equal to the disintegration kernel of `ρ` when evaluated on a measurable set. + +This theorem in the case of finite kernels is weaker than `eq_condKernel_of_measure_eq_compProd` +which asserts that the kernels are equal almost everywhere and not just on a given measurable +set. -/ +theorem eq_condKernel_of_measure_eq_compProd' (κ : kernel α Ω) [IsSFiniteKernel κ] + (hκ : ρ = ρ.fst ⊗ₘ κ) {s : Set Ω} (hs : MeasurableSet s) : + ∀ᵐ x ∂ρ.fst, κ x s = ρ.condKernel x s := by + refine' ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite + (kernel.measurable_coe κ hs) (kernel.measurable_coe ρ.condKernel hs) _ + intros t ht _ + conv_rhs => rw [set_lintegral_condKernel_eq_measure_prod _ ht hs, hκ] + simp only [Measure.compProd_apply (ht.prod hs), Set.mem_prod, ← lintegral_indicator _ ht] + congr; ext x + by_cases hx : x ∈ t + all_goals simp [hx] + +-- This lemma establishes uniqueness of the disintegration kernel on ℝ +lemma eq_condKernel_of_measure_eq_compProd_real (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] + (κ : kernel α ℝ) [IsFiniteKernel κ] (hκ : ρ = ρ.fst ⊗ₘ κ) : + ∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by + have huniv : ∀ᵐ x ∂ρ.fst, κ x Set.univ = ρ.condKernel x Set.univ := + eq_condKernel_of_measure_eq_compProd' ρ κ hκ MeasurableSet.univ + suffices : ∀ᵐ x ∂ρ.fst, ∀ ⦃t⦄, MeasurableSet t → κ x t = ρ.condKernel x t + · filter_upwards [this] with x hx + ext t ht; exact hx ht + apply MeasurableSpace.ae_induction_on_inter Real.borel_eq_generateFrom_Iic_rat + Real.isPiSystem_Iic_rat + · simp only [OuterMeasure.empty', Filter.eventually_true] + · simp only [iUnion_singleton_eq_range, mem_range, forall_exists_index, forall_apply_eq_imp_iff] + exact ae_all_iff.2 fun q => eq_condKernel_of_measure_eq_compProd' ρ κ hκ measurableSet_Iic + · filter_upwards [huniv] with x hxuniv t ht heq + rw [measure_compl ht <| measure_ne_top _ _, heq, hxuniv, measure_compl ht <| measure_ne_top _ _] + · refine' ae_of_all _ (fun x f hdisj hf heq => _) + rw [measure_iUnion hdisj hf, measure_iUnion hdisj hf] + exact tsum_congr heq + +/-- A finite kernel which satisfies the disintegration property is almost everywhere equal to the +disintegration kernel. -/ +theorem eq_condKernel_of_measure_eq_compProd (κ : kernel α Ω) [IsFiniteKernel κ] + (hκ : ρ = ρ.fst ⊗ₘ κ) : + ∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by +-- The idea is to transporting the question to `ℝ` from `Ω` using `exists_measurableEmbedding_real` +-- and then constructing a measure on `α × ℝ` using the obtained measurable embedding + obtain ⟨f, hf⟩ := exists_measurableEmbedding_real Ω + set ρ' : Measure (α × ℝ) := ρ.map (Prod.map id f) with hρ'def + have hρ' : ρ'.fst = ρ.fst + · ext s hs + rw [hρ'def, Measure.fst_apply, Measure.fst_apply, Measure.map_apply] + exacts [rfl, Measurable.prod measurable_fst <| hf.measurable.comp measurable_snd, + measurable_fst hs, hs, hs] + have hρ'' : ∀ᵐ x ∂ρ'.fst, kernel.map κ f hf.measurable x = ρ'.condKernel x + · refine' eq_condKernel_of_measure_eq_compProd_real ρ' (kernel.map κ f hf.measurable) _ + ext s hs + simp only [Measure.map_apply (measurable_id.prod_map hf.measurable) hs] + conv_lhs => congr; rw [hκ] + rw [Measure.compProd_apply hs, Measure.compProd_apply + (measurable_id.prod_map hf.measurable hs), (_ : (ρ.map (Prod.map id f)).fst = ρ.fst)] + · congr + ext x + simp only [Set.mem_preimage, Prod_map, id_eq, kernel.prodMkLeft_apply, kernel.map_apply] + rw [Measure.map_apply hf.measurable] + · rfl + · exact measurable_prod_mk_left hs + · rw [Measure.fst_map_prod_mk] + · simp only [Prod_map, id_eq]; rfl + · exact (hf.measurable.comp measurable_snd) + rw [hρ'] at hρ'' + suffices : ∀ᵐ x ∂ρ.fst, ∀ s, MeasurableSet s → + ((ρ.map (Prod.map id f)).condKernel x) s = (ρ.condKernel x) (f ⁻¹' s) + · filter_upwards [hρ'', this] with x hx h + rw [kernel.map_apply] at hx + ext s hs + rw [← Set.preimage_image_eq s hf.injective, + ← Measure.map_apply hf.measurable <| hf.measurableSet_image.2 hs, hx, + h _ <| hf.measurableSet_image.2 hs] + have hprod : (ρ.map (Prod.map id f)).fst = ρ.fst + · ext s hs + rw [Measure.fst_apply hs, + Measure.map_apply (measurable_id.prod_map hf.measurable) (measurable_fst hs), + ← Set.preimage_comp, Measure.fst_apply hs] + rfl + suffices : ρ.map (Prod.map id f) = + ((ρ.map (Prod.map id f)).fst ⊗ₘ (kernel.map (Measure.condKernel ρ) f hf.measurable)) + · have heq := eq_condKernel_of_measure_eq_compProd_real _ _ this + rw [hprod] at heq + filter_upwards [heq] with x hx s hs + rw [← hx, kernel.map_apply, Measure.map_apply hf.measurable hs] + ext s hs + simp only [hprod, Measure.compProd_apply hs, kernel.map_apply _ hf.measurable] + rw [Measure.map_apply (measurable_id.prod_map hf.measurable) hs, ← lintegral_condKernel_mem] + · simp only [mem_preimage, Prod_map, id_eq] + congr with a + rw [Measure.map_apply hf.measurable (measurable_prod_mk_left hs)] + rfl + · exact measurable_id.prod_map hf.measurable hs + +end ProbabilityTheory From c7492e9c98d121b03e3015fe42a5d5cafc6a4709 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 14 Feb 2024 15:24:31 +0100 Subject: [PATCH 013/129] generalize to finite kernel --- .../Probability/Kernel/Disintegration.lean | 23 ++-- .../Kernel/Disintegration/Basic.lean | 24 ++-- .../Kernel/Disintegration/KernelCDFBorel.lean | 119 +++++++++--------- .../Kernel/Disintegration/Unique.lean | 15 --- 4 files changed, 85 insertions(+), 96 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration.lean b/Mathlib/Probability/Kernel/Disintegration.lean index 1af8f1824d81b..0e681729a1dfe 100644 --- a/Mathlib/Probability/Kernel/Disintegration.lean +++ b/Mathlib/Probability/Kernel/Disintegration.lean @@ -178,7 +178,7 @@ variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E theorem _root_.MeasureTheory.AEStronglyMeasurable.integral_condKernel {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} (hf : AEStronglyMeasurable f ρ) : - AEStronglyMeasurable (fun x => ∫ y, f (x, y) ∂ρ.condKernel x) ρ.fst := by + AEStronglyMeasurable (fun x ↦ ∫ y, f (x, y) ∂ρ.condKernel x) ρ.fst := by rw [measure_eq_compProd ρ] at hf exact AEStronglyMeasurable.integral_kernel_compProd hf #align measure_theory.ae_strongly_measurable.integral_cond_kernel MeasureTheory.AEStronglyMeasurable.integral_condKernel @@ -223,7 +223,6 @@ namespace MeasureTheory We place these lemmas in the `MeasureTheory` namespace to enable dot notation. -/ - open ProbabilityTheory variable {α Ω E F : Type*} {mα : MeasurableSpace α} [MeasurableSpace Ω] @@ -232,38 +231,38 @@ variable {α Ω E F : Type*} {mα : MeasurableSpace α} [MeasurableSpace Ω] theorem AEStronglyMeasurable.ae_integrable_condKernel_iff {f : α × Ω → F} (hf : AEStronglyMeasurable f ρ) : - (∀ᵐ a ∂ρ.fst, Integrable (fun ω => f (a, ω)) (ρ.condKernel a)) ∧ - Integrable (fun a => ∫ ω, ‖f (a, ω)‖ ∂ρ.condKernel a) ρ.fst ↔ Integrable f ρ := by + (∀ᵐ a ∂ρ.fst, Integrable (fun ω ↦ f (a, ω)) (ρ.condKernel a)) ∧ + Integrable (fun a ↦ ∫ ω, ‖f (a, ω)‖ ∂ρ.condKernel a) ρ.fst ↔ Integrable f ρ := by rw [measure_eq_compProd ρ] at hf conv_rhs => rw [measure_eq_compProd ρ] rw [Measure.integrable_compProd_iff hf] #align measure_theory.ae_strongly_measurable.ae_integrable_cond_kernel_iff MeasureTheory.AEStronglyMeasurable.ae_integrable_condKernel_iff theorem Integrable.condKernel_ae {f : α × Ω → F} (hf_int : Integrable f ρ) : - ∀ᵐ a ∂ρ.fst, Integrable (fun ω => f (a, ω)) (ρ.condKernel a) := by + ∀ᵐ a ∂ρ.fst, Integrable (fun ω ↦ f (a, ω)) (ρ.condKernel a) := by have hf_ae : AEStronglyMeasurable f ρ := hf_int.1 rw [← hf_ae.ae_integrable_condKernel_iff] at hf_int exact hf_int.1 #align measure_theory.integrable.cond_kernel_ae MeasureTheory.Integrable.condKernel_ae theorem Integrable.integral_norm_condKernel {f : α × Ω → F} (hf_int : Integrable f ρ) : - Integrable (fun x => ∫ y, ‖f (x, y)‖ ∂ρ.condKernel x) ρ.fst := by + Integrable (fun x ↦ ∫ y, ‖f (x, y)‖ ∂ρ.condKernel x) ρ.fst := by have hf_ae : AEStronglyMeasurable f ρ := hf_int.1 rw [← hf_ae.ae_integrable_condKernel_iff] at hf_int exact hf_int.2 #align measure_theory.integrable.integral_norm_cond_kernel MeasureTheory.Integrable.integral_norm_condKernel theorem Integrable.norm_integral_condKernel {f : α × Ω → E} (hf_int : Integrable f ρ) : - Integrable (fun x => ‖∫ y, f (x, y) ∂ρ.condKernel x‖) ρ.fst := by - refine' hf_int.integral_norm_condKernel.mono hf_int.1.integral_condKernel.norm _ - refine' eventually_of_forall fun x => _ + Integrable (fun x ↦ ‖∫ y, f (x, y) ∂ρ.condKernel x‖) ρ.fst := by + refine hf_int.integral_norm_condKernel.mono hf_int.1.integral_condKernel.norm ?_ + refine eventually_of_forall fun x ↦ ?_ rw [norm_norm] - refine' (norm_integral_le_integral_norm _).trans_eq (Real.norm_of_nonneg _).symm - exact integral_nonneg_of_ae (eventually_of_forall fun y => norm_nonneg _) + refine (norm_integral_le_integral_norm _).trans_eq (Real.norm_of_nonneg ?_).symm + exact integral_nonneg_of_ae (eventually_of_forall fun y ↦ norm_nonneg _) #align measure_theory.integrable.norm_integral_cond_kernel MeasureTheory.Integrable.norm_integral_condKernel theorem Integrable.integral_condKernel {f : α × Ω → E} (hf_int : Integrable f ρ) : - Integrable (fun x => ∫ y, f (x, y) ∂ρ.condKernel x) ρ.fst := + Integrable (fun x ↦ ∫ y, f (x, y) ∂ρ.condKernel x) ρ.fst := (integrable_norm_iff hf_int.1.integral_condKernel).mp hf_int.norm_integral_condKernel #align measure_theory.integrable.integral_cond_kernel MeasureTheory.Integrable.integral_condKernel diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 494526d8908b8..4bc55df8c5132 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -135,44 +135,44 @@ end BorelSnd section StandardBorel noncomputable -def kernel.condKernelReal (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : kernel (α × ℝ) ℝ := +def kernel.condKernelReal (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : kernel (α × ℝ) ℝ := cdfKernel _ (isKernelCDF_mLimsupCDF κ) -instance (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : IsMarkovKernel (kernel.condKernelReal κ) := by +instance (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : IsMarkovKernel (kernel.condKernelReal κ) := by unfold kernel.condKernelReal; infer_instance -lemma kernel.eq_compProd_condKernelReal (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : +lemma kernel.eq_compProd_condKernelReal (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : κ = kernel.fst κ ⊗ₖ kernel.condKernelReal κ := kernel.eq_compProd_cdfKernel (isKernelCDF_mLimsupCDF κ) noncomputable -def condKernelBorelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel κ] : kernel (α × ℝ) Ω' := +def condKernelBorelAux (κ : kernel α (ℝ × Ω')) [IsFiniteKernel κ] : kernel (α × ℝ) Ω' := let f := measurableEmbedding_real Ω' let hf := measurableEmbedding_measurableEmbedding_real Ω' let κ' := kernel.map κ (Prod.map (id : ℝ → ℝ) f) (measurable_id.prod_map hf.measurable) condKernelBorelSnd κ (isKernelCDF_mLimsupCDF κ') -instance instIsMarkovKernel_condKernelBorelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel κ] : +instance instIsMarkovKernel_condKernelBorelAux (κ : kernel α (ℝ × Ω')) [IsFiniteKernel κ] : IsMarkovKernel (condKernelBorelAux κ) := by rw [condKernelBorelAux] infer_instance -lemma compProd_fst_condKernelBorelAux (κ : kernel α (ℝ × Ω')) [IsMarkovKernel κ] : +lemma compProd_fst_condKernelBorelAux (κ : kernel α (ℝ × Ω')) [IsFiniteKernel κ] : kernel.fst κ ⊗ₖ condKernelBorelAux κ = κ := by rw [condKernelBorelAux, compProd_fst_condKernelBorelSnd] noncomputable -def condKernelBorel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : kernel (α × Ω) Ω' := +def condKernelBorel (κ : kernel α (Ω × Ω')) [IsFiniteKernel κ] : kernel (α × Ω) Ω' := let f := measurableEmbedding_real Ω let hf := measurableEmbedding_measurableEmbedding_real Ω let κ' := kernel.map κ (Prod.map f (id : Ω' → Ω')) (hf.measurable.prod_map measurable_id) kernel.comap (condKernelBorelAux κ') (Prod.map (id : α → α) f) (measurable_id.prod_map hf.measurable) -instance instIsMarkovKernel_condKernelBorel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : +instance instIsMarkovKernel_condKernelBorel (κ : kernel α (Ω × Ω')) [IsFiniteKernel κ] : IsMarkovKernel (condKernelBorel κ) := by rw [condKernelBorel]; infer_instance -lemma compProd_fst_condKernelBorel (κ : kernel α (Ω × Ω')) [IsMarkovKernel κ] : +lemma compProd_fst_condKernelBorel (κ : kernel α (Ω × Ω')) [IsFiniteKernel κ] : kernel.fst κ ⊗ₖ condKernelBorel κ = κ := by let f := measurableEmbedding_real Ω let hf := measurableEmbedding_measurableEmbedding_real Ω @@ -345,7 +345,7 @@ instance instCountableOrStandardBorel_of_standardBorelSpace [h : StandardBorelSp open Classical in noncomputable def kernel.condKernel [h : CountableOrStandardBorel α β] - (κ : kernel α (β × Ω')) [IsMarkovKernel κ] : + (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : kernel (α × β) Ω' := if hα : Countable α ∧ MeasurableSingletonClass α then letI := hα.1; letI := hα.2; condKernelCountable κ @@ -353,13 +353,13 @@ def kernel.condKernel [h : CountableOrStandardBorel α β] letI := h.countableOrStandardBorel.resolve_left hα; condKernelBorel κ instance instIsMarkovKernel_condKernel [CountableOrStandardBorel α β] - (κ : kernel α (β × Ω')) [IsMarkovKernel κ] : + (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : IsMarkovKernel (kernel.condKernel κ) := by unfold kernel.condKernel split_ifs <;> infer_instance lemma compProd_fst_condKernel [hαβ : CountableOrStandardBorel α β] - (κ : kernel α (β × Ω')) [IsMarkovKernel κ] : + (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : kernel.fst κ ⊗ₖ kernel.condKernel κ = κ := by unfold kernel.condKernel split_ifs with h diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean index 711fab705fee4..e7c7ebfcf4872 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean @@ -267,20 +267,20 @@ lemma m_le_one (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h -lemma snorm_m_le_one (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] +lemma snorm_m_le (κ : kernel α (ℝ × β)) [IsFiniteKernel (kernel.fst κ)] (a : α) (s : Set β) (n : ℕ) : - snorm (M κ a s n) 1 (kernel.fst κ a) ≤ 1 := by + snorm (M κ a s n) 1 (kernel.fst κ a) ≤ kernel.fst κ a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun x ↦ ?_))).trans ?_ · simp only [Real.norm_eq_abs, abs_of_nonneg (m_nonneg κ a s n x), m_le_one κ a s n x] · simp -lemma integrable_m (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] +lemma integrable_m (κ : kernel α (ℝ × β)) [IsFiniteKernel (kernel.fst κ)] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : Integrable (M κ a s n) (kernel.fst κ a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ · exact measurable_m_right κ a hs n - · exact (snorm_m_le_one κ a s n).trans_lt ENNReal.one_lt_top + · exact (snorm_m_le κ a s n).trans_lt (measure_lt_top _ _) lemma set_integral_m_I (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) (k : ℤ) : @@ -321,7 +321,7 @@ lemma set_integral_m_I (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] rw [div_eq_mul_inv, mul_assoc, ENNReal.inv_mul_cancel h0, mul_one] exact measure_ne_top _ _ -lemma integral_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] +lemma integral_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : ∫ t, M κ a s n t ∂(kernel.fst κ a) = (κ a (univ ×ˢ s)).toReal := by rw [← integral_univ, ← iUnion_I n, iUnion_prod_const, measure_iUnion] @@ -336,7 +336,7 @@ lemma integral_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] congr with k rw [set_integral_m_I _ _ hs] -lemma set_integral_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] +lemma set_integral_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : ∫ t in A, M κ a s n t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal := by refine MeasurableSpace.induction_on_inter (m := ℱ n) (s := {s | ∃ k, s = I n k}) @@ -376,13 +376,13 @@ lemma set_integral_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] rw [Function.onFun, Set.disjoint_prod] exact Or.inl (hf_disj hij) -lemma set_integral_m_of_le (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] +lemma set_integral_m_of_le (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) {n m : ℕ} (hnm : n ≤ m) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : ∫ t in A, M κ a s m t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal := set_integral_m κ a hs m (ℱ.mono hnm A hA) -lemma condexp_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] (a : α) {s : Set β} +lemma condexp_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) {i j : ℕ} (hij : i ≤ j) : (kernel.fst κ a)[M κ a s j | ℱ i] =ᵐ[kernel.fst κ a] M κ a s i := by symm @@ -394,7 +394,7 @@ lemma condexp_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] (a : α) {s : S rw [set_integral_m κ a hs i ht, set_integral_m_of_le κ a hs hij ht] · exact StronglyMeasurable.aeStronglyMeasurable' (stronglyMeasurable_ℱ_m κ a s i) -lemma martingale_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] (a : α) {s : Set β} +lemma martingale_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) : Martingale (M κ a s) ℱ (kernel.fst κ a) := ⟨adapted_m κ a s, fun _ _ ↦ condexp_m κ a hs⟩ @@ -545,24 +545,26 @@ lemma tendsto_m_atTop_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKernel rw [← m_empty κ a n t] exact tendsto_m_atTop_empty_of_antitone κ a s hs hs_iInter hs_meas n t -lemma tendsto_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsMarkovKernel κ] +lemma tendsto_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsFiniteKernel κ] {s : Set β} (hs : MeasurableSet s) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun n ↦ M κ a s n t) atTop (𝓝 (ℱ.limitProcess (M κ a s) (kernel.fst κ a) t)) := by - refine Submartingale.ae_tendsto_limitProcess (martingale_m κ a hs).submartingale (R := 1) ?_ - intro n - rw [ENNReal.coe_one] - exact snorm_m_le_one κ a s n + refine Submartingale.ae_tendsto_limitProcess (martingale_m κ a hs).submartingale + (R := (kernel.fst κ a univ).toNNReal) (fun n ↦ ?_) + refine (snorm_m_le κ a s n).trans_eq ?_ + rw [ENNReal.coe_toNNReal] + exact measure_ne_top _ _ -lemma limitProcess_mem_L1 (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] +lemma limitProcess_mem_L1 (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) : Memℒp (ℱ.limitProcess (M κ a s) (kernel.fst κ a)) 1 (kernel.fst κ a) := by - refine Submartingale.memℒp_limitProcess (martingale_m κ a hs).submartingale (R := 1) ?_ - intro n - rw [ENNReal.coe_one] - exact snorm_m_le_one κ a s n + refine Submartingale.memℒp_limitProcess (martingale_m κ a hs).submartingale + (R := (kernel.fst κ a univ).toNNReal) (fun n ↦ ?_) + refine (snorm_m_le κ a s n).trans_eq ?_ + rw [ENNReal.coe_toNNReal] + exact measure_ne_top _ _ -lemma tendsto_snorm_one_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsMarkovKernel κ] +lemma tendsto_snorm_one_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsFiniteKernel κ] {s : Set β} (hs : MeasurableSet s) : Tendsto (fun n ↦ snorm (M κ a s n - ℱ.limitProcess (M κ a s) (kernel.fst κ a)) 1 (kernel.fst κ a)) @@ -596,7 +598,7 @@ lemma tendsto_snorm_restrict_zero {α β ι : Type*} {mα : MeasurableSpace α} · exact fun _ ↦ zero_le _ · exact fun _ ↦ snorm_restrict_le _ -lemma tendsto_snorm_one_restrict_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsMarkovKernel κ] +lemma tendsto_snorm_one_restrict_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsFiniteKernel κ] {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : Tendsto (fun n ↦ snorm (M κ a s n - ℱ.limitProcess (M κ a s) (kernel.fst κ a)) 1 ((kernel.fst κ a).restrict A)) @@ -607,12 +609,12 @@ noncomputable def MLimsup (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (t : ℝ) : ℝ := limsup (fun n ↦ M κ a s n t) atTop -lemma mLimsup_ae_eq_limitProcess (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] +lemma mLimsup_ae_eq_limitProcess (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) : MLimsup κ a s =ᵐ[kernel.fst κ a] ℱ.limitProcess (M κ a s) (kernel.fst κ a) := by filter_upwards [tendsto_m_limitProcess κ a hs] with t ht using ht.limsup_eq -lemma tendsto_m_mLimsup (κ : kernel α (ℝ × β)) (a : α) [IsMarkovKernel κ] +lemma tendsto_m_mLimsup (κ : kernel α (ℝ × β)) (a : α) [IsFiniteKernel κ] {s : Set β} (hs : MeasurableSet s) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun n ↦ M κ a s n t) atTop (𝓝 (MLimsup κ a s t)) := by @@ -679,21 +681,21 @@ lemma mLimsup_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) : simp_rw [ht] rw [limsup_const] -- should be simp -lemma snorm_mLimsup_le_one (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] +lemma snorm_mLimsup_le (κ : kernel α (ℝ × β)) [IsFiniteKernel (kernel.fst κ)] (a : α) (s : Set β) : - snorm (fun t ↦ MLimsup κ a s t) 1 (kernel.fst κ a) ≤ 1 := by + snorm (fun t ↦ MLimsup κ a s t) 1 (kernel.fst κ a) ≤ kernel.fst κ a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ · simp only [Real.norm_eq_abs, abs_of_nonneg (mLimsup_nonneg κ a s t), mLimsup_le_one κ a s t] · simp -lemma integrable_mLimsup (κ : kernel α (ℝ × β)) [IsMarkovKernel (kernel.fst κ)] +lemma integrable_mLimsup (κ : kernel α (ℝ × β)) [IsFiniteKernel (kernel.fst κ)] (a : α) {s : Set β} (hs : MeasurableSet s) : Integrable (fun t ↦ MLimsup κ a s t) (kernel.fst κ a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ · exact measurable_mLimsup_right κ hs a - · exact (snorm_mLimsup_le_one κ a s).trans_lt ENNReal.one_lt_top + · exact (snorm_mLimsup_le κ a s).trans_lt (measure_lt_top _ _) lemma tendsto_integral_of_L1' {ι G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] {μ : Measure α} @@ -729,7 +731,7 @@ lemma tendsto_set_integral_of_L1' {ι G : Type*} [NormedAddCommGroup G] [NormedS simp_rw [snorm_one_eq_lintegral_nnnorm, Pi.sub_apply] at hF exact hF -lemma tendsto_set_integral_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] +lemma tendsto_set_integral_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : Tendsto (fun i ↦ ∫ x in A, M κ a s i x ∂(kernel.fst κ) a) atTop (𝓝 (∫ x in A, MLimsup κ a s x ∂(kernel.fst κ) a)) := by @@ -741,7 +743,7 @@ lemma tendsto_set_integral_m (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] refine EventuallyEq.sub EventuallyEq.rfl ?_ exact (mLimsup_ae_eq_limitProcess κ a hs).symm -lemma set_integral_mLimsup_of_measurableSet (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] +lemma set_integral_mLimsup_of_measurableSet (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : ∫ t in A, MLimsup κ a s t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal := by @@ -759,12 +761,12 @@ lemma set_integral_mLimsup_of_measurableSet (κ : kernel α (ℝ × β)) [IsMark have h := tendsto_set_integral_m κ a hs A rw [h.limsup_eq] -lemma integral_mLimsup (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] +lemma integral_mLimsup (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) : ∫ t, MLimsup κ a s t ∂(kernel.fst κ a) = (κ a (univ ×ˢ s)).toReal := by rw [← integral_univ, set_integral_mLimsup_of_measurableSet κ a hs 0 MeasurableSet.univ] -lemma set_integral_mLimsup (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] +lemma set_integral_mLimsup (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set ℝ} (hA : MeasurableSet A) : ∫ t in A, MLimsup κ a s t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal := by have hA' : MeasurableSet[⨆ n, ℱ n] A := by rwa [iSup_ℱ] @@ -802,17 +804,18 @@ lemma set_integral_mLimsup (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] exact fun i ↦ (hf i).prod hs · rwa [iSup_ℱ] at hf -lemma tendsto_integral_mLimsup_of_monotone (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] +lemma tendsto_integral_mLimsup_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : - Tendsto (fun m ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop (𝓝 1) := by + Tendsto (fun m ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop (𝓝 (κ a univ).toReal) := by simp_rw [integral_mLimsup κ a (hs_meas _)] - rw [← ENNReal.one_toReal] - have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := 1) ?_ + have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := κ a univ) ?_ swap · rw [mem_nhds_iff] - refine ⟨Iio 2, fun x hx ↦ ne_top_of_lt (?_ : x < 2), isOpen_Iio, ENNReal.one_lt_two⟩ - simpa using hx + refine ⟨Iio (κ a univ + 1), fun x hx ↦ ne_top_of_lt (?_ : x < κ a univ + 1), isOpen_Iio, ?_⟩ + · simpa using hx + · simp only [mem_Iio] + exact ENNReal.lt_add_right (measure_ne_top _ _) one_ne_zero refine h_cont.tendsto.comp ?_ have h := tendsto_measure_iUnion (s := fun n ↦ univ ×ˢ s n) (μ := κ a) ?_ swap; · intro n m hnm x; simp only [mem_prod, mem_univ, true_and]; exact fun h ↦ hs hnm h @@ -820,7 +823,7 @@ lemma tendsto_integral_mLimsup_of_monotone (κ : kernel α (ℝ × β)) [IsMarko rw [← prod_iUnion, hs_iUnion] simp only [univ_prod_univ, measure_univ] -lemma tendsto_integral_mLimsup_of_antitone (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] +lemma tendsto_integral_mLimsup_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : Tendsto (fun m ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop (𝓝 0) := by @@ -934,7 +937,7 @@ theorem tendsto_atTop_atBot_iff_of_antitone {α β : Type*} Tendsto f atTop atBot ↔ ∀ b : β, ∃ a : α, f a ≤ b := @tendsto_atTop_atTop_iff_of_monotone _ βᵒᵈ _ _ _ _ hf -lemma tendsto_mLimsup_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] +lemma tendsto_mLimsup_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 1) := by @@ -970,7 +973,7 @@ lemma tendsto_mLimsup_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsMarko rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, ENNReal.ofReal_le_one, abs_of_nonneg (hF_nonneg _)] exact hF_le_one _ - _ < ⊤ := by simp only [lintegral_const, measure_univ, mul_one, ENNReal.one_lt_top] + _ < ⊤ := by simp only [lintegral_const, measure_univ, one_mul, measure_lt_top] -- it suffices to show that the limit `F` is 1 a.e. suffices ∀ᵐ t ∂(kernel.fst κ a), F t = 1 by filter_upwards [this] with t ht_eq @@ -989,12 +992,12 @@ lemma tendsto_mLimsup_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsMarko have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop (𝓝 (∫ _, 1 ∂(kernel.fst κ a))) := by - rw [integral_const, measure_univ] - simp only [ENNReal.one_toReal, smul_eq_mul, mul_one] + rw [integral_const, kernel.fst_apply' _ _ MeasurableSet.univ] + simp only [smul_eq_mul, mul_one] exact tendsto_integral_mLimsup_of_monotone κ a s hs hs_iUnion hs_meas exact tendsto_nhds_unique h_integral h_integral' -lemma tendsto_mLimsup_atTop_ae_of_antitone (κ : kernel α (ℝ × β)) [IsMarkovKernel κ] +lemma tendsto_mLimsup_atTop_ae_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 0) := by @@ -1031,7 +1034,9 @@ lemma tendsto_mLimsup_atTop_ae_of_antitone (κ : kernel α (ℝ × β)) [IsMarko rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, ENNReal.ofReal_le_one, abs_of_nonneg (hF_nonneg _)] exact hF_le_one _ - _ < ⊤ := by simp only [lintegral_const, measure_univ, mul_one, ENNReal.one_lt_top] + _ < ⊤ := by + simp only [lintegral_const, one_mul] + exact measure_lt_top _ _ -- it suffices to show that the limit `F` is 0 a.e. suffices ∀ᵐ t ∂(kernel.fst κ a), F t = 0 by filter_upwards [this] with t ht_eq @@ -1086,7 +1091,7 @@ theorem tendsto_nat_ceil_atTop {α : Type*} [LinearOrderedSemiring α] [FloorSem refine Nat.ceil_mono.tendsto_atTop_atTop (fun x ↦ ⟨x, ?_⟩) simp only [Nat.ceil_natCast, le_refl] -lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) : +lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ a t q) atTop (𝓝 1) := by suffices ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ a t n) atTop (𝓝 1) by filter_upwards [this] with t ht @@ -1109,7 +1114,7 @@ lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] filter_upwards [tendsto_mLimsup_atTop_ae_of_monotone κ a s hs hs_iUnion hs_meas] with x hx using hx -lemma tendsto_atBot_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) : +lemma tendsto_atBot_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ a t q) atBot (𝓝 0) := by suffices ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ a t (-q)) atTop (𝓝 0) by filter_upwards [this] with t ht @@ -1144,17 +1149,17 @@ lemma tendsto_atBot_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] rw [mLimsupIic] simp -lemma set_integral_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] +lemma set_integral_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) (q : ℚ) {A : Set ℝ} (hA : MeasurableSet A) : ∫ t in A, mLimsupIic κ a t q ∂(kernel.fst κ a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := set_integral_mLimsup κ a measurableSet_Iic hA -lemma integrable_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel (kernel.fst κ)] +lemma integrable_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel (kernel.fst κ)] (a : α) (q : ℚ) : Integrable (fun t ↦ mLimsupIic κ a t q) (kernel.fst κ a) := integrable_mLimsup _ _ measurableSet_Iic -lemma bddBelow_range_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel (kernel.fst κ)] +lemma bddBelow_range_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel (kernel.fst κ)] (a : α) (t : ℝ) (q : ℚ) : BddBelow (range fun (r : Ioi q) ↦ mLimsupIic κ a t r) := by refine ⟨0, ?_⟩ @@ -1162,12 +1167,12 @@ lemma bddBelow_range_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel (k rintro x ⟨y, rfl⟩ exact mLimsupIic_nonneg _ _ _ _ -lemma integrable_iInf_rat_gt_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) (q : ℚ) : +lemma integrable_iInf_rat_gt_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) (q : ℚ) : Integrable (fun t ↦ ⨅ r : Ioi q, mLimsupIic κ a t r) (kernel.fst κ a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ · exact measurable_iInf fun i ↦ measurable_mLimsupIic_right κ a i - refine (?_ : _ ≤ (1 : ℝ≥0∞)).trans_lt ENNReal.one_lt_top + refine (?_ : _ ≤ (kernel.fst κ a univ : ℝ≥0∞)).trans_lt (measure_lt_top _ _) refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ · rw [Real.norm_eq_abs, abs_of_nonneg] · refine ciInf_le_of_le ?_ ?_ ?_ @@ -1177,7 +1182,7 @@ lemma integrable_iInf_rat_gt_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovK · exact le_ciInf fun r ↦ mLimsupIic_nonneg κ a t r · simp -lemma set_integral_iInf_rat_gt_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] +lemma set_integral_iInf_rat_gt_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) (q : ℚ) {A : Set ℝ} (hA : MeasurableSet A) : ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ a t r ∂(kernel.fst κ a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by @@ -1201,7 +1206,7 @@ lemma set_integral_iInf_rat_gt_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarko · exact (integrable_iInf_rat_gt_mLimsupIic _ _ _).integrableOn · exact fun t ↦ le_ciInf (fun r ↦ monotone_mLimsupIic _ _ _ (le_of_lt r.prop)) -lemma iInf_rat_gt_mLimsupIic_eq (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) : +lemma iInf_rat_gt_mLimsupIic_eq (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), ∀ q : ℚ, ⨅ r : Ioi q, mLimsupIic κ a t r = mLimsupIic κ a t q := by rw [ae_all_iff] refine fun q ↦ ae_eq_of_forall_set_integral_eq_of_sigmaFinite (μ := kernel.fst κ a) ?_ ?_ ?_ @@ -1216,7 +1221,7 @@ end Iic_Q section Rat -lemma isRatStieltjesPoint_mLimsupIic_ae (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] (a : α) : +lemma isRatStieltjesPoint_mLimsupIic_ae (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), IsRatStieltjesPoint (fun p q ↦ mLimsupIic κ p.1 p.2 q) (a, t) := by filter_upwards [tendsto_atTop_mLimsupIic κ a, tendsto_atBot_mLimsupIic κ a, iInf_rat_gt_mLimsupIic_eq κ a] with t ht_top ht_bot ht_iInf @@ -1228,7 +1233,7 @@ lemma isRatStieltjesPoint_mLimsupIic_ae (κ : kernel α (ℝ × ℝ)) [IsMarkovK · exact ht_bot · exact ht_iInf -lemma isRatKernelCDF_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : +lemma isRatKernelCDF_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : IsRatKernelCDF (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2) κ (kernel.fst κ) where measurable := measurable_mLimsupIic κ isRatStieltjesPoint_ae := isRatStieltjesPoint_mLimsupIic_ae κ @@ -1240,10 +1245,10 @@ end Rat section KernelCDF noncomputable -def mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : α × ℝ → StieltjesFunction := +def mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : α × ℝ → StieltjesFunction := todo3 (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2) (isRatKernelCDF_mLimsupIic κ).measurable -lemma isKernelCDF_mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsMarkovKernel κ] : +lemma isKernelCDF_mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : IsKernelCDF (mLimsupCDF κ) κ (kernel.fst κ) := isKernelCDF_todo3 (isRatKernelCDF_mLimsupIic κ) diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index 9c398cf773354..eed85a2596fcf 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -8,21 +8,6 @@ import Mathlib.Probability.Kernel.Disintegration /-! # Disintegration of measures on product spaces -Let `ρ` be a finite measure on `α × Ω`, where `Ω` is a standard Borel space. In mathlib terms, `Ω` -verifies `[Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω]`. -Then there exists a kernel `ρ.condKernel : kernel α Ω` such that for any measurable set -`s : Set (α × Ω)`, `ρ s = ∫⁻ a, ρ.condKernel a {x | (a, x) ∈ s} ∂ρ.fst`. - -In terms of kernels, `ρ.condKernel` is such that for any measurable space `γ`, we -have a disintegration of the constant kernel from `γ` with value `ρ`: -`kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ (condKernel ρ))`, -where `ρ.fst` is the marginal measure of `ρ` on `α`. In particular, `ρ = ρ.fst ⊗ₘ ρ.condKernel`. - -In order to obtain a disintegration for any standard Borel space, we use that these spaces embed -measurably into `ℝ`: it then suffices to define a suitable kernel for `Ω = ℝ`. In the real case, -we define a conditional kernel by taking for each `a : α` the measure associated to the Stieltjes -function `condCDF ρ a` (the conditional cumulative distribution function). - ## Main statements * `ProbabilityTheory.eq_condKernel_of_measure_eq_compProd`: a.e. uniqueness of `Measure.condKernel` From 575a4c9e1b5795e19cdb58dfb01c560e6f09d081 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 14 Feb 2024 17:35:56 +0100 Subject: [PATCH 014/129] main proof done --- .../Kernel/Disintegration/KernelCDFBorel.lean | 34 +++++++++++++++++-- 1 file changed, 31 insertions(+), 3 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean index e7c7ebfcf4872..1a7c7b2f9ebd5 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean @@ -1,3 +1,4 @@ +import Mathlib.Analysis.SpecialFunctions.Log.Base import Mathlib.Probability.Martingale.Convergence import Mathlib.Probability.Kernel.Disintegration.BuildKernel @@ -159,9 +160,36 @@ lemma iInter_biUnion_I (x : ℝ) : ⋂ n, ⋃ (k) (_ : indexI n x ≤ k), I n k refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ · by_contra h_lt push_neg at h_lt - --have h_pos : ∀ i, 0 < (2 : ℝ) ^ i := fun i ↦ by positivity - --simp_rw [← div_eq_mul_inv, div_le_iff (h_pos _)] at h - sorry + have h_pos : ∀ i, 0 < (2 : ℝ) ^ i := fun i ↦ by positivity + simp_rw [← div_eq_mul_inv, div_le_iff (h_pos _)] at h + obtain ⟨i, hi⟩ : ∃ i, 1 < (x - t) * 2 ^ i := by + suffices ∃ i : ℝ, 1 ≤ (x - t) * 2 ^ i by + obtain ⟨i, hi⟩ := this + use ⌈i⌉₊ + 1 + refine hi.trans_lt ?_ + gcongr + · simp [h_lt] + · refine ((Real.rpow_lt_rpow_left_iff one_lt_two).mpr (?_ : i < ⌈i⌉₊ + 1)).trans_eq ?_ + · refine (Nat.le_ceil _).trans_lt ?_ + norm_num + · norm_cast + use Real.logb 2 ((x - t)⁻¹) + rw [Real.rpow_logb] + · rw [mul_inv_cancel] + rw [sub_ne_zero] + exact h_lt.ne' + · exact zero_lt_two + · simp + · simp [h_lt] + specialize h i + rw [mul_comm, mul_sub, lt_sub_iff_add_lt', mul_comm] at hi + have h' : ⌈x * 2 ^ i⌉ ≤ t * 2 ^ i + 1 := by + calc (⌈x * 2 ^ i⌉ : ℝ) ≤ ⌊x * 2 ^ i⌋ + 1 := by + exact mod_cast (Int.ceil_le_floor_add_one (x * 2 ^ i)) + _ ≤ t * 2 ^ i + 1 := by gcongr + have h'' : ↑⌈x * 2 ^ i⌉ < 2 ^ i * x := h'.trans_lt hi + rw [← not_le, mul_comm] at h'' + exact h'' (Int.le_ceil _) · intro n refine le_trans ?_ h rw [← div_eq_mul_inv, div_le_iff] From 8edfabd0432ffe09a8bdc3e28b45fd69410da695 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 15 Feb 2024 11:06:44 +0100 Subject: [PATCH 015/129] uniqueness of kernel.condexpKernel --- .../Probability/Kernel/Disintegration.lean | 24 +- .../Kernel/Disintegration/Basic.lean | 182 +++++++++++---- .../Kernel/Disintegration/Unique.lean | 221 ++++++++++++++---- 3 files changed, 313 insertions(+), 114 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration.lean b/Mathlib/Probability/Kernel/Disintegration.lean index 0e681729a1dfe..49b6efed11ce8 100644 --- a/Mathlib/Probability/Kernel/Disintegration.lean +++ b/Mathlib/Probability/Kernel/Disintegration.lean @@ -126,14 +126,14 @@ lemma condKernel_apply_of_ne_zero [MeasurableSingletonClass α] (measurableEmbedding_prod_mk_left x).comap_apply] simp [this, (measurableEmbedding_prod_mk_left x).comap_apply, hx] -theorem lintegral_condKernel_mem {s : Set (α × Ω)} (hs : MeasurableSet s) : +theorem Measure.lintegral_condKernel_mem {s : Set (α × Ω)} (hs : MeasurableSet s) : ∫⁻ a, ρ.condKernel a {x | (a, x) ∈ s} ∂ρ.fst = ρ s := by conv_rhs => rw [measure_eq_compProd ρ] simp_rw [Measure.compProd_apply hs] rfl -#align probability_theory.lintegral_cond_kernel_mem ProbabilityTheory.lintegral_condKernel_mem +#align probability_theory.lintegral_cond_kernel_mem ProbabilityTheory.Measure.lintegral_condKernel_mem -theorem set_lintegral_condKernel_eq_measure_prod {s : Set α} (hs : MeasurableSet s) {t : Set Ω} +theorem Measure.set_lintegral_condKernel_eq_measure_prod {s : Set α} (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) : ∫⁻ a in s, ρ.condKernel a t ∂ρ.fst = ρ (s ×ˢ t) := by have : ρ (s ×ˢ t) = (ρ.fst ⊗ₘ ρ.condKernel) (s ×ˢ t) := by congr; exact measure_eq_compProd ρ @@ -145,28 +145,28 @@ theorem set_lintegral_condKernel_eq_measure_prod {s : Set α} (hs : MeasurableSe by_cases ha : a ∈ s <;> simp [ha] simp_rw [this] rw [lintegral_indicator _ hs] -#align probability_theory.set_lintegral_cond_kernel_eq_measure_prod ProbabilityTheory.set_lintegral_condKernel_eq_measure_prod +#align probability_theory.set_lintegral_cond_kernel_eq_measure_prod ProbabilityTheory.Measure.set_lintegral_condKernel_eq_measure_prod -theorem lintegral_condKernel {f : α × Ω → ℝ≥0∞} (hf : Measurable f) : +theorem Measure.lintegral_condKernel {f : α × Ω → ℝ≥0∞} (hf : Measurable f) : ∫⁻ a, ∫⁻ ω, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫⁻ x, f x ∂ρ := by conv_rhs => rw [measure_eq_compProd ρ] rw [Measure.lintegral_compProd hf] #align probability_theory.lintegral_cond_kernel ProbabilityTheory.lintegral_condKernel -theorem set_lintegral_condKernel {f : α × Ω → ℝ≥0∞} (hf : Measurable f) {s : Set α} +theorem Measure.set_lintegral_condKernel {f : α × Ω → ℝ≥0∞} (hf : Measurable f) {s : Set α} (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) : ∫⁻ a in s, ∫⁻ ω in t, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫⁻ x in s ×ˢ t, f x ∂ρ := by conv_rhs => rw [measure_eq_compProd ρ] rw [Measure.set_lintegral_compProd hf hs ht] #align probability_theory.set_lintegral_cond_kernel ProbabilityTheory.set_lintegral_condKernel -theorem set_lintegral_condKernel_univ_right {f : α × Ω → ℝ≥0∞} (hf : Measurable f) {s : Set α} +theorem Measure.set_lintegral_condKernel_univ_right {f : α × Ω → ℝ≥0∞} (hf : Measurable f) {s : Set α} (hs : MeasurableSet s) : ∫⁻ a in s, ∫⁻ ω, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫⁻ x in s ×ˢ univ, f x ∂ρ := by rw [← set_lintegral_condKernel ρ hf hs MeasurableSet.univ]; simp_rw [Measure.restrict_univ] #align probability_theory.set_lintegral_cond_kernel_univ_right ProbabilityTheory.set_lintegral_condKernel_univ_right -theorem set_lintegral_condKernel_univ_left {f : α × Ω → ℝ≥0∞} (hf : Measurable f) {t : Set Ω} +theorem Measure.set_lintegral_condKernel_univ_left {f : α × Ω → ℝ≥0∞} (hf : Measurable f) {t : Set Ω} (ht : MeasurableSet t) : ∫⁻ a, ∫⁻ ω in t, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫⁻ x in univ ×ˢ t, f x ∂ρ := by rw [← set_lintegral_condKernel ρ hf MeasurableSet.univ ht]; simp_rw [Measure.restrict_univ] @@ -183,14 +183,14 @@ theorem _root_.MeasureTheory.AEStronglyMeasurable.integral_condKernel {ρ : Meas exact AEStronglyMeasurable.integral_kernel_compProd hf #align measure_theory.ae_strongly_measurable.integral_cond_kernel MeasureTheory.AEStronglyMeasurable.integral_condKernel -theorem integral_condKernel {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} +theorem Measure.integral_condKernel {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} (hf : Integrable f ρ) : ∫ a, ∫ x, f (a, x) ∂ρ.condKernel a ∂ρ.fst = ∫ ω, f ω ∂ρ := by conv_rhs => rw [measure_eq_compProd ρ] have hf': Integrable f (ρ.fst ⊗ₘ ρ.condKernel) := by rwa [measure_eq_compProd ρ] at hf rw [Measure.integral_compProd hf'] #align probability_theory.integral_cond_kernel ProbabilityTheory.integral_condKernel -theorem set_integral_condKernel {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} +theorem Measure.set_integral_condKernel {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} {s : Set α} (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) (hf : IntegrableOn f (s ×ˢ t) ρ) : ∫ a in s, ∫ ω in t, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫ x in s ×ˢ t, f x ∂ρ := by @@ -199,13 +199,13 @@ theorem set_integral_condKernel {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] { rwa [measure_eq_compProd ρ] at hf #align probability_theory.set_integral_cond_kernel ProbabilityTheory.set_integral_condKernel -theorem set_integral_condKernel_univ_right {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} +theorem Measure.set_integral_condKernel_univ_right {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} {s : Set α} (hs : MeasurableSet s) (hf : IntegrableOn f (s ×ˢ univ) ρ) : ∫ a in s, ∫ ω, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫ x in s ×ˢ univ, f x ∂ρ := by rw [← set_integral_condKernel hs MeasurableSet.univ hf]; simp_rw [Measure.restrict_univ] #align probability_theory.set_integral_cond_kernel_univ_right ProbabilityTheory.set_integral_condKernel_univ_right -theorem set_integral_condKernel_univ_left {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} +theorem Measure.set_integral_condKernel_univ_left {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} {t : Set Ω} (ht : MeasurableSet t) (hf : IntegrableOn f (univ ×ˢ t) ρ) : ∫ a, ∫ ω in t, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫ x in univ ×ˢ t, f x ∂ρ := by rw [← set_integral_condKernel MeasurableSet.univ ht hf]; simp_rw [Measure.restrict_univ] diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 4bc55df8c5132..90cad62e6d71c 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -24,28 +24,28 @@ variable {α β Ω Ω': Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace section BorelSnd noncomputable -def condKernelBorelSnd (κ : kernel α (β × Ω')) {f : α × β → StieltjesFunction} +def condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFunction} (hf : IsKernelCDF f - (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) - (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω').measurable)) - (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) - (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω').measurable)))) : - kernel (α × β) Ω' := - let e := measurableEmbedding_real Ω' - let he := measurableEmbedding_measurableEmbedding_real Ω' + (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) + (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)) + (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) + (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)))) : + kernel (α × β) Ω := + let e := measurableEmbedding_real Ω + let he := measurableEmbedding_measurableEmbedding_real Ω let x₀ := (range_nonempty e).choose kernel.comapRight (kernel.piecewise (measurableSet_eq_one hf he.measurableSet_range) (cdfKernel f hf) (kernel.deterministic (fun _ ↦ x₀) measurable_const)) he -instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω')) +instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFunction} (hf : IsKernelCDF f - (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) - (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω').measurable)) - (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) - (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω').measurable)))) : + (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) + (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)) + (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) + (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)))) : IsMarkovKernel (condKernelBorelSnd κ hf) := by rw [condKernelBorelSnd] refine kernel.IsMarkovKernel.comapRight _ _ fun a ↦ ?_ @@ -54,20 +54,20 @@ instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω')) · exact h_mem · classical rw [kernel.deterministic_apply' _ _ - (measurableEmbedding_measurableEmbedding_real Ω').measurableSet_range, + (measurableEmbedding_measurableEmbedding_real Ω).measurableSet_range, Set.indicator_apply, if_pos] - exact (range_nonempty (measurableEmbedding_real Ω')).choose_spec + exact (range_nonempty (measurableEmbedding_real Ω)).choose_spec -lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω')) [IsFiniteKernel κ] +lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω)) [IsFiniteKernel κ] {f : α × β → StieltjesFunction} (hf : IsKernelCDF f - (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) - (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω').measurable)) - (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω')) - (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω').measurable)))) : + (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) + (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)) + (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) + (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)))) : kernel.fst κ ⊗ₖ condKernelBorelSnd κ hf = κ := by - let e := measurableEmbedding_real Ω' - let he := measurableEmbedding_measurableEmbedding_real Ω' + let e := measurableEmbedding_real Ω + let he := measurableEmbedding_measurableEmbedding_real Ω let κ' := kernel.map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable) have h_prod_embed : MeasurableEmbedding (Prod.map (id : β → β) e) := MeasurableEmbedding.id.prod_mk he @@ -106,12 +106,12 @@ lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω')) [IsFiniteKern refine ae_cdfKernel_eq_one hf a he.measurableSet_range ?_ simp only [mem_compl_iff, mem_range, not_exists] rw [kernel.map_apply'] - · have h_empty : {a : β × Ω' | ∀ (x : Ω'), ¬e x = e a.2} = ∅ := by + · have h_empty : {a : β × Ω | ∀ (x : Ω), ¬e x = e a.2} = ∅ := by ext x simp only [mem_setOf_eq, mem_empty_iff_false, iff_false, not_forall, not_not, exists_apply_eq_apply] simp [h_empty] - · have : {x : β × ℝ | ∀ (y : Ω'), ¬ e y = x.2} = univ ×ˢ (range e)ᶜ := by + · have : {x : β × ℝ | ∀ (y : Ω), ¬ e y = x.2} = univ ×ˢ (range e)ᶜ := by ext x simp only [mem_setOf_eq, mem_prod, mem_univ, mem_compl_iff, mem_range, not_exists, true_and] rw [this] @@ -119,7 +119,7 @@ lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω')) [IsFiniteKern filter_upwards [h_ae c] with a ha rw [condKernelBorelSnd, kernel.comapRight_apply'] swap; · exact measurable_prod_mk_left ht - have h1 : {c : ℝ | (a, c) ∈ Prod.map id e '' t} = e '' {c : Ω' | (a, c) ∈ t} := by + have h1 : {c : ℝ | (a, c) ∈ Prod.map id e '' t} = e '' {c : Ω | (a, c) ∈ t} := by ext1 x simp only [Prod_map, id.def, mem_image, Prod.mk.inj_iff, Prod.exists, mem_setOf_eq] constructor @@ -146,18 +146,18 @@ lemma kernel.eq_compProd_condKernelReal (κ : kernel α (ℝ × ℝ)) [IsFiniteK kernel.eq_compProd_cdfKernel (isKernelCDF_mLimsupCDF κ) noncomputable -def condKernelBorelAux (κ : kernel α (ℝ × Ω')) [IsFiniteKernel κ] : kernel (α × ℝ) Ω' := - let f := measurableEmbedding_real Ω' - let hf := measurableEmbedding_measurableEmbedding_real Ω' +def condKernelBorelAux (κ : kernel α (ℝ × Ω)) [IsFiniteKernel κ] : kernel (α × ℝ) Ω := + let f := measurableEmbedding_real Ω + let hf := measurableEmbedding_measurableEmbedding_real Ω let κ' := kernel.map κ (Prod.map (id : ℝ → ℝ) f) (measurable_id.prod_map hf.measurable) condKernelBorelSnd κ (isKernelCDF_mLimsupCDF κ') -instance instIsMarkovKernel_condKernelBorelAux (κ : kernel α (ℝ × Ω')) [IsFiniteKernel κ] : +instance instIsMarkovKernel_condKernelBorelAux (κ : kernel α (ℝ × Ω)) [IsFiniteKernel κ] : IsMarkovKernel (condKernelBorelAux κ) := by rw [condKernelBorelAux] infer_instance -lemma compProd_fst_condKernelBorelAux (κ : kernel α (ℝ × Ω')) [IsFiniteKernel κ] : +lemma compProd_fst_condKernelBorelAux (κ : kernel α (ℝ × Ω)) [IsFiniteKernel κ] : kernel.fst κ ⊗ₖ condKernelBorelAux κ = κ := by rw [condKernelBorelAux, compProd_fst_condKernelBorelSnd] @@ -244,18 +244,18 @@ end Real section BorelSnd noncomputable -def condKernelUnitBorel (κ : kernel Unit (α × Ω')) [IsFiniteKernel κ] : kernel (Unit × α) Ω' := - let f := measurableEmbedding_real Ω' - let hf := measurableEmbedding_measurableEmbedding_real Ω' +def condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : kernel (Unit × α) Ω := + let f := measurableEmbedding_real Ω + let hf := measurableEmbedding_measurableEmbedding_real Ω let κ' := kernel.map κ (Prod.map (id : α → α) f) (measurable_id.prod_map hf.measurable) condKernelBorelSnd κ (isKernelCDF_condCDF (κ' ())) -instance instIsMarkovKernel_condKernelUnitBorel (κ : kernel Unit (α × Ω')) [IsFiniteKernel κ] : +instance instIsMarkovKernel_condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : IsMarkovKernel (condKernelUnitBorel κ) := by rw [condKernelUnitBorel] infer_instance -lemma compProd_fst_condKernelUnitBorel (κ : kernel Unit (α × Ω')) [IsFiniteKernel κ] : +lemma compProd_fst_condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : kernel.fst κ ⊗ₖ condKernelUnitBorel κ = κ := by rw [condKernelUnitBorel, compProd_fst_condKernelBorelSnd] @@ -268,26 +268,23 @@ section Measure /-- Conditional kernel of a measure on a product space: a Markov kernel such that `ρ = ρ.fst ⊗ₘ ρ.condKernel` (see `ProbabilityTheory.measure_eq_compProd`). -/ noncomputable -def _root_.MeasureTheory.Measure.condKernel (ρ : Measure (α × Ω')) [IsFiniteMeasure ρ] : - kernel α Ω' where - val a := condKernelUnitBorel (kernel.const Unit ρ) ((), a) - property := (condKernelUnitBorel (kernel.const Unit ρ)).property.comp measurable_prod_mk_left +def _root_.MeasureTheory.Measure.condKernel (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : + kernel α Ω := + kernel.comap (condKernelUnitBorel (kernel.const Unit ρ)) (fun a ↦ ((), a)) measurable_prod_mk_left #align measure_theory.measure.cond_kernel MeasureTheory.Measure.condKernel -lemma _root_.MeasureTheory.Measure.condKernel_apply (ρ : Measure (α × Ω')) [IsFiniteMeasure ρ] +lemma _root_.MeasureTheory.Measure.condKernel_apply (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] (a : α) : ρ.condKernel a = condKernelUnitBorel (kernel.const Unit ρ) ((), a) := rfl instance _root_.MeasureTheory.Measure.instIsMarkovKernel_condKernel - (ρ : Measure (α × Ω')) [IsFiniteMeasure ρ] : + (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : IsMarkovKernel ρ.condKernel := by - constructor - intro a - change IsProbabilityMeasure (condKernelUnitBorel (kernel.const Unit ρ) ((), a)) + rw [Measure.condKernel] infer_instance lemma _root_.MeasureTheory.Measure.compProd_fst_condKernel - (ρ : Measure (α × Ω')) [IsFiniteMeasure ρ] : + (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : ρ.fst ⊗ₘ ρ.condKernel = ρ := by have h1 : kernel.const Unit (Measure.fst ρ) = kernel.fst (kernel.const Unit ρ) := by ext @@ -306,21 +303,21 @@ section Countable variable [MeasurableSingletonClass α] [Countable α] noncomputable -def condKernelCountable (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : kernel (α × β) Ω' where +def condKernelCountable (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : kernel (α × β) Ω where val p := (κ p.1).condKernel p.2 property := by change Measurable ((fun q : β × α ↦ (κ q.2).condKernel q.1) ∘ Prod.swap) refine (measurable_from_prod_countable (fun a ↦ ?_)).comp measurable_swap exact kernel.measurable (κ a).condKernel -lemma condKernelCountable_apply (κ : kernel α (β × Ω')) [IsFiniteKernel κ] (p : α × β) : +lemma condKernelCountable_apply (κ : kernel α (β × Ω)) [IsFiniteKernel κ] (p : α × β) : condKernelCountable κ p = (κ p.1).condKernel p.2 := rfl -instance instIsMarkovKernel_condKernelCountable (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : +instance instIsMarkovKernel_condKernelCountable (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : IsMarkovKernel (condKernelCountable κ) := ⟨fun p ↦ (Measure.instIsMarkovKernel_condKernel (κ p.1)).isProbabilityMeasure p.2⟩ -lemma compProd_fst_condKernelCountable (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : +lemma compProd_fst_condKernelCountable (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : kernel.fst κ ⊗ₖ condKernelCountable κ = κ := by ext a s hs have h := (κ a).compProd_fst_condKernel @@ -352,13 +349,13 @@ def kernel.condKernel [h : CountableOrStandardBorel α β] else letI := h.countableOrStandardBorel.resolve_left hα; condKernelBorel κ -instance instIsMarkovKernel_condKernel [CountableOrStandardBorel α β] +instance kernel.instIsMarkovKernel_condKernel [CountableOrStandardBorel α β] (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : IsMarkovKernel (kernel.condKernel κ) := by unfold kernel.condKernel split_ifs <;> infer_instance -lemma compProd_fst_condKernel [hαβ : CountableOrStandardBorel α β] +lemma kernel.compProd_fst_condKernel [hαβ : CountableOrStandardBorel α β] (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : kernel.fst κ ⊗ₖ kernel.condKernel κ = κ := by unfold kernel.condKernel @@ -371,4 +368,87 @@ lemma compProd_fst_condKernel [hαβ : CountableOrStandardBorel α β] end CountableOrStandardBorel +section Lintegral + +variable [CountableOrStandardBorel α β] {κ : kernel α (β × Ω)} [IsFiniteKernel κ] + {f : β × Ω → ℝ≥0∞} + +lemma lintegral_condKernel_mem (a : α) {s : Set (β × Ω)} (hs : MeasurableSet s) : + ∫⁻ x, kernel.condKernel κ (a, x) {y | (x, y) ∈ s} ∂(kernel.fst κ a) = κ a s := by + conv_rhs => rw [← kernel.compProd_fst_condKernel κ] + simp_rw [kernel.compProd_apply _ _ _ hs] + +lemma set_lintegral_condKernel_eq_measure_prod (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set Ω} + (ht : MeasurableSet t) : + ∫⁻ b in s, kernel.condKernel κ (a, b) t ∂(kernel.fst κ a) = κ a (s ×ˢ t) := by + have : κ a (s ×ˢ t) = (kernel.fst κ ⊗ₖ kernel.condKernel κ) a (s ×ˢ t) := by + congr; exact (kernel.compProd_fst_condKernel κ).symm + rw [this, kernel.compProd_apply _ _ _ (hs.prod ht)] + classical + have : ∀ b, kernel.condKernel κ (a, b) {c | (b, c) ∈ s ×ˢ t} + = s.indicator (fun b ↦ kernel.condKernel κ (a, b) t) b := by + intro b + by_cases hb : b ∈ s <;> simp [hb] + simp_rw [this] + rw [lintegral_indicator _ hs] + +lemma lintegral_condKernel (hf : Measurable f) (a : α) : + ∫⁻ b, ∫⁻ ω, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) = ∫⁻ x, f x ∂(κ a) := by + conv_rhs => rw [← kernel.compProd_fst_condKernel κ] + rw [kernel.lintegral_compProd _ _ _ hf] + +lemma set_lintegral_condKernel (hf : Measurable f) (a : α) {s : Set β} + (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) : + ∫⁻ b in s, ∫⁻ ω in t, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) + = ∫⁻ x in s ×ˢ t, f x ∂(κ a) := by + conv_rhs => rw [← kernel.compProd_fst_condKernel κ] + rw [kernel.set_lintegral_compProd _ _ _ hf hs ht] + +lemma set_lintegral_condKernel_univ_right (hf : Measurable f) (a : α) {s : Set β} + (hs : MeasurableSet s) : + ∫⁻ b in s, ∫⁻ ω, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) + = ∫⁻ x in s ×ˢ univ, f x ∂(κ a) := by + rw [← set_lintegral_condKernel hf a hs MeasurableSet.univ]; simp_rw [Measure.restrict_univ] + +lemma set_lintegral_condKernel_univ_left (hf : Measurable f) (a : α) {t : Set Ω} + (ht : MeasurableSet t) : + ∫⁻ b, ∫⁻ ω in t, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) + = ∫⁻ x in univ ×ˢ t, f x ∂(κ a) := by + rw [← set_lintegral_condKernel hf a MeasurableSet.univ ht]; simp_rw [Measure.restrict_univ] + +end Lintegral + +section Integral + +variable [CountableOrStandardBorel α β] {κ : kernel α (β × Ω)} [IsFiniteKernel κ] + {E : Type*} {f : β × Ω → E} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] + +lemma integral_condKernel (a : α) (hf : Integrable f (κ a)) : + ∫ b, ∫ ω, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) = ∫ x, f x ∂(κ a) := by + conv_rhs => rw [← kernel.compProd_fst_condKernel κ] + rw [← kernel.compProd_fst_condKernel κ] at hf + rw [integral_compProd hf] + +lemma set_integral_condKernel (a : α) {s : Set β} (hs : MeasurableSet s) + {t : Set Ω} (ht : MeasurableSet t) (hf : IntegrableOn f (s ×ˢ t) (κ a)) : + ∫ b in s, ∫ ω in t, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) + = ∫ x in s ×ˢ t, f x ∂(κ a) := by + conv_rhs => rw [← kernel.compProd_fst_condKernel κ] + rw [← kernel.compProd_fst_condKernel κ] at hf + rw [set_integral_compProd hs ht hf] + +lemma set_integral_condKernel_univ_right (a : α) {s : Set β} (hs : MeasurableSet s) + (hf : IntegrableOn f (s ×ˢ univ) (κ a)) : + ∫ b in s, ∫ ω, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) + = ∫ x in s ×ˢ univ, f x ∂(κ a) := by + rw [← set_integral_condKernel a hs MeasurableSet.univ hf]; simp_rw [Measure.restrict_univ] + +lemma set_integral_condKernel_univ_left (a : α) {t : Set Ω} (ht : MeasurableSet t) + (hf : IntegrableOn f (univ ×ˢ t) (κ a)) : + ∫ b, ∫ ω in t, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) + = ∫ x in univ ×ˢ t, f x ∂(κ a) := by + rw [← set_integral_condKernel a MeasurableSet.univ ht hf]; simp_rw [Measure.restrict_univ] + +end Integral + end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index eed85a2596fcf..b9847b7832393 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -21,10 +21,12 @@ open scoped ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory -variable {α : Type*} {mα : MeasurableSpace α} +variable {α β Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] -variable {Ω : Type*} [MeasurableSpace Ω] [StandardBorelSpace Ω] - [Nonempty Ω] (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] +section Kernel + +variable [CountableOrStandardBorel α β] {ρ : kernel α (β × Ω)} [IsFiniteKernel ρ] /-! ### Uniqueness @@ -36,96 +38,213 @@ everywhere equal to the disintegration kernel of `ρ` when evaluated on a measur This theorem in the case of finite kernels is weaker than `eq_condKernel_of_measure_eq_compProd` which asserts that the kernels are equal almost everywhere and not just on a given measurable set. -/ -theorem eq_condKernel_of_measure_eq_compProd' (κ : kernel α Ω) [IsSFiniteKernel κ] +theorem eq_condKernel_of_measure_eq_compProd' {κ : kernel (α × β) Ω} [IsSFiniteKernel κ] + (hκ : ρ = kernel.fst ρ ⊗ₖ κ) {s : Set Ω} (hs : MeasurableSet s) (a : α) : + ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) s = kernel.condKernel ρ (a, x) s := by + refine ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite ?_ ?_ ?_ + · exact (kernel.measurable_coe κ hs).comp measurable_prod_mk_left + · exact (kernel.measurable_coe _ hs).comp measurable_prod_mk_left + intros t ht _ + conv_rhs => rw [set_lintegral_condKernel_eq_measure_prod _ ht hs, hκ] + simp only [kernel.compProd_apply _ _ _ (ht.prod hs), Set.mem_prod, ← lintegral_indicator _ ht] + congr with x + by_cases hx : x ∈ t + all_goals simp [hx] + +-- This lemma establishes uniqueness of the disintegration kernel on ℝ +lemma eq_condKernel_of_measure_eq_compProd_real {ρ : kernel α (β × ℝ)} [IsFiniteKernel ρ] + {κ : kernel (α × β) ℝ} [IsFiniteKernel κ] (hκ : ρ = kernel.fst ρ ⊗ₖ κ) (a : α) : + ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) = kernel.condKernel ρ (a, x) := by + have huniv : ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) Set.univ = kernel.condKernel ρ (a, x) Set.univ := + eq_condKernel_of_measure_eq_compProd' hκ MeasurableSet.univ a + suffices : ∀ᵐ x ∂(kernel.fst ρ a), + ∀ ⦃t⦄, MeasurableSet t → κ (a, x) t = kernel.condKernel ρ (a, x) t + · filter_upwards [this] with x hx + ext t ht; exact hx ht + apply MeasurableSpace.ae_induction_on_inter Real.borel_eq_generateFrom_Iic_rat + Real.isPiSystem_Iic_rat + · simp only [OuterMeasure.empty', Filter.eventually_true] + · simp only [iUnion_singleton_eq_range, mem_range, forall_exists_index, forall_apply_eq_imp_iff] + exact ae_all_iff.2 fun q => eq_condKernel_of_measure_eq_compProd' hκ measurableSet_Iic a + · filter_upwards [huniv] with x hxuniv t ht heq + rw [measure_compl ht <| measure_ne_top _ _, heq, hxuniv, measure_compl ht <| measure_ne_top _ _] + · refine ae_of_all _ (fun x f hdisj hf heq ↦ ?_) + rw [measure_iUnion hdisj hf, measure_iUnion hdisj hf] + exact tsum_congr heq + +/-- A finite kernel which satisfies the disintegration property is almost everywhere equal to the +disintegration kernel. -/ +theorem eq_condKernel_of_kernel_eq_compProd (κ : kernel (α × β) Ω) [IsFiniteKernel κ] + (hκ : ρ = kernel.fst ρ ⊗ₖ κ) (a : α) : + ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) = kernel.condKernel ρ (a, x) := by +-- The idea is to transporting the question to `ℝ` from `Ω` using `exists_measurableEmbedding_real` +-- and then constructing a measure on `α × ℝ` using the obtained measurable embedding + let f := measurableEmbedding_real Ω + let hf := measurableEmbedding_measurableEmbedding_real Ω + set ρ' : kernel α (β × ℝ) := kernel.map ρ (Prod.map id f) + (measurable_id.prod_map hf.measurable) with hρ'def + have hρ' : kernel.fst ρ' = kernel.fst ρ + · ext a s hs + rw [hρ'def, kernel.fst_apply' _ _ hs, kernel.fst_apply' _ _ hs, kernel.map_apply'] + exacts [rfl, measurable_fst hs] + have hρ'' : ∀ᵐ x ∂(kernel.fst ρ a), + kernel.map κ f hf.measurable (a, x) = kernel.condKernel ρ' (a, x) := by + rw [← hρ'] + refine eq_condKernel_of_measure_eq_compProd_real ?_ a + ext a s hs + conv_lhs => rw [hρ'def, hκ] + rw [kernel.map_apply' _ _ _ hs, kernel.compProd_apply _ _ _ hs, kernel.compProd_apply] + swap; · exact measurable_id.prod_map hf.measurable hs + congr with b + · congr with a t ht + rw [kernel.fst_apply' _ _ ht, kernel.fst_apply' _ _ ht, kernel.map_apply'] + exacts [rfl, measurable_fst ht] + · rw [kernel.map_apply'] + exacts [rfl, measurable_prod_mk_left hs] + suffices ∀ᵐ x ∂(kernel.fst ρ a), ∀ s, MeasurableSet s → + kernel.condKernel ρ' (a, x) s = kernel.condKernel ρ (a, x) (f ⁻¹' s) by + filter_upwards [hρ'', this] with x hx h + rw [kernel.map_apply] at hx + ext s hs + rw [← Set.preimage_image_eq s hf.injective, + ← Measure.map_apply hf.measurable <| hf.measurableSet_image.2 hs, hx, + h _ <| hf.measurableSet_image.2 hs] + suffices : ρ' = (kernel.fst ρ ⊗ₖ (kernel.map (kernel.condKernel ρ) f hf.measurable)) + · rw [← hρ'] at this + have heq := eq_condKernel_of_measure_eq_compProd_real this a + rw [hρ'] at heq + filter_upwards [heq] with x hx s hs + rw [← hx, kernel.map_apply, Measure.map_apply hf.measurable hs] + ext a s hs + rw [hρ'def, kernel.map_apply' _ _ _ hs, kernel.compProd_apply _ _ _ hs, + ← lintegral_condKernel_mem] + swap; · exact measurable_id.prod_map hf.measurable hs + congr with b + rw [kernel.map_apply'] + exacts [rfl, measurable_prod_mk_left hs] + +end Kernel + +section Measure + +variable {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] + +/-! ### Uniqueness + +This section will prove that the conditional kernel is unique almost everywhere. -/ + +/-- A s-finite kernel which satisfy the disintegration property of the given measure `ρ` is almost +everywhere equal to the disintegration kernel of `ρ` when evaluated on a measurable set. + +This theorem in the case of finite kernels is weaker than `eq_condKernel_of_measure_eq_compProd` +which asserts that the kernels are equal almost everywhere and not just on a given measurable +set. -/ +theorem Measure.eq_condKernel_of_measure_eq_compProd' (κ : kernel α Ω) [IsSFiniteKernel κ] (hκ : ρ = ρ.fst ⊗ₘ κ) {s : Set Ω} (hs : MeasurableSet s) : ∀ᵐ x ∂ρ.fst, κ x s = ρ.condKernel x s := by - refine' ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite - (kernel.measurable_coe κ hs) (kernel.measurable_coe ρ.condKernel hs) _ - intros t ht _ + refine ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite + (kernel.measurable_coe κ hs) (kernel.measurable_coe ρ.condKernel hs) (fun t ht _ ↦ ?_) conv_rhs => rw [set_lintegral_condKernel_eq_measure_prod _ ht hs, hκ] simp only [Measure.compProd_apply (ht.prod hs), Set.mem_prod, ← lintegral_indicator _ ht] - congr; ext x + congr with x by_cases hx : x ∈ t all_goals simp [hx] -- This lemma establishes uniqueness of the disintegration kernel on ℝ -lemma eq_condKernel_of_measure_eq_compProd_real (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] +lemma Measure.eq_condKernel_of_measure_eq_compProd_real {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] (κ : kernel α ℝ) [IsFiniteKernel κ] (hκ : ρ = ρ.fst ⊗ₘ κ) : ∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by have huniv : ∀ᵐ x ∂ρ.fst, κ x Set.univ = ρ.condKernel x Set.univ := - eq_condKernel_of_measure_eq_compProd' ρ κ hκ MeasurableSet.univ + eq_condKernel_of_measure_eq_compProd' κ hκ MeasurableSet.univ suffices : ∀ᵐ x ∂ρ.fst, ∀ ⦃t⦄, MeasurableSet t → κ x t = ρ.condKernel x t · filter_upwards [this] with x hx ext t ht; exact hx ht apply MeasurableSpace.ae_induction_on_inter Real.borel_eq_generateFrom_Iic_rat Real.isPiSystem_Iic_rat - · simp only [OuterMeasure.empty', Filter.eventually_true] + · simp · simp only [iUnion_singleton_eq_range, mem_range, forall_exists_index, forall_apply_eq_imp_iff] - exact ae_all_iff.2 fun q => eq_condKernel_of_measure_eq_compProd' ρ κ hκ measurableSet_Iic + exact ae_all_iff.2 fun q ↦ eq_condKernel_of_measure_eq_compProd' κ hκ measurableSet_Iic · filter_upwards [huniv] with x hxuniv t ht heq rw [measure_compl ht <| measure_ne_top _ _, heq, hxuniv, measure_compl ht <| measure_ne_top _ _] - · refine' ae_of_all _ (fun x f hdisj hf heq => _) + · refine ae_of_all _ (fun x f hdisj hf heq ↦ ?_) rw [measure_iUnion hdisj hf, measure_iUnion hdisj hf] exact tsum_congr heq /-- A finite kernel which satisfies the disintegration property is almost everywhere equal to the disintegration kernel. -/ -theorem eq_condKernel_of_measure_eq_compProd (κ : kernel α Ω) [IsFiniteKernel κ] +theorem Measure.eq_condKernel_of_measure_eq_compProd (κ : kernel α Ω) [IsFiniteKernel κ] (hκ : ρ = ρ.fst ⊗ₘ κ) : ∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by --- The idea is to transporting the question to `ℝ` from `Ω` using `exists_measurableEmbedding_real` --- and then constructing a measure on `α × ℝ` using the obtained measurable embedding - obtain ⟨f, hf⟩ := exists_measurableEmbedding_real Ω + -- The idea is to transporting the question to `ℝ` from `Ω` using + -- `exists_measurableEmbedding_real` and then constructing a measure on `α × ℝ` using + -- the obtained measurable embedding + let f := measurableEmbedding_real Ω + let hf := measurableEmbedding_measurableEmbedding_real Ω set ρ' : Measure (α × ℝ) := ρ.map (Prod.map id f) with hρ'def have hρ' : ρ'.fst = ρ.fst · ext s hs rw [hρ'def, Measure.fst_apply, Measure.fst_apply, Measure.map_apply] exacts [rfl, Measurable.prod measurable_fst <| hf.measurable.comp measurable_snd, measurable_fst hs, hs, hs] - have hρ'' : ∀ᵐ x ∂ρ'.fst, kernel.map κ f hf.measurable x = ρ'.condKernel x - · refine' eq_condKernel_of_measure_eq_compProd_real ρ' (kernel.map κ f hf.measurable) _ + have hρ'' : ∀ᵐ x ∂ρ.fst, kernel.map κ f hf.measurable x = ρ'.condKernel x + · rw [← hρ'] + refine eq_condKernel_of_measure_eq_compProd_real (kernel.map κ f hf.measurable) ?_ ext s hs - simp only [Measure.map_apply (measurable_id.prod_map hf.measurable) hs] - conv_lhs => congr; rw [hκ] - rw [Measure.compProd_apply hs, Measure.compProd_apply - (measurable_id.prod_map hf.measurable hs), (_ : (ρ.map (Prod.map id f)).fst = ρ.fst)] - · congr - ext x - simp only [Set.mem_preimage, Prod_map, id_eq, kernel.prodMkLeft_apply, kernel.map_apply] - rw [Measure.map_apply hf.measurable] - · rfl - · exact measurable_prod_mk_left hs - · rw [Measure.fst_map_prod_mk] - · simp only [Prod_map, id_eq]; rfl - · exact (hf.measurable.comp measurable_snd) - rw [hρ'] at hρ'' - suffices : ∀ᵐ x ∂ρ.fst, ∀ s, MeasurableSet s → - ((ρ.map (Prod.map id f)).condKernel x) s = (ρ.condKernel x) (f ⁻¹' s) + conv_lhs => rw [hρ'def, hκ] + rw [Measure.map_apply (measurable_id.prod_map hf.measurable) hs, hρ', + Measure.compProd_apply hs, Measure.compProd_apply (measurable_id.prod_map hf.measurable hs)] + congr with a + rw [kernel.map_apply'] + · rfl + · exact measurable_prod_mk_left hs + suffices : ∀ᵐ x ∂ρ.fst, ∀ s, MeasurableSet s → ρ'.condKernel x s = ρ.condKernel x (f ⁻¹' s) · filter_upwards [hρ'', this] with x hx h rw [kernel.map_apply] at hx ext s hs rw [← Set.preimage_image_eq s hf.injective, ← Measure.map_apply hf.measurable <| hf.measurableSet_image.2 hs, hx, h _ <| hf.measurableSet_image.2 hs] - have hprod : (ρ.map (Prod.map id f)).fst = ρ.fst - · ext s hs - rw [Measure.fst_apply hs, - Measure.map_apply (measurable_id.prod_map hf.measurable) (measurable_fst hs), - ← Set.preimage_comp, Measure.fst_apply hs] - rfl - suffices : ρ.map (Prod.map id f) = - ((ρ.map (Prod.map id f)).fst ⊗ₘ (kernel.map (Measure.condKernel ρ) f hf.measurable)) - · have heq := eq_condKernel_of_measure_eq_compProd_real _ _ this - rw [hprod] at heq + suffices : ρ.map (Prod.map id f) = (ρ.fst ⊗ₘ (kernel.map ρ.condKernel f hf.measurable)) + · rw [← hρ'] at this + have heq := eq_condKernel_of_measure_eq_compProd_real _ this + rw [hρ'] at heq filter_upwards [heq] with x hx s hs rw [← hx, kernel.map_apply, Measure.map_apply hf.measurable hs] ext s hs - simp only [hprod, Measure.compProd_apply hs, kernel.map_apply _ hf.measurable] - rw [Measure.map_apply (measurable_id.prod_map hf.measurable) hs, ← lintegral_condKernel_mem] - · simp only [mem_preimage, Prod_map, id_eq] - congr with a - rw [Measure.map_apply hf.measurable (measurable_prod_mk_left hs)] - rfl + conv_lhs => rw [← ρ.compProd_fst_condKernel] + rw [Measure.compProd_apply hs, Measure.map_apply (measurable_id.prod_map hf.measurable) hs, + Measure.compProd_apply] + · congr with a + rw [kernel.map_apply'] + · rfl + · exact measurable_prod_mk_left hs · exact measurable_id.prod_map hf.measurable hs +end Measure + +section KernelAndMeasure + +lemma kernel.condKernel_apply_eq_condKernel [CountableOrStandardBorel α β] + (κ : kernel α (β × Ω)) [IsFiniteKernel κ] (a : α) : + ∀ᵐ b ∂(kernel.fst κ a), kernel.condKernel κ (a, b) = (κ a).condKernel b := by + have : κ a = (κ a).fst + ⊗ₘ kernel.comap (kernel.condKernel κ) (fun b ↦ (a, b)) measurable_prod_mk_left := by + ext s hs + conv_lhs => rw [← compProd_fst_condKernel κ] + rw [Measure.compProd_apply hs, kernel.compProd_apply _ _ _ hs] + rfl + have h := Measure.eq_condKernel_of_measure_eq_compProd _ this + rw [kernel.fst_apply] + filter_upwards [h] with b hb + rw [← hb, kernel.comap_apply] + +lemma condKernel_const [CountableOrStandardBorel α β] (ρ : Measure (β × Ω)) [IsFiniteMeasure ρ] + (a : α) : + ∀ᵐ b ∂ρ.fst, kernel.condKernel (kernel.const α ρ) (a, b) = ρ.condKernel b := by + have h := kernel.condKernel_apply_eq_condKernel (kernel.const α ρ) a + simp_rw [kernel.fst_apply, kernel.const_apply] at h + filter_upwards [h] with b hb using hb + +end KernelAndMeasure + end ProbabilityTheory From 85f49529f7c8f00e85d630827e66720f96d5cc78 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 15 Feb 2024 12:15:33 +0100 Subject: [PATCH 016/129] add Integral file --- Mathlib/Probability/Kernel/CondCdf.lean | 12 +- .../Probability/Kernel/Disintegration.lean | 76 +----- .../Kernel/Disintegration/Basic.lean | 83 ------- .../Kernel/Disintegration/Integral.lean | 226 ++++++++++++++++++ .../Kernel/Disintegration/KernelCDFBorel.lean | 12 +- 5 files changed, 235 insertions(+), 174 deletions(-) create mode 100644 Mathlib/Probability/Kernel/Disintegration/Integral.lean diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index cb378b2f088f6..f0f511234f010 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -271,16 +271,8 @@ theorem tendsto_preCDF_atTop_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] have h_le_one := preCDF_le_one ρ -- `preCDF` has a limit a.e. have h_exists : ∀ᵐ a ∂ρ.fst, ∃ l, Tendsto (fun r => preCDF ρ r a) atTop (𝓝 l) := by - filter_upwards [h_mono, h_le_one] with a ha_mono ha_le_one - have h_tendsto : - Tendsto (fun r => preCDF ρ r a) atTop atTop ∨ - ∃ l, Tendsto (fun r => preCDF ρ r a) atTop (𝓝 l) := - tendsto_of_monotone ha_mono - cases' h_tendsto with h_absurd h_tendsto - · rw [Monotone.tendsto_atTop_atTop_iff ha_mono] at h_absurd - obtain ⟨r, hr⟩ := h_absurd 2 - exact absurd (hr.trans (ha_le_one r)) ENNReal.one_lt_two.not_le - · exact h_tendsto + filter_upwards [h_mono] with a ha_mono + exact ⟨_, tendsto_atTop_iSup ha_mono⟩ classical -- let `F` be the pointwise limit of `preCDF` where it exists, and 0 elsewhere. let F : α → ℝ≥0∞ := fun a => diff --git a/Mathlib/Probability/Kernel/Disintegration.lean b/Mathlib/Probability/Kernel/Disintegration.lean index 49b6efed11ce8..1a37d6ed08810 100644 --- a/Mathlib/Probability/Kernel/Disintegration.lean +++ b/Mathlib/Probability/Kernel/Disintegration.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne, Kexing Ying -/ -import Mathlib.Probability.Kernel.Disintegration.Basic +import Mathlib.Probability.Kernel.Disintegration.Integral #align_import probability.kernel.disintegration from "leanprover-community/mathlib"@"6315581f5650ffa2fbdbbbedc41243c8d7070981" @@ -126,52 +126,6 @@ lemma condKernel_apply_of_ne_zero [MeasurableSingletonClass α] (measurableEmbedding_prod_mk_left x).comap_apply] simp [this, (measurableEmbedding_prod_mk_left x).comap_apply, hx] -theorem Measure.lintegral_condKernel_mem {s : Set (α × Ω)} (hs : MeasurableSet s) : - ∫⁻ a, ρ.condKernel a {x | (a, x) ∈ s} ∂ρ.fst = ρ s := by - conv_rhs => rw [measure_eq_compProd ρ] - simp_rw [Measure.compProd_apply hs] - rfl -#align probability_theory.lintegral_cond_kernel_mem ProbabilityTheory.Measure.lintegral_condKernel_mem - -theorem Measure.set_lintegral_condKernel_eq_measure_prod {s : Set α} (hs : MeasurableSet s) {t : Set Ω} - (ht : MeasurableSet t) : ∫⁻ a in s, ρ.condKernel a t ∂ρ.fst = ρ (s ×ˢ t) := by - have : ρ (s ×ˢ t) = (ρ.fst ⊗ₘ ρ.condKernel) (s ×ˢ t) := by - congr; exact measure_eq_compProd ρ - rw [this, Measure.compProd_apply (hs.prod ht)] - classical - have : ∀ a, ρ.condKernel a (Prod.mk a ⁻¹' s ×ˢ t) - = s.indicator (fun a ↦ ρ.condKernel a t) a := by - intro a - by_cases ha : a ∈ s <;> simp [ha] - simp_rw [this] - rw [lintegral_indicator _ hs] -#align probability_theory.set_lintegral_cond_kernel_eq_measure_prod ProbabilityTheory.Measure.set_lintegral_condKernel_eq_measure_prod - -theorem Measure.lintegral_condKernel {f : α × Ω → ℝ≥0∞} (hf : Measurable f) : - ∫⁻ a, ∫⁻ ω, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫⁻ x, f x ∂ρ := by - conv_rhs => rw [measure_eq_compProd ρ] - rw [Measure.lintegral_compProd hf] -#align probability_theory.lintegral_cond_kernel ProbabilityTheory.lintegral_condKernel - -theorem Measure.set_lintegral_condKernel {f : α × Ω → ℝ≥0∞} (hf : Measurable f) {s : Set α} - (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) : - ∫⁻ a in s, ∫⁻ ω in t, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫⁻ x in s ×ˢ t, f x ∂ρ := by - conv_rhs => rw [measure_eq_compProd ρ] - rw [Measure.set_lintegral_compProd hf hs ht] -#align probability_theory.set_lintegral_cond_kernel ProbabilityTheory.set_lintegral_condKernel - -theorem Measure.set_lintegral_condKernel_univ_right {f : α × Ω → ℝ≥0∞} (hf : Measurable f) {s : Set α} - (hs : MeasurableSet s) : - ∫⁻ a in s, ∫⁻ ω, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫⁻ x in s ×ˢ univ, f x ∂ρ := by - rw [← set_lintegral_condKernel ρ hf hs MeasurableSet.univ]; simp_rw [Measure.restrict_univ] -#align probability_theory.set_lintegral_cond_kernel_univ_right ProbabilityTheory.set_lintegral_condKernel_univ_right - -theorem Measure.set_lintegral_condKernel_univ_left {f : α × Ω → ℝ≥0∞} (hf : Measurable f) {t : Set Ω} - (ht : MeasurableSet t) : - ∫⁻ a, ∫⁻ ω in t, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫⁻ x in univ ×ˢ t, f x ∂ρ := by - rw [← set_lintegral_condKernel ρ hf MeasurableSet.univ ht]; simp_rw [Measure.restrict_univ] -#align probability_theory.set_lintegral_cond_kernel_univ_left ProbabilityTheory.set_lintegral_condKernel_univ_left - section IntegralCondKernel variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] @@ -183,34 +137,6 @@ theorem _root_.MeasureTheory.AEStronglyMeasurable.integral_condKernel {ρ : Meas exact AEStronglyMeasurable.integral_kernel_compProd hf #align measure_theory.ae_strongly_measurable.integral_cond_kernel MeasureTheory.AEStronglyMeasurable.integral_condKernel -theorem Measure.integral_condKernel {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} - (hf : Integrable f ρ) : ∫ a, ∫ x, f (a, x) ∂ρ.condKernel a ∂ρ.fst = ∫ ω, f ω ∂ρ := by - conv_rhs => rw [measure_eq_compProd ρ] - have hf': Integrable f (ρ.fst ⊗ₘ ρ.condKernel) := by rwa [measure_eq_compProd ρ] at hf - rw [Measure.integral_compProd hf'] -#align probability_theory.integral_cond_kernel ProbabilityTheory.integral_condKernel - -theorem Measure.set_integral_condKernel {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} - {s : Set α} (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) - (hf : IntegrableOn f (s ×ˢ t) ρ) : - ∫ a in s, ∫ ω in t, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫ x in s ×ˢ t, f x ∂ρ := by - conv_rhs => rw [measure_eq_compProd ρ] - rw [Measure.set_integral_compProd hs ht] - rwa [measure_eq_compProd ρ] at hf -#align probability_theory.set_integral_cond_kernel ProbabilityTheory.set_integral_condKernel - -theorem Measure.set_integral_condKernel_univ_right {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} - {s : Set α} (hs : MeasurableSet s) (hf : IntegrableOn f (s ×ˢ univ) ρ) : - ∫ a in s, ∫ ω, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫ x in s ×ˢ univ, f x ∂ρ := by - rw [← set_integral_condKernel hs MeasurableSet.univ hf]; simp_rw [Measure.restrict_univ] -#align probability_theory.set_integral_cond_kernel_univ_right ProbabilityTheory.set_integral_condKernel_univ_right - -theorem Measure.set_integral_condKernel_univ_left {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} - {t : Set Ω} (ht : MeasurableSet t) (hf : IntegrableOn f (univ ×ˢ t) ρ) : - ∫ a, ∫ ω in t, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫ x in univ ×ˢ t, f x ∂ρ := by - rw [← set_integral_condKernel MeasurableSet.univ ht hf]; simp_rw [Measure.restrict_univ] -#align probability_theory.set_integral_cond_kernel_univ_left ProbabilityTheory.set_integral_condKernel_univ_left - end IntegralCondKernel end StandardBorel diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 90cad62e6d71c..14452483e8356 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -368,87 +368,4 @@ lemma kernel.compProd_fst_condKernel [hαβ : CountableOrStandardBorel α β] end CountableOrStandardBorel -section Lintegral - -variable [CountableOrStandardBorel α β] {κ : kernel α (β × Ω)} [IsFiniteKernel κ] - {f : β × Ω → ℝ≥0∞} - -lemma lintegral_condKernel_mem (a : α) {s : Set (β × Ω)} (hs : MeasurableSet s) : - ∫⁻ x, kernel.condKernel κ (a, x) {y | (x, y) ∈ s} ∂(kernel.fst κ a) = κ a s := by - conv_rhs => rw [← kernel.compProd_fst_condKernel κ] - simp_rw [kernel.compProd_apply _ _ _ hs] - -lemma set_lintegral_condKernel_eq_measure_prod (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set Ω} - (ht : MeasurableSet t) : - ∫⁻ b in s, kernel.condKernel κ (a, b) t ∂(kernel.fst κ a) = κ a (s ×ˢ t) := by - have : κ a (s ×ˢ t) = (kernel.fst κ ⊗ₖ kernel.condKernel κ) a (s ×ˢ t) := by - congr; exact (kernel.compProd_fst_condKernel κ).symm - rw [this, kernel.compProd_apply _ _ _ (hs.prod ht)] - classical - have : ∀ b, kernel.condKernel κ (a, b) {c | (b, c) ∈ s ×ˢ t} - = s.indicator (fun b ↦ kernel.condKernel κ (a, b) t) b := by - intro b - by_cases hb : b ∈ s <;> simp [hb] - simp_rw [this] - rw [lintegral_indicator _ hs] - -lemma lintegral_condKernel (hf : Measurable f) (a : α) : - ∫⁻ b, ∫⁻ ω, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) = ∫⁻ x, f x ∂(κ a) := by - conv_rhs => rw [← kernel.compProd_fst_condKernel κ] - rw [kernel.lintegral_compProd _ _ _ hf] - -lemma set_lintegral_condKernel (hf : Measurable f) (a : α) {s : Set β} - (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) : - ∫⁻ b in s, ∫⁻ ω in t, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) - = ∫⁻ x in s ×ˢ t, f x ∂(κ a) := by - conv_rhs => rw [← kernel.compProd_fst_condKernel κ] - rw [kernel.set_lintegral_compProd _ _ _ hf hs ht] - -lemma set_lintegral_condKernel_univ_right (hf : Measurable f) (a : α) {s : Set β} - (hs : MeasurableSet s) : - ∫⁻ b in s, ∫⁻ ω, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) - = ∫⁻ x in s ×ˢ univ, f x ∂(κ a) := by - rw [← set_lintegral_condKernel hf a hs MeasurableSet.univ]; simp_rw [Measure.restrict_univ] - -lemma set_lintegral_condKernel_univ_left (hf : Measurable f) (a : α) {t : Set Ω} - (ht : MeasurableSet t) : - ∫⁻ b, ∫⁻ ω in t, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) - = ∫⁻ x in univ ×ˢ t, f x ∂(κ a) := by - rw [← set_lintegral_condKernel hf a MeasurableSet.univ ht]; simp_rw [Measure.restrict_univ] - -end Lintegral - -section Integral - -variable [CountableOrStandardBorel α β] {κ : kernel α (β × Ω)} [IsFiniteKernel κ] - {E : Type*} {f : β × Ω → E} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] - -lemma integral_condKernel (a : α) (hf : Integrable f (κ a)) : - ∫ b, ∫ ω, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) = ∫ x, f x ∂(κ a) := by - conv_rhs => rw [← kernel.compProd_fst_condKernel κ] - rw [← kernel.compProd_fst_condKernel κ] at hf - rw [integral_compProd hf] - -lemma set_integral_condKernel (a : α) {s : Set β} (hs : MeasurableSet s) - {t : Set Ω} (ht : MeasurableSet t) (hf : IntegrableOn f (s ×ˢ t) (κ a)) : - ∫ b in s, ∫ ω in t, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) - = ∫ x in s ×ˢ t, f x ∂(κ a) := by - conv_rhs => rw [← kernel.compProd_fst_condKernel κ] - rw [← kernel.compProd_fst_condKernel κ] at hf - rw [set_integral_compProd hs ht hf] - -lemma set_integral_condKernel_univ_right (a : α) {s : Set β} (hs : MeasurableSet s) - (hf : IntegrableOn f (s ×ˢ univ) (κ a)) : - ∫ b in s, ∫ ω, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) - = ∫ x in s ×ˢ univ, f x ∂(κ a) := by - rw [← set_integral_condKernel a hs MeasurableSet.univ hf]; simp_rw [Measure.restrict_univ] - -lemma set_integral_condKernel_univ_left (a : α) {t : Set Ω} (ht : MeasurableSet t) - (hf : IntegrableOn f (univ ×ˢ t) (κ a)) : - ∫ b, ∫ ω in t, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) - = ∫ x in univ ×ˢ t, f x ∂(κ a) := by - rw [← set_integral_condKernel a MeasurableSet.univ ht hf]; simp_rw [Measure.restrict_univ] - -end Integral - end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean new file mode 100644 index 0000000000000..e2236ea775c20 --- /dev/null +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -0,0 +1,226 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Probability.Kernel.Disintegration.Basic + +/-! +# Integral + +## Main definitions + +* `FooBar` + +## Main statements + +* `fooBar_unique` + +## Notation + + + +## Implementation details + + + +## References + +* [F. Bar, *Quuxes*][bibkey] + +## Tags + +Foobars, barfoos +-/ + +open MeasureTheory ProbabilityTheory + +open scoped ENNReal + +namespace ProbabilityTheory + +variable {α β Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] + +section Lintegral + +variable [CountableOrStandardBorel α β] {κ : kernel α (β × Ω)} [IsFiniteKernel κ] + {f : β × Ω → ℝ≥0∞} + +lemma lintegral_condKernel_mem (a : α) {s : Set (β × Ω)} (hs : MeasurableSet s) : + ∫⁻ x, kernel.condKernel κ (a, x) {y | (x, y) ∈ s} ∂(kernel.fst κ a) = κ a s := by + conv_rhs => rw [← kernel.compProd_fst_condKernel κ] + simp_rw [kernel.compProd_apply _ _ _ hs] + +lemma set_lintegral_condKernel_eq_measure_prod (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set Ω} + (ht : MeasurableSet t) : + ∫⁻ b in s, kernel.condKernel κ (a, b) t ∂(kernel.fst κ a) = κ a (s ×ˢ t) := by + have : κ a (s ×ˢ t) = (kernel.fst κ ⊗ₖ kernel.condKernel κ) a (s ×ˢ t) := by + congr; exact (kernel.compProd_fst_condKernel κ).symm + rw [this, kernel.compProd_apply _ _ _ (hs.prod ht)] + classical + have : ∀ b, kernel.condKernel κ (a, b) {c | (b, c) ∈ s ×ˢ t} + = s.indicator (fun b ↦ kernel.condKernel κ (a, b) t) b := by + intro b + by_cases hb : b ∈ s <;> simp [hb] + simp_rw [this] + rw [lintegral_indicator _ hs] + +lemma lintegral_condKernel (hf : Measurable f) (a : α) : + ∫⁻ b, ∫⁻ ω, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) = ∫⁻ x, f x ∂(κ a) := by + conv_rhs => rw [← kernel.compProd_fst_condKernel κ] + rw [kernel.lintegral_compProd _ _ _ hf] + +lemma set_lintegral_condKernel (hf : Measurable f) (a : α) {s : Set β} + (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) : + ∫⁻ b in s, ∫⁻ ω in t, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) + = ∫⁻ x in s ×ˢ t, f x ∂(κ a) := by + conv_rhs => rw [← kernel.compProd_fst_condKernel κ] + rw [kernel.set_lintegral_compProd _ _ _ hf hs ht] + +lemma set_lintegral_condKernel_univ_right (hf : Measurable f) (a : α) {s : Set β} + (hs : MeasurableSet s) : + ∫⁻ b in s, ∫⁻ ω, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) + = ∫⁻ x in s ×ˢ Set.univ, f x ∂(κ a) := by + rw [← set_lintegral_condKernel hf a hs MeasurableSet.univ]; simp_rw [Measure.restrict_univ] + +lemma set_lintegral_condKernel_univ_left (hf : Measurable f) (a : α) {t : Set Ω} + (ht : MeasurableSet t) : + ∫⁻ b, ∫⁻ ω in t, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) + = ∫⁻ x in Set.univ ×ˢ t, f x ∂(κ a) := by + rw [← set_lintegral_condKernel hf a MeasurableSet.univ ht]; simp_rw [Measure.restrict_univ] + +end Lintegral + +section Integral + +variable [CountableOrStandardBorel α β] {κ : kernel α (β × Ω)} [IsFiniteKernel κ] + {E : Type*} {f : β × Ω → E} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] + +lemma integral_condKernel (a : α) (hf : Integrable f (κ a)) : + ∫ b, ∫ ω, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) = ∫ x, f x ∂(κ a) := by + conv_rhs => rw [← kernel.compProd_fst_condKernel κ] + rw [← kernel.compProd_fst_condKernel κ] at hf + rw [integral_compProd hf] + +lemma set_integral_condKernel (a : α) {s : Set β} (hs : MeasurableSet s) + {t : Set Ω} (ht : MeasurableSet t) (hf : IntegrableOn f (s ×ˢ t) (κ a)) : + ∫ b in s, ∫ ω in t, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) + = ∫ x in s ×ˢ t, f x ∂(κ a) := by + conv_rhs => rw [← kernel.compProd_fst_condKernel κ] + rw [← kernel.compProd_fst_condKernel κ] at hf + rw [set_integral_compProd hs ht hf] + +lemma set_integral_condKernel_univ_right (a : α) {s : Set β} (hs : MeasurableSet s) + (hf : IntegrableOn f (s ×ˢ Set.univ) (κ a)) : + ∫ b in s, ∫ ω, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) + = ∫ x in s ×ˢ Set.univ, f x ∂(κ a) := by + rw [← set_integral_condKernel a hs MeasurableSet.univ hf]; simp_rw [Measure.restrict_univ] + +lemma set_integral_condKernel_univ_left (a : α) {t : Set Ω} (ht : MeasurableSet t) + (hf : IntegrableOn f (Set.univ ×ˢ t) (κ a)) : + ∫ b, ∫ ω in t, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) + = ∫ x in Set.univ ×ˢ t, f x ∂(κ a) := by + rw [← set_integral_condKernel a MeasurableSet.univ ht hf]; simp_rw [Measure.restrict_univ] + +end Integral + +end ProbabilityTheory + +namespace MeasureTheory.Measure + +variable {α β Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] + +section Lintegral + +variable [CountableOrStandardBorel α β] {ρ : Measure (β × Ω)} [IsFiniteMeasure ρ] + {f : β × Ω → ℝ≥0∞} + +lemma lintegral_condKernel_mem {s : Set (β × Ω)} (hs : MeasurableSet s) : + ∫⁻ x, ρ.condKernel x {y | (x, y) ∈ s} ∂ρ.fst = ρ s := by + conv_rhs => rw [← compProd_fst_condKernel ρ] + simp_rw [compProd_apply hs] + rfl +#align probability_theory.lintegral_cond_kernel_mem MeasureTheory.Measure.lintegral_condKernel_mem + +lemma set_lintegral_condKernel_eq_measure_prod {s : Set β} (hs : MeasurableSet s) {t : Set Ω} + (ht : MeasurableSet t) : + ∫⁻ b in s, ρ.condKernel b t ∂ρ.fst = ρ (s ×ˢ t) := by + have : ρ (s ×ˢ t) = (ρ.fst ⊗ₘ ρ.condKernel) (s ×ˢ t) := by + congr; exact (compProd_fst_condKernel ρ).symm + rw [this, compProd_apply (hs.prod ht)] + classical + have : ∀ b, ρ.condKernel b (Prod.mk b ⁻¹' s ×ˢ t) + = s.indicator (fun b ↦ ρ.condKernel b t) b := by + intro b + by_cases hb : b ∈ s <;> simp [hb] + simp_rw [this] + rw [lintegral_indicator _ hs] +#align probability_theory.set_lintegral_cond_kernel_eq_measure_prod MeasureTheory.Measure.set_lintegral_condKernel_eq_measure_prod + +lemma lintegral_condKernel (hf : Measurable f) : + ∫⁻ b, ∫⁻ ω, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫⁻ x, f x ∂ρ := by + conv_rhs => rw [← compProd_fst_condKernel ρ] + rw [lintegral_compProd hf] +#align probability_theory.lintegral_cond_kernel MeasureTheory.Measure.lintegral_condKernel + +lemma set_lintegral_condKernel (hf : Measurable f) {s : Set β} + (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) : + ∫⁻ b in s, ∫⁻ ω in t, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst + = ∫⁻ x in s ×ˢ t, f x ∂ρ := by + conv_rhs => rw [← compProd_fst_condKernel ρ] + rw [set_lintegral_compProd hf hs ht] +#align probability_theory.set_lintegral_cond_kernel MeasureTheory.Measure.set_lintegral_condKernel + +lemma set_lintegral_condKernel_univ_right (hf : Measurable f) {s : Set β} + (hs : MeasurableSet s) : + ∫⁻ b in s, ∫⁻ ω, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst + = ∫⁻ x in s ×ˢ Set.univ, f x ∂ρ := by + rw [← set_lintegral_condKernel hf hs MeasurableSet.univ]; simp_rw [Measure.restrict_univ] +#align probability_theory.set_lintegral_cond_kernel_univ_right MeasureTheory.Measure.set_lintegral_condKernel_univ_right + +lemma set_lintegral_condKernel_univ_left (hf : Measurable f) {t : Set Ω} + (ht : MeasurableSet t) : + ∫⁻ b, ∫⁻ ω in t, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst + = ∫⁻ x in Set.univ ×ˢ t, f x ∂ρ := by + rw [← set_lintegral_condKernel hf MeasurableSet.univ ht]; simp_rw [Measure.restrict_univ] +#align probability_theory.set_lintegral_cond_kernel_univ_left MeasureTheory.Measure.set_lintegral_condKernel_univ_left + +end Lintegral + +section Integral + +variable {ρ : Measure (β × Ω)} [IsFiniteMeasure ρ] + {E : Type*} {f : β × Ω → E} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] + +lemma integral_condKernel (hf : Integrable f ρ) : + ∫ b, ∫ ω, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫ x, f x ∂ρ := by + conv_rhs => rw [← compProd_fst_condKernel ρ] + rw [← compProd_fst_condKernel ρ] at hf + rw [integral_compProd hf] +#align probability_theory.integral_cond_kernel MeasureTheory.Measure.integral_condKernel + +lemma set_integral_condKernel {s : Set β} (hs : MeasurableSet s) + {t : Set Ω} (ht : MeasurableSet t) (hf : IntegrableOn f (s ×ˢ t) ρ) : + ∫ b in s, ∫ ω in t, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫ x in s ×ˢ t, f x ∂ρ := by + conv_rhs => rw [← compProd_fst_condKernel ρ] + rw [← compProd_fst_condKernel ρ] at hf + rw [set_integral_compProd hs ht hf] +#align probability_theory.set_integral_cond_kernel MeasureTheory.Measure.set_integral_condKernel + +lemma set_integral_condKernel_univ_right {s : Set β} (hs : MeasurableSet s) + (hf : IntegrableOn f (s ×ˢ Set.univ) ρ) : + ∫ b in s, ∫ ω, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫ x in s ×ˢ Set.univ, f x ∂ρ := by + rw [← set_integral_condKernel hs MeasurableSet.univ hf]; simp_rw [Measure.restrict_univ] +#align probability_theory.set_integral_cond_kernel_univ_right MeasureTheory.Measure.set_integral_condKernel_univ_right + +lemma set_integral_condKernel_univ_left {t : Set Ω} (ht : MeasurableSet t) + (hf : IntegrableOn f (Set.univ ×ˢ t) ρ) : + ∫ b, ∫ ω in t, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫ x in Set.univ ×ˢ t, f x ∂ρ := by + rw [← set_integral_condKernel MeasurableSet.univ ht hf]; simp_rw [Measure.restrict_univ] +#align probability_theory.set_integral_cond_kernel_univ_left MeasureTheory.Measure.set_integral_condKernel_univ_left + +end Integral + +end MeasureTheory.Measure diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean index 1a7c7b2f9ebd5..0755bdc8943ae 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean @@ -11,6 +11,12 @@ namespace ProbabilityTheory variable {α Ω : Type*} {mα : MeasurableSpace α} [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] +lemma tendsto_atTop_atBot_iff_of_antitone {α β : Type*} + [Nonempty α] [SemilatticeSup α] [Preorder β] {f : α → β} + (hf : Antitone f) : + Tendsto f atTop atBot ↔ ∀ b : β, ∃ a : α, f a ≤ b := + @tendsto_atTop_atTop_iff_of_monotone _ βᵒᵈ _ _ _ _ hf + section Real section dissection_system @@ -959,12 +965,6 @@ lemma integral_tendsto_of_tendsto_of_antitone {μ : Measure α} {f : ℕ → α · filter_upwards [h_tendsto] with x hx exact hx.neg -theorem tendsto_atTop_atBot_iff_of_antitone {α β : Type*} - [Nonempty α] [SemilatticeSup α] [Preorder β] {f : α → β} - (hf : Antitone f) : - Tendsto f atTop atBot ↔ ∀ b : β, ∃ a : α, f a ≤ b := - @tendsto_atTop_atTop_iff_of_monotone _ βᵒᵈ _ _ _ _ hf - lemma tendsto_mLimsup_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : From 47b1d65d1a8a298c55780a65b23f085b6ecbcedf Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 15 Feb 2024 13:23:26 +0100 Subject: [PATCH 017/129] move results, fix --- Mathlib/Probability/Kernel/CondCdf.lean | 5 - Mathlib/Probability/Kernel/CondDistrib.lean | 11 +- .../Probability/Kernel/Disintegration.lean | 195 ------------------ .../Kernel/Disintegration/Basic.lean | 98 +++++++++ .../Kernel/Disintegration/Integral.lean | 69 ++++++- .../Kernel/Disintegration/Unique.lean | 4 +- 6 files changed, 173 insertions(+), 209 deletions(-) delete mode 100644 Mathlib/Probability/Kernel/Disintegration.lean diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index f0f511234f010..ff6acd4ea20fb 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -34,11 +34,6 @@ and limit 1 at +∞, and such that for all `x : ℝ`, `a ↦ cond_cdf ρ a x` is The construction of the conditional cdf in this file follows the proof of Theorem 3.4 in [O. Kallenberg, Foundations of modern probability][kallenberg2021]. -## TODO - -* The conditional cdf can be used to define the cdf of a real measure by using the - conditional cdf of `(measure.dirac unit.star).prod μ : measure (unit × ℝ)`. - -/ diff --git a/Mathlib/Probability/Kernel/CondDistrib.lean b/Mathlib/Probability/Kernel/CondDistrib.lean index 3cbd488e58506..ae027619eaab0 100644 --- a/Mathlib/Probability/Kernel/CondDistrib.lean +++ b/Mathlib/Probability/Kernel/CondDistrib.lean @@ -3,7 +3,8 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.Disintegration +import Mathlib.Probability.Kernel.Disintegration.Integral +import Mathlib.Probability.Kernel.Disintegration.Unique import Mathlib.Probability.Notation #align_import probability.kernel.cond_distrib from "leanprover-community/mathlib"@"00abe0695d8767201e6d008afa22393978bb324d" @@ -74,7 +75,7 @@ variable {mβ : MeasurableSpace β} {s : Set Ω} {t : Set β} {f : β × Ω → lemma condDistrib_apply_of_ne_zero [MeasurableSingletonClass β] (hY : Measurable Y) (x : β) (hX : μ.map X {x} ≠ 0) (s : Set Ω) : condDistrib Y X μ x s = (μ.map X {x})⁻¹ * μ.map (fun a => (X a, Y a)) ({x} ×ˢ s) := by - rw [condDistrib, condKernel_apply_of_ne_zero _ s] + rw [condDistrib, Measure.condKernel_apply_of_ne_zero _ s] · rw [Measure.fst_map_prod_mk hY] · rwa [Measure.fst_map_prod_mk hY] @@ -125,7 +126,7 @@ theorem condDistrib_ae_eq_of_measure_eq_compProd (hX : Measurable X) (hY : Measu rw [Measure.map_apply hX hs, Measure.fst_apply hs, Measure.map_apply] exacts [rfl, Measurable.prod hX hY, measurable_fst hs] rw [heq, condDistrib] - refine' eq_condKernel_of_measure_eq_compProd _ _ _ + refine Measure.eq_condKernel_of_measure_eq_compProd _ ?_ convert hκ exact heq.symm @@ -202,7 +203,7 @@ theorem set_lintegral_preimage_condDistrib (hX : Measurable X) (hY : AEMeasurabl -- (`rw` does not see that the two forms are defeq) conv_lhs => arg 2; change (fun a => ((condDistrib Y X μ) a) s) ∘ X rw [lintegral_comp (kernel.measurable_coe _ hs) hX, condDistrib, ← Measure.restrict_map hX ht, ← - Measure.fst_map_prod_mk₀ hY, set_lintegral_condKernel_eq_measure_prod _ ht hs, + Measure.fst_map_prod_mk₀ hY, Measure.set_lintegral_condKernel_eq_measure_prod ht hs, Measure.map_apply_of_aemeasurable (hX.aemeasurable.prod_mk hY) (ht.prod hs), mk_preimage_prod] #align probability_theory.set_lintegral_preimage_cond_distrib ProbabilityTheory.set_lintegral_preimage_condDistrib @@ -248,7 +249,7 @@ theorem condexp_prod_ae_eq_integral_condDistrib' [NormedSpace ℝ F] [CompleteSp · rw [← Measure.restrict_map hX ht] exact (hf_int.1.integral_condDistrib_map hY).restrict rw [← Measure.restrict_map hX ht, ← Measure.fst_map_prod_mk₀ hY, condDistrib, - set_integral_condKernel_univ_right ht hf_int.integrableOn, + Measure.set_integral_condKernel_univ_right ht hf_int.integrableOn, set_integral_map (ht.prod MeasurableSet.univ) hf_int.1 (hX.aemeasurable.prod_mk hY), mk_preimage_prod, preimage_univ, inter_univ] · exact aestronglyMeasurable'_integral_condDistrib hX.aemeasurable hY hf_int.1 diff --git a/Mathlib/Probability/Kernel/Disintegration.lean b/Mathlib/Probability/Kernel/Disintegration.lean deleted file mode 100644 index 1a37d6ed08810..0000000000000 --- a/Mathlib/Probability/Kernel/Disintegration.lean +++ /dev/null @@ -1,195 +0,0 @@ -/- -Copyright (c) 2023 Rémy Degenne. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Rémy Degenne, Kexing Ying --/ -import Mathlib.Probability.Kernel.Disintegration.Integral - -#align_import probability.kernel.disintegration from "leanprover-community/mathlib"@"6315581f5650ffa2fbdbbbedc41243c8d7070981" - -/-! -# Disintegration of measures on product spaces - -Let `ρ` be a finite measure on `α × Ω`, where `Ω` is a standard Borel space. In mathlib terms, `Ω` -verifies `[Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω]`. -Then there exists a kernel `ρ.condKernel : kernel α Ω` such that for any measurable set -`s : Set (α × Ω)`, `ρ s = ∫⁻ a, ρ.condKernel a {x | (a, x) ∈ s} ∂ρ.fst`. - -In terms of kernels, `ρ.condKernel` is such that for any measurable space `γ`, we -have a disintegration of the constant kernel from `γ` with value `ρ`: -`kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ (condKernel ρ))`, -where `ρ.fst` is the marginal measure of `ρ` on `α`. In particular, `ρ = ρ.fst ⊗ₘ ρ.condKernel`. - -In order to obtain a disintegration for any standard Borel space, we use that these spaces embed -measurably into `ℝ`: it then suffices to define a suitable kernel for `Ω = ℝ`. In the real case, -we define a conditional kernel by taking for each `a : α` the measure associated to the Stieltjes -function `condCDF ρ a` (the conditional cumulative distribution function). - -## Main definitions - -* `MeasureTheory.Measure.condKernel ρ : kernel α Ω`: conditional kernel described above. - -## Main statements - -* `ProbabilityTheory.lintegral_condKernel`: - `∫⁻ a, ∫⁻ ω, f (a, ω) ∂(ρ.condKernel a) ∂ρ.fst = ∫⁻ x, f x ∂ρ` -* `ProbabilityTheory.lintegral_condKernel_mem`: - `∫⁻ a, ρ.condKernel a {x | (a, x) ∈ s} ∂ρ.fst = ρ s` -* `ProbabilityTheory.kernel.const_eq_compProd`: - `kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ ρ.condKernel)` -* `ProbabilityTheory.measure_eq_compProd`: `ρ = ρ.fst ⊗ₘ ρ.condKernel` -* `ProbabilityTheory.eq_condKernel_of_measure_eq_compProd`: a.e. uniqueness of `condKernel` - --/ - - -open MeasureTheory Set Filter - -open scoped ENNReal MeasureTheory Topology ProbabilityTheory - -namespace ProbabilityTheory - -variable {α : Type*} {mα : MeasurableSpace α} - -#noalign probability_theory.cond_kernel_real -#noalign probability_theory.cond_kernel_real_Iic -#noalign probability_theory.set_lintegral_cond_kernel_real_Iic -#noalign probability_theory.set_lintegral_cond_kernel_real_univ -#noalign probability_theory.lintegral_cond_kernel_real_univ -#noalign probability_theory.set_lintegral_cond_kernel_real_prod -#noalign probability_theory.lintegral_cond_kernel_real_mem -#noalign probability_theory.kernel.const_eq_comp_prod_real -#noalign probability_theory.measure_eq_comp_prod_real -#noalign probability_theory.lintegral_cond_kernel_real -#noalign probability_theory.ae_cond_kernel_real_eq_one - -section StandardBorel - -/-! ### Disintegration of measures on standard Borel spaces - -Since every standard Borel space embeds measurably into `ℝ`, we can generalize the disintegration -property on `ℝ` to all these spaces. -/ - - -variable {Ω : Type*} [MeasurableSpace Ω] [StandardBorelSpace Ω] - [Nonempty Ω] (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] - -#noalign probability_theory.exists_cond_kernel -#noalign probability_theory.cond_kernel_def -#noalign probability_theory.kernel.const_unit_eq_comp_prod - -/-- **Disintegration** of finite product measures on `α × Ω`, where `Ω` is standard Borel. Such a -measure can be written as the composition-product of the constant kernel with value `ρ.fst` -(marginal measure over `α`) and a Markov kernel from `α` to `Ω`. We call that Markov kernel -`ProbabilityTheory.condKernel ρ`. -/ -theorem measure_eq_compProd : ρ = ρ.fst ⊗ₘ ρ.condKernel := by - rw [Measure.compProd_fst_condKernel] -#align probability_theory.measure_eq_comp_prod ProbabilityTheory.measure_eq_compProd - -#noalign probability_theory.kernel.const_eq_comp_prod - -/-- Auxiliary lemma for `condKernel_apply_of_ne_zero`. -/ -lemma condKernel_apply_of_ne_zero_of_measurableSet [MeasurableSingletonClass α] - {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] - {x : α} (hx : ρ.fst {x} ≠ 0) {s : Set Ω} (hs : MeasurableSet s) : - ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s) := by - nth_rewrite 3 [measure_eq_compProd ρ] - rw [Measure.compProd_apply (measurableSet_prod.mpr (Or.inl ⟨measurableSet_singleton x, hs⟩))] - classical - have : ∀ a, ρ.condKernel a (Prod.mk a ⁻¹' {x} ×ˢ s) - = ({x} : Set α).indicator (fun a ↦ ρ.condKernel a s) a := by - intro a - by_cases hax : a = x - · simp only [hax, Set.singleton_prod, Set.mem_singleton_iff, Set.indicator_of_mem] - congr with y - simp - · simp only [Set.singleton_prod, Set.mem_singleton_iff, hax, not_false_eq_true, - Set.indicator_of_not_mem] - have : Prod.mk a ⁻¹' (Prod.mk x '' s) = ∅ := by - ext y - simp [Ne.symm hax] - simp only [this, measure_empty] - simp_rw [this] - rw [MeasureTheory.lintegral_indicator _ (measurableSet_singleton x)] - simp only [Measure.restrict_singleton, lintegral_smul_measure, lintegral_dirac] - rw [← mul_assoc, ENNReal.inv_mul_cancel hx (measure_ne_top ρ.fst _), one_mul] - -/-- If the singleton `{x}` has non-zero mass for `ρ.fst`, then for all `s : Set Ω`, -`ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s)` . -/ -lemma condKernel_apply_of_ne_zero [MeasurableSingletonClass α] - {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {x : α} (hx : ρ.fst {x} ≠ 0) - (s : Set Ω) : - ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s) := by - have : ρ.condKernel x s = ((ρ.fst {x})⁻¹ • ρ).comap (fun (y : Ω) ↦ (x, y)) s := by - congr 2 with s hs - simp [condKernel_apply_of_ne_zero_of_measurableSet hx hs, - (measurableEmbedding_prod_mk_left x).comap_apply] - simp [this, (measurableEmbedding_prod_mk_left x).comap_apply, hx] - -section IntegralCondKernel - -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] - -theorem _root_.MeasureTheory.AEStronglyMeasurable.integral_condKernel {ρ : Measure (α × Ω)} - [IsFiniteMeasure ρ] {f : α × Ω → E} (hf : AEStronglyMeasurable f ρ) : - AEStronglyMeasurable (fun x ↦ ∫ y, f (x, y) ∂ρ.condKernel x) ρ.fst := by - rw [measure_eq_compProd ρ] at hf - exact AEStronglyMeasurable.integral_kernel_compProd hf -#align measure_theory.ae_strongly_measurable.integral_cond_kernel MeasureTheory.AEStronglyMeasurable.integral_condKernel - -end IntegralCondKernel - -end StandardBorel - -end ProbabilityTheory - -namespace MeasureTheory - -/-! ### Integrability - -We place these lemmas in the `MeasureTheory` namespace to enable dot notation. -/ - -open ProbabilityTheory - -variable {α Ω E F : Type*} {mα : MeasurableSpace α} [MeasurableSpace Ω] - [StandardBorelSpace Ω] [Nonempty Ω] [NormedAddCommGroup E] [NormedSpace ℝ E] - [CompleteSpace E] [NormedAddCommGroup F] {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] - -theorem AEStronglyMeasurable.ae_integrable_condKernel_iff {f : α × Ω → F} - (hf : AEStronglyMeasurable f ρ) : - (∀ᵐ a ∂ρ.fst, Integrable (fun ω ↦ f (a, ω)) (ρ.condKernel a)) ∧ - Integrable (fun a ↦ ∫ ω, ‖f (a, ω)‖ ∂ρ.condKernel a) ρ.fst ↔ Integrable f ρ := by - rw [measure_eq_compProd ρ] at hf - conv_rhs => rw [measure_eq_compProd ρ] - rw [Measure.integrable_compProd_iff hf] -#align measure_theory.ae_strongly_measurable.ae_integrable_cond_kernel_iff MeasureTheory.AEStronglyMeasurable.ae_integrable_condKernel_iff - -theorem Integrable.condKernel_ae {f : α × Ω → F} (hf_int : Integrable f ρ) : - ∀ᵐ a ∂ρ.fst, Integrable (fun ω ↦ f (a, ω)) (ρ.condKernel a) := by - have hf_ae : AEStronglyMeasurable f ρ := hf_int.1 - rw [← hf_ae.ae_integrable_condKernel_iff] at hf_int - exact hf_int.1 -#align measure_theory.integrable.cond_kernel_ae MeasureTheory.Integrable.condKernel_ae - -theorem Integrable.integral_norm_condKernel {f : α × Ω → F} (hf_int : Integrable f ρ) : - Integrable (fun x ↦ ∫ y, ‖f (x, y)‖ ∂ρ.condKernel x) ρ.fst := by - have hf_ae : AEStronglyMeasurable f ρ := hf_int.1 - rw [← hf_ae.ae_integrable_condKernel_iff] at hf_int - exact hf_int.2 -#align measure_theory.integrable.integral_norm_cond_kernel MeasureTheory.Integrable.integral_norm_condKernel - -theorem Integrable.norm_integral_condKernel {f : α × Ω → E} (hf_int : Integrable f ρ) : - Integrable (fun x ↦ ‖∫ y, f (x, y) ∂ρ.condKernel x‖) ρ.fst := by - refine hf_int.integral_norm_condKernel.mono hf_int.1.integral_condKernel.norm ?_ - refine eventually_of_forall fun x ↦ ?_ - rw [norm_norm] - refine (norm_integral_le_integral_norm _).trans_eq (Real.norm_of_nonneg ?_).symm - exact integral_nonneg_of_ae (eventually_of_forall fun y ↦ norm_nonneg _) -#align measure_theory.integrable.norm_integral_cond_kernel MeasureTheory.Integrable.norm_integral_condKernel - -theorem Integrable.integral_condKernel {f : α × Ω → E} (hf_int : Integrable f ρ) : - Integrable (fun x ↦ ∫ y, f (x, y) ∂ρ.condKernel x) ρ.fst := - (integrable_norm_iff hf_int.1.integral_condKernel).mp hf_int.norm_integral_condKernel -#align measure_theory.integrable.integral_cond_kernel MeasureTheory.Integrable.integral_condKernel - -end MeasureTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 14452483e8356..1e7f8f5526ec2 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -9,12 +9,61 @@ import Mathlib.Probability.Kernel.Disintegration.KernelCDFBorel /-! # Disintegration of kernels and measures +Let `ρ` be a finite measure on `α × Ω`, where `Ω` is a standard Borel space. In mathlib terms, `Ω` +verifies `[Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω]`. +Then there exists a kernel `ρ.condKernel : kernel α Ω` such that for any measurable set +`s : Set (α × Ω)`, `ρ s = ∫⁻ a, ρ.condKernel a {x | (a, x) ∈ s} ∂ρ.fst`. + +In terms of kernels, `ρ.condKernel` is such that for any measurable space `γ`, we +have a disintegration of the constant kernel from `γ` with value `ρ`: +`kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ (condKernel ρ))`, +where `ρ.fst` is the marginal measure of `ρ` on `α`. In particular, `ρ = ρ.fst ⊗ₘ ρ.condKernel`. + +In order to obtain a disintegration for any standard Borel space, we use that these spaces embed +measurably into `ℝ`: it then suffices to define a suitable kernel for `Ω = ℝ`. In the real case, +we define a conditional kernel by taking for each `a : α` the measure associated to the Stieltjes +function `condCDF ρ a` (the conditional cumulative distribution function). + +## Main definitions + +* `MeasureTheory.Measure.condKernel ρ : kernel α Ω`: conditional kernel described above. + +## Main statements + +* `ProbabilityTheory.lintegral_condKernel`: + `∫⁻ a, ∫⁻ ω, f (a, ω) ∂(ρ.condKernel a) ∂ρ.fst = ∫⁻ x, f x ∂ρ` +* `ProbabilityTheory.lintegral_condKernel_mem`: + `∫⁻ a, ρ.condKernel a {x | (a, x) ∈ s} ∂ρ.fst = ρ s` +* `ProbabilityTheory.kernel.const_eq_compProd`: + `kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ ρ.condKernel)` +* `ProbabilityTheory.measure_eq_compProd`: `ρ = ρ.fst ⊗ₘ ρ.condKernel` + -/ +#align_import probability.kernel.disintegration from "leanprover-community/mathlib"@"6315581f5650ffa2fbdbbbedc41243c8d7070981" + open MeasureTheory Set Filter open scoped ENNReal MeasureTheory Topology ProbabilityTheory + +#noalign probability_theory.cond_kernel_real +#noalign probability_theory.cond_kernel_real_Iic +#noalign probability_theory.set_lintegral_cond_kernel_real_Iic +#noalign probability_theory.set_lintegral_cond_kernel_real_univ +#noalign probability_theory.lintegral_cond_kernel_real_univ +#noalign probability_theory.set_lintegral_cond_kernel_real_prod +#noalign probability_theory.lintegral_cond_kernel_real_mem +#noalign probability_theory.kernel.const_eq_comp_prod_real +#noalign probability_theory.measure_eq_comp_prod_real +#noalign probability_theory.lintegral_cond_kernel_real +#noalign probability_theory.ae_cond_kernel_real_eq_one +#noalign probability_theory.exists_cond_kernel +#noalign probability_theory.cond_kernel_def +#noalign probability_theory.kernel.const_unit_eq_comp_prod +#noalign probability_theory.kernel.const_eq_comp_prod + + namespace ProbabilityTheory variable {α β Ω Ω': Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} @@ -23,6 +72,13 @@ variable {α β Ω Ω': Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace section BorelSnd +/-! ### Disintegration of kenrels on standard Borel spaces + +Since every standard Borel space embeds measurably into `ℝ`, we can generalize a disintegration +property on `ℝ` to all these spaces. + +On `ℝ`, we get disintegration by constructing a map `f` with the property `IsKernelCDF f`. -/ + noncomputable def condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFunction} (hf : IsKernelCDF f @@ -265,6 +321,8 @@ end Unit section Measure +variable {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] + /-- Conditional kernel of a measure on a product space: a Markov kernel such that `ρ = ρ.fst ⊗ₘ ρ.condKernel` (see `ProbabilityTheory.measure_eq_compProd`). -/ noncomputable @@ -283,6 +341,9 @@ instance _root_.MeasureTheory.Measure.instIsMarkovKernel_condKernel rw [Measure.condKernel] infer_instance +/-- **Disintegration** of finite product measures on `α × Ω`, where `Ω` is standard Borel. Such a +measure can be written as the composition-product of `ρ.fst` (marginal measure over `α`) and +a Markov kernel from `α` to `Ω`. We call that Markov kernel `ρ.condKernel`. -/ lemma _root_.MeasureTheory.Measure.compProd_fst_condKernel (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : ρ.fst ⊗ₘ ρ.condKernel = ρ := by @@ -295,6 +356,43 @@ lemma _root_.MeasureTheory.Measure.compProd_fst_condKernel simp only [kernel.prodMkLeft_apply, Measure.condKernel_apply] rw [Measure.compProd, h1, h2, compProd_fst_condKernelUnitBorel] simp +#align probability_theory.measure_eq_comp_prod MeasureTheory.Measure.compProd_fst_condKernel + +/-- Auxiliary lemma for `condKernel_apply_of_ne_zero`. -/ +lemma _root_.MeasureTheory.Measure.condKernel_apply_of_ne_zero_of_measurableSet + [MeasurableSingletonClass α] {x : α} (hx : ρ.fst {x} ≠ 0) {s : Set Ω} (hs : MeasurableSet s) : + ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s) := by + nth_rewrite 3 [← ρ.compProd_fst_condKernel] + rw [Measure.compProd_apply (measurableSet_prod.mpr (Or.inl ⟨measurableSet_singleton x, hs⟩))] + classical + have : ∀ a, ρ.condKernel a (Prod.mk a ⁻¹' {x} ×ˢ s) + = ({x} : Set α).indicator (fun a ↦ ρ.condKernel a s) a := by + intro a + by_cases hax : a = x + · simp only [hax, Set.singleton_prod, Set.mem_singleton_iff, Set.indicator_of_mem] + congr with y + simp + · simp only [Set.singleton_prod, Set.mem_singleton_iff, hax, not_false_eq_true, + Set.indicator_of_not_mem] + have : Prod.mk a ⁻¹' (Prod.mk x '' s) = ∅ := by + ext y + simp [Ne.symm hax] + simp only [this, measure_empty] + simp_rw [this] + rw [MeasureTheory.lintegral_indicator _ (measurableSet_singleton x)] + simp only [Measure.restrict_singleton, lintegral_smul_measure, lintegral_dirac] + rw [← mul_assoc, ENNReal.inv_mul_cancel hx (measure_ne_top ρ.fst _), one_mul] + +/-- If the singleton `{x}` has non-zero mass for `ρ.fst`, then for all `s : Set Ω`, +`ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s)` . -/ +lemma _root_.MeasureTheory.Measure.condKernel_apply_of_ne_zero [MeasurableSingletonClass α] + {x : α} (hx : ρ.fst {x} ≠ 0) (s : Set Ω) : + ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s) := by + have : ρ.condKernel x s = ((ρ.fst {x})⁻¹ • ρ).comap (fun (y : Ω) ↦ (x, y)) s := by + congr 2 with s hs + simp [Measure.condKernel_apply_of_ne_zero_of_measurableSet hx hs, + (measurableEmbedding_prod_mk_left x).comap_apply] + simp [this, (measurableEmbedding_prod_mk_left x).comap_apply, hx] end Measure diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index e2236ea775c20..d307ad50af939 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -52,8 +52,8 @@ lemma lintegral_condKernel_mem (a : α) {s : Set (β × Ω)} (hs : MeasurableSet conv_rhs => rw [← kernel.compProd_fst_condKernel κ] simp_rw [kernel.compProd_apply _ _ _ hs] -lemma set_lintegral_condKernel_eq_measure_prod (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set Ω} - (ht : MeasurableSet t) : +lemma set_lintegral_condKernel_eq_measure_prod (a : α) {s : Set β} (hs : MeasurableSet s) + {t : Set Ω} (ht : MeasurableSet t) : ∫⁻ b in s, kernel.condKernel κ (a, b) t ∂(kernel.fst κ a) = κ a (s ×ˢ t) := by have : κ a (s ×ˢ t) = (kernel.fst κ ⊗ₖ kernel.condKernel κ) a (s ×ˢ t) := by congr; exact (kernel.compProd_fst_condKernel κ).symm @@ -97,6 +97,13 @@ section Integral variable [CountableOrStandardBorel α β] {κ : kernel α (β × Ω)} [IsFiniteKernel κ] {E : Type*} {f : β × Ω → E} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] +lemma _root_.MeasureTheory.AEStronglyMeasurable.integral_kernel_condKernel (a : α) + (hf : AEStronglyMeasurable f (κ a)) : + AEStronglyMeasurable (fun x ↦ ∫ y, f (x, y) ∂(kernel.condKernel κ (a, x))) + (kernel.fst κ a) := by + rw [← kernel.compProd_fst_condKernel κ] at hf + exact AEStronglyMeasurable.integral_kernel_compProd hf + lemma integral_condKernel (a : α) (hf : Integrable f (κ a)) : ∫ b, ∫ ω, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a) = ∫ x, f x ∂(κ a) := by conv_rhs => rw [← kernel.compProd_fst_condKernel κ] @@ -194,6 +201,13 @@ section Integral variable {ρ : Measure (β × Ω)} [IsFiniteMeasure ρ] {E : Type*} {f : β × Ω → E} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] +lemma _root_.MeasureTheory.AEStronglyMeasurable.integral_condKernel + (hf : AEStronglyMeasurable f ρ) : + AEStronglyMeasurable (fun x ↦ ∫ y, f (x, y) ∂ρ.condKernel x) ρ.fst := by + rw [← ρ.compProd_fst_condKernel] at hf + exact AEStronglyMeasurable.integral_kernel_compProd hf +#align measure_theory.ae_strongly_measurable.integral_cond_kernel MeasureTheory.AEStronglyMeasurable.integral_condKernel + lemma integral_condKernel (hf : Integrable f ρ) : ∫ b, ∫ ω, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫ x, f x ∂ρ := by conv_rhs => rw [← compProd_fst_condKernel ρ] @@ -224,3 +238,54 @@ lemma set_integral_condKernel_univ_left {t : Set Ω} (ht : MeasurableSet t) end Integral end MeasureTheory.Measure + +namespace MeasureTheory + +/-! ### Integrability + +We place these lemmas in the `MeasureTheory` namespace to enable dot notation. -/ + +open ProbabilityTheory + +variable {α Ω E F : Type*} {mα : MeasurableSpace α} [MeasurableSpace Ω] + [StandardBorelSpace Ω] [Nonempty Ω] [NormedAddCommGroup E] [NormedSpace ℝ E] + [CompleteSpace E] [NormedAddCommGroup F] {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] + +theorem AEStronglyMeasurable.ae_integrable_condKernel_iff {f : α × Ω → F} + (hf : AEStronglyMeasurable f ρ) : + (∀ᵐ a ∂ρ.fst, Integrable (fun ω ↦ f (a, ω)) (ρ.condKernel a)) ∧ + Integrable (fun a ↦ ∫ ω, ‖f (a, ω)‖ ∂ρ.condKernel a) ρ.fst ↔ Integrable f ρ := by + rw [← ρ.compProd_fst_condKernel] at hf + conv_rhs => rw [← ρ.compProd_fst_condKernel] + rw [Measure.integrable_compProd_iff hf] +#align measure_theory.ae_strongly_measurable.ae_integrable_cond_kernel_iff MeasureTheory.AEStronglyMeasurable.ae_integrable_condKernel_iff + +theorem Integrable.condKernel_ae {f : α × Ω → F} (hf_int : Integrable f ρ) : + ∀ᵐ a ∂ρ.fst, Integrable (fun ω ↦ f (a, ω)) (ρ.condKernel a) := by + have hf_ae : AEStronglyMeasurable f ρ := hf_int.1 + rw [← hf_ae.ae_integrable_condKernel_iff] at hf_int + exact hf_int.1 +#align measure_theory.integrable.cond_kernel_ae MeasureTheory.Integrable.condKernel_ae + +theorem Integrable.integral_norm_condKernel {f : α × Ω → F} (hf_int : Integrable f ρ) : + Integrable (fun x ↦ ∫ y, ‖f (x, y)‖ ∂ρ.condKernel x) ρ.fst := by + have hf_ae : AEStronglyMeasurable f ρ := hf_int.1 + rw [← hf_ae.ae_integrable_condKernel_iff] at hf_int + exact hf_int.2 +#align measure_theory.integrable.integral_norm_cond_kernel MeasureTheory.Integrable.integral_norm_condKernel + +theorem Integrable.norm_integral_condKernel {f : α × Ω → E} (hf_int : Integrable f ρ) : + Integrable (fun x ↦ ‖∫ y, f (x, y) ∂ρ.condKernel x‖) ρ.fst := by + refine hf_int.integral_norm_condKernel.mono hf_int.1.integral_condKernel.norm ?_ + refine Filter.eventually_of_forall fun x ↦ ?_ + rw [norm_norm] + refine (norm_integral_le_integral_norm _).trans_eq (Real.norm_of_nonneg ?_).symm + exact integral_nonneg_of_ae (Filter.eventually_of_forall fun y ↦ norm_nonneg _) +#align measure_theory.integrable.norm_integral_cond_kernel MeasureTheory.Integrable.norm_integral_condKernel + +theorem Integrable.integral_condKernel {f : α × Ω → E} (hf_int : Integrable f ρ) : + Integrable (fun x ↦ ∫ y, f (x, y) ∂ρ.condKernel x) ρ.fst := + (integrable_norm_iff hf_int.1.integral_condKernel).mp hf_int.norm_integral_condKernel +#align measure_theory.integrable.integral_cond_kernel MeasureTheory.Integrable.integral_condKernel + +end MeasureTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index b9847b7832393..41ffd83354046 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ -import Mathlib.Probability.Kernel.Disintegration +import Mathlib.Probability.Kernel.Disintegration.Integral /-! # Disintegration of measures on product spaces @@ -144,7 +144,7 @@ theorem Measure.eq_condKernel_of_measure_eq_compProd' (κ : kernel α Ω) [IsSFi ∀ᵐ x ∂ρ.fst, κ x s = ρ.condKernel x s := by refine ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite (kernel.measurable_coe κ hs) (kernel.measurable_coe ρ.condKernel hs) (fun t ht _ ↦ ?_) - conv_rhs => rw [set_lintegral_condKernel_eq_measure_prod _ ht hs, hκ] + conv_rhs => rw [Measure.set_lintegral_condKernel_eq_measure_prod ht hs, hκ] simp only [Measure.compProd_apply (ht.prod hs), Set.mem_prod, ← lintegral_indicator _ ht] congr with x by_cases hx : x ∈ t From 76961e94f73dc102b4504fec0e5582ba0a5a6775 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 15 Feb 2024 13:30:24 +0100 Subject: [PATCH 018/129] docstring for Integral.lean --- .../Kernel/Disintegration/Integral.lean | 28 ++++++------------- 1 file changed, 9 insertions(+), 19 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index d307ad50af939..432948403de51 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -6,31 +6,21 @@ Authors: Rémy Degenne import Mathlib.Probability.Kernel.Disintegration.Basic /-! -# Integral +# Lebesgue and Bochner integrals of conditional kernels -## Main definitions - -* `FooBar` +Integrals of `ProbabilityTheory.kernel.condKernel` and `MeasureTheory.Measure.condKerenl`. ## Main statements -* `fooBar_unique` - -## Notation - - - -## Implementation details - - - -## References - -* [F. Bar, *Quuxes*][bibkey] +* `ProbabilityTheory.set_integral_condKernel`: the integral + `∫ b in s, ∫ ω in t, f (b, ω) ∂(kernel.condKernel κ (a, b)) ∂(kernel.fst κ a)` is equal to + `∫ x in s ×ˢ t, f x ∂(κ a)`. +* `MeasureTheory.Measure.set_integral_condKernel`: + `∫ b in s, ∫ ω in t, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫ x in s ×ˢ t, f x ∂ρ` -## Tags +Corresponding statements for the Lebesgue integral and/or without the sets `s` and `t` are also +provided. -Foobars, barfoos -/ open MeasureTheory ProbabilityTheory From 9d97cb16749aaef2fc87a6ba602320968e0318c4 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 15 Feb 2024 17:46:27 +0100 Subject: [PATCH 019/129] move/extract lemmas --- Mathlib/Order/Filter/AtTopBot.lean | 12 + Mathlib/Order/LiminfLimsup.lean | 2 + Mathlib/Probability/Kernel/CondCdf.lean | 27 +- .../Kernel/Disintegration/BuildKernel.lean | 116 ++--- .../Kernel/Disintegration/KernelCDFBorel.lean | 458 +++++++++--------- .../Kernel/Disintegration/StieltjesReal.lean | 126 ++--- 6 files changed, 392 insertions(+), 349 deletions(-) diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index d584fddde36ce..fabc76e07d924 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -1377,18 +1377,30 @@ theorem tendsto_atBot_atBot_of_monotone [Preorder α] [Preorder β] {f : α → let ⟨a, ha⟩ := h b; mem_of_superset (mem_atBot a) fun _a' ha' => le_trans (hf ha') ha #align filter.tendsto_at_bot_at_bot_of_monotone Filter.tendsto_atBot_atBot_of_monotone +theorem tendsto_atBot_atTop_of_antitone [Preorder α] [Preorder β] {f : α → β} (hf : Antitone f) + (h : ∀ b, ∃ a, b ≤ f a) : Tendsto f atBot atTop := + @tendsto_atBot_atBot_of_monotone _ βᵒᵈ _ _ _ hf h + theorem tendsto_atTop_atTop_iff_of_monotone [Nonempty α] [SemilatticeSup α] [Preorder β] {f : α → β} (hf : Monotone f) : Tendsto f atTop atTop ↔ ∀ b : β, ∃ a : α, b ≤ f a := tendsto_atTop_atTop.trans <| forall_congr' fun _ => exists_congr fun a => ⟨fun h => h a (le_refl a), fun h _a' ha' => le_trans h <| hf ha'⟩ #align filter.tendsto_at_top_at_top_iff_of_monotone Filter.tendsto_atTop_atTop_iff_of_monotone +theorem tendsto_atTop_atBot_iff_of_antitone [Nonempty α] [SemilatticeSup α] [Preorder β] {f : α → β} + (hf : Antitone f) : Tendsto f atTop atBot ↔ ∀ b : β, ∃ a : α, f a ≤ b := + @tendsto_atTop_atTop_iff_of_monotone _ βᵒᵈ _ _ _ _ hf + theorem tendsto_atBot_atBot_iff_of_monotone [Nonempty α] [SemilatticeInf α] [Preorder β] {f : α → β} (hf : Monotone f) : Tendsto f atBot atBot ↔ ∀ b : β, ∃ a : α, f a ≤ b := tendsto_atBot_atBot.trans <| forall_congr' fun _ => exists_congr fun a => ⟨fun h => h a (le_refl a), fun h _a' ha' => le_trans (hf ha') h⟩ #align filter.tendsto_at_bot_at_bot_iff_of_monotone Filter.tendsto_atBot_atBot_iff_of_monotone +theorem tendsto_atBot_atTop_iff_of_antitone [Nonempty α] [SemilatticeInf α] [Preorder β] {f : α → β} + (hf : Antitone f) : Tendsto f atBot atTop ↔ ∀ b : β, ∃ a : α, b ≤ f a := + @tendsto_atBot_atBot_iff_of_monotone _ βᵒᵈ _ _ _ _ hf + alias _root_.Monotone.tendsto_atTop_atTop := tendsto_atTop_atTop_of_monotone #align monotone.tendsto_at_top_at_top Monotone.tendsto_atTop_atTop diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 74b135ea9525f..f5c16fff032a7 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -663,11 +663,13 @@ theorem liminf_congr {α : Type*} [ConditionallyCompleteLattice β] {f : Filter limsup_congr (β := βᵒᵈ) h #align filter.liminf_congr Filter.liminf_congr +@[simp] theorem limsup_const {α : Type*} [ConditionallyCompleteLattice β] {f : Filter α} [NeBot f] (b : β) : limsup (fun _ => b) f = b := by simpa only [limsup_eq, eventually_const] using csInf_Ici #align filter.limsup_const Filter.limsup_const +@[simp] theorem liminf_const {α : Type*} [ConditionallyCompleteLattice β] {f : Filter α} [NeBot f] (b : β) : liminf (fun _ => b) f = b := limsup_const (β := βᵒᵈ) b diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index ff6acd4ea20fb..359b954ddf564 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -468,41 +468,42 @@ lemma isRatKernelCDF_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : /-- Conditional cdf of the measure given the value on `α`, as a Stieltjes function. -/ noncomputable def condCDF (ρ : Measure (α × ℝ)) (a : α) : StieltjesFunction := - todo3 (fun a r ↦ (preCDF ρ r a).toReal) (fun _ ↦ measurable_preCDF.ennreal_toReal) a + stieltjesOfMeasurableRat (fun a r ↦ (preCDF ρ r a).toReal) + (fun _ ↦ measurable_preCDF.ennreal_toReal) a #align probability_theory.cond_cdf ProbabilityTheory.condCDF -lemma condCDF_eq_todo3_unit_prod (ρ : Measure (α × ℝ)) (a : α) : - condCDF ρ a = todo3 (fun (p : Unit × α) r ↦ (preCDF ρ r p.2).toReal) +lemma condCDF_eq_stieltjesOfMeasurableRat_unit_prod (ρ : Measure (α × ℝ)) (a : α) : + condCDF ρ a = stieltjesOfMeasurableRat (fun (p : Unit × α) r ↦ (preCDF ρ r p.2).toReal) (fun _ ↦ measurable_preCDF.ennreal_toReal.comp measurable_snd) ((), a) := by ext x - rw [condCDF, ← todo3_unit_prod] + rw [condCDF, ← stieltjesOfMeasurableRat_unit_prod] #noalign probability_theory.cond_cdf_eq_cond_cdf_rat /-- The conditional cdf is non-negative for all `a : α`. -/ theorem condCDF_nonneg (ρ : Measure (α × ℝ)) (a : α) (r : ℝ) : 0 ≤ condCDF ρ a r := - todo3_nonneg _ a r + stieltjesOfMeasurableRat_nonneg _ a r #align probability_theory.cond_cdf_nonneg ProbabilityTheory.condCDF_nonneg /-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ theorem condCDF_le_one (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : condCDF ρ a x ≤ 1 := - todo3_le_one _ _ _ + stieltjesOfMeasurableRat_le_one _ _ _ #align probability_theory.cond_cdf_le_one ProbabilityTheory.condCDF_le_one /-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ theorem tendsto_condCDF_atBot (ρ : Measure (α × ℝ)) (a : α) : - Tendsto (condCDF ρ a) atBot (𝓝 0) := tendsto_todo3_atBot _ _ + Tendsto (condCDF ρ a) atBot (𝓝 0) := tendsto_stieltjesOfMeasurableRat_atBot _ _ #align probability_theory.tendsto_cond_cdf_at_bot ProbabilityTheory.tendsto_condCDF_atBot /-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ theorem tendsto_condCDF_atTop (ρ : Measure (α × ℝ)) (a : α) : - Tendsto (condCDF ρ a) atTop (𝓝 1) := tendsto_todo3_atTop _ _ + Tendsto (condCDF ρ a) atTop (𝓝 1) := tendsto_stieltjesOfMeasurableRat_atTop _ _ #align probability_theory.tendsto_cond_cdf_at_top ProbabilityTheory.tendsto_condCDF_atTop theorem condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) : (fun a ↦ condCDF ρ a r) =ᵐ[ρ.fst] fun a ↦ (preCDF ρ r a).toReal := by filter_upwards [isRatStieltjesPoint_ae ρ] with a ha - rw [condCDF, todo3_eq, toCDFLike_of_isRatStieltjesPoint ha] + rw [condCDF, stieltjesOfMeasurableRat_eq, toCDFLike_of_isRatStieltjesPoint ha] #align probability_theory.cond_cdf_ae_eq ProbabilityTheory.condCDF_ae_eq theorem ofReal_condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) : @@ -514,7 +515,7 @@ theorem ofReal_condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r /-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ theorem measurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : Measurable fun a => condCDF ρ a x := - measurable_todo3 _ _ + measurable_stieltjesOfMeasurableRat _ _ #align probability_theory.measurable_cond_cdf ProbabilityTheory.measurable_condCDF #noalign probability_theory.set_lintegral_cond_cdf_rat @@ -522,12 +523,12 @@ theorem measurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : Measurable fun theorem set_lintegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) {s : Set α} (hs : MeasurableSet s) : ∫⁻ a in s, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x) := by - have h := set_lintegral_todo3 (isRatKernelCDF_preCDF ρ) () x hs + have h := set_lintegral_stieltjesOfMeasurableRat (isRatKernelCDF_preCDF ρ) () x hs simp only [kernel.const_apply] at h rw [← h] congr with a congr - exact condCDF_eq_todo3_unit_prod _ _ + exact condCDF_eq_stieltjesOfMeasurableRat_unit_prod _ _ #align probability_theory.set_lintegral_cond_cdf ProbabilityTheory.set_lintegral_condCDF theorem lintegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) : @@ -537,7 +538,7 @@ theorem lintegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : /-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ theorem stronglyMeasurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : - StronglyMeasurable fun a => condCDF ρ a x := stronglyMeasurable_todo3 _ _ + StronglyMeasurable fun a => condCDF ρ a x := stronglyMeasurable_stieltjesOfMeasurableRat _ _ #align probability_theory.strongly_measurable_cond_cdf ProbabilityTheory.stronglyMeasurable_condCDF theorem integrable_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) : diff --git a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean index bec5c1621a242..13a2838cedce8 100644 --- a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean @@ -38,7 +38,7 @@ namespace ProbabilityTheory variable {α β : Type*} [MeasurableSpace α] -section todo3 +section stieltjesOfMeasurableRat variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} {f : α × β → ℚ → ℝ} {μ : kernel α (β × ℝ)} {ν : kernel α β} @@ -50,32 +50,35 @@ structure IsRatKernelCDF (f : α × β → ℚ → ℝ) (μ : kernel α (β × (isCDF (a : α) {s : Set β} (_hs : MeasurableSet s) (q : ℚ) : ∫ t in s, f (a, t) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal) -lemma todo3_ae_eq (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) : - (fun t ↦ todo3 f hf.measurable (a, t) q) =ᵐ[ν a] fun t ↦ f (a, t) q := by +lemma stieltjesOfMeasurableRat_ae_eq (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) : + (fun t ↦ stieltjesOfMeasurableRat f hf.measurable (a, t) q) =ᵐ[ν a] fun t ↦ f (a, t) q := by filter_upwards [hf.isRatStieltjesPoint_ae a] with a ha - rw [todo3_eq, toCDFLike_of_isRatStieltjesPoint ha] + rw [stieltjesOfMeasurableRat_eq, toCDFLike_of_isRatStieltjesPoint ha] -lemma set_integral_todo3_rat (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) +lemma set_integral_stieltjesOfMeasurableRat_rat (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : - ∫ t in s, todo3 f hf.measurable (a, t) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal := by + ∫ t in s, stieltjesOfMeasurableRat f hf.measurable (a, t) q ∂(ν a) + = (μ a (s ×ˢ Iic (q : ℝ))).toReal := by rw [set_integral_congr_ae hs (g := fun t ↦ f (a, t) q) ?_, hf.isCDF a hs] - filter_upwards [todo3_ae_eq hf a q] with b hb using fun _ ↦ hb + filter_upwards [stieltjesOfMeasurableRat_ae_eq hf a q] with b hb using fun _ ↦ hb -lemma set_lintegral_todo3_rat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma set_lintegral_stieltjesOfMeasurableRat_rat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ t in s, ENNReal.ofReal (todo3 f hf.measurable (a, t) q) ∂(ν a) = μ a (s ×ˢ Iic (q : ℝ)) := by + ∫⁻ t in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) q) ∂(ν a) + = μ a (s ×ˢ Iic (q : ℝ)) := by rw [← ofReal_integral_eq_lintegral_ofReal] - · rw [set_integral_todo3_rat hf a q hs, ENNReal.ofReal_toReal] + · rw [set_integral_stieltjesOfMeasurableRat_rat hf a q hs, ENNReal.ofReal_toReal] exact measure_ne_top _ _ · refine Integrable.restrict ?_ - rw [integrable_congr (todo3_ae_eq hf a q)] + rw [integrable_congr (stieltjesOfMeasurableRat_ae_eq hf a q)] exact hf.integrable a q - · exact ae_of_all _ (fun x ↦ todo3_nonneg _ _ _) + · exact ae_of_all _ (fun x ↦ stieltjesOfMeasurableRat_nonneg _ _ _) -lemma set_lintegral_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ t in s, ENNReal.ofReal (todo3 f hf.measurable (a, t) x) ∂(ν a) = μ a (s ×ˢ Iic x) := by - -- We have the result for `x : ℚ` thanks to `set_lintegral_todo3_rat`. + ∫⁻ t in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) x) ∂(ν a) + = μ a (s ×ˢ Iic x) := by + -- We have the result for `x : ℚ` thanks to `set_lintegral_stieltjesOfMeasurableRat_rat`. -- We use the equality `condCDF ρ a x = ⨅ r : {r' : ℚ // x < r'}, condCDF ρ a r` and a monotone -- convergence argument to extend it to the reals. by_cases hρ_zero : (ν a).restrict s = 0 @@ -91,11 +94,11 @@ lemma set_lintegral_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) simpa [measure_ne_top] using this rw [← hf.isCDF a hs q] simp [hρ_zero] - have h : ∫⁻ t in s, ENNReal.ofReal (todo3 f hf.measurable (a, t) x) ∂(ν a) + have h : ∫⁻ t in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) x) ∂(ν a) = ∫⁻ t in s, ⨅ r : { r' : ℚ // x < r' }, - ENNReal.ofReal (todo3 f hf.measurable (a, t) r) ∂(ν a) := by + ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) r) ∂(ν a) := by congr with t : 1 - simp_rw [← measure_todo3_Iic] + simp_rw [← measure_stieltjesOfMeasurableRat_Iic] rw [← measure_iInter_eq_iInf] · congr with y : 1 simp only [mem_Iic, mem_iInter, Subtype.forall] @@ -115,15 +118,15 @@ lemma set_lintegral_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) rw [h, lintegral_iInf_directed_of_measurable hρ_zero fun q : { r' : ℚ // x < ↑r' } ↦ ?_] rotate_left · intro b - rw [set_lintegral_todo3_rat hf a _ hs] + rw [set_lintegral_stieltjesOfMeasurableRat_rat hf a _ hs] exact measure_ne_top _ _ · refine Monotone.directed_ge fun i j hij t ↦ ?_ - simp_rw [← measure_todo3_Iic] + simp_rw [← measure_stieltjesOfMeasurableRat_Iic] refine measure_mono (Iic_subset_Iic.mpr ?_) exact mod_cast hij · refine Measurable.ennreal_ofReal ?_ - exact (measurable_todo3 hf.measurable _).comp measurable_prod_mk_left - simp_rw [set_lintegral_todo3_rat hf _ _ hs] + exact (measurable_stieltjesOfMeasurableRat hf.measurable _).comp measurable_prod_mk_left + simp_rw [set_lintegral_stieltjesOfMeasurableRat_rat hf _ _ hs] rw [← measure_iInter_eq_iInf] · rw [← prod_iInter] congr with y @@ -135,41 +138,47 @@ lemma set_lintegral_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) exact mod_cast hij · exact ⟨h_nonempty.some, measure_ne_top _ _⟩ -lemma lintegral_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) (x : ℝ) : - ∫⁻ t, ENNReal.ofReal (todo3 f hf.measurable (a, t) x) ∂(ν a) = μ a (univ ×ˢ Iic x) := by - rw [← set_lintegral_univ, set_lintegral_todo3 hf _ _ MeasurableSet.univ] - -lemma integrable_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) (x : ℝ) : - Integrable (fun t ↦ todo3 f hf.measurable (a, t) x) (ν a) := by - have : (fun t ↦ todo3 f hf.measurable (a, t) x) - = fun t ↦ (ENNReal.ofReal (todo3 f hf.measurable (a, t) x)).toReal := by +lemma lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) + (a : α) (x : ℝ) : + ∫⁻ t, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) x) ∂(ν a) + = μ a (univ ×ˢ Iic x) := by + rw [← set_lintegral_univ, set_lintegral_stieltjesOfMeasurableRat hf _ _ MeasurableSet.univ] + +lemma integrable_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) + (a : α) (x : ℝ) : + Integrable (fun t ↦ stieltjesOfMeasurableRat f hf.measurable (a, t) x) (ν a) := by + have : (fun t ↦ stieltjesOfMeasurableRat f hf.measurable (a, t) x) + = fun t ↦ (ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) x)).toReal := by ext t rw [ENNReal.toReal_ofReal] - exact todo3_nonneg _ _ _ + exact stieltjesOfMeasurableRat_nonneg _ _ _ rw [this] refine integrable_toReal_of_lintegral_ne_top ?_ ?_ · refine (Measurable.ennreal_ofReal ?_).aemeasurable - exact (measurable_todo3 hf.measurable x).comp measurable_prod_mk_left - · rw [lintegral_todo3 hf] + exact (measurable_stieltjesOfMeasurableRat hf.measurable x).comp measurable_prod_mk_left + · rw [lintegral_stieltjesOfMeasurableRat hf] exact measure_ne_top _ _ -lemma set_integral_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma set_integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : - ∫ t in s, todo3 f hf.measurable (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal := by + ∫ t in s, stieltjesOfMeasurableRat f hf.measurable (a, t) x ∂(ν a) + = (μ a (s ×ˢ Iic x)).toReal := by rw [← ENNReal.ofReal_eq_ofReal_iff, ENNReal.ofReal_toReal] rotate_left · exact measure_ne_top _ _ - · exact set_integral_nonneg hs (fun _ _ ↦ todo3_nonneg _ _ _) + · exact set_integral_nonneg hs (fun _ _ ↦ stieltjesOfMeasurableRat_nonneg _ _ _) · exact ENNReal.toReal_nonneg - rw [ofReal_integral_eq_lintegral_ofReal, set_lintegral_todo3 hf _ _ hs] - · exact (integrable_todo3 hf _ _).restrict - · exact ae_of_all _ (fun _ ↦ todo3_nonneg _ _ _) + rw [ofReal_integral_eq_lintegral_ofReal, set_lintegral_stieltjesOfMeasurableRat hf _ _ hs] + · exact (integrable_stieltjesOfMeasurableRat hf _ _).restrict + · exact ae_of_all _ (fun _ ↦ stieltjesOfMeasurableRat_nonneg _ _ _) -lemma integral_todo3 [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) (a : α) (x : ℝ) : - ∫ t, todo3 f hf.measurable (a, t) x ∂(ν a) = (μ a (univ ×ˢ Iic x)).toReal := by - rw [← integral_univ, set_integral_todo3 hf _ _ MeasurableSet.univ] +lemma integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) + (a : α) (x : ℝ) : + ∫ t, stieltjesOfMeasurableRat f hf.measurable (a, t) x ∂(ν a) + = (μ a (univ ×ˢ Iic x)).toReal := by + rw [← integral_univ, set_integral_stieltjesOfMeasurableRat hf _ _ MeasurableSet.univ] -end todo3 +end stieltjesOfMeasurableRat section IsKernelCDF @@ -200,13 +209,14 @@ lemma IsKernelCDF.set_lintegral [IsFiniteKernel μ] (ae_of_all _ (fun _ ↦ hf.nonneg _ _)), hf.set_integral a hs x, ENNReal.ofReal_toReal] exact measure_ne_top _ _ -lemma isKernelCDF_todo3 {f : α × β → ℚ → ℝ} (hf : IsRatKernelCDF f μ ν) [IsFiniteKernel μ] : - IsKernelCDF (todo3 f hf.measurable) μ ν where - measurable := measurable_todo3 hf.measurable - integrable := integrable_todo3 hf - tendsto_atTop_one := tendsto_todo3_atTop hf.measurable - tendsto_atBot_zero := tendsto_todo3_atBot hf.measurable - set_integral a _ hs x := set_integral_todo3 hf a x hs +lemma isKernelCDF_stieltjesOfMeasurableRat {f : α × β → ℚ → ℝ} (hf : IsRatKernelCDF f μ ν) + [IsFiniteKernel μ] : + IsKernelCDF (stieltjesOfMeasurableRat f hf.measurable) μ ν where + measurable := measurable_stieltjesOfMeasurableRat hf.measurable + integrable := integrable_stieltjesOfMeasurableRat hf + tendsto_atTop_one := tendsto_stieltjesOfMeasurableRat_atTop hf.measurable + tendsto_atBot_zero := tendsto_stieltjesOfMeasurableRat_atBot hf.measurable + set_integral a _ hs x := set_integral_stieltjesOfMeasurableRat hf a x hs end IsKernelCDF @@ -268,12 +278,6 @@ lemma cdfKernel_Iic (p : α × β) (x : ℝ) : rw [cdfKernel_apply p, (f p).measure_Iic (hf.tendsto_atBot_zero p)] simp ---lemma cdfKernel_unit_prod_Iic (a : α) (x : ℝ) : --- cdfKernel (fun p : Unit × α ↦ f p.2) (fun q ↦ (hf q).comp measurable_snd) ((), a) (Iic x) --- = cdfKernel f hf a (Iic x) := by --- simp only [cdfKernel_Iic] --- rw [todo3_unit_prod hf a] - end kernel section diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean index 0755bdc8943ae..face9bc639381 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean @@ -1,23 +1,204 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ import Mathlib.Analysis.SpecialFunctions.Log.Base import Mathlib.Probability.Martingale.Convergence import Mathlib.Probability.Kernel.Disintegration.BuildKernel +/-! +# KernelCDFBorel + +## Main definitions + +* `FooBar` + +## Main statements + +* `fooBar_unique` + +## Notation + + + +## Implementation details + + +issue with the following: joint measurability + +def M' (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : ℝ≥0∞ := + (((κ a).restrict (univ ×ˢ s)).fst.trim (ℱ.le n)).rnDeriv (((kernel.fst κ a)).trim (ℱ.le n)) t + + +## References + +* [F. Bar, *Quuxes*][bibkey] + +## Tags + +Foobars, barfoos +-/ + + open MeasureTheory Set Filter open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory -variable {α Ω : Type*} {mα : MeasurableSpace α} - [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] +variable {α β : Type*} {mα : MeasurableSpace α} + +section Move + +lemma snorm_restrict_le [NormedAddCommGroup β] {p : ℝ≥0∞} {f : α → β} {μ : Measure α} (s : Set α) : + snorm f p (μ.restrict s) ≤ snorm f p μ := + snorm_mono_measure f Measure.restrict_le_self + +lemma tendsto_snorm_restrict_zero {α β ι : Type*} {mα : MeasurableSpace α} [NormedAddCommGroup β] + {p : ℝ≥0∞} {f : ι → α → β} {μ : Measure α} {l : Filter ι} + (h : Tendsto (fun n ↦ snorm (f n) p μ) l (𝓝 0)) (s : Set α) : + Tendsto (fun n ↦ snorm (f n) p (μ.restrict s)) l (𝓝 0) := by + refine tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds h ?_ ?_ + · exact fun _ ↦ zero_le _ + · exact fun _ ↦ snorm_restrict_le _ + +lemma tendsto_integral_of_L1' {ι G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] + {μ : Measure α} + (f : α → G) (hfi : Integrable f μ) + {F : ι → α → G} {l : Filter ι} + (hFi : ∀ᶠ i in l, Integrable (F i) μ) + (hF : Tendsto (fun i ↦ snorm (F i - f) 1 μ) l (𝓝 0)) : + Tendsto (fun i ↦ ∫ x, F i x ∂μ) l (𝓝 (∫ x, f x ∂μ)) := by + refine tendsto_integral_of_L1 f hfi hFi ?_ + simp_rw [snorm_one_eq_lintegral_nnnorm, Pi.sub_apply] at hF + exact hF + +lemma tendsto_set_integral_of_L1 {ι G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] + {μ : Measure α} + (f : α → G) (hfi : Integrable f μ) + {F : ι → α → G} {l : Filter ι} + (hFi : ∀ᶠ i in l, Integrable (F i) μ) + (hF : Tendsto (fun i ↦ ∫⁻ x, ‖F i x - f x‖₊ ∂μ) l (𝓝 0)) (s : Set α) : + Tendsto (fun i ↦ ∫ x in s, F i x ∂μ) l (𝓝 (∫ x in s, f x ∂μ)) := by + refine tendsto_integral_of_L1 f hfi.restrict ?_ ?_ + · filter_upwards [hFi] with i hi using hi.restrict + · simp_rw [← snorm_one_eq_lintegral_nnnorm] at hF ⊢ + exact tendsto_snorm_restrict_zero hF s + +lemma tendsto_set_integral_of_L1' {ι G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] + {μ : Measure α} + (f : α → G) (hfi : Integrable f μ) + {F : ι → α → G} {l : Filter ι} + (hFi : ∀ᶠ i in l, Integrable (F i) μ) + (hF : Tendsto (fun i ↦ snorm (F i - f) 1 μ) l (𝓝 0)) (s : Set α) : + Tendsto (fun i ↦ ∫ x in s, F i x ∂μ) l (𝓝 (∫ x in s, f x ∂μ)) := by + refine tendsto_set_integral_of_L1 f hfi hFi ?_ s + simp_rw [snorm_one_eq_lintegral_nnnorm, Pi.sub_apply] at hF + exact hF + +lemma ae_eq_of_integral_eq_of_ae_le {μ : Measure α} {f g : α → ℝ} (hf : Integrable f μ) + (hg : Integrable g μ) (h_le : f ≤ᵐ[μ] g) (h_eq : ∫ a, f a ∂μ = ∫ a, g a ∂μ) : + f =ᵐ[μ] g := by + suffices g - f =ᵐ[μ] 0 by + filter_upwards [this] with a ha + symm + simpa only [Pi.sub_apply, Pi.zero_apply, sub_eq_zero] using ha + have h_eq' : ∫ a, (g - f) a ∂μ = 0 := by + simp_rw [Pi.sub_apply] + rwa [integral_sub hg hf, sub_eq_zero, eq_comm] + rwa [integral_eq_zero_iff_of_nonneg_ae _ (hg.sub hf)] at h_eq' + filter_upwards [h_le] with a ha + simpa + +lemma integral_tendsto_of_tendsto_of_monotone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} + (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Monotone fun n ↦ f n x) + (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : + Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by + let f' := fun n x ↦ f n x - f 0 x + have hf'_nonneg : ∀ᵐ x ∂μ, ∀ n, 0 ≤ f' n x := by + filter_upwards [h_mono] with a ha n + simp [ha (zero_le n)] + have hf'_meas : ∀ n, Integrable (f' n) μ := fun n ↦ (hf n).sub (hf 0) + suffices Tendsto (fun n ↦ ∫ x, f' n x ∂μ) atTop (𝓝 (∫ x, (F - f 0) x ∂μ)) by + rw [integral_sub' hF (hf 0)] at this + have h_sub : ∀ n, ∫ x, f' n x ∂μ = ∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ := by + intro n + simp only + rw [integral_sub (hf n) (hf 0)] + simp_rw [h_sub] at this + have h1 : (fun n ↦ ∫ x, f n x ∂μ) + = fun n ↦ (∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by ext n; abel + have h2 : ∫ x, F x ∂μ = (∫ x, F x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by abel + rw [h1, h2] + exact this.add tendsto_const_nhds + have hF_ge : 0 ≤ᵐ[μ] fun x ↦ (F - f 0) x := by + filter_upwards [h_tendsto, h_mono] with x hx_tendsto hx_mono + simp only [Pi.zero_apply, Pi.sub_apply, sub_nonneg] + exact ge_of_tendsto' hx_tendsto (fun n ↦ hx_mono (zero_le _)) + rw [ae_all_iff] at hf'_nonneg + simp_rw [integral_eq_lintegral_of_nonneg_ae (hf'_nonneg _) (hf'_meas _).1] + rw [integral_eq_lintegral_of_nonneg_ae hF_ge (hF.1.sub (hf 0).1)] + have h_cont := ENNReal.continuousOn_toReal.continuousAt + (x := ∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ) ?_ + swap + · rw [mem_nhds_iff] + refine ⟨Iio (∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ + 1), ?_, isOpen_Iio, ?_⟩ + · intro x + simp only [Pi.sub_apply, mem_Iio, ne_eq, mem_setOf_eq] + exact ne_top_of_lt + · simp only [Pi.sub_apply, mem_Iio] + refine ENNReal.lt_add_right ?_ one_ne_zero + rw [← ofReal_integral_eq_lintegral_ofReal] + · exact ENNReal.ofReal_ne_top + · exact hF.sub (hf 0) + · exact hF_ge + refine h_cont.tendsto.comp ?_ + refine lintegral_tendsto_of_tendsto_of_monotone ?_ ?_ ?_ + · exact fun n ↦ ((hf n).sub (hf 0)).aemeasurable.ennreal_ofReal + · filter_upwards [h_mono] with x hx + intro n m hnm + refine ENNReal.ofReal_le_ofReal ?_ + simp only [tsub_le_iff_right, sub_add_cancel] + exact hx hnm + · filter_upwards [h_tendsto] with x hx + refine (ENNReal.continuous_ofReal.tendsto _).comp ?_ + simp only [Pi.sub_apply] + exact Tendsto.sub hx tendsto_const_nhds + +lemma integral_tendsto_of_tendsto_of_antitone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} + (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Antitone fun n ↦ f n x) + (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : + Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by + suffices Tendsto (fun n ↦ ∫ x, -f n x ∂μ) atTop (𝓝 (∫ x, -F x ∂μ)) by + suffices Tendsto (fun n ↦ ∫ x, - -f n x ∂μ) atTop (𝓝 (∫ x, - -F x ∂μ)) by + simp_rw [neg_neg] at this + exact this + convert this.neg <;> rw [integral_neg] + refine integral_tendsto_of_tendsto_of_monotone (fun n ↦ (hf n).neg) hF.neg ?_ ?_ + · filter_upwards [h_mono] with x hx + intro n m hnm + simp only [neg_le_neg_iff] + exact hx hnm + · filter_upwards [h_tendsto] with x hx + exact hx.neg + +lemma tendsto_nat_ceil_atTop {α : Type*} [LinearOrderedSemiring α] [FloorSemiring α] : + Tendsto (fun x : α ↦ ⌈x⌉₊) atTop atTop := by + refine Nat.ceil_mono.tendsto_atTop_atTop (fun x ↦ ⟨x, ?_⟩) + simp only [Nat.ceil_natCast, le_refl] + +lemma isCoboundedUnder_le_of_eventually_le {ι α : Type*} [Preorder α] (l : Filter ι) [NeBot l] + {f : ι → α} {x : α} (hf : ∀ᶠ i in l, x ≤ f i) : + IsCoboundedUnder (· ≤ ·) l f := + IsBoundedUnder.isCoboundedUnder_le ⟨x, hf⟩ -lemma tendsto_atTop_atBot_iff_of_antitone {α β : Type*} - [Nonempty α] [SemilatticeSup α] [Preorder β] {f : α → β} - (hf : Antitone f) : - Tendsto f atTop atBot ↔ ∀ b : β, ∃ a : α, f a ≤ b := - @tendsto_atTop_atTop_iff_of_monotone _ βᵒᵈ _ _ _ _ hf +lemma isCoboundedUnder_le_of_le {ι α : Type*} [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} + {x : α} (hf : ∀ i, x ≤ f i) : + IsCoboundedUnder (· ≤ ·) l f := + isCoboundedUnder_le_of_eventually_le l (eventually_of_forall hf) -section Real +end Move section dissection_system @@ -91,26 +272,7 @@ lemma I_succ_union (n : ℕ) (k : ℤ) : I (n+1) (2 * k) ∪ I (n+1) (2 * k + 1) norm_num rw [one_div, mul_add, ← mul_assoc, inv_mul_cancel two_ne_zero, one_mul] --- todo : `Filtration` should be renamed to `filtration` -def ℱ : Filtration ℕ (borel ℝ) where - seq := fun n ↦ MeasurableSpace.generateFrom {s | ∃ k, s = I n k} - mono' := by - refine monotone_nat_of_le_succ ?_ - intro n - refine MeasurableSpace.generateFrom_le fun s ⟨k, hs⟩ ↦ ?_ - rw [hs, ← I_succ_union n k] - refine MeasurableSet.union ?_ ?_ - · exact MeasurableSpace.measurableSet_generateFrom ⟨2 * k, rfl⟩ - · exact MeasurableSpace.measurableSet_generateFrom ⟨2 * k + 1, rfl⟩ - le' := fun n ↦ by - refine MeasurableSpace.generateFrom_le fun s ⟨k, hs⟩ ↦ ?_ - rw [hs] - exact measurableSet_I n k - -lemma measurableSet_ℱ_I (n : ℕ) (k : ℤ) : MeasurableSet[ℱ n] (I n k) := - MeasurableSpace.measurableSet_generateFrom ⟨k, rfl⟩ - -noncomputable def indexI (n : ℕ) (t : ℝ) : ℤ := Int.floor (t * 2^n) +noncomputable def indexI (n : ℕ) (t : ℝ) : ℤ := Int.floor (t * 2 ^ n) lemma mem_I_indexI (n : ℕ) (t : ℝ) : t ∈ I n (indexI n t) := by rw [indexI, I] @@ -135,15 +297,6 @@ lemma indexI_of_mem (n : ℕ) (k : ℤ) (t : ℝ) (ht : t ∈ I n k) : indexI n lemma mem_I_iff_indexI (n : ℕ) (k : ℤ) (t : ℝ) : t ∈ I n k ↔ indexI n t = k := ⟨fun h ↦ indexI_of_mem n k t h, fun h ↦ h ▸ mem_I_indexI n t⟩ -lemma measurable_indexI (n : ℕ) : Measurable[ℱ n] (indexI n) := by - unfold indexI - refine @measurable_to_countable' ℤ ℝ _ _ (ℱ n) _ (fun k ↦ ?_) - have : (fun t ↦ ⌊t * (2 : ℝ) ^ n⌋) ⁻¹' {k} = I n k := by - ext t - simp only [mem_I_iff_floor, mem_preimage, mem_singleton_iff] - rw [this] - exact measurableSet_ℱ_I n k - lemma iUnion_I (n : ℕ) : ⋃ k, I n k = univ := by ext x simp only [mem_iUnion, mem_univ, iff_true] @@ -163,7 +316,7 @@ lemma iUnion_ge_I (n : ℕ) (t : ℝ) : lemma iInter_biUnion_I (x : ℝ) : ⋂ n, ⋃ (k) (_ : indexI n x ≤ k), I n k = Ici x := by ext t simp [iUnion_ge_I] - refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + refine ⟨fun h ↦ ?_, fun h n ↦ le_trans ?_ h⟩ · by_contra h_lt push_neg at h_lt have h_pos : ∀ i, 0 < (2 : ℝ) ^ i := fun i ↦ by positivity @@ -171,8 +324,7 @@ lemma iInter_biUnion_I (x : ℝ) : ⋂ n, ⋃ (k) (_ : indexI n x ≤ k), I n k obtain ⟨i, hi⟩ : ∃ i, 1 < (x - t) * 2 ^ i := by suffices ∃ i : ℝ, 1 ≤ (x - t) * 2 ^ i by obtain ⟨i, hi⟩ := this - use ⌈i⌉₊ + 1 - refine hi.trans_lt ?_ + refine ⟨⌈i⌉₊ + 1, hi.trans_lt ?_⟩ gcongr · simp [h_lt] · refine ((Real.rpow_lt_rpow_left_iff one_lt_two).mpr (?_ : i < ⌈i⌉₊ + 1)).trans_eq ?_ @@ -196,12 +348,38 @@ lemma iInter_biUnion_I (x : ℝ) : ⋂ n, ⋃ (k) (_ : indexI n x ≤ k), I n k have h'' : ↑⌈x * 2 ^ i⌉ < 2 ^ i * x := h'.trans_lt hi rw [← not_le, mul_comm] at h'' exact h'' (Int.le_ceil _) - · intro n - refine le_trans ?_ h - rw [← div_eq_mul_inv, div_le_iff] + · rw [← div_eq_mul_inv, div_le_iff] · exact Int.floor_le (x * 2 ^ n) · positivity +-- todo : `Filtration` should be renamed to `filtration` +def ℱ : Filtration ℕ (borel ℝ) where + seq := fun n ↦ MeasurableSpace.generateFrom {s | ∃ k, s = I n k} + mono' := by + refine monotone_nat_of_le_succ ?_ + intro n + refine MeasurableSpace.generateFrom_le fun s ⟨k, hs⟩ ↦ ?_ + rw [hs, ← I_succ_union n k] + refine MeasurableSet.union ?_ ?_ + · exact MeasurableSpace.measurableSet_generateFrom ⟨2 * k, rfl⟩ + · exact MeasurableSpace.measurableSet_generateFrom ⟨2 * k + 1, rfl⟩ + le' := fun n ↦ by + refine MeasurableSpace.generateFrom_le fun s ⟨k, hs⟩ ↦ ?_ + rw [hs] + exact measurableSet_I n k + +lemma measurableSet_ℱ_I (n : ℕ) (k : ℤ) : MeasurableSet[ℱ n] (I n k) := + MeasurableSpace.measurableSet_generateFrom ⟨k, rfl⟩ + +lemma measurable_indexI (n : ℕ) : Measurable[ℱ n] (indexI n) := by + unfold indexI + refine @measurable_to_countable' ℤ ℝ _ _ (ℱ n) _ (fun k ↦ ?_) + have : (fun t ↦ ⌊t * (2 : ℝ) ^ n⌋) ⁻¹' {k} = I n k := by + ext t + simp only [mem_I_iff_floor, mem_preimage, mem_singleton_iff] + rw [this] + exact measurableSet_ℱ_I n k + lemma iSup_ℱ : ⨆ n, ℱ n = borel ℝ := by refine le_antisymm ?_ ?_ · rw [iSup_le_iff] @@ -216,12 +394,7 @@ lemma iSup_ℱ : ⨆ n, ℱ n = borel ℝ := by end dissection_system -variable {β : Type*} [MeasurableSpace β] - --- issue with the following: joint measurability ---noncomputable ---def M' (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : ℝ≥0∞ := --- (((κ a).restrict (univ ×ˢ s)).fst.trim (ℱ.le n)).rnDeriv (((kernel.fst κ a)).trim (ℱ.le n)) t +variable [MeasurableSpace β] noncomputable def M (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : ℝ := @@ -301,8 +474,7 @@ lemma m_le_one (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h -lemma snorm_m_le (κ : kernel α (ℝ × β)) [IsFiniteKernel (kernel.fst κ)] - (a : α) (s : Set β) (n : ℕ) : +lemma snorm_m_le (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) : snorm (M κ a s n) 1 (kernel.fst κ a) ≤ kernel.fst κ a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun x ↦ ?_))).trans ?_ · simp only [Real.norm_eq_abs, abs_of_nonneg (m_nonneg κ a s n x), m_le_one κ a s n x] @@ -620,18 +792,6 @@ lemma tendsto_snorm_one_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [Is simp · simp -lemma snorm_restrict_le [NormedAddCommGroup β] {p : ℝ≥0∞} {f : α → β} {μ : Measure α} (s : Set α) : - snorm f p (μ.restrict s) ≤ snorm f p μ := - snorm_mono_measure f Measure.restrict_le_self - -lemma tendsto_snorm_restrict_zero {α β ι : Type*} {mα : MeasurableSpace α} [NormedAddCommGroup β] - {p : ℝ≥0∞} {f : ι → α → β} {μ : Measure α} {l : Filter ι} - (h : Tendsto (fun n ↦ snorm (f n) p μ) l (𝓝 0)) (s : Set α) : - Tendsto (fun n ↦ snorm (f n) p (μ.restrict s)) l (𝓝 0) := by - refine tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds h ?_ ?_ - · exact fun _ ↦ zero_le _ - · exact fun _ ↦ snorm_restrict_le _ - lemma tendsto_snorm_one_restrict_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsFiniteKernel κ] {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : Tendsto (fun n ↦ snorm (M κ a s n - ℱ.limitProcess (M κ a s) (kernel.fst κ a)) 1 @@ -675,35 +835,19 @@ lemma mLimsup_mono_set (κ : kernel α (ℝ × β)) (a : α) {s s' : Set β} (h rw [MLimsup, MLimsup] refine limsup_le_limsup ?_ ?_ ?_ · exact eventually_of_forall (fun n ↦ m_mono_set κ a h n t) - · -- todo: extract lemma (of find it) - refine ⟨0, ?_⟩ - simp only [eventually_map, eventually_atTop, ge_iff_le, forall_exists_index] - intro x n h - specialize h n le_rfl - exact (m_nonneg _ _ _ _ _).trans h - · -- todo: extract lemma (of find it) - refine ⟨1, ?_⟩ - simp only [eventually_map, eventually_atTop, ge_iff_le] - exact ⟨0, fun n _ ↦ m_le_one _ _ _ _ _⟩ + · exact isCoboundedUnder_le_of_le atTop (fun i ↦ m_nonneg _ _ _ _ _) + · exact isBoundedUnder_of ⟨1, fun n ↦ m_le_one _ _ _ _ _⟩ lemma mLimsup_nonneg (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (t : ℝ) : 0 ≤ MLimsup κ a s t := by refine le_limsup_of_frequently_le ?_ ?_ · exact frequently_of_forall (fun n ↦ m_nonneg _ _ _ _ _) - · -- todo: extract lemma (of find it) - refine ⟨1, ?_⟩ - simp only [eventually_map, eventually_atTop, ge_iff_le] - exact ⟨0, fun n _ ↦ m_le_one _ _ _ _ _⟩ + · exact isBoundedUnder_of ⟨1, fun n ↦ m_le_one _ _ _ _ _⟩ lemma mLimsup_le_one (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (t : ℝ) : MLimsup κ a s t ≤ 1 := by refine limsup_le_of_le ?_ ?_ - · -- todo: extract lemma (of find it) - refine ⟨0, ?_⟩ - simp only [eventually_map, eventually_atTop, ge_iff_le, forall_exists_index] - intro x n h - specialize h n le_rfl - exact (m_nonneg _ _ _ _ _).trans h + · exact isCoboundedUnder_le_of_le atTop (fun i ↦ m_nonneg _ _ _ _ _) · exact eventually_of_forall (fun n ↦ m_le_one _ _ _ _ _) lemma mLimsup_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) : @@ -712,10 +856,9 @@ lemma mLimsup_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) : rw [← ae_all_iff] at h filter_upwards [h] with t ht rw [MLimsup] - simp_rw [ht] - rw [limsup_const] -- should be simp + simp [ht] -lemma snorm_mLimsup_le (κ : kernel α (ℝ × β)) [IsFiniteKernel (kernel.fst κ)] +lemma snorm_mLimsup_le (κ : kernel α (ℝ × β)) (a : α) (s : Set β) : snorm (fun t ↦ MLimsup κ a s t) 1 (kernel.fst κ a) ≤ kernel.fst κ a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ @@ -731,40 +874,6 @@ lemma integrable_mLimsup (κ : kernel α (ℝ × β)) [IsFiniteKernel (kernel.fs · exact measurable_mLimsup_right κ hs a · exact (snorm_mLimsup_le κ a s).trans_lt (measure_lt_top _ _) -lemma tendsto_integral_of_L1' {ι G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] - {μ : Measure α} - (f : α → G) (hfi : Integrable f μ) - {F : ι → α → G} {l : Filter ι} - (hFi : ∀ᶠ i in l, Integrable (F i) μ) - (hF : Tendsto (fun i ↦ snorm (F i - f) 1 μ) l (𝓝 0)) : - Tendsto (fun i ↦ ∫ x, F i x ∂μ) l (𝓝 (∫ x, f x ∂μ)) := by - refine tendsto_integral_of_L1 f hfi hFi ?_ - simp_rw [snorm_one_eq_lintegral_nnnorm, Pi.sub_apply] at hF - exact hF - -lemma tendsto_set_integral_of_L1 {ι G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] - {μ : Measure α} - (f : α → G) (hfi : Integrable f μ) - {F : ι → α → G} {l : Filter ι} - (hFi : ∀ᶠ i in l, Integrable (F i) μ) - (hF : Tendsto (fun i ↦ ∫⁻ x, ‖F i x - f x‖₊ ∂μ) l (𝓝 0)) (s : Set α) : - Tendsto (fun i ↦ ∫ x in s, F i x ∂μ) l (𝓝 (∫ x in s, f x ∂μ)) := by - refine tendsto_integral_of_L1 f hfi.restrict ?_ ?_ - · filter_upwards [hFi] with i hi using hi.restrict - · simp_rw [← snorm_one_eq_lintegral_nnnorm] at hF ⊢ - exact tendsto_snorm_restrict_zero hF s - -lemma tendsto_set_integral_of_L1' {ι G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] - {μ : Measure α} - (f : α → G) (hfi : Integrable f μ) - {F : ι → α → G} {l : Filter ι} - (hFi : ∀ᶠ i in l, Integrable (F i) μ) - (hF : Tendsto (fun i ↦ snorm (F i - f) 1 μ) l (𝓝 0)) (s : Set α) : - Tendsto (fun i ↦ ∫ x in s, F i x ∂μ) l (𝓝 (∫ x in s, f x ∂μ)) := by - refine tendsto_set_integral_of_L1 f hfi hFi ?_ s - simp_rw [snorm_one_eq_lintegral_nnnorm, Pi.sub_apply] at hF - exact hF - lemma tendsto_set_integral_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : Tendsto (fun i ↦ ∫ x in A, M κ a s i x ∂(kernel.fst κ) a) atTop @@ -879,92 +988,6 @@ lemma tendsto_integral_mLimsup_of_antitone (κ : kernel α (ℝ × β)) [IsFinit rw [← prod_iInter, hs_iInter] simp only [ne_eq, prod_empty, OuterMeasure.empty', forall_exists_index] -lemma ae_eq_of_integral_eq_of_ae_le {μ : Measure α} {f g : α → ℝ} (hf : Integrable f μ) - (hg : Integrable g μ) (h_le : f ≤ᵐ[μ] g) (h_eq : ∫ a, f a ∂μ = ∫ a, g a ∂μ) : - f =ᵐ[μ] g := by - suffices g - f =ᵐ[μ] 0 by - filter_upwards [this] with a ha - symm - simpa only [Pi.sub_apply, Pi.zero_apply, sub_eq_zero] using ha - have h_eq' : ∫ a, (g - f) a ∂μ = 0 := by - simp_rw [Pi.sub_apply] - rwa [integral_sub hg hf, sub_eq_zero, eq_comm] - rwa [integral_eq_zero_iff_of_nonneg_ae _ (hg.sub hf)] at h_eq' - filter_upwards [h_le] with a ha - simpa - -lemma integral_tendsto_of_tendsto_of_monotone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} - (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Monotone fun n ↦ f n x) - (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : - Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by - let f' := fun n x ↦ f n x - f 0 x - have hf'_nonneg : ∀ᵐ x ∂μ, ∀ n, 0 ≤ f' n x := by - filter_upwards [h_mono] with a ha n - simp [ha (zero_le n)] - have hf'_meas : ∀ n, Integrable (f' n) μ := fun n ↦ (hf n).sub (hf 0) - suffices Tendsto (fun n ↦ ∫ x, f' n x ∂μ) atTop (𝓝 (∫ x, (F - f 0) x ∂μ)) by - rw [integral_sub' hF (hf 0)] at this - have h_sub : ∀ n, ∫ x, f' n x ∂μ = ∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ := by - intro n - simp only - rw [integral_sub (hf n) (hf 0)] - simp_rw [h_sub] at this - have h1 : (fun n ↦ ∫ x, f n x ∂μ) - = fun n ↦ (∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by ext n; abel - have h2 : ∫ x, F x ∂μ = (∫ x, F x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by abel - rw [h1, h2] - exact this.add tendsto_const_nhds - have hF_ge : 0 ≤ᵐ[μ] fun x ↦ (F - f 0) x := by - filter_upwards [h_tendsto, h_mono] with x hx_tendsto hx_mono - simp only [Pi.zero_apply, Pi.sub_apply, sub_nonneg] - exact ge_of_tendsto' hx_tendsto (fun n ↦ hx_mono (zero_le _)) - rw [ae_all_iff] at hf'_nonneg - simp_rw [integral_eq_lintegral_of_nonneg_ae (hf'_nonneg _) (hf'_meas _).1] - rw [integral_eq_lintegral_of_nonneg_ae hF_ge (hF.1.sub (hf 0).1)] - have h_cont := ENNReal.continuousOn_toReal.continuousAt - (x := ∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ) ?_ - swap - · rw [mem_nhds_iff] - refine ⟨Iio (∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ + 1), ?_, isOpen_Iio, ?_⟩ - · intro x - simp only [Pi.sub_apply, mem_Iio, ne_eq, mem_setOf_eq] - exact ne_top_of_lt - · simp only [Pi.sub_apply, mem_Iio] - refine ENNReal.lt_add_right ?_ one_ne_zero - rw [← ofReal_integral_eq_lintegral_ofReal] - · exact ENNReal.ofReal_ne_top - · exact hF.sub (hf 0) - · exact hF_ge - refine h_cont.tendsto.comp ?_ - refine lintegral_tendsto_of_tendsto_of_monotone ?_ ?_ ?_ - · exact fun n ↦ ((hf n).sub (hf 0)).aemeasurable.ennreal_ofReal - · filter_upwards [h_mono] with x hx - intro n m hnm - refine ENNReal.ofReal_le_ofReal ?_ - simp only [tsub_le_iff_right, sub_add_cancel] - exact hx hnm - · filter_upwards [h_tendsto] with x hx - refine (ENNReal.continuous_ofReal.tendsto _).comp ?_ - simp only [Pi.sub_apply] - exact Tendsto.sub hx tendsto_const_nhds - -lemma integral_tendsto_of_tendsto_of_antitone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} - (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Antitone fun n ↦ f n x) - (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : - Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by - suffices Tendsto (fun n ↦ ∫ x, -f n x ∂μ) atTop (𝓝 (∫ x, -F x ∂μ)) by - suffices Tendsto (fun n ↦ ∫ x, - -f n x ∂μ) atTop (𝓝 (∫ x, - -F x ∂μ)) by - simp_rw [neg_neg] at this - exact this - convert this.neg <;> rw [integral_neg] - refine integral_tendsto_of_tendsto_of_monotone (fun n ↦ (hf n).neg) hF.neg ?_ ?_ - · filter_upwards [h_mono] with x hx - intro n m hnm - simp only [neg_le_neg_iff] - exact hx hnm - · filter_upwards [h_tendsto] with x hx - exact hx.neg - lemma tendsto_mLimsup_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : @@ -1114,11 +1137,6 @@ lemma mLimsupIic_nonneg (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : lemma mLimsupIic_le_one (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : mLimsupIic κ a t q ≤ 1 := mLimsup_le_one κ a _ t -theorem tendsto_nat_ceil_atTop {α : Type*} [LinearOrderedSemiring α] [FloorSemiring α] : - Tendsto (fun x : α ↦ ⌈x⌉₊) atTop atTop := by - refine Nat.ceil_mono.tendsto_atTop_atTop (fun x ↦ ⟨x, ?_⟩) - simp only [Nat.ceil_natCast, le_refl] - lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ a t q) atTop (𝓝 1) := by suffices ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ a t n) atTop (𝓝 1) by @@ -1187,8 +1205,7 @@ lemma integrable_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel (kerne Integrable (fun t ↦ mLimsupIic κ a t q) (kernel.fst κ a) := integrable_mLimsup _ _ measurableSet_Iic -lemma bddBelow_range_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel (kernel.fst κ)] - (a : α) (t : ℝ) (q : ℚ) : +lemma bddBelow_range_mLimsupIic (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : BddBelow (range fun (r : Ioi q) ↦ mLimsupIic κ a t r) := by refine ⟨0, ?_⟩ rw [mem_lowerBounds] @@ -1247,8 +1264,6 @@ lemma iInf_rat_gt_mLimsupIic_eq (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ end Iic_Q -section Rat - lemma isRatStieltjesPoint_mLimsupIic_ae (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), IsRatStieltjesPoint (fun p q ↦ mLimsupIic κ p.1 p.2 q) (a, t) := by filter_upwards [tendsto_atTop_mLimsupIic κ a, tendsto_atBot_mLimsupIic κ a, @@ -1268,20 +1283,13 @@ lemma isRatKernelCDF_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ integrable := integrable_mLimsupIic κ isCDF := fun _ _ hs _ ↦ set_integral_mLimsupIic _ _ _ hs -end Rat - -section KernelCDF - noncomputable def mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : α × ℝ → StieltjesFunction := - todo3 (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2) (isRatKernelCDF_mLimsupIic κ).measurable + stieltjesOfMeasurableRat (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2) + (isRatKernelCDF_mLimsupIic κ).measurable lemma isKernelCDF_mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : IsKernelCDF (mLimsupCDF κ) κ (kernel.fst κ) := - isKernelCDF_todo3 (isRatKernelCDF_mLimsupIic κ) - -end KernelCDF - -end Real + isKernelCDF_stieltjesOfMeasurableRat (isRatKernelCDF_mLimsupIic κ) end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean b/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean index 9120601ff26a8..909246fdf43a3 100644 --- a/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean +++ b/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean @@ -225,15 +225,16 @@ variable {f : α → ℚ → ℝ} (hf : IsCDFLike f) /-- Conditional cdf of the measure given the value on `α`, as a plain function. This is an auxiliary definition used to define `cond_cdf`. -/ -noncomputable irreducible_def todo1 (f : α → ℚ → ℝ) : α → ℝ → ℝ := +noncomputable irreducible_def IsCDFLike.stieltjesFunctionAux (f : α → ℚ → ℝ) : α → ℝ → ℝ := fun a t ↦ ⨅ r : { r' : ℚ // t < r' }, f a r -lemma todo1_def' (f : α → ℚ → ℝ) (a : α) : - todo1 f a = fun (t : ℝ) ↦ ⨅ r : { r' : ℚ // t < r' }, f a r := by ext t; exact todo1_def f a t +lemma IsCDFLike.stieltjesFunctionAux_def' (f : α → ℚ → ℝ) (a : α) : + IsCDFLike.stieltjesFunctionAux f a = fun (t : ℝ) ↦ ⨅ r : { r' : ℚ // t < r' }, f a r := by + ext t; exact IsCDFLike.stieltjesFunctionAux_def f a t -lemma todo1_eq (a : α) (r : ℚ) : - todo1 f a r = f a r := by - rw [← hf.iInf_rat_gt_eq a r, todo1] +lemma IsCDFLike.stieltjesFunctionAux_eq (a : α) (r : ℚ) : + IsCDFLike.stieltjesFunctionAux f a r = f a r := by + rw [← hf.iInf_rat_gt_eq a r, IsCDFLike.stieltjesFunctionAux] refine Equiv.iInf_congr ?_ ?_ · exact { toFun := fun t ↦ ⟨t.1, mod_cast t.2⟩ @@ -243,65 +244,69 @@ lemma todo1_eq (a : α) (r : ℚ) : · intro t simp only [Equiv.coe_fn_mk, Subtype.coe_mk] -lemma todo1_unit_prod (a : α) : - todo1 (fun (p : Unit × α) ↦ f p.2) ((), a) = todo1 f a := by simp_rw [todo1_def'] +lemma IsCDFLike.stieltjesFunctionAux_unit_prod (a : α) : + IsCDFLike.stieltjesFunctionAux (fun (p : Unit × α) ↦ f p.2) ((), a) = + IsCDFLike.stieltjesFunctionAux f a := by simp_rw [IsCDFLike.stieltjesFunctionAux_def'] -lemma todo1_nonneg (a : α) (r : ℝ) : 0 ≤ todo1 f a r := by +lemma IsCDFLike.stieltjesFunctionAux_nonneg (a : α) (r : ℝ) : + 0 ≤ IsCDFLike.stieltjesFunctionAux f a r := by have : Nonempty { r' : ℚ // r < ↑r' } := by obtain ⟨r, hrx⟩ := exists_rat_gt r exact ⟨⟨r, hrx⟩⟩ - rw [todo1_def] + rw [IsCDFLike.stieltjesFunctionAux_def] exact le_ciInf fun r' ↦ hf.nonneg a _ lemma bddBelow_range_gt (a : α) (x : ℝ) : BddBelow (range fun r : { r' : ℚ // x < ↑r' } ↦ f a r) := by refine ⟨0, fun z ↦ ?_⟩; rintro ⟨u, rfl⟩; exact hf.nonneg a _ -lemma monotone_todo1 (a : α) : Monotone (todo1 f a) := by +lemma IsCDFLike.monotone_stieltjesFunctionAux (a : α) : + Monotone (IsCDFLike.stieltjesFunctionAux f a) := by intro x y hxy have : Nonempty { r' : ℚ // y < ↑r' } := by obtain ⟨r, hrx⟩ := exists_rat_gt y exact ⟨⟨r, hrx⟩⟩ - simp_rw [todo1_def] + simp_rw [IsCDFLike.stieltjesFunctionAux_def] refine le_ciInf fun r ↦ (ciInf_le ?_ ?_).trans_eq ?_ · exact bddBelow_range_gt hf a x · exact ⟨r.1, hxy.trans_lt r.prop⟩ · rfl -lemma continuousWithinAt_todo1_Ici (a : α) (x : ℝ) : - ContinuousWithinAt (todo1 f a) (Ici x) x := by +lemma IsCDFLike.continuousWithinAt_stieltjesFunctionAux_Ici (a : α) (x : ℝ) : + ContinuousWithinAt (IsCDFLike.stieltjesFunctionAux f a) (Ici x) x := by rw [← continuousWithinAt_Ioi_iff_Ici] - convert Monotone.tendsto_nhdsWithin_Ioi (monotone_todo1 hf a) x + convert Monotone.tendsto_nhdsWithin_Ioi (monotone_stieltjesFunctionAux hf a) x rw [sInf_image'] - have h' : ⨅ r : Ioi x, todo1 f a r = ⨅ r : { r' : ℚ // x < r' }, todo1 f a r := by - refine Real.iInf_Ioi_eq_iInf_rat_gt x ?_ (monotone_todo1 hf a) + have h' : ⨅ r : Ioi x, stieltjesFunctionAux f a r + = ⨅ r : { r' : ℚ // x < r' }, stieltjesFunctionAux f a r := by + refine Real.iInf_Ioi_eq_iInf_rat_gt x ?_ (monotone_stieltjesFunctionAux hf a) refine ⟨0, fun z ↦ ?_⟩ rintro ⟨u, -, rfl⟩ - exact todo1_nonneg hf a u + exact stieltjesFunctionAux_nonneg hf a u have h'' : - ⨅ r : { r' : ℚ // x < r' }, todo1 f a r = + ⨅ r : { r' : ℚ // x < r' }, stieltjesFunctionAux f a r = ⨅ r : { r' : ℚ // x < r' }, f a r := by congr with r - exact todo1_eq hf a r + exact stieltjesFunctionAux_eq hf a r rw [h', h'', ContinuousWithinAt] congr! - rw [todo1_def] + rw [stieltjesFunctionAux_def] /-! ### Conditional cdf -/ /-- Conditional cdf of the measure given the value on `α`, as a Stieltjes function. -/ noncomputable def IsCDFLike.stieltjesFunction (a : α) : StieltjesFunction where - toFun := todo1 f a - mono' := monotone_todo1 hf a - right_continuous' x := continuousWithinAt_todo1_Ici hf a x + toFun := stieltjesFunctionAux f a + mono' := monotone_stieltjesFunctionAux hf a + right_continuous' x := continuousWithinAt_stieltjesFunctionAux_Ici hf a x lemma IsCDFLike.stieltjesFunction_eq (a : α) (r : ℚ) : hf.stieltjesFunction a r = f a r := - todo1_eq hf a r + stieltjesFunctionAux_eq hf a r /-- The conditional cdf is non-negative for all `a : α`. -/ lemma IsCDFLike.stieltjesFunction_nonneg (a : α) (r : ℝ) : 0 ≤ hf.stieltjesFunction a r := - todo1_nonneg hf a r + stieltjesFunctionAux_nonneg hf a r /-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ lemma IsCDFLike.stieltjesFunction_le_one (a : α) (x : ℝ) : hf.stieltjesFunction a x ≤ 1 := by @@ -405,66 +410,77 @@ end Measure end IsCDFLike.stieltjesFunction -section todo3 +section stieltjesOfMeasurableRat variable {f : α → ℚ → ℝ} noncomputable -def todo3 (f : α → ℚ → ℝ) (hf : ∀ q, Measurable fun a ↦ f a q) : α → StieltjesFunction := +def stieltjesOfMeasurableRat (f : α → ℚ → ℝ) (hf : ∀ q, Measurable fun a ↦ f a q) : + α → StieltjesFunction := (isCDFLike_toCDFLike hf).stieltjesFunction -lemma todo3_eq (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℚ) : - todo3 f hf a r = toCDFLike f a r := IsCDFLike.stieltjesFunction_eq _ a r +lemma stieltjesOfMeasurableRat_eq (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℚ) : + stieltjesOfMeasurableRat f hf a r = toCDFLike f a r := IsCDFLike.stieltjesFunction_eq _ a r -lemma todo3_unit_prod (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : - todo3 (fun (p : Unit × α) ↦ f p.2) (fun q ↦ (hf q).comp measurable_snd) ((), a) - = todo3 f hf a := by - simp_rw [todo3,IsCDFLike.stieltjesFunction, ← todo1_unit_prod a] +lemma stieltjesOfMeasurableRat_unit_prod (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : + stieltjesOfMeasurableRat (fun (p : Unit × α) ↦ f p.2) + (fun q ↦ (hf q).comp measurable_snd) ((), a) + = stieltjesOfMeasurableRat f hf a := by + simp_rw [stieltjesOfMeasurableRat,IsCDFLike.stieltjesFunction, + ← IsCDFLike.stieltjesFunctionAux_unit_prod a] congr with x congr 1 with p : 1 cases p with | mk _ b => rw [← toCDFLike_unit_prod b] /-- The conditional cdf is non-negative for all `a : α`. -/ -lemma todo3_nonneg (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℝ) : - 0 ≤ todo3 f hf a r := IsCDFLike.stieltjesFunction_nonneg _ a r +lemma stieltjesOfMeasurableRat_nonneg (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℝ) : + 0 ≤ stieltjesOfMeasurableRat f hf a r := IsCDFLike.stieltjesFunction_nonneg _ a r /-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ -lemma todo3_le_one (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : - todo3 f hf a x ≤ 1 := IsCDFLike.stieltjesFunction_le_one _ a x +lemma stieltjesOfMeasurableRat_le_one (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : + stieltjesOfMeasurableRat f hf a x ≤ 1 := IsCDFLike.stieltjesFunction_le_one _ a x /-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ -lemma tendsto_todo3_atBot (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : - Tendsto (todo3 f hf a) atBot (𝓝 0) := IsCDFLike.tendsto_stieltjesFunction_atBot _ a +lemma tendsto_stieltjesOfMeasurableRat_atBot (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : + Tendsto (stieltjesOfMeasurableRat f hf a) atBot (𝓝 0) := + IsCDFLike.tendsto_stieltjesFunction_atBot _ a /-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ -lemma tendsto_todo3_atTop (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : - Tendsto (todo3 f hf a) atTop (𝓝 1) := IsCDFLike.tendsto_stieltjesFunction_atTop _ a +lemma tendsto_stieltjesOfMeasurableRat_atTop (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : + Tendsto (stieltjesOfMeasurableRat f hf a) atTop (𝓝 1) := + IsCDFLike.tendsto_stieltjesFunction_atTop _ a /-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ -lemma measurable_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : - Measurable fun a ↦ todo3 f hf a x := IsCDFLike.measurable_stieltjesFunction _ x +lemma measurable_stieltjesOfMeasurableRat (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : + Measurable fun a ↦ stieltjesOfMeasurableRat f hf a x := + IsCDFLike.measurable_stieltjesFunction _ x /-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ -lemma stronglyMeasurable_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : - StronglyMeasurable fun a ↦ todo3 f hf a x := IsCDFLike.stronglyMeasurable_stieltjesFunction _ x +lemma stronglyMeasurable_stieltjesOfMeasurableRat (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : + StronglyMeasurable fun a ↦ stieltjesOfMeasurableRat f hf a x := + IsCDFLike.stronglyMeasurable_stieltjesFunction _ x section Measure -lemma measure_todo3_Iic (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : - (todo3 f hf a).measure (Iic x) = ENNReal.ofReal (todo3 f hf a x) := +lemma measure_stieltjesOfMeasurableRat_Iic (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : + (stieltjesOfMeasurableRat f hf a).measure (Iic x) + = ENNReal.ofReal (stieltjesOfMeasurableRat f hf a x) := IsCDFLike.measure_stieltjesFunction_Iic _ _ _ -lemma measure_todo3_univ (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : - (todo3 f hf a).measure univ = 1 := IsCDFLike.measure_stieltjesFunction_univ _ _ +lemma measure_stieltjesOfMeasurableRat_univ (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : + (stieltjesOfMeasurableRat f hf a).measure univ = 1 := + IsCDFLike.measure_stieltjesFunction_univ _ _ -instance instIsProbabilityMeasure_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : - IsProbabilityMeasure (todo3 f hf a).measure := +instance instIsProbabilityMeasure_stieltjesOfMeasurableRat + (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : + IsProbabilityMeasure (stieltjesOfMeasurableRat f hf a).measure := IsCDFLike.instIsProbabilityMeasure_stieltjesFunction _ _ -lemma measurable_measure_todo3 (hf : ∀ q, Measurable fun a ↦ f a q) : - Measurable fun a ↦ (todo3 f hf a).measure := IsCDFLike.measurable_measure_stieltjesFunction _ +lemma measurable_measure_stieltjesOfMeasurableRat (hf : ∀ q, Measurable fun a ↦ f a q) : + Measurable fun a ↦ (stieltjesOfMeasurableRat f hf a).measure := + IsCDFLike.measurable_measure_stieltjesFunction _ end Measure -end todo3 +end stieltjesOfMeasurableRat From cfde47b3fdf9b48c9c46b94c8e1a197b2bcb23a5 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 15 Feb 2024 19:39:13 +0100 Subject: [PATCH 020/129] various minor --- .../Function/LpSeminorm/Basic.lean | 4 + Mathlib/MeasureTheory/Integral/Bochner.lean | 52 +++++- Mathlib/Probability/Kernel/CondCdf.lean | 30 ++-- .../Kernel/Disintegration/AuxLemmas.lean | 129 +++++++++++++++ .../Kernel/Disintegration/Basic.lean | 4 +- .../Kernel/Disintegration/BuildKernel.lean | 25 +-- .../Kernel/Disintegration/KernelCDFBorel.lean | 151 ------------------ .../Kernel/Disintegration/StieltjesReal.lean | 43 +---- .../Kernel/Disintegration/Unique.lean | 6 +- 9 files changed, 207 insertions(+), 237 deletions(-) create mode 100644 Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 226ede2dc81b5..c0f26e4e8f830 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -622,6 +622,10 @@ theorem Memℒp.mono_measure {f : α → E} (hμν : ν ≤ μ) (hf : Memℒp f ⟨hf.1.mono_measure hμν, (snorm_mono_measure f hμν).trans_lt hf.2⟩ #align measure_theory.mem_ℒp.mono_measure MeasureTheory.Memℒp.mono_measure +lemma snorm_restrict_le {p : ℝ≥0∞} {f : α → F} {μ : Measure α} (s : Set α) : + snorm f p (μ.restrict s) ≤ snorm f p μ := + snorm_mono_measure f Measure.restrict_le_self + theorem Memℒp.restrict (s : Set α) {f : α → E} (hf : Memℒp f p μ) : Memℒp f p (μ.restrict s) := hf.mono_measure Measure.restrict_le_self #align measure_theory.mem_ℒp.restrict MeasureTheory.Memℒp.restrict diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 5c68e8c618b4a..15dd90476120f 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1018,6 +1018,43 @@ theorem tendsto_integral_of_L1 {ι} (f : α → G) (hfi : Integrable f μ) {F : set_option linter.uppercaseLean3 false in #align measure_theory.tendsto_integral_of_L1 MeasureTheory.tendsto_integral_of_L1 +/-- If `F i → f` in `L1`, then `∫ x, F i x ∂μ → ∫ x, f x ∂μ`. -/ +lemma tendsto_integral_of_L1' {ι} (f : α → G) (hfi : Integrable f μ) {F : ι → α → G} {l : Filter ι} + (hFi : ∀ᶠ i in l, Integrable (F i) μ) (hF : Tendsto (fun i ↦ snorm (F i - f) 1 μ) l (𝓝 0)) : + Tendsto (fun i ↦ ∫ x, F i x ∂μ) l (𝓝 (∫ x, f x ∂μ)) := by + refine tendsto_integral_of_L1 f hfi hFi ?_ + simp_rw [snorm_one_eq_lintegral_nnnorm, Pi.sub_apply] at hF + exact hF + +lemma tendsto_snorm_restrict_zero {β ι : Type*} [NormedAddCommGroup β] + {p : ℝ≥0∞} {f : ι → α → β} {l : Filter ι} + (h : Tendsto (fun n ↦ snorm (f n) p μ) l (𝓝 0)) (s : Set α) : + Tendsto (fun n ↦ snorm (f n) p (μ.restrict s)) l (𝓝 0) := by + refine tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds h ?_ ?_ + · exact fun _ ↦ zero_le _ + · exact fun _ ↦ snorm_restrict_le _ + +/-- If `F i → f` in `L1`, then `∫ x in s, F i x ∂μ → ∫ x in s, f x ∂μ`. -/ +lemma tendsto_set_integral_of_L1 {ι} (f : α → G) (hfi : Integrable f μ) {F : ι → α → G} + {l : Filter ι} + (hFi : ∀ᶠ i in l, Integrable (F i) μ) (hF : Tendsto (fun i ↦ ∫⁻ x, ‖F i x - f x‖₊ ∂μ) l (𝓝 0)) + (s : Set α) : + Tendsto (fun i ↦ ∫ x in s, F i x ∂μ) l (𝓝 (∫ x in s, f x ∂μ)) := by + refine tendsto_integral_of_L1 f hfi.restrict ?_ ?_ + · filter_upwards [hFi] with i hi using hi.restrict + · simp_rw [← snorm_one_eq_lintegral_nnnorm] at hF ⊢ + exact tendsto_snorm_restrict_zero hF s + +/-- If `F i → f` in `L1`, then `∫ x in s, F i x ∂μ → ∫ x in s, f x ∂μ`. -/ +lemma tendsto_set_integral_of_L1' {ι} (f : α → G) (hfi : Integrable f μ) {F : ι → α → G} + {l : Filter ι} + (hFi : ∀ᶠ i in l, Integrable (F i) μ) (hF : Tendsto (fun i ↦ snorm (F i - f) 1 μ) l (𝓝 0)) + (s : Set α) : + Tendsto (fun i ↦ ∫ x in s, F i x ∂μ) l (𝓝 (∫ x in s, f x ∂μ)) := by + refine tendsto_set_integral_of_L1 f hfi hFi ?_ s + simp_rw [snorm_one_eq_lintegral_nnnorm, Pi.sub_apply] at hF + exact hF + /-- **Lebesgue dominated convergence theorem** provides sufficient conditions under which almost everywhere convergence of a sequence of functions implies the convergence of their integrals. We could weaken the condition `bound_integrable` to require `HasFiniteIntegral bound μ` instead @@ -1270,9 +1307,22 @@ theorem integral_eq_zero_iff_of_nonneg_ae {f : α → ℝ} (hf : 0 ≤ᵐ[μ] f) · rw [← hf.le_iff_eq, Filter.EventuallyEq, Filter.EventuallyLE] simp only [Pi.zero_apply, ofReal_eq_zero] · exact (ENNReal.measurable_ofReal.comp_aemeasurable hfi.1.aemeasurable) - #align measure_theory.integral_eq_zero_iff_of_nonneg_ae MeasureTheory.integral_eq_zero_iff_of_nonneg_ae +lemma ae_eq_of_integral_eq_of_ae_le {f g : α → ℝ} (hf : Integrable f μ) + (hg : Integrable g μ) (h_le : f ≤ᵐ[μ] g) (h_eq : ∫ a, f a ∂μ = ∫ a, g a ∂μ) : + f =ᵐ[μ] g := by + suffices g - f =ᵐ[μ] 0 by + filter_upwards [this] with a ha + symm + simpa only [Pi.sub_apply, Pi.zero_apply, sub_eq_zero] using ha + have h_eq' : ∫ a, (g - f) a ∂μ = 0 := by + simp_rw [Pi.sub_apply] + rwa [integral_sub hg hf, sub_eq_zero, eq_comm] + rwa [integral_eq_zero_iff_of_nonneg_ae _ (hg.sub hf)] at h_eq' + filter_upwards [h_le] with a ha + simpa + theorem integral_eq_zero_iff_of_nonneg {f : α → ℝ} (hf : 0 ≤ f) (hfi : Integrable f μ) : ∫ x, f x ∂μ = 0 ↔ f =ᵐ[μ] 0 := integral_eq_zero_iff_of_nonneg_ae (eventually_of_forall hf) hfi diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index 359b954ddf564..c527f5b638bfd 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -11,23 +11,23 @@ import Mathlib.Probability.Kernel.Disintegration.BuildKernel /-! # Conditional cumulative distribution function -Given `ρ : measure (α × ℝ)`, we define the conditional cumulative distribution function -(conditional cdf) of `ρ`. It is a function `cond_cdf ρ : α → ℝ → ℝ` such that if `ρ` is a finite -measure, then for all `a : α` `cond_cdf ρ a` is monotone and right-continuous with limit 0 at -∞ -and limit 1 at +∞, and such that for all `x : ℝ`, `a ↦ cond_cdf ρ a x` is measurable. For all +Given `ρ : Measure (α × ℝ)`, we define the conditional cumulative distribution function +(conditional cdf) of `ρ`. It is a function `condCDF ρ : α → ℝ → ℝ` such that if `ρ` is a finite +measure, then for all `a : α` `condCDF ρ a` is monotone and right-continuous with limit 0 at -∞ +and limit 1 at +∞, and such that for all `x : ℝ`, `a ↦ condCDF ρ a x` is measurable. For all `x : ℝ` and measurable set `s`, that function satisfies -`∫⁻ a in s, ennreal.of_real (cond_cdf ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x)`. +`∫⁻ a in s, ennreal.of_real (condCDF ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x)`. ## Main definitions -* `probability_theory.cond_cdf ρ : α → stieltjes_function`: the conditional cdf of - `ρ : measure (α × ℝ)`. A `stieltjes_function` is a function `ℝ → ℝ` which is monotone and +* `ProbabilityTheory.condCDF ρ : α → StieltjesFunction`: the conditional cdf of + `ρ : Measure (α × ℝ)`. A `StieltjesFunction` is a function `ℝ → ℝ` which is monotone and right-continuous. ## Main statements -* `probability_theory.set_lintegral_cond_cdf`: for all `a : α` and `x : ℝ`, all measurable set `s`, - `∫⁻ a in s, ennreal.of_real (cond_cdf ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x)`. +* `ProbabilityTheory.set_lintegral_condCDF`: for all `a : α` and `x : ℝ`, all measurable set `s`, + `∫⁻ a in s, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x)`. ## References @@ -154,13 +154,11 @@ attribute [local instance] MeasureTheory.Measure.IsFiniteMeasure.IicSnd /-! ### Auxiliary definitions -We build towards the definition of `probability_theory.cond_cdf`. We first define -`probability_theory.pre_cdf`, a function defined on `α × ℚ` with the properties of a cdf almost -everywhere. We then introduce `probability_theory.cond_cdf_rat`, a function on `α × ℚ` which has -the properties of a cdf for all `a : α`. We finally extend to `ℝ`. -/ +We build towards the definition of `ProbabilityTheory.cond_cdf`. We first define +`ProbabilityTheory.preCDF`, a function defined on `α × ℚ` with the properties of a cdf almost +everywhere. -/ - -/-- `pre_cdf` is the Radon-Nikodym derivative of `ρ.IicSnd` with respect to `ρ.fst` at each +/-- `preCDF` is the Radon-Nikodym derivative of `ρ.IicSnd` with respect to `ρ.fst` at each `r : ℚ`. This function `ℚ → α → ℝ≥0∞` is such that for almost all `a : α`, the function `ℚ → ℝ≥0∞` satisfies the properties of a cdf (monotone with limit 0 at -∞ and 1 at +∞, right-continuous). @@ -321,7 +319,7 @@ theorem tendsto_preCDF_atTop_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] theorem tendsto_preCDF_atBot_zero (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : ∀ᵐ a ∂ρ.fst, Tendsto (fun r => preCDF ρ r a) atBot (𝓝 0) := by -- We show first that `preCDF` has a limit in ℝ≥0∞ almost everywhere. - -- We then show that the integral of `pre_cdf` tends to 0, and that it also tends + -- We then show that the integral of `preCDF` tends to 0, and that it also tends -- to the integral of the limit. Since the limit has integral 0, it is equal to 0 a.e. suffices ∀ᵐ a ∂ρ.fst, Tendsto (fun r => preCDF ρ (-r) a) atTop (𝓝 0) by filter_upwards [this] with a ha diff --git a/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean b/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean new file mode 100644 index 0000000000000..daec77bca6d2d --- /dev/null +++ b/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean @@ -0,0 +1,129 @@ +import Mathlib.MeasureTheory.Constructions.Polish +import Mathlib.MeasureTheory.Integral.Bochner + +open Filter Set MeasureTheory + +open scoped Topology ENNReal + +variable {α β : Type*} {mα : MeasurableSpace α} + +theorem Real.iUnion_Iic_rat : ⋃ r : ℚ, Iic (r : ℝ) = univ := by + ext1 x + simp only [mem_iUnion, mem_Iic, mem_univ, iff_true_iff] + obtain ⟨r, hr⟩ := exists_rat_gt x + exact ⟨r, hr.le⟩ +#align real.Union_Iic_rat Real.iUnion_Iic_rat + +theorem Real.iInter_Iic_rat : ⋂ r : ℚ, Iic (r : ℝ) = ∅ := by + ext1 x + simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false_iff, not_forall, not_le] + exact exists_rat_lt x +#align real.Inter_Iic_rat Real.iInter_Iic_rat + +lemma measurableSet_tendsto_nhds {β γ ι : Type*} [MeasurableSpace β] + [TopologicalSpace γ] [PolishSpace γ] [MeasurableSpace γ] + [hγ : OpensMeasurableSpace γ] [Countable ι] {l : Filter ι} + [l.IsCountablyGenerated] {f : ι → β → γ} (hf : ∀ i, Measurable (f i)) (c : γ) : + MeasurableSet { x | Tendsto (fun n ↦ f n x) l (𝓝 c) } := sorry + +lemma measurableSet_tendsto_fun {β γ ι : Type*} [MeasurableSpace β] + [TopologicalSpace γ] [PolishSpace γ] [MeasurableSpace γ] + [hγ : OpensMeasurableSpace γ] [Countable ι] {l : Filter ι} + [l.IsCountablyGenerated] {f : ι → β → γ} (hf : ∀ i, Measurable (f i)) {g : β → γ} + (hg : Measurable g) : + MeasurableSet { x | Tendsto (fun n ↦ f n x) l (𝓝 (g x)) } := by + letI := upgradePolishSpace γ + have : { x | Tendsto (fun n ↦ f n x) l (𝓝 (g x)) } + = { x | Tendsto (fun n ↦ dist (f n x) (g x)) l (𝓝 0) } := by + ext x + simp only [Set.mem_setOf_eq] + rw [tendsto_iff_dist_tendsto_zero] + rw [this] + exact measurableSet_tendsto_nhds (fun n ↦ (hf n).dist hg) 0 + +lemma integral_tendsto_of_tendsto_of_monotone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} + (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Monotone fun n ↦ f n x) + (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : + Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by + let f' := fun n x ↦ f n x - f 0 x + have hf'_nonneg : ∀ᵐ x ∂μ, ∀ n, 0 ≤ f' n x := by + filter_upwards [h_mono] with a ha n + simp [ha (zero_le n)] + have hf'_meas : ∀ n, Integrable (f' n) μ := fun n ↦ (hf n).sub (hf 0) + suffices Tendsto (fun n ↦ ∫ x, f' n x ∂μ) atTop (𝓝 (∫ x, (F - f 0) x ∂μ)) by + rw [integral_sub' hF (hf 0)] at this + have h_sub : ∀ n, ∫ x, f' n x ∂μ = ∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ := by + intro n + simp only + rw [integral_sub (hf n) (hf 0)] + simp_rw [h_sub] at this + have h1 : (fun n ↦ ∫ x, f n x ∂μ) + = fun n ↦ (∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by ext n; abel + have h2 : ∫ x, F x ∂μ = (∫ x, F x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by abel + rw [h1, h2] + exact this.add tendsto_const_nhds + have hF_ge : 0 ≤ᵐ[μ] fun x ↦ (F - f 0) x := by + filter_upwards [h_tendsto, h_mono] with x hx_tendsto hx_mono + simp only [Pi.zero_apply, Pi.sub_apply, sub_nonneg] + exact ge_of_tendsto' hx_tendsto (fun n ↦ hx_mono (zero_le _)) + rw [ae_all_iff] at hf'_nonneg + simp_rw [integral_eq_lintegral_of_nonneg_ae (hf'_nonneg _) (hf'_meas _).1] + rw [integral_eq_lintegral_of_nonneg_ae hF_ge (hF.1.sub (hf 0).1)] + have h_cont := ENNReal.continuousOn_toReal.continuousAt + (x := ∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ) ?_ + swap + · rw [mem_nhds_iff] + refine ⟨Iio (∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ + 1), ?_, isOpen_Iio, ?_⟩ + · intro x + simp only [Pi.sub_apply, mem_Iio, ne_eq, mem_setOf_eq] + exact ne_top_of_lt + · simp only [Pi.sub_apply, mem_Iio] + refine ENNReal.lt_add_right ?_ one_ne_zero + rw [← ofReal_integral_eq_lintegral_ofReal] + · exact ENNReal.ofReal_ne_top + · exact hF.sub (hf 0) + · exact hF_ge + refine h_cont.tendsto.comp ?_ + refine lintegral_tendsto_of_tendsto_of_monotone ?_ ?_ ?_ + · exact fun n ↦ ((hf n).sub (hf 0)).aemeasurable.ennreal_ofReal + · filter_upwards [h_mono] with x hx + intro n m hnm + refine ENNReal.ofReal_le_ofReal ?_ + simp only [tsub_le_iff_right, sub_add_cancel] + exact hx hnm + · filter_upwards [h_tendsto] with x hx + refine (ENNReal.continuous_ofReal.tendsto _).comp ?_ + simp only [Pi.sub_apply] + exact Tendsto.sub hx tendsto_const_nhds + +lemma integral_tendsto_of_tendsto_of_antitone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} + (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Antitone fun n ↦ f n x) + (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : + Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by + suffices Tendsto (fun n ↦ ∫ x, -f n x ∂μ) atTop (𝓝 (∫ x, -F x ∂μ)) by + suffices Tendsto (fun n ↦ ∫ x, - -f n x ∂μ) atTop (𝓝 (∫ x, - -F x ∂μ)) by + simp_rw [neg_neg] at this + exact this + convert this.neg <;> rw [integral_neg] + refine integral_tendsto_of_tendsto_of_monotone (fun n ↦ (hf n).neg) hF.neg ?_ ?_ + · filter_upwards [h_mono] with x hx + intro n m hnm + simp only [neg_le_neg_iff] + exact hx hnm + · filter_upwards [h_tendsto] with x hx + exact hx.neg + +lemma tendsto_nat_ceil_atTop {α : Type*} [LinearOrderedSemiring α] [FloorSemiring α] : + Tendsto (fun x : α ↦ ⌈x⌉₊) atTop atTop := by + refine Nat.ceil_mono.tendsto_atTop_atTop (fun x ↦ ⟨x, ?_⟩) + simp only [Nat.ceil_natCast, le_refl] + +lemma isCoboundedUnder_le_of_eventually_le {ι α : Type*} [Preorder α] (l : Filter ι) [NeBot l] + {f : ι → α} {x : α} (hf : ∀ᶠ i in l, x ≤ f i) : + IsCoboundedUnder (· ≤ ·) l f := + IsBoundedUnder.isCoboundedUnder_le ⟨x, hf⟩ + +lemma isCoboundedUnder_le_of_le {ι α : Type*} [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} + {x : α} (hf : ∀ i, x ≤ f i) : + IsCoboundedUnder (· ≤ ·) l f := + isCoboundedUnder_le_of_eventually_le l (eventually_of_forall hf) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 1e7f8f5526ec2..4d17140a125c7 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -91,7 +91,7 @@ def condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFu let he := measurableEmbedding_measurableEmbedding_real Ω let x₀ := (range_nonempty e).choose kernel.comapRight - (kernel.piecewise (measurableSet_eq_one hf he.measurableSet_range) + (kernel.piecewise (measurableSet_cdfKernel_eq_one hf he.measurableSet_range) (cdfKernel f hf) (kernel.deterministic (fun _ ↦ x₀) measurable_const)) he @@ -324,7 +324,7 @@ section Measure variable {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] /-- Conditional kernel of a measure on a product space: a Markov kernel such that -`ρ = ρ.fst ⊗ₘ ρ.condKernel` (see `ProbabilityTheory.measure_eq_compProd`). -/ +`ρ = ρ.fst ⊗ₘ ρ.condKernel` (see `MeasureTheory.Measure.compProd_fst_condKernel`). -/ noncomputable def _root_.MeasureTheory.Measure.condKernel (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : kernel α Ω := diff --git a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean index 13a2838cedce8..c0c82c8fecdb0 100644 --- a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean @@ -5,6 +5,7 @@ Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.Disintegration.StieltjesReal import Mathlib.Probability.Kernel.MeasureCompProd +import Mathlib.Probability.Kernel.Disintegration.AuxLemmas /-! @@ -15,24 +16,6 @@ import Mathlib.Probability.Kernel.MeasureCompProd open MeasureTheory Set Filter TopologicalSpace open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory -section AuxLemmasToBeMoved - -variable {α β ι : Type*} - -theorem Real.iUnion_Iic_rat : ⋃ r : ℚ, Iic (r : ℝ) = univ := by - ext1 x - simp only [mem_iUnion, mem_Iic, mem_univ, iff_true_iff] - obtain ⟨r, hr⟩ := exists_rat_gt x - exact ⟨r, hr.le⟩ -#align real.Union_Iic_rat Real.iUnion_Iic_rat - -theorem Real.iInter_Iic_rat : ⋂ r : ℚ, Iic (r : ℝ) = ∅ := by - ext1 x - simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false_iff, not_forall, not_le] - exact exists_rat_lt x -#align real.Inter_Iic_rat Real.iInter_Iic_rat - -end AuxLemmasToBeMoved namespace ProbabilityTheory @@ -79,8 +62,7 @@ lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKern ∫⁻ t in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) x) ∂(ν a) = μ a (s ×ˢ Iic x) := by -- We have the result for `x : ℚ` thanks to `set_lintegral_stieltjesOfMeasurableRat_rat`. - -- We use the equality `condCDF ρ a x = ⨅ r : {r' : ℚ // x < r'}, condCDF ρ a r` and a monotone - -- convergence argument to extend it to the reals. + -- We use a monotone convergence argument to extend it to the reals. by_cases hρ_zero : (ν a).restrict s = 0 · rw [hρ_zero, lintegral_zero_measure] have ⟨q, hq⟩ := exists_rat_gt x @@ -193,7 +175,6 @@ structure IsKernelCDF (f : α × β → StieltjesFunction) (μ : kernel α (β (tendsto_atBot_zero (p : α × β) : Tendsto (f p) atBot (𝓝 0)) (set_integral (a : α) {s : Set β} (_hs : MeasurableSet s) (x : ℝ) : ∫ t in s, f (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal) --- todo: nonneg and le_one are consequences of tendsto_atTop_one and tendsto_atBot_zero lemma IsKernelCDF.nonneg (hf : IsKernelCDF f μ ν) (p : α × β) (x : ℝ) : 0 ≤ f p x := Monotone.le_of_tendsto (f p).mono (hf.tendsto_atBot_zero p) x @@ -451,7 +432,7 @@ lemma ae_cdfKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsKerne change cdfKernel f hf (a, t) sᶜ = 0 at ht rwa [prob_compl_eq_zero_iff hs] at ht -lemma measurableSet_eq_one (hf : IsKernelCDF f μ ν) {s : Set ℝ} (hs : MeasurableSet s) : +lemma measurableSet_cdfKernel_eq_one (hf : IsKernelCDF f μ ν) {s : Set ℝ} (hs : MeasurableSet s) : MeasurableSet {p | cdfKernel f hf p s = 1} := (kernel.measurable_coe _ hs) (measurableSet_singleton 1) diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean index face9bc639381..b802372247ea6 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean @@ -49,157 +49,6 @@ namespace ProbabilityTheory variable {α β : Type*} {mα : MeasurableSpace α} -section Move - -lemma snorm_restrict_le [NormedAddCommGroup β] {p : ℝ≥0∞} {f : α → β} {μ : Measure α} (s : Set α) : - snorm f p (μ.restrict s) ≤ snorm f p μ := - snorm_mono_measure f Measure.restrict_le_self - -lemma tendsto_snorm_restrict_zero {α β ι : Type*} {mα : MeasurableSpace α} [NormedAddCommGroup β] - {p : ℝ≥0∞} {f : ι → α → β} {μ : Measure α} {l : Filter ι} - (h : Tendsto (fun n ↦ snorm (f n) p μ) l (𝓝 0)) (s : Set α) : - Tendsto (fun n ↦ snorm (f n) p (μ.restrict s)) l (𝓝 0) := by - refine tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds h ?_ ?_ - · exact fun _ ↦ zero_le _ - · exact fun _ ↦ snorm_restrict_le _ - -lemma tendsto_integral_of_L1' {ι G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] - {μ : Measure α} - (f : α → G) (hfi : Integrable f μ) - {F : ι → α → G} {l : Filter ι} - (hFi : ∀ᶠ i in l, Integrable (F i) μ) - (hF : Tendsto (fun i ↦ snorm (F i - f) 1 μ) l (𝓝 0)) : - Tendsto (fun i ↦ ∫ x, F i x ∂μ) l (𝓝 (∫ x, f x ∂μ)) := by - refine tendsto_integral_of_L1 f hfi hFi ?_ - simp_rw [snorm_one_eq_lintegral_nnnorm, Pi.sub_apply] at hF - exact hF - -lemma tendsto_set_integral_of_L1 {ι G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] - {μ : Measure α} - (f : α → G) (hfi : Integrable f μ) - {F : ι → α → G} {l : Filter ι} - (hFi : ∀ᶠ i in l, Integrable (F i) μ) - (hF : Tendsto (fun i ↦ ∫⁻ x, ‖F i x - f x‖₊ ∂μ) l (𝓝 0)) (s : Set α) : - Tendsto (fun i ↦ ∫ x in s, F i x ∂μ) l (𝓝 (∫ x in s, f x ∂μ)) := by - refine tendsto_integral_of_L1 f hfi.restrict ?_ ?_ - · filter_upwards [hFi] with i hi using hi.restrict - · simp_rw [← snorm_one_eq_lintegral_nnnorm] at hF ⊢ - exact tendsto_snorm_restrict_zero hF s - -lemma tendsto_set_integral_of_L1' {ι G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] - {μ : Measure α} - (f : α → G) (hfi : Integrable f μ) - {F : ι → α → G} {l : Filter ι} - (hFi : ∀ᶠ i in l, Integrable (F i) μ) - (hF : Tendsto (fun i ↦ snorm (F i - f) 1 μ) l (𝓝 0)) (s : Set α) : - Tendsto (fun i ↦ ∫ x in s, F i x ∂μ) l (𝓝 (∫ x in s, f x ∂μ)) := by - refine tendsto_set_integral_of_L1 f hfi hFi ?_ s - simp_rw [snorm_one_eq_lintegral_nnnorm, Pi.sub_apply] at hF - exact hF - -lemma ae_eq_of_integral_eq_of_ae_le {μ : Measure α} {f g : α → ℝ} (hf : Integrable f μ) - (hg : Integrable g μ) (h_le : f ≤ᵐ[μ] g) (h_eq : ∫ a, f a ∂μ = ∫ a, g a ∂μ) : - f =ᵐ[μ] g := by - suffices g - f =ᵐ[μ] 0 by - filter_upwards [this] with a ha - symm - simpa only [Pi.sub_apply, Pi.zero_apply, sub_eq_zero] using ha - have h_eq' : ∫ a, (g - f) a ∂μ = 0 := by - simp_rw [Pi.sub_apply] - rwa [integral_sub hg hf, sub_eq_zero, eq_comm] - rwa [integral_eq_zero_iff_of_nonneg_ae _ (hg.sub hf)] at h_eq' - filter_upwards [h_le] with a ha - simpa - -lemma integral_tendsto_of_tendsto_of_monotone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} - (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Monotone fun n ↦ f n x) - (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : - Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by - let f' := fun n x ↦ f n x - f 0 x - have hf'_nonneg : ∀ᵐ x ∂μ, ∀ n, 0 ≤ f' n x := by - filter_upwards [h_mono] with a ha n - simp [ha (zero_le n)] - have hf'_meas : ∀ n, Integrable (f' n) μ := fun n ↦ (hf n).sub (hf 0) - suffices Tendsto (fun n ↦ ∫ x, f' n x ∂μ) atTop (𝓝 (∫ x, (F - f 0) x ∂μ)) by - rw [integral_sub' hF (hf 0)] at this - have h_sub : ∀ n, ∫ x, f' n x ∂μ = ∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ := by - intro n - simp only - rw [integral_sub (hf n) (hf 0)] - simp_rw [h_sub] at this - have h1 : (fun n ↦ ∫ x, f n x ∂μ) - = fun n ↦ (∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by ext n; abel - have h2 : ∫ x, F x ∂μ = (∫ x, F x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by abel - rw [h1, h2] - exact this.add tendsto_const_nhds - have hF_ge : 0 ≤ᵐ[μ] fun x ↦ (F - f 0) x := by - filter_upwards [h_tendsto, h_mono] with x hx_tendsto hx_mono - simp only [Pi.zero_apply, Pi.sub_apply, sub_nonneg] - exact ge_of_tendsto' hx_tendsto (fun n ↦ hx_mono (zero_le _)) - rw [ae_all_iff] at hf'_nonneg - simp_rw [integral_eq_lintegral_of_nonneg_ae (hf'_nonneg _) (hf'_meas _).1] - rw [integral_eq_lintegral_of_nonneg_ae hF_ge (hF.1.sub (hf 0).1)] - have h_cont := ENNReal.continuousOn_toReal.continuousAt - (x := ∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ) ?_ - swap - · rw [mem_nhds_iff] - refine ⟨Iio (∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ + 1), ?_, isOpen_Iio, ?_⟩ - · intro x - simp only [Pi.sub_apply, mem_Iio, ne_eq, mem_setOf_eq] - exact ne_top_of_lt - · simp only [Pi.sub_apply, mem_Iio] - refine ENNReal.lt_add_right ?_ one_ne_zero - rw [← ofReal_integral_eq_lintegral_ofReal] - · exact ENNReal.ofReal_ne_top - · exact hF.sub (hf 0) - · exact hF_ge - refine h_cont.tendsto.comp ?_ - refine lintegral_tendsto_of_tendsto_of_monotone ?_ ?_ ?_ - · exact fun n ↦ ((hf n).sub (hf 0)).aemeasurable.ennreal_ofReal - · filter_upwards [h_mono] with x hx - intro n m hnm - refine ENNReal.ofReal_le_ofReal ?_ - simp only [tsub_le_iff_right, sub_add_cancel] - exact hx hnm - · filter_upwards [h_tendsto] with x hx - refine (ENNReal.continuous_ofReal.tendsto _).comp ?_ - simp only [Pi.sub_apply] - exact Tendsto.sub hx tendsto_const_nhds - -lemma integral_tendsto_of_tendsto_of_antitone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} - (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Antitone fun n ↦ f n x) - (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : - Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by - suffices Tendsto (fun n ↦ ∫ x, -f n x ∂μ) atTop (𝓝 (∫ x, -F x ∂μ)) by - suffices Tendsto (fun n ↦ ∫ x, - -f n x ∂μ) atTop (𝓝 (∫ x, - -F x ∂μ)) by - simp_rw [neg_neg] at this - exact this - convert this.neg <;> rw [integral_neg] - refine integral_tendsto_of_tendsto_of_monotone (fun n ↦ (hf n).neg) hF.neg ?_ ?_ - · filter_upwards [h_mono] with x hx - intro n m hnm - simp only [neg_le_neg_iff] - exact hx hnm - · filter_upwards [h_tendsto] with x hx - exact hx.neg - -lemma tendsto_nat_ceil_atTop {α : Type*} [LinearOrderedSemiring α] [FloorSemiring α] : - Tendsto (fun x : α ↦ ⌈x⌉₊) atTop atTop := by - refine Nat.ceil_mono.tendsto_atTop_atTop (fun x ↦ ⟨x, ?_⟩) - simp only [Nat.ceil_natCast, le_refl] - -lemma isCoboundedUnder_le_of_eventually_le {ι α : Type*} [Preorder α] (l : Filter ι) [NeBot l] - {f : ι → α} {x : α} (hf : ∀ᶠ i in l, x ≤ f i) : - IsCoboundedUnder (· ≤ ·) l f := - IsBoundedUnder.isCoboundedUnder_le ⟨x, hf⟩ - -lemma isCoboundedUnder_le_of_le {ι α : Type*} [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} - {x : α} (hf : ∀ i, x ≤ f i) : - IsCoboundedUnder (· ≤ ·) l f := - isCoboundedUnder_le_of_eventually_le l (eventually_of_forall hf) - -end Move - section dissection_system def I (n : ℕ) (k : ℤ) : Set ℝ := Set.Ico (k * (2⁻¹ : ℝ) ^ n) ((k + 1) * ((2 : ℝ) ^ n)⁻¹) diff --git a/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean b/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean index 909246fdf43a3..48a0767a7d814 100644 --- a/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean +++ b/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean @@ -9,9 +9,10 @@ import Mathlib.MeasureTheory.Measure.GiryMonad import Mathlib.MeasureTheory.Measure.Stieltjes import Mathlib.Analysis.Normed.Order.Lattice import Mathlib.MeasureTheory.Function.StronglyMeasurable.Basic +import Mathlib.Probability.Kernel.Disintegration.AuxLemmas /-! - +# Cumulative distributions functions of Markov kernels -/ @@ -20,27 +21,6 @@ open MeasureTheory Set Filter TopologicalSpace open scoped NNReal ENNReal MeasureTheory Topology -lemma measurableSet_tendsto_nhds {β γ ι : Type*} [MeasurableSpace β] - [TopologicalSpace γ] [PolishSpace γ] [MeasurableSpace γ] - [hγ : OpensMeasurableSpace γ] [Countable ι] {l : Filter ι} - [l.IsCountablyGenerated] {f : ι → β → γ} (hf : ∀ i, Measurable (f i)) (c : γ) : - MeasurableSet { x | Tendsto (fun n ↦ f n x) l (𝓝 c) } := sorry - -lemma measurableSet_tendsto_fun {β γ ι : Type*} [MeasurableSpace β] - [TopologicalSpace γ] [PolishSpace γ] [MeasurableSpace γ] - [hγ : OpensMeasurableSpace γ] [Countable ι] {l : Filter ι} - [l.IsCountablyGenerated] {f : ι → β → γ} (hf : ∀ i, Measurable (f i)) {g : β → γ} - (hg : Measurable g) : - MeasurableSet { x | Tendsto (fun n ↦ f n x) l (𝓝 (g x)) } := by - letI := upgradePolishSpace γ - have : { x | Tendsto (fun n ↦ f n x) l (𝓝 (g x)) } - = { x | Tendsto (fun n ↦ dist (f n x) (g x)) l (𝓝 0) } := by - ext x - simp only [mem_setOf_eq] - rw [tendsto_iff_dist_tendsto_zero] - rw [this] - exact measurableSet_tendsto_nhds (fun n ↦ (hf n).dist hg) 0 - namespace ProbabilityTheory variable {α β ι : Type*} [MeasurableSpace α] @@ -223,8 +203,6 @@ section IsCDFLike.stieltjesFunction variable {f : α → ℚ → ℝ} (hf : IsCDFLike f) -/-- Conditional cdf of the measure given the value on `α`, as a plain function. This is an auxiliary -definition used to define `cond_cdf`. -/ noncomputable irreducible_def IsCDFLike.stieltjesFunctionAux (f : α → ℚ → ℝ) : α → ℝ → ℝ := fun a t ↦ ⨅ r : { r' : ℚ // t < r' }, f a r @@ -292,10 +270,6 @@ lemma IsCDFLike.continuousWithinAt_stieltjesFunctionAux_Ici (a : α) (x : ℝ) congr! rw [stieltjesFunctionAux_def] -/-! ### Conditional cdf -/ - - -/-- Conditional cdf of the measure given the value on `α`, as a Stieltjes function. -/ noncomputable def IsCDFLike.stieltjesFunction (a : α) : StieltjesFunction where toFun := stieltjesFunctionAux f a mono' := monotone_stieltjesFunctionAux hf a @@ -304,11 +278,9 @@ noncomputable def IsCDFLike.stieltjesFunction (a : α) : StieltjesFunction where lemma IsCDFLike.stieltjesFunction_eq (a : α) (r : ℚ) : hf.stieltjesFunction a r = f a r := stieltjesFunctionAux_eq hf a r -/-- The conditional cdf is non-negative for all `a : α`. -/ lemma IsCDFLike.stieltjesFunction_nonneg (a : α) (r : ℝ) : 0 ≤ hf.stieltjesFunction a r := stieltjesFunctionAux_nonneg hf a r -/-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ lemma IsCDFLike.stieltjesFunction_le_one (a : α) (x : ℝ) : hf.stieltjesFunction a x ≤ 1 := by obtain ⟨r, hrx⟩ := exists_rat_gt x rw [← StieltjesFunction.iInf_rat_gt_eq] @@ -316,7 +288,6 @@ lemma IsCDFLike.stieltjesFunction_le_one (a : α) (x : ℝ) : hf.stieltjesFuncti refine ciInf_le_of_le (bddBelow_range_gt hf a x) ?_ (hf.le_one _ _) exact ⟨r, hrx⟩ -/-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ lemma IsCDFLike.tendsto_stieltjesFunction_atBot (a : α) : Tendsto (hf.stieltjesFunction a) atBot (𝓝 0) := by have h_exists : ∀ x : ℝ, ∃ q : ℚ, x < q ∧ ↑q < x + 1 := fun x ↦ exists_rat_btwn (lt_add_one x) @@ -333,7 +304,6 @@ lemma IsCDFLike.tendsto_stieltjesFunction_atBot (a : α) : rw [Function.comp_apply, ← stieltjesFunction_eq hf] exact (hf.stieltjesFunction a).mono (h_exists x).choose_spec.1.le -/-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ lemma IsCDFLike.tendsto_stieltjesFunction_atTop (a : α) : Tendsto (hf.stieltjesFunction a) atTop (𝓝 1) := by have h_exists : ∀ x : ℝ, ∃ q : ℚ, x - 1 < q ∧ ↑q < x := fun x ↦ exists_rat_btwn (sub_one_lt x) @@ -350,7 +320,6 @@ lemma IsCDFLike.tendsto_stieltjesFunction_atTop (a : α) : rw [Function.comp_apply, ← stieltjesFunction_eq hf] exact (hf.stieltjesFunction a).mono (le_of_lt (h_exists x).choose_spec.2) -/-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ lemma IsCDFLike.measurable_stieltjesFunction (x : ℝ) : Measurable fun a ↦ hf.stieltjesFunction a x := by have : (fun a ↦ hf.stieltjesFunction a x) = fun a ↦ ⨅ r : { r' : ℚ // x < r' }, f a ↑r := by @@ -361,7 +330,6 @@ lemma IsCDFLike.measurable_stieltjesFunction (x : ℝ) : rw [this] exact measurable_iInf (fun q ↦ hf.measurable q) -/-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ lemma IsCDFLike.stronglyMeasurable_stieltjesFunction (x : ℝ) : StronglyMeasurable fun a ↦ hf.stieltjesFunction a x := (measurable_stieltjesFunction hf x).stronglyMeasurable @@ -383,7 +351,6 @@ instance IsCDFLike.instIsProbabilityMeasure_stieltjesFunction (a : α) : IsProbabilityMeasure (hf.stieltjesFunction a).measure := ⟨measure_stieltjesFunction_univ hf a⟩ -/-- The function `a ↦ (condCDF ρ a).measure` is measurable. -/ lemma IsCDFLike.measurable_measure_stieltjesFunction : Measurable fun a ↦ (hf.stieltjesFunction a).measure := by rw [Measure.measurable_measure] @@ -433,30 +400,24 @@ lemma stieltjesOfMeasurableRat_unit_prod (hf : ∀ q, Measurable fun a ↦ f a q cases p with | mk _ b => rw [← toCDFLike_unit_prod b] -/-- The conditional cdf is non-negative for all `a : α`. -/ lemma stieltjesOfMeasurableRat_nonneg (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℝ) : 0 ≤ stieltjesOfMeasurableRat f hf a r := IsCDFLike.stieltjesFunction_nonneg _ a r -/-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ lemma stieltjesOfMeasurableRat_le_one (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : stieltjesOfMeasurableRat f hf a x ≤ 1 := IsCDFLike.stieltjesFunction_le_one _ a x -/-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ lemma tendsto_stieltjesOfMeasurableRat_atBot (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : Tendsto (stieltjesOfMeasurableRat f hf a) atBot (𝓝 0) := IsCDFLike.tendsto_stieltjesFunction_atBot _ a -/-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ lemma tendsto_stieltjesOfMeasurableRat_atTop (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : Tendsto (stieltjesOfMeasurableRat f hf a) atTop (𝓝 1) := IsCDFLike.tendsto_stieltjesFunction_atTop _ a -/-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ lemma measurable_stieltjesOfMeasurableRat (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : Measurable fun a ↦ stieltjesOfMeasurableRat f hf a x := IsCDFLike.measurable_stieltjesFunction _ x -/-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ lemma stronglyMeasurable_stieltjesOfMeasurableRat (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : StronglyMeasurable fun a ↦ stieltjesOfMeasurableRat f hf a x := IsCDFLike.stronglyMeasurable_stieltjesFunction _ x diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index 41ffd83354046..d096646caa25e 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -195,8 +195,7 @@ theorem Measure.eq_condKernel_of_measure_eq_compProd (κ : kernel α Ω) [IsFini Measure.compProd_apply hs, Measure.compProd_apply (measurable_id.prod_map hf.measurable hs)] congr with a rw [kernel.map_apply'] - · rfl - · exact measurable_prod_mk_left hs + exacts [rfl, measurable_prod_mk_left hs] suffices : ∀ᵐ x ∂ρ.fst, ∀ s, MeasurableSet s → ρ'.condKernel x s = ρ.condKernel x (f ⁻¹' s) · filter_upwards [hρ'', this] with x hx h rw [kernel.map_apply] at hx @@ -216,8 +215,7 @@ theorem Measure.eq_condKernel_of_measure_eq_compProd (κ : kernel α Ω) [IsFini Measure.compProd_apply] · congr with a rw [kernel.map_apply'] - · rfl - · exact measurable_prod_mk_left hs + exacts [rfl, measurable_prod_mk_left hs] · exact measurable_id.prod_map hf.measurable hs end Measure From be152bb958c5122c8b9ff0c7400ed6bf72264856 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 15 Feb 2024 19:47:22 +0100 Subject: [PATCH 021/129] rename files --- Mathlib/Probability/Cdf.lean | 2 +- Mathlib/Probability/Kernel/Disintegration/Basic.lean | 2 +- .../Kernel/Disintegration/{BuildKernel.lean => CDFKernel.lean} | 2 +- Mathlib/Probability/Kernel/{ => Disintegration}/CondCdf.lean | 2 +- Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean | 2 +- .../{StieltjesReal.lean => MeasurableStieltjes.lean} | 0 6 files changed, 5 insertions(+), 5 deletions(-) rename Mathlib/Probability/Kernel/Disintegration/{BuildKernel.lean => CDFKernel.lean} (99%) rename Mathlib/Probability/Kernel/{ => Disintegration}/CondCdf.lean (99%) rename Mathlib/Probability/Kernel/Disintegration/{StieltjesReal.lean => MeasurableStieltjes.lean} (100%) diff --git a/Mathlib/Probability/Cdf.lean b/Mathlib/Probability/Cdf.lean index e67943ae3cac2..6f64b28d7f57a 100644 --- a/Mathlib/Probability/Cdf.lean +++ b/Mathlib/Probability/Cdf.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.CondCdf +import Mathlib.Probability.Kernel.Disintegration.CondCdf /-! # Cumulative distribution function of a real probability measure diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 4d17140a125c7..88cb1b9e3f190 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.CondCdf +import Mathlib.Probability.Kernel.Disintegration.CondCdf import Mathlib.Probability.Kernel.Disintegration.KernelCDFBorel /-! diff --git a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean similarity index 99% rename from Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean rename to Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean index c0c82c8fecdb0..6f9f73007cf91 100644 --- a/Mathlib/Probability/Kernel/Disintegration/BuildKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.Disintegration.StieltjesReal +import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes import Mathlib.Probability.Kernel.MeasureCompProd import Mathlib.Probability.Kernel.Disintegration.AuxLemmas diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean similarity index 99% rename from Mathlib/Probability/Kernel/CondCdf.lean rename to Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index c527f5b638bfd..384214034556c 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.MeasureTheory.Decomposition.RadonNikodym -import Mathlib.Probability.Kernel.Disintegration.BuildKernel +import Mathlib.Probability.Kernel.Disintegration.CDFKernel #align_import probability.kernel.cond_cdf from "leanprover-community/mathlib"@"3b88f4005dc2e28d42f974cc1ce838f0dafb39b8" diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean index b802372247ea6..84df3c3fcd09d 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean @@ -5,7 +5,7 @@ Authors: Rémy Degenne -/ import Mathlib.Analysis.SpecialFunctions.Log.Base import Mathlib.Probability.Martingale.Convergence -import Mathlib.Probability.Kernel.Disintegration.BuildKernel +import Mathlib.Probability.Kernel.Disintegration.CDFKernel /-! # KernelCDFBorel diff --git a/Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean similarity index 100% rename from Mathlib/Probability/Kernel/Disintegration/StieltjesReal.lean rename to Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean From 7544dacac172e3e99b6e2ee36726b93435e653ac Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 15 Feb 2024 19:49:23 +0100 Subject: [PATCH 022/129] fix space --- .../Probability/Kernel/Disintegration/MeasurableStieltjes.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean index 48a0767a7d814..588339717332a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean +++ b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean @@ -151,7 +151,7 @@ lemma inf_gt_rat_defaultRatCDF (t : ℚ) : exact h.trans (mem_Ioi.mp x.prop).le lemma measurable_defaultRatCDF (α : Type*) [MeasurableSpace α] (q : ℚ) : - Measurable (fun (_ : α) ↦ defaultRatCDF q) := measurable_const + Measurable (fun (_ : α) ↦ defaultRatCDF q) := measurable_const lemma isCDFLike_defaultRatCDF (α : Type*) [MeasurableSpace α] : IsCDFLike (fun (_ : α) (q : ℚ) ↦ defaultRatCDF q) where From be68022a18f2e0a08602b18f5c2fb5946c98ef26 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 16 Feb 2024 11:36:16 +0100 Subject: [PATCH 023/129] docstrings in MeasurableStieltjes --- .../Kernel/Disintegration/CDFKernel.lean | 10 +- .../Kernel/Disintegration/CondCdf.lean | 23 +- .../Kernel/Disintegration/KernelCDFBorel.lean | 17 +- .../Disintegration/MeasurableStieltjes.lean | 342 ++++++++++-------- 4 files changed, 215 insertions(+), 177 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean index 6f9f73007cf91..af6dd8420fdc4 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean @@ -27,22 +27,22 @@ variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} {f : α × β → ℚ → ℝ} {μ : kernel α (β × ℝ)} {ν : kernel α β} structure IsRatKernelCDF (f : α × β → ℚ → ℝ) (μ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := - (measurable (q : ℚ) : Measurable fun p ↦ f p q) + (measurable : Measurable f) (isRatStieltjesPoint_ae (a : α) : ∀ᵐ t ∂(ν a), IsRatStieltjesPoint f (a, t)) (integrable (a : α) (q : ℚ) : Integrable (fun t ↦ f (a, t) q) (ν a)) - (isCDF (a : α) {s : Set β} (_hs : MeasurableSet s) (q : ℚ) : + (set_integral (a : α) {s : Set β} (_hs : MeasurableSet s) (q : ℚ) : ∫ t in s, f (a, t) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal) lemma stieltjesOfMeasurableRat_ae_eq (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) : (fun t ↦ stieltjesOfMeasurableRat f hf.measurable (a, t) q) =ᵐ[ν a] fun t ↦ f (a, t) q := by filter_upwards [hf.isRatStieltjesPoint_ae a] with a ha - rw [stieltjesOfMeasurableRat_eq, toCDFLike_of_isRatStieltjesPoint ha] + rw [stieltjesOfMeasurableRat_eq, toRatCDF_of_isRatStieltjesPoint ha] lemma set_integral_stieltjesOfMeasurableRat_rat (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : ∫ t in s, stieltjesOfMeasurableRat f hf.measurable (a, t) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal := by - rw [set_integral_congr_ae hs (g := fun t ↦ f (a, t) q) ?_, hf.isCDF a hs] + rw [set_integral_congr_ae hs (g := fun t ↦ f (a, t) q) ?_, hf.set_integral a hs] filter_upwards [stieltjesOfMeasurableRat_ae_eq hf a q] with b hb using fun _ ↦ hb lemma set_lintegral_stieltjesOfMeasurableRat_rat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) @@ -74,7 +74,7 @@ lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKern suffices (μ a (s ×ˢ Iic (q : ℝ))).toReal = 0 by rw [ENNReal.toReal_eq_zero_iff] at this simpa [measure_ne_top] using this - rw [← hf.isCDF a hs q] + rw [← hf.set_integral a hs q] simp [hρ_zero] have h : ∫⁻ t in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) x) ∂(ν a) = ∫⁻ t in s, ⨅ r : { r' : ℚ // x < r' }, diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index 384214034556c..d176c0c854549 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -173,6 +173,11 @@ theorem measurable_preCDF {ρ : Measure (α × ℝ)} {r : ℚ} : Measurable (pre Measure.measurable_rnDeriv _ _ #align probability_theory.measurable_pre_cdf ProbabilityTheory.measurable_preCDF +lemma measurable_preCDF' {ρ : Measure (α × ℝ)} : + Measurable fun a r ↦ ENNReal.toReal (preCDF ρ r a) := by + rw [measurable_pi_iff] + exact fun _ ↦ measurable_preCDF.ennreal_toReal + theorem withDensity_preCDF (ρ : Measure (α × ℝ)) (r : ℚ) [IsFiniteMeasure ρ] : ρ.fst.withDensity (preCDF ρ r) = ρ.IicSnd r := Measure.absolutelyContinuous_iff_withDensity_rnDeriv_eq.mp (Measure.IicSnd_ac_fst ρ r) @@ -425,10 +430,6 @@ lemma isRatStieltjesPoint_ae (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : ((h2 r).trans_lt ENNReal.one_lt_top).ne rw [ENNReal.toReal_le_toReal (h_ne_top _) (h_ne_top _)] exact h1 hrr' - · exact fun _ ↦ ENNReal.toReal_nonneg - · refine fun r ↦ ENNReal.toReal_le_of_le_ofReal zero_le_one ?_ - rw [ENNReal.ofReal_one] - exact h2 r · rw [← ENNReal.one_toReal, ENNReal.tendsto_toReal_iff] · exact h3 · exact fun r ↦ ((h2 r).trans_lt ENNReal.one_lt_top).ne @@ -455,24 +456,22 @@ theorem integrable_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : lemma isRatKernelCDF_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : IsRatKernelCDF (fun p r ↦ (preCDF ρ r p.2).toReal) (kernel.const Unit ρ) (kernel.const Unit ρ.fst) where - measurable q := measurable_preCDF.ennreal_toReal.comp measurable_snd + measurable := measurable_preCDF'.comp measurable_snd isRatStieltjesPoint_ae a := by filter_upwards [isRatStieltjesPoint_ae ρ] with a ha - exact ⟨ha.mono, ha.nonneg, ha.le_one, ha.tendsto_atTop_one, ha.tendsto_atBot_zero, - ha.iInf_rat_gt_eq⟩ + exact ⟨ha.mono, ha.tendsto_atTop_one, ha.tendsto_atBot_zero, ha.iInf_rat_gt_eq⟩ integrable _ q := integrable_preCDF ρ q - isCDF a s hs q := by rw [kernel.const_apply, kernel.const_apply, set_integral_preCDF_fst _ _ hs, + set_integral a s hs q := by rw [kernel.const_apply, kernel.const_apply, set_integral_preCDF_fst _ _ hs, Measure.IicSnd_apply _ _ hs] /-- Conditional cdf of the measure given the value on `α`, as a Stieltjes function. -/ noncomputable def condCDF (ρ : Measure (α × ℝ)) (a : α) : StieltjesFunction := - stieltjesOfMeasurableRat (fun a r ↦ (preCDF ρ r a).toReal) - (fun _ ↦ measurable_preCDF.ennreal_toReal) a + stieltjesOfMeasurableRat (fun a r ↦ (preCDF ρ r a).toReal) measurable_preCDF' a #align probability_theory.cond_cdf ProbabilityTheory.condCDF lemma condCDF_eq_stieltjesOfMeasurableRat_unit_prod (ρ : Measure (α × ℝ)) (a : α) : condCDF ρ a = stieltjesOfMeasurableRat (fun (p : Unit × α) r ↦ (preCDF ρ r p.2).toReal) - (fun _ ↦ measurable_preCDF.ennreal_toReal.comp measurable_snd) ((), a) := by + (measurable_preCDF'.comp measurable_snd) ((), a) := by ext x rw [condCDF, ← stieltjesOfMeasurableRat_unit_prod] @@ -501,7 +500,7 @@ theorem tendsto_condCDF_atTop (ρ : Measure (α × ℝ)) (a : α) : theorem condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) : (fun a ↦ condCDF ρ a r) =ᵐ[ρ.fst] fun a ↦ (preCDF ρ r a).toReal := by filter_upwards [isRatStieltjesPoint_ae ρ] with a ha - rw [condCDF, stieltjesOfMeasurableRat_eq, toCDFLike_of_isRatStieltjesPoint ha] + rw [condCDF, stieltjesOfMeasurableRat_eq, toRatCDF_of_isRatStieltjesPoint ha] #align probability_theory.cond_cdf_ae_eq ProbabilityTheory.condCDF_ae_eq theorem ofReal_condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) : diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean index 84df3c3fcd09d..4b57be9d64e80 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean @@ -965,9 +965,10 @@ section Iic_Q noncomputable def mLimsupIic (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : ℝ := MLimsup κ a (Set.Iic q) t -lemma measurable_mLimsupIic (κ : kernel α (ℝ × ℝ)) (q : ℚ) : - Measurable (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2 q) := - measurable_mLimsup κ measurableSet_Iic +lemma measurable_mLimsupIic (κ : kernel α (ℝ × ℝ)) : + Measurable (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2) := by + rw [measurable_pi_iff] + exact fun _ ↦ measurable_mLimsup κ measurableSet_Iic lemma measurable_mLimsupIic_right (κ : kernel α (ℝ × ℝ)) (a : α) (q : ℚ) : Measurable (fun t ↦ mLimsupIic κ a t q) := @@ -1117,20 +1118,14 @@ lemma isRatStieltjesPoint_mLimsupIic_ae (κ : kernel α (ℝ × ℝ)) [IsFiniteK ∀ᵐ t ∂(kernel.fst κ a), IsRatStieltjesPoint (fun p q ↦ mLimsupIic κ p.1 p.2 q) (a, t) := by filter_upwards [tendsto_atTop_mLimsupIic κ a, tendsto_atBot_mLimsupIic κ a, iInf_rat_gt_mLimsupIic_eq κ a] with t ht_top ht_bot ht_iInf - constructor - · exact monotone_mLimsupIic κ a t - · exact mLimsupIic_nonneg κ a t - · exact mLimsupIic_le_one κ a t - · exact ht_top - · exact ht_bot - · exact ht_iInf + exact ⟨monotone_mLimsupIic κ a t, ht_top, ht_bot, ht_iInf⟩ lemma isRatKernelCDF_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : IsRatKernelCDF (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2) κ (kernel.fst κ) where measurable := measurable_mLimsupIic κ isRatStieltjesPoint_ae := isRatStieltjesPoint_mLimsupIic_ae κ integrable := integrable_mLimsupIic κ - isCDF := fun _ _ hs _ ↦ set_integral_mLimsupIic _ _ _ hs + set_integral := fun _ _ hs _ ↦ set_integral_mLimsupIic _ _ _ hs noncomputable def mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : α × ℝ → StieltjesFunction := diff --git a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean index 588339717332a..a348e13464c6a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean +++ b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean @@ -14,6 +14,29 @@ import Mathlib.Probability.Kernel.Disintegration.AuxLemmas /-! # Cumulative distributions functions of Markov kernels +We provide tools to build a measurable function `α → StieltjesFunction` with limits 0 at -∞ and 1 at ++∞ for all `a : α` from a measurable function `f : α → ℚ → ℝ`. + +This is possible if `f a : ℚ → ℝ` satisfies a package of properties for all `a`: monotonicity, +limits at +-∞ at a continuity property. We define `IsRatStieltjesPoint f a` to state that this is +the case at `a` and define the property `IsRatCDF f` that `f` is measurable and +`IsRatStieltjesPoint f a` for all `a`. +The function `α → StieltjesFunction` obtained by extending `f` by continuity is then called +`IsRatCDF.stieltjesFunction`. + +In applications, we will only have `IsRatStieltjesPoint f a` almost surely with respect to some +measure. We thus define a modification +`toRatCDF (f : α → ℚ → ℝ) := fun a q ↦ if IsRatStieltjesPoint f a then f a q else defaultRatCDF q`, +which satisfies the property `IsRatCDF (toRatCDF f)`. + +Finally, we define `stieltjesOfMeasurableRat`, composition of `toRatCDF` and +`IsRatCDF.stieltjesFunction`. + +## Main definitions + +* `stieltjesOfMeasurableRat`: turn a measurable function `f : α → ℚ → ℝ` into a measurable + function `α → StieltjesFunction`. + -/ @@ -25,14 +48,14 @@ namespace ProbabilityTheory variable {α β ι : Type*} [MeasurableSpace α] -section IsCDFLike +section IsRatCDF variable {f : α → ℚ → ℝ} +/-- `a : α` is a Stieltjes point for `f : α → ℚ → ℝ` if `f a` is monotone with limit 0 at -∞ +and 1 at +∞ and satisfies a continuity property. -/ structure IsRatStieltjesPoint (f : α → ℚ → ℝ) (a : α) : Prop where mono : Monotone (f a) - nonneg : ∀ q, 0 ≤ f a q - le_one : ∀ q, f a q ≤ 1 tendsto_atTop_one : Tendsto (f a) atTop (𝓝 1) tendsto_atBot_zero : Tendsto (f a) atBot (𝓝 0) iInf_rat_gt_eq : ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t @@ -40,13 +63,10 @@ structure IsRatStieltjesPoint (f : α → ℚ → ℝ) (a : α) : Prop where lemma isRatStieltjesPoint_unit_prod_iff (f : α → ℚ → ℝ) (a : α) : IsRatStieltjesPoint (fun p : Unit × α ↦ f p.2) ((), a) ↔ IsRatStieltjesPoint f a := by - refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ - · exact ⟨h.mono, h.nonneg, h.le_one, h.tendsto_atTop_one, h.tendsto_atBot_zero, - h.iInf_rat_gt_eq⟩ - · exact ⟨h.mono, h.nonneg, h.le_one, h.tendsto_atTop_one, h.tendsto_atBot_zero, - h.iInf_rat_gt_eq⟩ + constructor <;> + exact fun h ↦ ⟨h.mono, h.tendsto_atTop_one, h.tendsto_atBot_zero, h.iInf_rat_gt_eq⟩ -lemma measurableSet_isRatStieltjesPoint (hf : ∀ q, Measurable (fun a ↦ f a q)) : +lemma measurableSet_isRatStieltjesPoint (hf : Measurable f) : MeasurableSet {a | IsRatStieltjesPoint f a} := by have h1 : MeasurableSet {a | Monotone (f a)} := by change MeasurableSet {a | ∀ q r (_ : q ≤ r), f a q ≤ f a r} @@ -54,47 +74,72 @@ lemma measurableSet_isRatStieltjesPoint (hf : ∀ q, Measurable (fun a ↦ f a q refine MeasurableSet.iInter (fun q ↦ ?_) refine MeasurableSet.iInter (fun r ↦ ?_) refine MeasurableSet.iInter (fun _ ↦ ?_) - exact measurableSet_le (hf q) (hf r) - have h2 : MeasurableSet {a | ∀ q, 0 ≤ f a q} := by - simp_rw [Set.setOf_forall] - exact MeasurableSet.iInter (fun q ↦ measurableSet_le measurable_const (hf q)) - have h3 : MeasurableSet {a | ∀ q, f a q ≤ 1} := by - simp_rw [Set.setOf_forall] - exact MeasurableSet.iInter (fun q ↦ measurableSet_le (hf q) measurable_const) - have h4 : MeasurableSet {a | Tendsto (f a) atTop (𝓝 1)} := - measurableSet_tendsto_nhds (fun q ↦ hf q) 1 - have h5 : MeasurableSet {a | Tendsto (f a) atBot (𝓝 0)} := - measurableSet_tendsto_nhds (fun q ↦ hf q) 0 - have h6 : MeasurableSet {a | ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t} := by + exact measurableSet_le hf.eval hf.eval + have h2 : MeasurableSet {a | Tendsto (f a) atTop (𝓝 1)} := + measurableSet_tendsto_nhds (fun q ↦ hf.eval) 1 + have h3 : MeasurableSet {a | Tendsto (f a) atBot (𝓝 0)} := + measurableSet_tendsto_nhds (fun q ↦ hf.eval) 0 + have h4 : MeasurableSet {a | ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t} := by rw [Set.setOf_forall] refine MeasurableSet.iInter (fun q ↦ ?_) - exact measurableSet_eq_fun (measurable_iInf fun _ ↦ hf _) (hf _) + exact measurableSet_eq_fun (measurable_iInf fun _ ↦ hf.eval) hf.eval suffices {a | IsRatStieltjesPoint f a} - = ({a | Monotone (f a)} ∩ {a | ∀ (q : ℚ), 0 ≤ f a q} ∩ {a | ∀ (q : ℚ), f a q ≤ 1} - ∩ {a | Tendsto (f a) atTop (𝓝 1)} ∩ {a | Tendsto (f a) atBot (𝓝 0)} ∩ - {a | ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t}) by + = ({a | Monotone (f a)} ∩ {a | Tendsto (f a) atTop (𝓝 1)} ∩ {a | Tendsto (f a) atBot (𝓝 0)} + ∩ {a | ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t}) by rw [this] - exact ((((h1.inter h2).inter h3).inter h4).inter h5).inter h6 + exact (((h1.inter h2).inter h3).inter h4) ext a simp only [mem_setOf_eq, mem_inter_iff] refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ - · exact ⟨⟨⟨⟨⟨h.mono, h.nonneg⟩, h.le_one⟩, h.tendsto_atTop_one⟩, h.tendsto_atBot_zero⟩, - h.iInf_rat_gt_eq⟩ - · exact ⟨h.1.1.1.1.1, h.1.1.1.1.2, h.1.1.1.2, h.1.1.2, h.1.2, h.2⟩ + · exact ⟨⟨⟨h.mono, h.tendsto_atTop_one⟩, h.tendsto_atBot_zero⟩, h.iInf_rat_gt_eq⟩ + · exact ⟨h.1.1.1, h.1.1.2, h.1.2, h.2⟩ + +lemma IsRatStieltjesPoint.ite {f g : α → ℚ → ℝ} {a : α} (p : α → Prop) [DecidablePred p] + (hf : p a → IsRatStieltjesPoint f a) (hg : ¬ p a → IsRatStieltjesPoint g a): + IsRatStieltjesPoint (fun a ↦ if p a then f a else g a) a where + mono := by split_ifs with h; exacts [(hf h).mono, (hg h).mono] + tendsto_atTop_one := by + split_ifs with h; exacts [(hf h).tendsto_atTop_one, (hg h).tendsto_atTop_one] + tendsto_atBot_zero := by + split_ifs with h; exacts [(hf h).tendsto_atBot_zero, (hg h).tendsto_atBot_zero] + iInf_rat_gt_eq := by split_ifs with h; exacts [(hf h).iInf_rat_gt_eq, (hg h).iInf_rat_gt_eq] + + +/-- A function `f : α → ℚ → ℝ` is a (kernel) rational cumulative distribution function if it is +measurable in the first argument and if `f a` satisfies a list of properties for all `a : α`: +monotonicity between 0 at -∞ and 1 at +∞ and a form of continuity. + +A function with these properties can be extended to a measurable function `α → StieltjesFunction`. +See `ProbabilityTheory.IsRatCDF.stieltjesFunction`. +-/ +structure IsRatCDF (f : α → ℚ → ℝ) : Prop where + isRatStieltjesPoint : ∀ a, IsRatStieltjesPoint f a + measurable : Measurable f + +lemma IsRatCDF.nonneg {f : α → ℚ → ℝ} (hf : IsRatCDF f) (a : α) (q : ℚ) : 0 ≤ f a q := + Monotone.le_of_tendsto (hf.isRatStieltjesPoint a).mono + (hf.isRatStieltjesPoint a).tendsto_atBot_zero q + +lemma IsRatCDF.le_one {f : α → ℚ → ℝ} (hf : IsRatCDF f) (a : α) (q : ℚ) : f a q ≤ 1 := + Monotone.ge_of_tendsto (hf.isRatStieltjesPoint a).mono + (hf.isRatStieltjesPoint a).tendsto_atTop_one q -structure IsCDFLike (f : α → ℚ → ℝ) : Prop where - mono : ∀ a, Monotone (f a) - nonneg : ∀ a q, 0 ≤ f a q - le_one : ∀ a q, f a q ≤ 1 - tendsto_atTop_one : ∀ a, Tendsto (f a) atTop (𝓝 1) - tendsto_atBot_zero : ∀ a, Tendsto (f a) atBot (𝓝 0) - iInf_rat_gt_eq : ∀ a, ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t - measurable : ∀ q, Measurable (fun a ↦ f a q) +lemma IsRatCDF.tendsto_atTop_one {f : α → ℚ → ℝ} (hf : IsRatCDF f) (a : α) : + Tendsto (f a) atTop (𝓝 1) := (hf.isRatStieltjesPoint a).tendsto_atTop_one -end IsCDFLike +lemma IsRatCDF.tendsto_atBot_zero {f : α → ℚ → ℝ} (hf : IsRatCDF f) (a : α) : + Tendsto (f a) atBot (𝓝 0) := (hf.isRatStieltjesPoint a).tendsto_atBot_zero + +lemma IsRatCDF.iInf_rat_gt_eq {f : α → ℚ → ℝ} (hf : IsRatCDF f) (a : α) (q : ℚ) : + ⨅ r : Ioi q, f a r = f a q := (hf.isRatStieltjesPoint a).iInf_rat_gt_eq q + +end IsRatCDF section DefaultRatCDF +/-- A function with the property `IsRatCDF`. +Used in a `piecewise` construction to convert a function which only satisfies the properties +defining `IsRatCDF` almost everywhere into a true `IsRatCDF`. -/ def defaultRatCDF (q : ℚ) := if q < 0 then (0 : ℝ) else 1 lemma monotone_defaultRatCDF : Monotone defaultRatCDF := by @@ -150,69 +195,69 @@ lemma inf_gt_rat_defaultRatCDF (t : ℚ) : rw [not_lt] at h ⊢ exact h.trans (mem_Ioi.mp x.prop).le -lemma measurable_defaultRatCDF (α : Type*) [MeasurableSpace α] (q : ℚ) : - Measurable (fun (_ : α) ↦ defaultRatCDF q) := measurable_const +lemma isRatStieltjesPoint_defaultRatCDF (a : α) : + IsRatStieltjesPoint (fun (_ : α) ↦ defaultRatCDF) a where + mono := monotone_defaultRatCDF + tendsto_atTop_one := tendsto_defaultRatCDF_atTop + tendsto_atBot_zero := tendsto_defaultRatCDF_atBot + iInf_rat_gt_eq := inf_gt_rat_defaultRatCDF -lemma isCDFLike_defaultRatCDF (α : Type*) [MeasurableSpace α] : - IsCDFLike (fun (_ : α) (q : ℚ) ↦ defaultRatCDF q) where - mono _ := monotone_defaultRatCDF - nonneg _ := defaultRatCDF_nonneg - le_one _ := defaultRatCDF_le_one - tendsto_atBot_zero _ := tendsto_defaultRatCDF_atBot - tendsto_atTop_one _ := tendsto_defaultRatCDF_atTop - iInf_rat_gt_eq _ := inf_gt_rat_defaultRatCDF - measurable := measurable_defaultRatCDF α +lemma IsRatCDF_defaultRatCDF (α : Type*) [MeasurableSpace α] : + IsRatCDF (fun (_ : α) (q : ℚ) ↦ defaultRatCDF q) where + isRatStieltjesPoint := isRatStieltjesPoint_defaultRatCDF + measurable := measurable_const end DefaultRatCDF -section ToCDFLike +section ToRatCDF variable {f : α → ℚ → ℝ} open Classical in +/-- Turn a function `f : α → ℚ → ℝ` into another with the property `IsRatCDF`. +Mainly useful when `f` satisfies the property `IsRatStieltjesPoint f a` almost everywhere. -/ noncomputable -def toCDFLike (f : α → ℚ → ℝ) : α → ℚ → ℝ := fun a q ↦ - if IsRatStieltjesPoint f a then f a q else defaultRatCDF q - -lemma toCDFLike_of_isRatStieltjesPoint {a : α} (h : IsRatStieltjesPoint f a) (q : ℚ) : - toCDFLike f a q = f a q := by - unfold toCDFLike; simp [h] - -lemma isCDFLike_toCDFLike (hf : ∀ q, Measurable fun a ↦ f a q) : - IsCDFLike (toCDFLike f) where - mono a := by unfold toCDFLike; split_ifs with h; exacts [h.mono, monotone_defaultRatCDF] - nonneg a := by unfold toCDFLike; split_ifs with h; exacts [h.nonneg, defaultRatCDF_nonneg] - le_one a := by unfold toCDFLike; split_ifs with h; exacts [h.le_one, defaultRatCDF_le_one] - tendsto_atTop_one a := by - unfold toCDFLike; split_ifs with h; exacts [h.tendsto_atTop_one, tendsto_defaultRatCDF_atTop] - tendsto_atBot_zero a := by - unfold toCDFLike; split_ifs with h; exacts [h.tendsto_atBot_zero, tendsto_defaultRatCDF_atBot] - iInf_rat_gt_eq a := by - unfold toCDFLike; split_ifs with h; exacts [h.iInf_rat_gt_eq, inf_gt_rat_defaultRatCDF] - measurable q := - Measurable.ite (measurableSet_isRatStieltjesPoint hf) (hf q) (measurable_defaultRatCDF α q) - -lemma toCDFLike_unit_prod (a : α) : - toCDFLike (fun (p : Unit × α) ↦ f p.2) ((), a) = toCDFLike f a := by - unfold toCDFLike +def toRatCDF (f : α → ℚ → ℝ) : α → ℚ → ℝ := fun a ↦ + if IsRatStieltjesPoint f a then f a else defaultRatCDF + +lemma toRatCDF_of_isRatStieltjesPoint {a : α} (h : IsRatStieltjesPoint f a) (q : ℚ) : + toRatCDF f a q = f a q := by + rw [toRatCDF, if_pos h] + +lemma measurable_toRatCDF (hf : Measurable f) : Measurable (toRatCDF f) := + Measurable.ite (measurableSet_isRatStieltjesPoint hf) hf measurable_const + +lemma IsRatCDF_toRatCDF (hf : Measurable f) : + IsRatCDF (toRatCDF f) where + isRatStieltjesPoint a := by + classical + exact IsRatStieltjesPoint.ite (IsRatStieltjesPoint f) id + (fun _ ↦ isRatStieltjesPoint_defaultRatCDF a) + measurable := measurable_toRatCDF hf + +lemma toRatCDF_unit_prod (a : α) : + toRatCDF (fun (p : Unit × α) ↦ f p.2) ((), a) = toRatCDF f a := by + unfold toRatCDF rw [isRatStieltjesPoint_unit_prod_iff] -end ToCDFLike +end ToRatCDF -section IsCDFLike.stieltjesFunction +section IsRatCDF.stieltjesFunction -variable {f : α → ℚ → ℝ} (hf : IsCDFLike f) +variable {f : α → ℚ → ℝ} (hf : IsRatCDF f) -noncomputable irreducible_def IsCDFLike.stieltjesFunctionAux (f : α → ℚ → ℝ) : α → ℝ → ℝ := - fun a t ↦ ⨅ r : { r' : ℚ // t < r' }, f a r +/-- Auxiliary definition for `IsRatCDF.stieltjesFunction`: turn `f : α → ℚ → ℝ` into a function +`α → ℝ → ℝ` by assigning to `f a x` the infimum of `f a q` over `q : ℚ` with `x < q`. -/ +noncomputable irreducible_def IsRatCDF.stieltjesFunctionAux (f : α → ℚ → ℝ) : α → ℝ → ℝ := + fun a x ↦ ⨅ q : { q' : ℚ // x < q' }, f a q -lemma IsCDFLike.stieltjesFunctionAux_def' (f : α → ℚ → ℝ) (a : α) : - IsCDFLike.stieltjesFunctionAux f a = fun (t : ℝ) ↦ ⨅ r : { r' : ℚ // t < r' }, f a r := by - ext t; exact IsCDFLike.stieltjesFunctionAux_def f a t +lemma IsRatCDF.stieltjesFunctionAux_def' (f : α → ℚ → ℝ) (a : α) : + IsRatCDF.stieltjesFunctionAux f a = fun (t : ℝ) ↦ ⨅ r : { r' : ℚ // t < r' }, f a r := by + ext t; exact IsRatCDF.stieltjesFunctionAux_def f a t -lemma IsCDFLike.stieltjesFunctionAux_eq (a : α) (r : ℚ) : - IsCDFLike.stieltjesFunctionAux f a r = f a r := by - rw [← hf.iInf_rat_gt_eq a r, IsCDFLike.stieltjesFunctionAux] +lemma IsRatCDF.stieltjesFunctionAux_eq (a : α) (r : ℚ) : + IsRatCDF.stieltjesFunctionAux f a r = f a r := by + rw [← hf.iInf_rat_gt_eq a r, IsRatCDF.stieltjesFunctionAux] refine Equiv.iInf_congr ?_ ?_ · exact { toFun := fun t ↦ ⟨t.1, mod_cast t.2⟩ @@ -222,36 +267,32 @@ lemma IsCDFLike.stieltjesFunctionAux_eq (a : α) (r : ℚ) : · intro t simp only [Equiv.coe_fn_mk, Subtype.coe_mk] -lemma IsCDFLike.stieltjesFunctionAux_unit_prod (a : α) : - IsCDFLike.stieltjesFunctionAux (fun (p : Unit × α) ↦ f p.2) ((), a) = - IsCDFLike.stieltjesFunctionAux f a := by simp_rw [IsCDFLike.stieltjesFunctionAux_def'] +lemma IsRatCDF.stieltjesFunctionAux_unit_prod (a : α) : + IsRatCDF.stieltjesFunctionAux (fun (p : Unit × α) ↦ f p.2) ((), a) = + IsRatCDF.stieltjesFunctionAux f a := by simp_rw [IsRatCDF.stieltjesFunctionAux_def'] -lemma IsCDFLike.stieltjesFunctionAux_nonneg (a : α) (r : ℝ) : - 0 ≤ IsCDFLike.stieltjesFunctionAux f a r := by +lemma IsRatCDF.stieltjesFunctionAux_nonneg (a : α) (r : ℝ) : + 0 ≤ IsRatCDF.stieltjesFunctionAux f a r := by have : Nonempty { r' : ℚ // r < ↑r' } := by obtain ⟨r, hrx⟩ := exists_rat_gt r exact ⟨⟨r, hrx⟩⟩ - rw [IsCDFLike.stieltjesFunctionAux_def] + rw [IsRatCDF.stieltjesFunctionAux_def] exact le_ciInf fun r' ↦ hf.nonneg a _ -lemma bddBelow_range_gt (a : α) (x : ℝ) : - BddBelow (range fun r : { r' : ℚ // x < ↑r' } ↦ f a r) := by - refine ⟨0, fun z ↦ ?_⟩; rintro ⟨u, rfl⟩; exact hf.nonneg a _ - -lemma IsCDFLike.monotone_stieltjesFunctionAux (a : α) : - Monotone (IsCDFLike.stieltjesFunctionAux f a) := by +lemma IsRatCDF.monotone_stieltjesFunctionAux (a : α) : + Monotone (IsRatCDF.stieltjesFunctionAux f a) := by intro x y hxy have : Nonempty { r' : ℚ // y < ↑r' } := by obtain ⟨r, hrx⟩ := exists_rat_gt y exact ⟨⟨r, hrx⟩⟩ - simp_rw [IsCDFLike.stieltjesFunctionAux_def] + simp_rw [IsRatCDF.stieltjesFunctionAux_def] refine le_ciInf fun r ↦ (ciInf_le ?_ ?_).trans_eq ?_ - · exact bddBelow_range_gt hf a x + · refine ⟨0, fun z ↦ ?_⟩; rintro ⟨u, rfl⟩; exact hf.nonneg a _ · exact ⟨r.1, hxy.trans_lt r.prop⟩ · rfl -lemma IsCDFLike.continuousWithinAt_stieltjesFunctionAux_Ici (a : α) (x : ℝ) : - ContinuousWithinAt (IsCDFLike.stieltjesFunctionAux f a) (Ici x) x := by +lemma IsRatCDF.continuousWithinAt_stieltjesFunctionAux_Ici (a : α) (x : ℝ) : + ContinuousWithinAt (IsRatCDF.stieltjesFunctionAux f a) (Ici x) x := by rw [← continuousWithinAt_Ioi_iff_Ici] convert Monotone.tendsto_nhdsWithin_Ioi (monotone_stieltjesFunctionAux hf a) x rw [sInf_image'] @@ -270,25 +311,28 @@ lemma IsCDFLike.continuousWithinAt_stieltjesFunctionAux_Ici (a : α) (x : ℝ) congr! rw [stieltjesFunctionAux_def] -noncomputable def IsCDFLike.stieltjesFunction (a : α) : StieltjesFunction where +/-- Extend a function `f : α → ℚ → ℝ` with property `IsRatCDF` from `ℚ` to `ℝ`, to a function +`α → StieltjesFunction`. -/ +noncomputable def IsRatCDF.stieltjesFunction (a : α) : StieltjesFunction where toFun := stieltjesFunctionAux f a mono' := monotone_stieltjesFunctionAux hf a right_continuous' x := continuousWithinAt_stieltjesFunctionAux_Ici hf a x -lemma IsCDFLike.stieltjesFunction_eq (a : α) (r : ℚ) : hf.stieltjesFunction a r = f a r := +lemma IsRatCDF.stieltjesFunction_eq (a : α) (r : ℚ) : hf.stieltjesFunction a r = f a r := stieltjesFunctionAux_eq hf a r -lemma IsCDFLike.stieltjesFunction_nonneg (a : α) (r : ℝ) : 0 ≤ hf.stieltjesFunction a r := +lemma IsRatCDF.stieltjesFunction_nonneg (a : α) (r : ℝ) : 0 ≤ hf.stieltjesFunction a r := stieltjesFunctionAux_nonneg hf a r -lemma IsCDFLike.stieltjesFunction_le_one (a : α) (x : ℝ) : hf.stieltjesFunction a x ≤ 1 := by +lemma IsRatCDF.stieltjesFunction_le_one (a : α) (x : ℝ) : hf.stieltjesFunction a x ≤ 1 := by obtain ⟨r, hrx⟩ := exists_rat_gt x rw [← StieltjesFunction.iInf_rat_gt_eq] - simp_rw [IsCDFLike.stieltjesFunction_eq] - refine ciInf_le_of_le (bddBelow_range_gt hf a x) ?_ (hf.le_one _ _) - exact ⟨r, hrx⟩ + simp_rw [IsRatCDF.stieltjesFunction_eq] + refine ciInf_le_of_le ?_ ?_ (hf.le_one _ _) + · refine ⟨0, fun z ↦ ?_⟩; rintro ⟨u, rfl⟩; exact hf.nonneg a _ + · exact ⟨r, hrx⟩ -lemma IsCDFLike.tendsto_stieltjesFunction_atBot (a : α) : +lemma IsRatCDF.tendsto_stieltjesFunction_atBot (a : α) : Tendsto (hf.stieltjesFunction a) atBot (𝓝 0) := by have h_exists : ∀ x : ℝ, ∃ q : ℚ, x < q ∧ ↑q < x + 1 := fun x ↦ exists_rat_btwn (lt_add_one x) let qs : ℝ → ℚ := fun x ↦ (h_exists x).choose @@ -304,7 +348,7 @@ lemma IsCDFLike.tendsto_stieltjesFunction_atBot (a : α) : rw [Function.comp_apply, ← stieltjesFunction_eq hf] exact (hf.stieltjesFunction a).mono (h_exists x).choose_spec.1.le -lemma IsCDFLike.tendsto_stieltjesFunction_atTop (a : α) : +lemma IsRatCDF.tendsto_stieltjesFunction_atTop (a : α) : Tendsto (hf.stieltjesFunction a) atTop (𝓝 1) := by have h_exists : ∀ x : ℝ, ∃ q : ℚ, x - 1 < q ∧ ↑q < x := fun x ↦ exists_rat_btwn (sub_one_lt x) let qs : ℝ → ℚ := fun x ↦ (h_exists x).choose @@ -320,7 +364,7 @@ lemma IsCDFLike.tendsto_stieltjesFunction_atTop (a : α) : rw [Function.comp_apply, ← stieltjesFunction_eq hf] exact (hf.stieltjesFunction a).mono (le_of_lt (h_exists x).choose_spec.2) -lemma IsCDFLike.measurable_stieltjesFunction (x : ℝ) : +lemma IsRatCDF.measurable_stieltjesFunction (x : ℝ) : Measurable fun a ↦ hf.stieltjesFunction a x := by have : (fun a ↦ hf.stieltjesFunction a x) = fun a ↦ ⨅ r : { r' : ℚ // x < r' }, f a ↑r := by ext1 a @@ -328,30 +372,30 @@ lemma IsCDFLike.measurable_stieltjesFunction (x : ℝ) : congr with q rw [stieltjesFunction_eq] rw [this] - exact measurable_iInf (fun q ↦ hf.measurable q) + exact measurable_iInf (fun q ↦ hf.measurable.eval) -lemma IsCDFLike.stronglyMeasurable_stieltjesFunction (x : ℝ) : +lemma IsRatCDF.stronglyMeasurable_stieltjesFunction (x : ℝ) : StronglyMeasurable fun a ↦ hf.stieltjesFunction a x := (measurable_stieltjesFunction hf x).stronglyMeasurable section Measure -lemma IsCDFLike.measure_stieltjesFunction_Iic (a : α) (x : ℝ) : +lemma IsRatCDF.measure_stieltjesFunction_Iic (a : α) (x : ℝ) : (hf.stieltjesFunction a).measure (Iic x) = ENNReal.ofReal (hf.stieltjesFunction a x) := by rw [← sub_zero (hf.stieltjesFunction a x)] exact (hf.stieltjesFunction a).measure_Iic (tendsto_stieltjesFunction_atBot hf a) _ -lemma IsCDFLike.measure_stieltjesFunction_univ (a : α) : +lemma IsRatCDF.measure_stieltjesFunction_univ (a : α) : (hf.stieltjesFunction a).measure univ = 1 := by rw [← ENNReal.ofReal_one, ← sub_zero (1 : ℝ)] exact StieltjesFunction.measure_univ _ (tendsto_stieltjesFunction_atBot hf a) (tendsto_stieltjesFunction_atTop hf a) -instance IsCDFLike.instIsProbabilityMeasure_stieltjesFunction (a : α) : +instance IsRatCDF.instIsProbabilityMeasure_stieltjesFunction (a : α) : IsProbabilityMeasure (hf.stieltjesFunction a).measure := ⟨measure_stieltjesFunction_univ hf a⟩ -lemma IsCDFLike.measurable_measure_stieltjesFunction : +lemma IsRatCDF.measurable_measure_stieltjesFunction : Measurable fun a ↦ (hf.stieltjesFunction a).measure := by rw [Measure.measurable_measure] refine fun s hs ↦ MeasurableSpace.induction_on_inter @@ -375,72 +419,72 @@ lemma IsCDFLike.measurable_measure_stieltjesFunction : end Measure -end IsCDFLike.stieltjesFunction +end IsRatCDF.stieltjesFunction section stieltjesOfMeasurableRat variable {f : α → ℚ → ℝ} +/-- Turn a measurable function `f : α → ℚ → ℝ` into a measurable function `α → StieltjesFunction`. +Composition of `toRatCDF` and `IsRatCDF.stieltjesFunction`. -/ noncomputable -def stieltjesOfMeasurableRat (f : α → ℚ → ℝ) (hf : ∀ q, Measurable fun a ↦ f a q) : - α → StieltjesFunction := - (isCDFLike_toCDFLike hf).stieltjesFunction +def stieltjesOfMeasurableRat (f : α → ℚ → ℝ) (hf : Measurable f) : α → StieltjesFunction := + (IsRatCDF_toRatCDF hf).stieltjesFunction -lemma stieltjesOfMeasurableRat_eq (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℚ) : - stieltjesOfMeasurableRat f hf a r = toCDFLike f a r := IsCDFLike.stieltjesFunction_eq _ a r +lemma stieltjesOfMeasurableRat_eq (hf : Measurable f) (a : α) (r : ℚ) : + stieltjesOfMeasurableRat f hf a r = toRatCDF f a r := IsRatCDF.stieltjesFunction_eq _ a r -lemma stieltjesOfMeasurableRat_unit_prod (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : - stieltjesOfMeasurableRat (fun (p : Unit × α) ↦ f p.2) - (fun q ↦ (hf q).comp measurable_snd) ((), a) +lemma stieltjesOfMeasurableRat_unit_prod (hf : Measurable f) (a : α) : + stieltjesOfMeasurableRat (fun (p : Unit × α) ↦ f p.2) (hf.comp measurable_snd) ((), a) = stieltjesOfMeasurableRat f hf a := by - simp_rw [stieltjesOfMeasurableRat,IsCDFLike.stieltjesFunction, - ← IsCDFLike.stieltjesFunctionAux_unit_prod a] + simp_rw [stieltjesOfMeasurableRat,IsRatCDF.stieltjesFunction, + ← IsRatCDF.stieltjesFunctionAux_unit_prod a] congr with x congr 1 with p : 1 cases p with - | mk _ b => rw [← toCDFLike_unit_prod b] + | mk _ b => rw [← toRatCDF_unit_prod b] -lemma stieltjesOfMeasurableRat_nonneg (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (r : ℝ) : - 0 ≤ stieltjesOfMeasurableRat f hf a r := IsCDFLike.stieltjesFunction_nonneg _ a r +lemma stieltjesOfMeasurableRat_nonneg (hf : Measurable f) (a : α) (r : ℝ) : + 0 ≤ stieltjesOfMeasurableRat f hf a r := IsRatCDF.stieltjesFunction_nonneg _ a r -lemma stieltjesOfMeasurableRat_le_one (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : - stieltjesOfMeasurableRat f hf a x ≤ 1 := IsCDFLike.stieltjesFunction_le_one _ a x +lemma stieltjesOfMeasurableRat_le_one (hf : Measurable f) (a : α) (x : ℝ) : + stieltjesOfMeasurableRat f hf a x ≤ 1 := IsRatCDF.stieltjesFunction_le_one _ a x -lemma tendsto_stieltjesOfMeasurableRat_atBot (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : +lemma tendsto_stieltjesOfMeasurableRat_atBot (hf : Measurable f) (a : α) : Tendsto (stieltjesOfMeasurableRat f hf a) atBot (𝓝 0) := - IsCDFLike.tendsto_stieltjesFunction_atBot _ a + IsRatCDF.tendsto_stieltjesFunction_atBot _ a -lemma tendsto_stieltjesOfMeasurableRat_atTop (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : +lemma tendsto_stieltjesOfMeasurableRat_atTop (hf : Measurable f) (a : α) : Tendsto (stieltjesOfMeasurableRat f hf a) atTop (𝓝 1) := - IsCDFLike.tendsto_stieltjesFunction_atTop _ a + IsRatCDF.tendsto_stieltjesFunction_atTop _ a -lemma measurable_stieltjesOfMeasurableRat (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : +lemma measurable_stieltjesOfMeasurableRat (hf : Measurable f) (x : ℝ) : Measurable fun a ↦ stieltjesOfMeasurableRat f hf a x := - IsCDFLike.measurable_stieltjesFunction _ x + IsRatCDF.measurable_stieltjesFunction _ x -lemma stronglyMeasurable_stieltjesOfMeasurableRat (hf : ∀ q, Measurable fun a ↦ f a q) (x : ℝ) : +lemma stronglyMeasurable_stieltjesOfMeasurableRat (hf : Measurable f) (x : ℝ) : StronglyMeasurable fun a ↦ stieltjesOfMeasurableRat f hf a x := - IsCDFLike.stronglyMeasurable_stieltjesFunction _ x + IsRatCDF.stronglyMeasurable_stieltjesFunction _ x section Measure -lemma measure_stieltjesOfMeasurableRat_Iic (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) (x : ℝ) : +lemma measure_stieltjesOfMeasurableRat_Iic (hf : Measurable f) (a : α) (x : ℝ) : (stieltjesOfMeasurableRat f hf a).measure (Iic x) = ENNReal.ofReal (stieltjesOfMeasurableRat f hf a x) := - IsCDFLike.measure_stieltjesFunction_Iic _ _ _ + IsRatCDF.measure_stieltjesFunction_Iic _ _ _ -lemma measure_stieltjesOfMeasurableRat_univ (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : +lemma measure_stieltjesOfMeasurableRat_univ (hf : Measurable f) (a : α) : (stieltjesOfMeasurableRat f hf a).measure univ = 1 := - IsCDFLike.measure_stieltjesFunction_univ _ _ + IsRatCDF.measure_stieltjesFunction_univ _ _ instance instIsProbabilityMeasure_stieltjesOfMeasurableRat - (hf : ∀ q, Measurable fun a ↦ f a q) (a : α) : + (hf : Measurable f) (a : α) : IsProbabilityMeasure (stieltjesOfMeasurableRat f hf a).measure := - IsCDFLike.instIsProbabilityMeasure_stieltjesFunction _ _ + IsRatCDF.instIsProbabilityMeasure_stieltjesFunction _ _ -lemma measurable_measure_stieltjesOfMeasurableRat (hf : ∀ q, Measurable fun a ↦ f a q) : +lemma measurable_measure_stieltjesOfMeasurableRat (hf : Measurable f) : Measurable fun a ↦ (stieltjesOfMeasurableRat f hf a).measure := - IsCDFLike.measurable_measure_stieltjesFunction _ + IsRatCDF.measurable_measure_stieltjesFunction _ end Measure From 7f42f9fd1c6080120308a448f67d4796b5b7910e Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 16 Feb 2024 13:54:49 +0100 Subject: [PATCH 024/129] add docstring --- .../Kernel/Disintegration/Basic.lean | 61 ++++++++++++------- .../Disintegration/MeasurableStieltjes.lean | 16 ++--- 2 files changed, 48 insertions(+), 29 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 88cb1b9e3f190..b0fc4351b8789 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -9,34 +9,51 @@ import Mathlib.Probability.Kernel.Disintegration.KernelCDFBorel /-! # Disintegration of kernels and measures -Let `ρ` be a finite measure on `α × Ω`, where `Ω` is a standard Borel space. In mathlib terms, `Ω` -verifies `[Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω]`. -Then there exists a kernel `ρ.condKernel : kernel α Ω` such that for any measurable set -`s : Set (α × Ω)`, `ρ s = ∫⁻ a, ρ.condKernel a {x | (a, x) ∈ s} ∂ρ.fst`. - -In terms of kernels, `ρ.condKernel` is such that for any measurable space `γ`, we -have a disintegration of the constant kernel from `γ` with value `ρ`: -`kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ (condKernel ρ))`, -where `ρ.fst` is the marginal measure of `ρ` on `α`. In particular, `ρ = ρ.fst ⊗ₘ ρ.condKernel`. - -In order to obtain a disintegration for any standard Borel space, we use that these spaces embed -measurably into `ℝ`: it then suffices to define a suitable kernel for `Ω = ℝ`. In the real case, -we define a conditional kernel by taking for each `a : α` the measure associated to the Stieltjes -function `condCDF ρ a` (the conditional cumulative distribution function). +Let `κ : kernel α (β × Ω)` be a finite kernel, where `Ω` is a standard Borel space. Then if `α` is +countable or `β` is a standard Borel space, there exists a `kernel (α × β) Ω`, called conditional +kernel and denoted by `kernel.condKernel κ` such that `κ = kernel.fst κ ⊗ₖ kernel.condKernel κ`. +We also define a conditional kernel for a measure `ρ : Measure (β × Ω)`, where `Ω` is a standard +Borel space. This is a `kernel β Ω` denoted by `ρ.condKernel` such that `ρ = ρ.fst ⊗ₘ ρ.condKernel`. + +In order to obtain a disintegration for any standard Borel space `Ω`, we use that these spaces embed +measurably into `ℝ`: it then suffices to define a suitable kernel for `Ω = ℝ`. + +For `κ : kernel α (β × ℝ)`, the construction of the conditional kernel proceeds as follows: +* Build a measurable function `f : (α × β) → ℚ → ℝ` such that for all measurable sets + `s` and all `q : ℚ`, `∫ x in s, f (a, x) q ∂(kernel.fst κ a) = (κ a (s ×ˢ Iic (q : ℝ))).toReal`. + We restrict to `Q` here to be able to prove the measurability. +* Extend that function to `(α × β) → StieltjesFunction`. See the file `MeasurableStieltjes.lean`. +* Finally obtain from the measurable Stieltjes function a measure on `ℝ` for each element of `α × β` + in a measurable way: we have obtained a `kernel (α × β) ℝ`. + See the file `CDFKernel.lean` for that step. + +The first step (building the measurable function on `ℚ`) is done differently depending on whether +`α` is countable or not. +* If `α` is countable, we can provide for each `a : α` a function `f : β → ℚ → ℝ` and proceed as + above to obtain a `kernel β ℝ`. Since `α` is countable, measurability is not an issue and we can + put those together into a `kernel (α × β) ℝ`. The construction of that `f` is done in + the `CondCDF.lean` file. +* If `α` is not countable, we can't proceed separately for each `a : α` and have to build a function + `f : α × β → ℚ → ℝ` which is measurable on the product. We are able to do so if `β` is standard + borel by reducing to the case `β = ℝ`. See the file `KernelCDFBorel.lean`. + +We define a class `CountableOrStandardBorel α β` which encodes the property +`(Countable α ∧ MeasurableSingletonClass α) ∨ StandardBorelSpace β`. The conditional kernel is +defined under that assumption. + +Properties of integrals involving `kernel.condKernel` are collated in the file `Integral.lean`. +The conditional kernel is unique (almost everywhere w.r.t. `kernel.fst κ`): this is proved in the +file `Unique.lean`. ## Main definitions -* `MeasureTheory.Measure.condKernel ρ : kernel α Ω`: conditional kernel described above. +* `ProbabilityTheory.kernel.condKernel κ : kernel (α × β) Ω`: conditional kernel described above. +* `MeasureTheory.Measure.condKernel ρ : kernel β Ω`: conditional kernel of a measure. ## Main statements -* `ProbabilityTheory.lintegral_condKernel`: - `∫⁻ a, ∫⁻ ω, f (a, ω) ∂(ρ.condKernel a) ∂ρ.fst = ∫⁻ x, f x ∂ρ` -* `ProbabilityTheory.lintegral_condKernel_mem`: - `∫⁻ a, ρ.condKernel a {x | (a, x) ∈ s} ∂ρ.fst = ρ s` -* `ProbabilityTheory.kernel.const_eq_compProd`: - `kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ ρ.condKernel)` -* `ProbabilityTheory.measure_eq_compProd`: `ρ = ρ.fst ⊗ₘ ρ.condKernel` +* `ProbabilityTheory.kernel.compProd_fst_condKernel`: `kernel.fst κ ⊗ₖ kernel.condKernel κ = κ` +* `MeasureTheory.Measure.compProd_fst_condKernel`: `ρ.fst ⊗ₘ ρ.condKernel = ρ` -/ diff --git a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean index a348e13464c6a..22470d6ff7d0a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean +++ b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean @@ -15,17 +15,19 @@ import Mathlib.Probability.Kernel.Disintegration.AuxLemmas # Cumulative distributions functions of Markov kernels We provide tools to build a measurable function `α → StieltjesFunction` with limits 0 at -∞ and 1 at -+∞ for all `a : α` from a measurable function `f : α → ℚ → ℝ`. ++∞ for all `a : α` from a measurable function `f : α → ℚ → ℝ`. The reason for going through `ℚ` +instead of defining directly a Stieltjes function is that since `ℚ` is countable, building a +measurable function is much easier. -This is possible if `f a : ℚ → ℝ` satisfies a package of properties for all `a`: monotonicity, -limits at +-∞ at a continuity property. We define `IsRatStieltjesPoint f a` to state that this is -the case at `a` and define the property `IsRatCDF f` that `f` is measurable and +This construction will be possible if `f a : ℚ → ℝ` satisfies a package of properties for all `a`: +monotonicity, limits at +-∞ at a continuity property. We define `IsRatStieltjesPoint f a` to state +that this is the case at `a` and define the property `IsRatCDF f` that `f` is measurable and `IsRatStieltjesPoint f a` for all `a`. -The function `α → StieltjesFunction` obtained by extending `f` by continuity is then called -`IsRatCDF.stieltjesFunction`. +The function `α → StieltjesFunction` obtained by extending `f` by continuity from the right is then +called `IsRatCDF.stieltjesFunction`. In applications, we will only have `IsRatStieltjesPoint f a` almost surely with respect to some -measure. We thus define a modification +measure. We thus define `toRatCDF (f : α → ℚ → ℝ) := fun a q ↦ if IsRatStieltjesPoint f a then f a q else defaultRatCDF q`, which satisfies the property `IsRatCDF (toRatCDF f)`. From 75e3389785f479f2233d39c8bf1cac5f7fea1ba1 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 16 Feb 2024 13:58:07 +0100 Subject: [PATCH 025/129] Mathlib.lean --- Mathlib.lean | 9 +++++++-- 1 file changed, 7 insertions(+), 2 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 24bc64ae0988c..ee96b0276107c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3106,10 +3106,15 @@ import Mathlib.Probability.Independence.ZeroOne import Mathlib.Probability.Integration import Mathlib.Probability.Kernel.Basic import Mathlib.Probability.Kernel.Composition -import Mathlib.Probability.Kernel.CondCdf import Mathlib.Probability.Kernel.CondDistrib import Mathlib.Probability.Kernel.Condexp -import Mathlib.Probability.Kernel.Disintegration +import Mathlib.Probability.Kernel.Disintegration.Basic +import Mathlib.Probability.Kernel.Disintegration.CDFKernel +import Mathlib.Probability.Kernel.Disintegration.CondCdf +import Mathlib.Probability.Kernel.Disintegration.integral +import Mathlib.Probability.Kernel.Disintegration.KernelCDFBorel +import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes +import Mathlib.Probability.Kernel.Disintegration.Unique import Mathlib.Probability.Kernel.IntegralCompProd import Mathlib.Probability.Kernel.Invariance import Mathlib.Probability.Kernel.MeasurableIntegral From 2de05e6a7fbe61fcf94f616f13dec89beb6def9c Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 16 Feb 2024 14:02:10 +0100 Subject: [PATCH 026/129] fix --- Mathlib.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib.lean b/Mathlib.lean index ee96b0276107c..82786ebd7f15f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3111,7 +3111,7 @@ import Mathlib.Probability.Kernel.Condexp import Mathlib.Probability.Kernel.Disintegration.Basic import Mathlib.Probability.Kernel.Disintegration.CDFKernel import Mathlib.Probability.Kernel.Disintegration.CondCdf -import Mathlib.Probability.Kernel.Disintegration.integral +import Mathlib.Probability.Kernel.Disintegration.Integral import Mathlib.Probability.Kernel.Disintegration.KernelCDFBorel import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes import Mathlib.Probability.Kernel.Disintegration.Unique From ef5d2374e5a17ddef77f33a855516dd7e6ea4723 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sat, 17 Feb 2024 10:44:42 +0100 Subject: [PATCH 027/129] change KernelCDFBorel to prepare for kernel Radon-Nikodym --- .../Kernel/Disintegration/KernelCDFBorel.lean | 1008 +++++++++-------- 1 file changed, 550 insertions(+), 458 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean index 4b57be9d64e80..c6ee8e8bdc2a4 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean @@ -245,20 +245,25 @@ end dissection_system variable [MeasurableSpace β] +section M + +variable {κ : kernel α (ℝ × β)} {ν : kernel α ℝ} + noncomputable -def M (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : ℝ := - (κ a (I n (indexI n t) ×ˢ s) / kernel.fst κ a (I n (indexI n t))).toReal +def M (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : ℝ := + (κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t))).toReal -lemma m_def (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) : - M κ a s n = fun t ↦ (κ a (I n (indexI n t) ×ˢ s) / kernel.fst κ a (I n (indexI n t))).toReal := +lemma m_def (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) : + M κ ν a s n = fun t ↦ (κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t))).toReal := rfl -lemma measurable_m_aux (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : +lemma measurable_m_aux (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) + {s : Set β} (hs : MeasurableSet s) (n : ℕ) : Measurable (fun (p : α × ℝ) ↦ - κ p.1 (I n (indexI n p.2) ×ˢ s) / kernel.fst κ p.1 (I n (indexI n p.2))) := by + κ p.1 (I n (indexI n p.2) ×ˢ s) / ν p.1 (I n (indexI n p.2))) := by change Measurable ((fun (p : α × ℤ) ↦ - κ p.1 (I n p.2 ×ˢ s) / kernel.fst κ p.1 (I n p.2)) ∘ (fun (p : α × ℝ) ↦ (p.1, indexI n p.2))) - have h1 : Measurable (fun (p : α × ℤ) ↦ κ p.1 (I n p.2 ×ˢ s) / kernel.fst κ p.1 (I n p.2)) := by + κ p.1 (I n p.2 ×ˢ s) / ν p.1 (I n p.2)) ∘ (fun (p : α × ℝ) ↦ (p.1, indexI n p.2))) + have h1 : Measurable (fun (p : α × ℤ) ↦ κ p.1 (I n p.2 ×ˢ s) / ν p.1 (I n p.2)) := by refine Measurable.div ?_ ?_ · have h_swap : Measurable fun (p : ℤ × α) ↦ κ p.2 (I n p.1 ×ˢ s) := by refine measurable_uncurry_of_continuous_of_measurable @@ -267,118 +272,128 @@ lemma measurable_m_aux (κ : kernel α (ℝ × β)) {s : Set β} (hs : Measurabl · exact fun _ ↦ kernel.measurable_coe _ ((measurableSet_I _ _).prod hs) change Measurable ((fun (p : ℤ × α) ↦ κ p.2 (I n p.1 ×ˢ s)) ∘ Prod.swap) exact h_swap.comp measurable_swap - · have h_swap : Measurable fun (p : ℤ × α) ↦ kernel.fst κ p.2 (I n p.1) := by + · have h_swap : Measurable fun (p : ℤ × α) ↦ ν p.2 (I n p.1) := by refine measurable_uncurry_of_continuous_of_measurable - (u := fun k a ↦ kernel.fst κ a (I n k)) ?_ ?_ + (u := fun k a ↦ ν a (I n k)) ?_ ?_ · exact fun _ ↦ continuous_bot · exact fun _ ↦ kernel.measurable_coe _ (measurableSet_I _ _) - change Measurable ((fun (p : ℤ × α) ↦ kernel.fst κ p.2 (I n p.1)) ∘ Prod.swap) + change Measurable ((fun (p : ℤ × α) ↦ ν p.2 (I n p.1)) ∘ Prod.swap) exact h_swap.comp measurable_swap refine h1.comp (measurable_fst.prod_mk ?_) exact (Measurable.mono (measurable_indexI n) (ℱ.le n) le_rfl).comp measurable_snd -lemma measurable_m (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : - Measurable (fun (p : α × ℝ) ↦ M κ p.1 s n p.2) := - (measurable_m_aux κ hs n).ennreal_toReal +lemma measurable_m (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) + {s : Set β} (hs : MeasurableSet s) (n : ℕ) : + Measurable (fun (p : α × ℝ) ↦ M κ ν p.1 s n p.2) := + (measurable_m_aux κ ν hs n).ennreal_toReal -lemma measurable_m_left (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) (n : ℕ) (t : ℝ) : - Measurable (fun a ↦ M κ a s n t) := - (measurable_m κ hs n).comp (measurable_id.prod_mk measurable_const) +lemma measurable_m_left (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) + {s : Set β} (hs : MeasurableSet s) (n : ℕ) (t : ℝ) : + Measurable (fun a ↦ M κ ν a s n t) := + (measurable_m κ ν hs n).comp (measurable_id.prod_mk measurable_const) -lemma measurable_m_right (κ : kernel α (ℝ × β)) {s : Set β} (a : α) (hs : MeasurableSet s) (n : ℕ) : - Measurable (M κ a s n) := - (measurable_m κ hs n).comp (measurable_const.prod_mk measurable_id) +lemma measurable_m_right (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) + {s : Set β} (a : α) (hs : MeasurableSet s) (n : ℕ) : + Measurable (M κ ν a s n) := + (measurable_m κ ν hs n).comp (measurable_const.prod_mk measurable_id) -lemma measurable_ℱ_m (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) : - Measurable[ℱ n] (M κ a s n) := by +lemma measurable_ℱ_m (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) : + Measurable[ℱ n] (M κ ν a s n) := by rw [m_def] refine @Measurable.ennreal_toReal _ (ℱ n) _ ?_ refine Measurable.div ?_ ?_ · change Measurable[ℱ n] ((fun k ↦ κ a (I n k ×ˢ s)) ∘ (indexI n)) refine Measurable.comp ?_ (measurable_indexI n) exact measurable_of_countable _ - · change Measurable[ℱ n] ((fun k ↦ (kernel.fst κ) a (I n k)) ∘ (indexI n)) + · change Measurable[ℱ n] ((fun k ↦ ν a (I n k)) ∘ (indexI n)) refine Measurable.comp ?_ (measurable_indexI n) exact measurable_of_countable _ -lemma stronglyMeasurable_ℱ_m (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) : - StronglyMeasurable[ℱ n] (M κ a s n) := - (measurable_ℱ_m κ a s n).stronglyMeasurable +lemma stronglyMeasurable_ℱ_m (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) : + StronglyMeasurable[ℱ n] (M κ ν a s n) := + (measurable_ℱ_m κ ν a s n).stronglyMeasurable -lemma adapted_m (κ : kernel α (ℝ × β)) (a : α) (s : Set β) : Adapted ℱ (M κ a s) := - stronglyMeasurable_ℱ_m κ a s +lemma adapted_m (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) : + Adapted ℱ (M κ ν a s) := + stronglyMeasurable_ℱ_m κ ν a s -lemma m_nonneg (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : - 0 ≤ M κ a s n t := +lemma m_nonneg (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : + 0 ≤ M κ ν a s n t := ENNReal.toReal_nonneg -lemma m_le_one (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : - M κ a s n t ≤ 1 := by +lemma m_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : + M κ ν a s n t ≤ 1 := by rw [M] refine ENNReal.toReal_le_of_le_ofReal zero_le_one ?_ rw [ENNReal.ofReal_one] refine ENNReal.div_le_of_le_mul ?_ - rw [one_mul, kernel.fst_apply' _ _ (measurableSet_I _ _)] - refine measure_mono (fun x ↦ ?_) - simp only [mem_prod, mem_setOf_eq, and_imp] - exact fun h _ ↦ h - -lemma snorm_m_le (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) : - snorm (M κ a s n) 1 (kernel.fst κ a) ≤ kernel.fst κ a univ := by + rw [one_mul] + calc κ a (I n (indexI n t) ×ˢ s) ≤ kernel.fst κ a (I n (indexI n t)) := by + rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] + refine measure_mono (fun x ↦ ?_) + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + _ ≤ ν a (I n (indexI n t)) := hκν a _ (measurableSet_I _ _) + +lemma snorm_m_le (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (n : ℕ) : + snorm (M κ ν a s n) 1 (ν a) ≤ ν a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun x ↦ ?_))).trans ?_ - · simp only [Real.norm_eq_abs, abs_of_nonneg (m_nonneg κ a s n x), m_le_one κ a s n x] + · simp only [Real.norm_eq_abs, abs_of_nonneg (m_nonneg κ ν a s n x), m_le_one hκν a s n x] · simp -lemma integrable_m (κ : kernel α (ℝ × β)) [IsFiniteKernel (kernel.fst κ)] +lemma integrable_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : - Integrable (M κ a s n) (kernel.fst κ a) := by + Integrable (M κ ν a s n) (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ - · exact measurable_m_right κ a hs n - · exact (snorm_m_le κ a s n).trans_lt (measure_lt_top _ _) + · exact measurable_m_right κ ν a hs n + · exact (snorm_m_le hκν a s n).trans_lt (measure_lt_top _ _) -lemma set_integral_m_I (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma set_integral_m_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) (k : ℤ) : - ∫ t in I n k, M κ a s n t ∂(kernel.fst κ a) = (κ a (I n k ×ˢ s)).toReal := by + ∫ t in I n k, M κ ν a s n t ∂(ν a) = (κ a (I n k ×ˢ s)).toReal := by simp_rw [M] rw [integral_toReal] rotate_left · refine Measurable.aemeasurable ?_ - have h := measurable_m_aux κ hs n + have h := measurable_m_aux κ ν hs n change Measurable - ((fun (p : α × ℝ) ↦ κ p.1 (I n (indexI n p.2) ×ˢ s) / kernel.fst κ p.1 (I n (indexI n p.2))) + ((fun (p : α × ℝ) ↦ κ p.1 (I n (indexI n p.2) ×ˢ s) / ν p.1 (I n (indexI n p.2))) ∘ (fun t ↦ (a, t))) exact h.comp measurable_prod_mk_left · refine ae_of_all _ (fun t ↦ ?_) - by_cases h0 : kernel.fst κ a (I n (indexI n t)) = 0 + by_cases h0 : ν a (I n (indexI n t)) = 0 · suffices κ a (I n (indexI n t) ×ˢ s) = 0 by simp [h0, this] - rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0 - refine measure_mono_null (fun x ↦ ?_) h0 + have h0' : kernel.fst κ a (I n (indexI n t)) = 0 := + le_antisymm ((hκν a _ (measurableSet_I _ _)).trans h0.le) zero_le' + rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0' + refine measure_mono_null (fun x ↦ ?_) h0' simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h · refine ENNReal.div_lt_top ?_ h0 exact measure_ne_top _ _ congr - have : ∫⁻ t in I n k, κ a (I n (indexI n t) ×ˢ s) - / kernel.fst κ a (I n (indexI n t)) ∂(kernel.fst κ) a - = ∫⁻ _ in I n k, κ a (I n k ×ˢ s) / kernel.fst κ a (I n k) ∂(kernel.fst κ) a := by + have : ∫⁻ t in I n k, κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t)) ∂(ν a) + = ∫⁻ _ in I n k, κ a (I n k ×ˢ s) / ν a (I n k) ∂(ν a) := by refine set_lintegral_congr_fun (measurableSet_I _ _) (ae_of_all _ (fun t ht ↦ ?_)) rw [indexI_of_mem _ _ _ ht] rw [this] simp only [lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter] - by_cases h0 : kernel.fst κ a (I n k) = 0 + by_cases h0 : ν a (I n k) = 0 · simp only [h0, mul_zero] - rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0 - refine (measure_mono_null ?_ h0).symm + have h0' : kernel.fst κ a (I n k) = 0 := + le_antisymm ((hκν a _ (measurableSet_I _ _)).trans h0.le) zero_le' + rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0' + refine (measure_mono_null ?_ h0').symm intro p simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h rw [div_eq_mul_inv, mul_assoc, ENNReal.inv_mul_cancel h0, mul_one] exact measure_ne_top _ _ -lemma integral_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : - ∫ t, M κ a s n t ∂(kernel.fst κ a) = (κ a (univ ×ˢ s)).toReal := by + ∫ t, M κ ν a s n t ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by rw [← integral_univ, ← iUnion_I n, iUnion_prod_const, measure_iUnion] rotate_left · intro i j hij @@ -386,16 +401,16 @@ lemma integral_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] exact Or.inl (pairwise_disjoint_I n hij) · exact fun k ↦ (measurableSet_I n k).prod hs rw [integral_iUnion (measurableSet_I n) (pairwise_disjoint_I n) - (integrable_m κ a hs n).integrableOn] + (integrable_m hκν a hs n).integrableOn] rw [ENNReal.tsum_toReal_eq (fun _ ↦ measure_ne_top _ _)] congr with k - rw [set_integral_m_I _ _ hs] + rw [set_integral_m_I hκν _ hs] -lemma set_integral_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : - ∫ t in A, M κ a s n t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal := by + ∫ t in A, M κ ν a s n t ∂(ν a) = (κ a (A ×ˢ s)).toReal := by refine MeasurableSpace.induction_on_inter (m := ℱ n) (s := {s | ∃ k, s = I n k}) - (C := fun A ↦ ∫ t in A, M κ a s n t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal) rfl + (C := fun A ↦ ∫ t in A, M κ ν a s n t ∂(ν a) = (κ a (A ×ˢ s)).toReal) rfl ?_ ?_ ?_ ?_ ?_ hA · rintro s ⟨i, rfl⟩ t ⟨j, rfl⟩ hst refine ⟨i, ?_⟩ @@ -407,11 +422,11 @@ lemma set_integral_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] rwa [Function.onFun, disjoint_iff_inter_eq_empty] at h_disj · simp · rintro _ ⟨k, rfl⟩ - rw [set_integral_m_I _ _ hs] + rw [set_integral_m_I hκν _ hs] · intro A hA hA_eq have hA' : MeasurableSet A := ℱ.le _ _ hA - have h := integral_add_compl hA' (integrable_m κ a hs n) - rw [hA_eq, integral_m κ a hs] at h + have h := integral_add_compl hA' (integrable_m hκν a hs n) + rw [hA_eq, integral_m hκν a hs] at h have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by rw [prod_diff_prod, compl_eq_univ_diff] simp @@ -422,7 +437,7 @@ lemma set_integral_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] exact measure_mono (by intro x; simp) · intro f hf_disj hf h_eq - rw [integral_iUnion (fun i ↦ ℱ.le n _ (hf i)) hf_disj (integrable_m κ _ hs _).integrableOn] + rw [integral_iUnion (fun i ↦ ℱ.le n _ (hf i)) hf_disj (integrable_m hκν _ hs _).integrableOn] simp_rw [h_eq] rw [iUnion_prod_const, measure_iUnion _ (fun i ↦ (ℱ.le n _ (hf i)).prod hs)] · rw [ENNReal.tsum_toReal_eq] @@ -431,156 +446,124 @@ lemma set_integral_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] rw [Function.onFun, Set.disjoint_prod] exact Or.inl (hf_disj hij) -lemma set_integral_m_of_le (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma set_integral_m_of_le (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {n m : ℕ} (hnm : n ≤ m) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : - ∫ t in A, M κ a s m t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal := - set_integral_m κ a hs m (ℱ.mono hnm A hA) + ∫ t in A, M κ ν a s m t ∂(ν a) = (κ a (A ×ˢ s)).toReal := + set_integral_m hκν a hs m (ℱ.mono hnm A hA) -lemma condexp_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} +lemma condexp_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) {i j : ℕ} (hij : i ≤ j) : - (kernel.fst κ a)[M κ a s j | ℱ i] =ᵐ[kernel.fst κ a] M κ a s i := by + (ν a)[M κ ν a s j | ℱ i] =ᵐ[ν a] M κ ν a s i := by symm refine ae_eq_condexp_of_forall_set_integral_eq ?_ ?_ ?_ ?_ ?_ - · exact integrable_m κ a hs j + · exact integrable_m hκν a hs j · refine fun t _ _ ↦ Integrable.integrableOn ?_ - exact integrable_m κ _ hs _ + exact integrable_m hκν _ hs _ · intro t ht _ - rw [set_integral_m κ a hs i ht, set_integral_m_of_le κ a hs hij ht] - · exact StronglyMeasurable.aeStronglyMeasurable' (stronglyMeasurable_ℱ_m κ a s i) + rw [set_integral_m hκν a hs i ht, set_integral_m_of_le hκν a hs hij ht] + · exact StronglyMeasurable.aeStronglyMeasurable' (stronglyMeasurable_ℱ_m κ ν a s i) -lemma martingale_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) {s : Set β} - (hs : MeasurableSet s) : - Martingale (M κ a s) ℱ (kernel.fst κ a) := - ⟨adapted_m κ a s, fun _ _ ↦ condexp_m κ a hs⟩ - -lemma m_mono_set (κ : kernel α (ℝ × β)) (a : α) {s s' : Set β} (h : s ⊆ s') (n : ℕ) (t : ℝ) : - M κ a s n t ≤ M κ a s' n t := by - have h_ne_top : ∀ s, κ a (I n (indexI n t) ×ˢ s) / kernel.fst κ a (I n (indexI n t)) ≠ ⊤ := by +lemma martingale_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) : + Martingale (M κ ν a s) ℱ (ν a) := + ⟨adapted_m κ ν a s, fun _ _ ↦ condexp_m hκν a hs⟩ + +lemma m_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) {s s' : Set β} (h : s ⊆ s') + (n : ℕ) (t : ℝ) : + M κ ν a s n t ≤ M κ ν a s' n t := by + unfold M + by_cases h0 : ν a (I n (indexI n t)) = 0 + · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] + simp + have h_ne_top : ∀ s, κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t)) ≠ ⊤ := by intro s rw [ne_eq, ENNReal.div_eq_top] - push_neg - simp_rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] - constructor - · refine fun h h0 ↦ h (measure_mono_null (fun x ↦ ?_) h0) - simp only [mem_prod, mem_setOf_eq, and_imp] - exact fun h _ ↦ h - · refine fun h_top ↦ eq_top_mono (measure_mono (fun x ↦ ?_)) h_top - simp only [mem_prod, mem_setOf_eq, and_imp] - exact fun h _ ↦ h - rw [M, M, ENNReal.toReal_le_toReal] + simp only [ne_eq, h0, and_false, false_or, not_and, not_not] + refine fun h_top ↦ eq_top_mono ?_ h_top + calc κ a (I n (indexI n t) ×ˢ s) ≤ kernel.fst κ a (I n (indexI n t)) := by + rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] + refine measure_mono (fun x ↦ ?_) + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + _ ≤ ν a (I n (indexI n t)) := hκν a _ (measurableSet_I _ _) + rw [ENNReal.toReal_le_toReal] · gcongr rw [prod_subset_prod_iff] simp [subset_rfl, h] · exact h_ne_top s · exact h_ne_top s' -lemma m_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ) (t : ℝ) : - M κ a univ n t = if kernel.fst κ a (I n (indexI n t)) = 0 then 0 else 1 := by - rw [M] - by_cases h : kernel.fst κ a (I n (indexI n t)) = 0 - · simp [h] - by_cases h' : κ a (I n (indexI n t) ×ˢ univ) = 0 - · simp [h'] - · rw [ENNReal.div_zero h'] - simp - · simp only [h, ite_false] - rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] - have : I n (indexI n t) ×ˢ univ = {p : ℝ × β | p.1 ∈ I n (indexI n t)} := by - ext x - simp - rw [this, ENNReal.div_self] - · simp - · rwa [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h - · exact measure_ne_top _ _ - -lemma m_empty (κ : kernel α (ℝ × β)) (a : α) (n : ℕ) (t : ℝ) : - M κ a ∅ n t = 0 := by +lemma m_mono_kernel_left {κ' : kernel α (ℝ × β)} (hκκ' : κ ≤ κ') (hκ'ν : kernel.fst κ' ≤ ν) + (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) (t : ℝ) : + M κ ν a s n t ≤ M κ' ν a s n t := by + unfold M + by_cases h0 : ν a (I n (indexI n t)) = 0 + · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] + simp + have h_le : κ' a (I n (indexI n t) ×ˢ s) ≤ ν a (I n (indexI n t)) := by + calc κ' a (I n (indexI n t) ×ˢ s) ≤ kernel.fst κ' a (I n (indexI n t)) := by + rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] + refine measure_mono (fun x ↦ ?_) + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + _ ≤ ν a (I n (indexI n t)) := hκ'ν a _ (measurableSet_I _ _) + rw [ENNReal.toReal_le_toReal] + · gcongr + exact hκκ' _ _ ((measurableSet_I _ _).prod hs) + · rw [ne_eq, ENNReal.div_eq_top] + simp only [ne_eq, h0, and_false, false_or, not_and, not_not] + refine fun h_top ↦ eq_top_mono ?_ h_top + exact (hκκ' _ _ ((measurableSet_I _ _).prod hs)).trans h_le + · rw [ne_eq, ENNReal.div_eq_top] + simp only [ne_eq, h0, and_false, false_or, not_and, not_not] + exact fun h_top ↦ eq_top_mono h_le h_top + +lemma m_antitone_kernel_right {ν' : kernel α ℝ} (hνν' : ν ≤ ν') (hκν : kernel.fst κ ≤ ν) + (a : α) (s : Set β) (n : ℕ) (t : ℝ) : + M κ ν' a s n t ≤ M κ ν a s n t := by + unfold M + have h_le : κ a (I n (indexI n t) ×ˢ s) ≤ ν a (I n (indexI n t)) := by + calc κ a (I n (indexI n t) ×ˢ s) ≤ kernel.fst κ a (I n (indexI n t)) := by + rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] + refine measure_mono (fun x ↦ ?_) + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + _ ≤ ν a (I n (indexI n t)) := hκν a _ (measurableSet_I _ _) + by_cases h0 : ν a (I n (indexI n t)) = 0 + · suffices κ a (I n (indexI n t) ×ˢ s) = 0 by + simp only [this, ENNReal.zero_div, ENNReal.zero_toReal, h0, le_refl] + exact le_antisymm (h_le.trans h0.le) zero_le' + have h0' : ν' a (I n (indexI n t)) ≠ 0 := by + refine fun h ↦ h0 (le_antisymm (le_trans ?_ h.le) zero_le') + exact hνν' _ _ (measurableSet_I _ _) + rw [ENNReal.toReal_le_toReal] + · gcongr + exact hνν' _ _ (measurableSet_I _ _) + · simp only [ne_eq, ENNReal.div_eq_top, h0', and_false, false_or, not_and, not_not] + refine fun h_top ↦ eq_top_mono ?_ h_top + exact h_le.trans (hνν' _ _ (measurableSet_I _ _)) + · simp only [ne_eq, ENNReal.div_eq_top, h0, and_false, false_or, not_and, not_not] + exact fun h_top ↦ eq_top_mono h_le h_top + +lemma m_empty (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (n : ℕ) (t : ℝ) : + M κ ν a ∅ n t = 0 := by rw [M] simp -lemma m_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ) : - ∀ᵐ t ∂(kernel.fst κ a), M κ a univ n t = 1 := by - rw [ae_iff] - have : {t | ¬ M κ a univ n t = 1} ⊆ {t | kernel.fst κ a (I n (indexI n t)) = 0} := by - intro t ht - simp only [mem_setOf_eq] at ht ⊢ - rw [m_univ] at ht - simpa using ht - refine measure_mono_null this ?_ - have : {t | kernel.fst κ a (I n (indexI n t)) = 0} - ⊆ ⋃ (k) (_ : kernel.fst κ a (I n k) = 0), I n k := by - intro t - simp only [mem_setOf_eq, mem_iUnion, exists_prop] - intro ht - exact ⟨indexI n t, ht, mem_I_indexI _ _⟩ - refine measure_mono_null this ?_ - rw [measure_iUnion_null] - intro i - simp - -lemma tendsto_m_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) - (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) (t : ℝ) : - Tendsto (fun m ↦ M κ a (s m) n t) atTop (𝓝 (M κ a univ n t)) := by - simp_rw [M] - refine (ENNReal.tendsto_toReal ?_).comp ?_ - · rw [ne_eq, ENNReal.div_eq_top] - push_neg - simp_rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] - constructor - · refine fun h h0 ↦ h (measure_mono_null (fun x ↦ ?_) h0) - simp only [mem_prod, mem_setOf_eq, and_imp] - exact fun h _ ↦ h - · refine fun h_top ↦ eq_top_mono (measure_mono (fun x ↦ ?_)) h_top - simp only [mem_prod, mem_setOf_eq, and_imp] - exact fun h _ ↦ h - by_cases h0 : kernel.fst κ a (I n (indexI n t)) = 0 - · rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0 ⊢ - suffices ∀ m, κ a (I n (indexI n t) ×ˢ s m) = 0 by - simp only [this, h0, ENNReal.zero_div, tendsto_const_nhds_iff] - suffices κ a (I n (indexI n t) ×ˢ univ) = 0 by simp only [this, ENNReal.zero_div] - convert h0 - ext x - simp only [mem_prod, mem_univ, and_true, mem_setOf_eq] - refine fun m ↦ measure_mono_null (fun x ↦ ?_) h0 - simp only [mem_prod, mem_setOf_eq, and_imp] - exact fun h _ ↦ h - refine ENNReal.Tendsto.div_const ?_ ?_ - · have h := tendsto_measure_iUnion (μ := κ a) (s := fun m ↦ I n (indexI n t) ×ˢ s m) ?_ - swap - · intro m m' hmm' - simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] - exact Or.inl <| hs hmm' - convert h - rw [← prod_iUnion, hs_iUnion] - · exact Or.inr h0 - -lemma tendsto_m_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] - (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) : - ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ M κ a (s m) n t) atTop (𝓝 1) := by - filter_upwards [m_univ_ae κ a n] with t ht - rw [← ht] - exact tendsto_m_atTop_univ_of_monotone κ a s hs hs_iUnion n t - -lemma tendsto_m_atTop_empty_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma tendsto_m_atTop_empty_of_antitone (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) (n : ℕ) (t : ℝ) : - Tendsto (fun m ↦ M κ a (s m) n t) atTop (𝓝 (M κ a ∅ n t)) := by + Tendsto (fun m ↦ M κ ν a (s m) n t) atTop (𝓝 (M κ ν a ∅ n t)) := by simp_rw [M] + by_cases h0 : ν a (I n (indexI n t)) = 0 + · simp_rw [h0, ENNReal.toReal_div] + simp refine (ENNReal.tendsto_toReal ?_).comp ?_ · rw [ne_eq, ENNReal.div_eq_top] push_neg simp - by_cases h0 : kernel.fst κ a (I n (indexI n t)) = 0 - · simp only [h0, prod_empty, OuterMeasure.empty', ENNReal.zero_div] - have : ∀ m, (κ a) (I n (indexI n t) ×ˢ s m) = 0 := by - intro m - rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0 - refine measure_mono_null ?_ h0 - intro x hx - rw [mem_prod] at hx - exact hx.1 - simp [this] refine ENNReal.Tendsto.div_const ?_ ?_ · have h := tendsto_measure_iInter (μ := κ a) (s := fun m ↦ I n (indexI n t) ×ˢ s m) ?_ ?_ ?_ · convert h @@ -593,187 +576,180 @@ lemma tendsto_m_atTop_empty_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKe · simp only [prod_empty, OuterMeasure.empty', ne_eq, not_true_eq_false, false_or, h0, not_false_iff] -lemma tendsto_m_atTop_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma tendsto_m_atTop_of_antitone (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) (n : ℕ) (t : ℝ) : - Tendsto (fun m ↦ M κ a (s m) n t) atTop (𝓝 0) := by - rw [← m_empty κ a n t] - exact tendsto_m_atTop_empty_of_antitone κ a s hs hs_iInter hs_meas n t + Tendsto (fun m ↦ M κ ν a (s m) n t) atTop (𝓝 0) := by + rw [← m_empty κ ν a n t] + exact tendsto_m_atTop_empty_of_antitone κ ν a s hs hs_iInter hs_meas n t -lemma tendsto_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsFiniteKernel κ] - {s : Set β} (hs : MeasurableSet s) : - ∀ᵐ t ∂(kernel.fst κ a), - Tendsto (fun n ↦ M κ a s n t) atTop (𝓝 (ℱ.limitProcess (M κ a s) (kernel.fst κ a) t)) := by - refine Submartingale.ae_tendsto_limitProcess (martingale_m κ a hs).submartingale - (R := (kernel.fst κ a univ).toNNReal) (fun n ↦ ?_) - refine (snorm_m_le κ a s n).trans_eq ?_ +lemma tendsto_m_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) : + ∀ᵐ t ∂(ν a), + Tendsto (fun n ↦ M κ ν a s n t) atTop (𝓝 (ℱ.limitProcess (M κ ν a s) (ν a) t)) := by + refine Submartingale.ae_tendsto_limitProcess (martingale_m hκν a hs).submartingale + (R := (ν a univ).toNNReal) (fun n ↦ ?_) + refine (snorm_m_le hκν a s n).trans_eq ?_ rw [ENNReal.coe_toNNReal] exact measure_ne_top _ _ -lemma limitProcess_mem_L1 (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma limitProcess_mem_L1 (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Memℒp (ℱ.limitProcess (M κ a s) (kernel.fst κ a)) 1 (kernel.fst κ a) := by - refine Submartingale.memℒp_limitProcess (martingale_m κ a hs).submartingale - (R := (kernel.fst κ a univ).toNNReal) (fun n ↦ ?_) - refine (snorm_m_le κ a s n).trans_eq ?_ + Memℒp (ℱ.limitProcess (M κ ν a s) (ν a)) 1 (ν a) := by + refine Submartingale.memℒp_limitProcess (martingale_m hκν a hs).submartingale + (R := (ν a univ).toNNReal) (fun n ↦ ?_) + refine (snorm_m_le hκν a s n).trans_eq ?_ rw [ENNReal.coe_toNNReal] exact measure_ne_top _ _ -lemma tendsto_snorm_one_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsFiniteKernel κ] - {s : Set β} (hs : MeasurableSet s) : +lemma tendsto_snorm_one_m_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : Tendsto - (fun n ↦ snorm (M κ a s n - ℱ.limitProcess (M κ a s) (kernel.fst κ a)) 1 (kernel.fst κ a)) + (fun n ↦ snorm (M κ ν a s n - ℱ.limitProcess (M κ ν a s) (ν a)) 1 (ν a)) atTop (𝓝 0) := by refine Submartingale.tendsto_snorm_one_limitProcess ?_ ?_ - · exact (martingale_m κ a hs).submartingale + · exact (martingale_m hκν a hs).submartingale · refine uniformIntegrable_of le_rfl ENNReal.one_ne_top ?_ ?_ - · exact fun n ↦ (measurable_m_right κ a hs n).aestronglyMeasurable + · exact fun n ↦ (measurable_m_right κ ν a hs n).aestronglyMeasurable · intro ε _ refine ⟨2, fun n ↦ ?_⟩ refine le_of_eq_of_le ?_ (?_ : 0 ≤ ENNReal.ofReal ε) - · have : {x | 2 ≤ ‖M κ a s n x‖₊} = ∅ := by + · have : {x | 2 ≤ ‖M κ ν a s n x‖₊} = ∅ := by ext x simp only [mem_setOf_eq, mem_empty_iff_false, iff_false, not_le] refine (?_ : _ ≤ (1 : ℝ≥0)).trans_lt one_lt_two - rw [Real.nnnorm_of_nonneg (m_nonneg _ _ _ _ _)] - exact mod_cast (m_le_one _ _ _ _ _) + rw [Real.nnnorm_of_nonneg (m_nonneg _ _ _ _ _ _)] + exact mod_cast (m_le_one hκν _ _ _ _) rw [this] simp · simp -lemma tendsto_snorm_one_restrict_m_limitProcess (κ : kernel α (ℝ × β)) (a : α) [IsFiniteKernel κ] - {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : - Tendsto (fun n ↦ snorm (M κ a s n - ℱ.limitProcess (M κ a s) (kernel.fst κ a)) 1 - ((kernel.fst κ a).restrict A)) +lemma tendsto_snorm_one_restrict_m_limitProcess [IsFiniteKernel κ] [IsFiniteKernel ν] + (hκν : kernel.fst κ ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : + Tendsto (fun n ↦ snorm (M κ ν a s n - ℱ.limitProcess (M κ ν a s) (ν a)) 1 ((ν a).restrict A)) atTop (𝓝 0) := - tendsto_snorm_restrict_zero (tendsto_snorm_one_m_limitProcess κ a hs) A + tendsto_snorm_restrict_zero (tendsto_snorm_one_m_limitProcess hκν a hs) A noncomputable -def MLimsup (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (t : ℝ) : ℝ := - limsup (fun n ↦ M κ a s n t) atTop +def MLimsup (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (t : ℝ) : ℝ := + limsup (fun n ↦ M κ ν a s n t) atTop -lemma mLimsup_ae_eq_limitProcess (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma mLimsup_ae_eq_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - MLimsup κ a s =ᵐ[kernel.fst κ a] ℱ.limitProcess (M κ a s) (kernel.fst κ a) := by - filter_upwards [tendsto_m_limitProcess κ a hs] with t ht using ht.limsup_eq + MLimsup κ ν a s =ᵐ[ν a] ℱ.limitProcess (M κ ν a s) (ν a) := by + filter_upwards [tendsto_m_limitProcess hκν a hs] with t ht using ht.limsup_eq -lemma tendsto_m_mLimsup (κ : kernel α (ℝ × β)) (a : α) [IsFiniteKernel κ] +lemma tendsto_m_mLimsup (hκν : kernel.fst κ ≤ ν) (a : α) [IsFiniteKernel κ] [IsFiniteKernel ν] {s : Set β} (hs : MeasurableSet s) : - ∀ᵐ t ∂(kernel.fst κ a), - Tendsto (fun n ↦ M κ a s n t) atTop (𝓝 (MLimsup κ a s t)) := by - filter_upwards [tendsto_m_limitProcess κ a hs, mLimsup_ae_eq_limitProcess κ a hs] with t h1 h2 + ∀ᵐ t ∂(ν a), + Tendsto (fun n ↦ M κ ν a s n t) atTop (𝓝 (MLimsup κ ν a s t)) := by + filter_upwards [tendsto_m_limitProcess hκν a hs, mLimsup_ae_eq_limitProcess hκν a hs] with t h1 h2 rw [h2] exact h1 -lemma measurable_mLimsup (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) : - Measurable (fun (p : α × ℝ) ↦ MLimsup κ p.1 s p.2) := - measurable_limsup (fun n ↦ measurable_m κ hs n) - -lemma measurable_mLimsup_left (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) (t : ℝ) : - Measurable (fun a ↦ MLimsup κ a s t) := by - change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ p.1 s p.2) ∘ (fun a ↦ (a, t))) - exact (measurable_mLimsup κ hs).comp (measurable_id.prod_mk measurable_const) - -lemma measurable_mLimsup_right (κ : kernel α (ℝ × β)) {s : Set β} (hs : MeasurableSet s) (a : α) : - Measurable (MLimsup κ a s) := by - change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ p.1 s p.2) ∘ (fun t ↦ (a, t))) - exact (measurable_mLimsup κ hs).comp (measurable_const.prod_mk measurable_id) - -lemma mLimsup_mono_set (κ : kernel α (ℝ × β)) (a : α) {s s' : Set β} (h : s ⊆ s') (t : ℝ) : - MLimsup κ a s t ≤ MLimsup κ a s' t := by +lemma measurable_mLimsup (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) + {s : Set β} (hs : MeasurableSet s) : + Measurable (fun (p : α × ℝ) ↦ MLimsup κ ν p.1 s p.2) := + measurable_limsup (fun n ↦ measurable_m κ ν hs n) + +lemma measurable_mLimsup_left (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) + {s : Set β} (hs : MeasurableSet s) (t : ℝ) : + Measurable (fun a ↦ MLimsup κ ν a s t) := by + change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ ν p.1 s p.2) ∘ (fun a ↦ (a, t))) + exact (measurable_mLimsup κ ν hs).comp (measurable_id.prod_mk measurable_const) + +lemma measurable_mLimsup_right (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) + {s : Set β} (hs : MeasurableSet s) (a : α) : + Measurable (MLimsup κ ν a s) := by + change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ ν p.1 s p.2) ∘ (fun t ↦ (a, t))) + exact (measurable_mLimsup κ ν hs).comp (measurable_const.prod_mk measurable_id) + +lemma mLimsup_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) {s s' : Set β} (h : s ⊆ s') (t : ℝ) : + MLimsup κ ν a s t ≤ MLimsup κ ν a s' t := by rw [MLimsup, MLimsup] refine limsup_le_limsup ?_ ?_ ?_ - · exact eventually_of_forall (fun n ↦ m_mono_set κ a h n t) - · exact isCoboundedUnder_le_of_le atTop (fun i ↦ m_nonneg _ _ _ _ _) - · exact isBoundedUnder_of ⟨1, fun n ↦ m_le_one _ _ _ _ _⟩ + · exact eventually_of_forall (fun n ↦ m_mono_set hκν a h n t) + · exact isCoboundedUnder_le_of_le atTop (fun i ↦ m_nonneg _ _ _ _ _ _) + · exact isBoundedUnder_of ⟨1, fun n ↦ m_le_one hκν _ _ _ _⟩ -lemma mLimsup_nonneg (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (t : ℝ) : - 0 ≤ MLimsup κ a s t := by +lemma mLimsup_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (t : ℝ) : + 0 ≤ MLimsup κ ν a s t := by refine le_limsup_of_frequently_le ?_ ?_ - · exact frequently_of_forall (fun n ↦ m_nonneg _ _ _ _ _) - · exact isBoundedUnder_of ⟨1, fun n ↦ m_le_one _ _ _ _ _⟩ + · exact frequently_of_forall (fun n ↦ m_nonneg _ _ _ _ _ _) + · exact isBoundedUnder_of ⟨1, fun n ↦ m_le_one hκν _ _ _ _⟩ -lemma mLimsup_le_one (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (t : ℝ) : - MLimsup κ a s t ≤ 1 := by +lemma mLimsup_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (t : ℝ) : + MLimsup κ ν a s t ≤ 1 := by refine limsup_le_of_le ?_ ?_ - · exact isCoboundedUnder_le_of_le atTop (fun i ↦ m_nonneg _ _ _ _ _) - · exact eventually_of_forall (fun n ↦ m_le_one _ _ _ _ _) + · exact isCoboundedUnder_le_of_le atTop (fun i ↦ m_nonneg _ _ _ _ _ _) + · exact eventually_of_forall (fun n ↦ m_le_one hκν _ _ _ _) -lemma mLimsup_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), MLimsup κ a Set.univ t = 1 := by - have h := m_univ_ae κ a - rw [← ae_all_iff] at h - filter_upwards [h] with t ht - rw [MLimsup] - simp [ht] - -lemma snorm_mLimsup_le (κ : kernel α (ℝ × β)) - (a : α) (s : Set β) : - snorm (fun t ↦ MLimsup κ a s t) 1 (kernel.fst κ a) ≤ kernel.fst κ a univ := by +lemma snorm_mLimsup_le (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) : + snorm (fun t ↦ MLimsup κ ν a s t) 1 (ν a) ≤ ν a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ - · simp only [Real.norm_eq_abs, abs_of_nonneg (mLimsup_nonneg κ a s t), - mLimsup_le_one κ a s t] + · simp only [Real.norm_eq_abs, abs_of_nonneg (mLimsup_nonneg hκν a s t), + mLimsup_le_one hκν a s t] · simp -lemma integrable_mLimsup (κ : kernel α (ℝ × β)) [IsFiniteKernel (kernel.fst κ)] +lemma integrable_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Integrable (fun t ↦ MLimsup κ a s t) (kernel.fst κ a) := by + Integrable (fun t ↦ MLimsup κ ν a s t) (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ - · exact measurable_mLimsup_right κ hs a - · exact (snorm_mLimsup_le κ a s).trans_lt (measure_lt_top _ _) + · exact measurable_mLimsup_right κ ν hs a + · exact (snorm_mLimsup_le hκν a s).trans_lt (measure_lt_top _ _) -lemma tendsto_set_integral_m (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma tendsto_set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : - Tendsto (fun i ↦ ∫ x in A, M κ a s i x ∂(kernel.fst κ) a) atTop - (𝓝 (∫ x in A, MLimsup κ a s x ∂(kernel.fst κ) a)) := by - refine tendsto_set_integral_of_L1' (μ := kernel.fst κ a) (MLimsup κ a s) - (integrable_mLimsup κ a hs) (F := fun i t ↦ M κ a s i t) (l := atTop) - (eventually_of_forall (fun n ↦ integrable_m _ _ hs _)) ?_ A - refine (tendsto_congr fun n ↦ ?_).mp (tendsto_snorm_one_m_limitProcess κ a hs) + Tendsto (fun i ↦ ∫ x in A, M κ ν a s i x ∂(ν a)) atTop + (𝓝 (∫ x in A, MLimsup κ ν a s x ∂(ν a))) := by + refine tendsto_set_integral_of_L1' (μ := ν a) (MLimsup κ ν a s) + (integrable_mLimsup hκν a hs) (F := fun i t ↦ M κ ν a s i t) (l := atTop) + (eventually_of_forall (fun n ↦ integrable_m hκν _ hs _)) ?_ A + refine (tendsto_congr fun n ↦ ?_).mp (tendsto_snorm_one_m_limitProcess hκν a hs) refine snorm_congr_ae ?_ refine EventuallyEq.sub EventuallyEq.rfl ?_ - exact (mLimsup_ae_eq_limitProcess κ a hs).symm + exact (mLimsup_ae_eq_limitProcess hκν a hs).symm -lemma set_integral_mLimsup_of_measurableSet (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] - (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) - {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : - ∫ t in A, MLimsup κ a s t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal := by - suffices ∫ t in A, MLimsup κ a s t ∂(kernel.fst κ a) = ∫ t in A, M κ a s n t ∂(kernel.fst κ a) by +lemma set_integral_mLimsup_of_measurableSet (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : + ∫ t in A, MLimsup κ ν a s t ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + suffices ∫ t in A, MLimsup κ ν a s t ∂(ν a) = ∫ t in A, M κ ν a s n t ∂(ν a) by rw [this] - exact set_integral_m _ _ hs _ hA - suffices ∫ t in A, MLimsup κ a s t ∂(kernel.fst κ a) - = limsup (fun i ↦ ∫ t in A, M κ a s i t ∂(kernel.fst κ a)) atTop by - rw [this, ← limsup_const (α := ℕ) (f := atTop) (∫ t in A, M κ a s n t ∂(kernel.fst κ a)), + exact set_integral_m hκν _ hs _ hA + suffices ∫ t in A, MLimsup κ ν a s t ∂(ν a) + = limsup (fun i ↦ ∫ t in A, M κ ν a s i t ∂(ν a)) atTop by + rw [this, ← limsup_const (α := ℕ) (f := atTop) (∫ t in A, M κ ν a s n t ∂(ν a)), limsup_congr] simp only [eventually_atTop, ge_iff_le] refine ⟨n, fun m hnm ↦ ?_⟩ - rw [set_integral_m_of_le _ _ hs hnm hA, set_integral_m _ _ hs _ hA] + rw [set_integral_m_of_le hκν _ hs hnm hA, set_integral_m hκν _ hs _ hA] -- use L1 convergence - have h := tendsto_set_integral_m κ a hs A + have h := tendsto_set_integral_m hκν a hs A rw [h.limsup_eq] -lemma integral_mLimsup (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma integral_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - ∫ t, MLimsup κ a s t ∂(kernel.fst κ a) = (κ a (univ ×ˢ s)).toReal := by - rw [← integral_univ, set_integral_mLimsup_of_measurableSet κ a hs 0 MeasurableSet.univ] + ∫ t, MLimsup κ ν a s t ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by + rw [← integral_univ, set_integral_mLimsup_of_measurableSet hκν a hs 0 MeasurableSet.univ] -lemma set_integral_mLimsup (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma set_integral_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set ℝ} (hA : MeasurableSet A) : - ∫ t in A, MLimsup κ a s t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal := by + ∫ t in A, MLimsup κ ν a s t ∂(ν a) = (κ a (A ×ˢ s)).toReal := by have hA' : MeasurableSet[⨆ n, ℱ n] A := by rwa [iSup_ℱ] refine MeasurableSpace.induction_on_inter (m := ⨆ n, ℱ n) - (C := fun A ↦ ∫ t in A, MLimsup κ a s t ∂(kernel.fst κ a) = (κ a (A ×ˢ s)).toReal) + (C := fun A ↦ ∫ t in A, MLimsup κ ν a s t ∂(ν a) = (κ a (A ×ˢ s)).toReal) (MeasurableSpace.measurableSpace_iSup_eq ℱ) ?_ ?_ ?_ ?_ ?_ hA' · rintro s ⟨n, hs⟩ t ⟨m, ht⟩ _ exact ⟨max n m, (ℱ.mono (le_max_left n m) _ hs).inter (ℱ.mono (le_max_right n m) _ ht)⟩ · simp · intro A ⟨n, hA⟩ - exact set_integral_mLimsup_of_measurableSet κ a hs n hA + exact set_integral_mLimsup_of_measurableSet hκν a hs n hA · intro A hA hA_eq rw [iSup_ℱ] at hA - have h := integral_add_compl hA (integrable_mLimsup κ a hs) - rw [hA_eq, integral_mLimsup κ a hs] at h + have h := integral_add_compl hA (integrable_mLimsup hκν a hs) + rw [hA_eq, integral_mLimsup hκν a hs] at h have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by rw [prod_diff_prod, compl_eq_univ_diff] simp @@ -784,7 +760,7 @@ lemma set_integral_mLimsup (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] exact measure_mono (by intro x; simp) · intro f hf_disj hf h_eq - rw [integral_iUnion _ hf_disj (integrable_mLimsup _ _ hs).integrableOn] + rw [integral_iUnion _ hf_disj (integrable_mLimsup hκν _ hs).integrableOn] · simp_rw [h_eq] rw [← ENNReal.tsum_toReal_eq (fun _ ↦ measure_ne_top _ _)] congr @@ -796,11 +772,12 @@ lemma set_integral_mLimsup (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] exact fun i ↦ (hf i).prod hs · rwa [iSup_ℱ] at hf -lemma tendsto_integral_mLimsup_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma tendsto_integral_mLimsup_of_monotone (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : - Tendsto (fun m ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop (𝓝 (κ a univ).toReal) := by - simp_rw [integral_mLimsup κ a (hs_meas _)] + Tendsto (fun m ↦ ∫ t, MLimsup κ ν a (s m) t ∂(ν a)) atTop (𝓝 (κ a univ).toReal) := by + simp_rw [integral_mLimsup hκν a (hs_meas _)] have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := κ a univ) ?_ swap · rw [mem_nhds_iff] @@ -815,11 +792,12 @@ lemma tendsto_integral_mLimsup_of_monotone (κ : kernel α (ℝ × β)) [IsFinit rw [← prod_iUnion, hs_iUnion] simp only [univ_prod_univ, measure_univ] -lemma tendsto_integral_mLimsup_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma tendsto_integral_mLimsup_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : - Tendsto (fun m ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop (𝓝 0) := by - simp_rw [integral_mLimsup κ a (hs_meas _)] + Tendsto (fun m ↦ ∫ t, MLimsup κ ν a (s m) t ∂(ν a)) atTop (𝓝 0) := by + simp_rw [integral_mLimsup hκν a (hs_meas _)] rw [← ENNReal.zero_toReal] have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := 0) ?_ swap @@ -837,166 +815,274 @@ lemma tendsto_integral_mLimsup_of_antitone (κ : kernel α (ℝ × β)) [IsFinit rw [← prod_iInter, hs_iInter] simp only [ne_eq, prod_empty, OuterMeasure.empty', forall_exists_index] -lemma tendsto_mLimsup_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] - (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) +lemma tendsto_mLimsup_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] + (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : - ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 1) := by - have h_mono : ∀ t, Monotone (fun m ↦ MLimsup κ a (s m) t) := - fun t n m hnm ↦ mLimsup_mono_set κ a (hs hnm) t - have h_le_one : ∀ m t, MLimsup κ a (s m) t ≤ 1 := fun m t ↦ mLimsup_le_one κ a (s m) t + ∀ᵐ t ∂(ν a), Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 0) := by + have h_anti : ∀ t, Antitone (fun m ↦ MLimsup κ ν a (s m) t) := + fun t n m hnm ↦ mLimsup_mono_set hκν a (hs hnm) t + have h_le_one : ∀ m t, MLimsup κ ν a (s m) t ≤ 1 := fun m t ↦ mLimsup_le_one hκν a (s m) t -- for all `t`, `fun m ↦ MLimsup κ a (s m) t` has a limit - have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 l) := by + have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 l) := by intro t - have h_tendsto : Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop atTop ∨ - ∃ l, Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 l) := - tendsto_of_monotone (h_mono t) + have h_tendsto : Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop atBot ∨ + ∃ l, Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 l) := + tendsto_of_antitone (h_anti t) cases' h_tendsto with h_absurd h_tendsto - · rw [tendsto_atTop_atTop_iff_of_monotone (h_mono t)] at h_absurd - obtain ⟨r, hr⟩ := h_absurd 2 - exact absurd (hr.trans (h_le_one r t)) one_lt_two.not_le + · rw [tendsto_atTop_atBot_iff_of_antitone (h_anti t)] at h_absurd + obtain ⟨r, hr⟩ := h_absurd (-1) + have h_nonneg := mLimsup_nonneg hκν a (s r) t + linarith · exact h_tendsto -- let `F` be the pointwise limit of `fun m ↦ MLimsup κ a (s m) t` for all `t` let F : ℝ → ℝ := fun t ↦ (h_exists t).choose - have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 (F t)) := + have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 (F t)) := fun t ↦ (h_exists t).choose_spec have hF_nonneg : ∀ t, 0 ≤ F t := - fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ mLimsup_nonneg κ a (s m) t) + fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ mLimsup_nonneg hκν a (s m) t) have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) - have hF_int : Integrable F (kernel.fst κ a) := by + have hF_int : Integrable F (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨?_, ?_⟩ · refine aestronglyMeasurable_of_tendsto_ae atTop (fun n ↦ ?_) (ae_of_all _ hF_tendsto) - exact (measurable_mLimsup_right κ (hs_meas _) a).aestronglyMeasurable + exact (measurable_mLimsup_right κ ν (hs_meas _) a).aestronglyMeasurable · rw [snorm_one_eq_lintegral_nnnorm] - calc ∫⁻ x, ‖F x‖₊ ∂(kernel.fst κ a) ≤ ∫⁻ _, 1 ∂(kernel.fst κ a) := by + calc ∫⁻ x, ‖F x‖₊ ∂(ν a) ≤ ∫⁻ _, 1 ∂(ν a) := by refine lintegral_mono (fun x ↦ ?_) rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, ENNReal.ofReal_le_one, abs_of_nonneg (hF_nonneg _)] exact hF_le_one _ - _ < ⊤ := by simp only [lintegral_const, measure_univ, one_mul, measure_lt_top] - -- it suffices to show that the limit `F` is 1 a.e. - suffices ∀ᵐ t ∂(kernel.fst κ a), F t = 1 by + _ < ⊤ := by + simp only [lintegral_const, one_mul] + exact measure_lt_top _ _ + -- it suffices to show that the limit `F` is 0 a.e. + suffices ∀ᵐ t ∂(ν a), F t = 0 by filter_upwards [this] with t ht_eq rw [← ht_eq] exact hF_tendsto t - -- since `F` is at most 1, proving that its integral is the same as the integral of 1 will tell - -- us that `F` is 1 a.e. - refine ae_eq_of_integral_eq_of_ae_le hF_int (integrable_const _) (ae_of_all _ hF_le_one) ?_ + -- since `F` is nonnegative, proving that its integral is 0 is sufficient to get that + -- `F` is 0 a.e. + suffices ∀ᵐ (t : ℝ) ∂(ν a), 0 = F t by filter_upwards [this] with a ha; simp [ha] + refine ae_eq_of_integral_eq_of_ae_le (integrable_const _) hF_int (ae_of_all _ hF_nonneg) ?_ have h_integral : - Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop - (𝓝 (∫ t, F t ∂(kernel.fst κ a))) := by - refine integral_tendsto_of_tendsto_of_monotone ?_ hF_int ?_ ?_ - · exact fun n ↦ integrable_mLimsup _ _ (hs_meas n) - · exact ae_of_all _ h_mono + Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a (s m) t ∂(ν a)) atTop (𝓝 (∫ t, F t ∂(ν a))) := by + refine integral_tendsto_of_tendsto_of_antitone ?_ hF_int ?_ ?_ + · exact fun n ↦ integrable_mLimsup hκν _ (hs_meas n) + · exact ae_of_all _ h_anti · exact ae_of_all _ hF_tendsto have h_integral' : - Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop - (𝓝 (∫ _, 1 ∂(kernel.fst κ a))) := by - rw [integral_const, kernel.fst_apply' _ _ MeasurableSet.univ] - simp only [smul_eq_mul, mul_one] - exact tendsto_integral_mLimsup_of_monotone κ a s hs hs_iUnion hs_meas - exact tendsto_nhds_unique h_integral h_integral' + Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a (s m) t ∂(ν a)) atTop + (𝓝 (∫ _, 0 ∂(ν a))) := by + simp only [integral_zero] + exact tendsto_integral_mLimsup_of_antitone hκν a s hs hs_iInter hs_meas + exact (tendsto_nhds_unique h_integral h_integral').symm -lemma tendsto_mLimsup_atTop_ae_of_antitone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] - (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) +section UnivFst + +lemma m_univ [IsFiniteKernel κ] (a : α) (n : ℕ) (t : ℝ) : + M κ (kernel.fst κ) a univ n t = if kernel.fst κ a (I n (indexI n t)) = 0 then 0 else 1 := by + rw [M] + by_cases h : kernel.fst κ a (I n (indexI n t)) = 0 + · simp [h] + by_cases h' : κ a (I n (indexI n t) ×ˢ univ) = 0 + · simp [h'] + · rw [ENNReal.div_zero h'] + simp + · simp only [h, ite_false] + rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] + have : I n (indexI n t) ×ˢ univ = {p : ℝ × β | p.1 ∈ I n (indexI n t)} := by + ext x + simp + rw [this, ENNReal.div_self] + · simp + · rwa [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h + · exact measure_ne_top _ _ + +lemma m_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ) : + ∀ᵐ t ∂(kernel.fst κ a), M κ (kernel.fst κ) a univ n t = 1 := by + rw [ae_iff] + have : {t | ¬ M κ (kernel.fst κ) a univ n t = 1} + ⊆ {t | kernel.fst κ a (I n (indexI n t)) = 0} := by + intro t ht + simp only [mem_setOf_eq] at ht ⊢ + rw [m_univ] at ht + simpa using ht + refine measure_mono_null this ?_ + have : {t | kernel.fst κ a (I n (indexI n t)) = 0} + ⊆ ⋃ (k) (_ : kernel.fst κ a (I n k) = 0), I n k := by + intro t + simp only [mem_setOf_eq, mem_iUnion, exists_prop] + intro ht + exact ⟨indexI n t, ht, mem_I_indexI _ _⟩ + refine measure_mono_null this ?_ + rw [measure_iUnion_null] + intro i + simp + +lemma tendsto_m_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) + (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) (t : ℝ) : + Tendsto (fun m ↦ M κ (kernel.fst κ) a (s m) n t) atTop + (𝓝 (M κ (kernel.fst κ) a univ n t)) := by + simp_rw [M] + refine (ENNReal.tendsto_toReal ?_).comp ?_ + · rw [ne_eq, ENNReal.div_eq_top] + push_neg + simp_rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] + constructor + · refine fun h h0 ↦ h (measure_mono_null (fun x ↦ ?_) h0) + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + · refine fun h_top ↦ eq_top_mono (measure_mono (fun x ↦ ?_)) h_top + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + by_cases h0 : kernel.fst κ a (I n (indexI n t)) = 0 + · rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0 ⊢ + suffices ∀ m, κ a (I n (indexI n t) ×ˢ s m) = 0 by + simp only [this, h0, ENNReal.zero_div, tendsto_const_nhds_iff] + suffices κ a (I n (indexI n t) ×ˢ univ) = 0 by simp only [this, ENNReal.zero_div] + convert h0 + ext x + simp only [mem_prod, mem_univ, and_true, mem_setOf_eq] + refine fun m ↦ measure_mono_null (fun x ↦ ?_) h0 + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + refine ENNReal.Tendsto.div_const ?_ ?_ + · have h := tendsto_measure_iUnion (μ := κ a) (s := fun m ↦ I n (indexI n t) ×ˢ s m) ?_ + swap + · intro m m' hmm' + simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] + exact Or.inl <| hs hmm' + convert h + rw [← prod_iUnion, hs_iUnion] + · exact Or.inr h0 + +lemma tendsto_m_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] + (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) : + ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ M κ (kernel.fst κ) a (s m) n t) atTop (𝓝 1) := by + filter_upwards [m_univ_ae κ a n] with t ht + rw [← ht] + exact tendsto_m_atTop_univ_of_monotone κ a s hs hs_iUnion n t + +lemma mLimsup_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) : + ∀ᵐ t ∂(kernel.fst κ a), MLimsup κ (kernel.fst κ) a Set.univ t = 1 := by + have h := m_univ_ae κ a + rw [← ae_all_iff] at h + filter_upwards [h] with t ht + rw [MLimsup] + simp [ht] + +lemma tendsto_mLimsup_atTop_ae_of_monotone [IsFiniteKernel κ] + (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : - ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 0) := by - have h_anti : ∀ t, Antitone (fun m ↦ MLimsup κ a (s m) t) := - fun t n m hnm ↦ mLimsup_mono_set κ a (hs hnm) t - have h_le_one : ∀ m t, MLimsup κ a (s m) t ≤ 1 := fun m t ↦ mLimsup_le_one κ a (s m) t + ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ (kernel.fst κ) a (s m) t) atTop (𝓝 1) := by + let ν := kernel.fst κ + have h_mono : ∀ t, Monotone (fun m ↦ MLimsup κ (kernel.fst κ) a (s m) t) := + fun t n m hnm ↦ mLimsup_mono_set le_rfl a (hs hnm) t + have h_le_one : ∀ m t, MLimsup κ ν a (s m) t ≤ 1 := fun m t ↦ mLimsup_le_one le_rfl a (s m) t -- for all `t`, `fun m ↦ MLimsup κ a (s m) t` has a limit - have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 l) := by + have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 l) := by intro t - have h_tendsto : Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop atBot ∨ - ∃ l, Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 l) := - tendsto_of_antitone (h_anti t) + have h_tendsto : Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop atTop ∨ + ∃ l, Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 l) := + tendsto_of_monotone (h_mono t) cases' h_tendsto with h_absurd h_tendsto - · rw [tendsto_atTop_atBot_iff_of_antitone (h_anti t)] at h_absurd - obtain ⟨r, hr⟩ := h_absurd (-1) - have h_nonneg := mLimsup_nonneg κ a (s r) t - linarith + · rw [tendsto_atTop_atTop_iff_of_monotone (h_mono t)] at h_absurd + obtain ⟨r, hr⟩ := h_absurd 2 + exact absurd (hr.trans (h_le_one r t)) one_lt_two.not_le · exact h_tendsto -- let `F` be the pointwise limit of `fun m ↦ MLimsup κ a (s m) t` for all `t` let F : ℝ → ℝ := fun t ↦ (h_exists t).choose - have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ a (s m) t) atTop (𝓝 (F t)) := + have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 (F t)) := fun t ↦ (h_exists t).choose_spec have hF_nonneg : ∀ t, 0 ≤ F t := - fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ mLimsup_nonneg κ a (s m) t) + fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ mLimsup_nonneg le_rfl a (s m) t) have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) - have hF_int : Integrable F (kernel.fst κ a) := by + have hF_int : Integrable F (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨?_, ?_⟩ · refine aestronglyMeasurable_of_tendsto_ae atTop (fun n ↦ ?_) (ae_of_all _ hF_tendsto) - exact (measurable_mLimsup_right κ (hs_meas _) a).aestronglyMeasurable + exact (measurable_mLimsup_right κ ν (hs_meas _) a).aestronglyMeasurable · rw [snorm_one_eq_lintegral_nnnorm] - calc ∫⁻ x, ‖F x‖₊ ∂(kernel.fst κ a) ≤ ∫⁻ _, 1 ∂(kernel.fst κ a) := by + calc ∫⁻ x, ‖F x‖₊ ∂(ν a) ≤ ∫⁻ _, 1 ∂(ν a) := by refine lintegral_mono (fun x ↦ ?_) rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, ENNReal.ofReal_le_one, abs_of_nonneg (hF_nonneg _)] exact hF_le_one _ - _ < ⊤ := by - simp only [lintegral_const, one_mul] - exact measure_lt_top _ _ - -- it suffices to show that the limit `F` is 0 a.e. - suffices ∀ᵐ t ∂(kernel.fst κ a), F t = 0 by + _ < ⊤ := by simp only [lintegral_const, measure_univ, one_mul, measure_lt_top] + -- it suffices to show that the limit `F` is 1 a.e. + suffices ∀ᵐ t ∂(ν a), F t = 1 by filter_upwards [this] with t ht_eq rw [← ht_eq] exact hF_tendsto t - -- since `F` is nonnegative, proving that its integral is 0 is sufficient to get that - -- `F` is 0 a.e. - suffices ∀ᵐ (t : ℝ) ∂(kernel.fst κ) a, 0 = F t by filter_upwards [this] with a ha; simp [ha] - refine ae_eq_of_integral_eq_of_ae_le (integrable_const _) hF_int (ae_of_all _ hF_nonneg) ?_ + -- since `F` is at most 1, proving that its integral is the same as the integral of 1 will tell + -- us that `F` is 1 a.e. + refine ae_eq_of_integral_eq_of_ae_le hF_int (integrable_const _) (ae_of_all _ hF_le_one) ?_ have h_integral : - Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop - (𝓝 (∫ t, F t ∂(kernel.fst κ a))) := by - refine integral_tendsto_of_tendsto_of_antitone ?_ hF_int ?_ ?_ - · exact fun n ↦ integrable_mLimsup _ _ (hs_meas n) - · exact ae_of_all _ h_anti + Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a (s m) t ∂(ν a)) atTop + (𝓝 (∫ t, F t ∂(ν a))) := by + refine integral_tendsto_of_tendsto_of_monotone ?_ hF_int ?_ ?_ + · exact fun n ↦ integrable_mLimsup le_rfl _ (hs_meas n) + · exact ae_of_all _ h_mono · exact ae_of_all _ hF_tendsto have h_integral' : - Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ a (s m) t ∂(kernel.fst κ a)) atTop - (𝓝 (∫ _, 0 ∂(kernel.fst κ a))) := by - simp only [integral_zero] - exact tendsto_integral_mLimsup_of_antitone κ a s hs hs_iInter hs_meas - exact (tendsto_nhds_unique h_integral h_integral').symm + Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a (s m) t ∂(ν a)) atTop + (𝓝 (∫ _, 1 ∂(ν a))) := by + rw [integral_const] + simp only [smul_eq_mul, mul_one] + rw [kernel.fst_apply' _ _ MeasurableSet.univ] + exact tendsto_integral_mLimsup_of_monotone le_rfl a s hs hs_iUnion hs_meas + exact tendsto_nhds_unique h_integral h_integral' + +end UnivFst + +end M + +variable {κ : kernel α (ℝ × ℝ)} {ν : kernel α ℝ} section Iic_Q noncomputable -def mLimsupIic (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : ℝ := MLimsup κ a (Set.Iic q) t +def mLimsupIic (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) (a : α) (t : ℝ) (q : ℚ) : ℝ := + MLimsup κ ν a (Set.Iic q) t -lemma measurable_mLimsupIic (κ : kernel α (ℝ × ℝ)) : - Measurable (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2) := by +lemma measurable_mLimsupIic (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) : + Measurable (fun p : α × ℝ ↦ mLimsupIic κ ν p.1 p.2) := by rw [measurable_pi_iff] - exact fun _ ↦ measurable_mLimsup κ measurableSet_Iic + exact fun _ ↦ measurable_mLimsup κ ν measurableSet_Iic -lemma measurable_mLimsupIic_right (κ : kernel α (ℝ × ℝ)) (a : α) (q : ℚ) : - Measurable (fun t ↦ mLimsupIic κ a t q) := - measurable_mLimsup_right _ measurableSet_Iic _ +lemma measurable_mLimsupIic_right (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) (a : α) (q : ℚ) : + Measurable (fun t ↦ mLimsupIic κ ν a t q) := + measurable_mLimsup_right _ _ measurableSet_Iic _ -lemma monotone_mLimsupIic (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) : Monotone (mLimsupIic κ a t) := by +lemma monotone_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) : + Monotone (mLimsupIic κ ν a t) := by intro q r hqr rw [mLimsupIic, mLimsupIic] - refine mLimsup_mono_set κ a ?_ t + refine mLimsup_mono_set hκν a ?_ t refine Iic_subset_Iic.mpr ?_ exact_mod_cast hqr -lemma mLimsupIic_nonneg (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : 0 ≤ mLimsupIic κ a t q := - mLimsup_nonneg κ a _ t +lemma mLimsupIic_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : + 0 ≤ mLimsupIic κ ν a t q := + mLimsup_nonneg hκν a _ t -lemma mLimsupIic_le_one (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : mLimsupIic κ a t q ≤ 1 := - mLimsup_le_one κ a _ t +lemma mLimsupIic_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : + mLimsupIic κ ν a t q ≤ 1 := + mLimsup_le_one hκν a _ t lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ a t q) atTop (𝓝 1) := by - suffices ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ a t n) atTop (𝓝 1) by + ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ (kernel.fst κ) a t q) atTop (𝓝 1) := by + let ν := kernel.fst κ + suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ ν a t n) atTop (𝓝 1) by filter_upwards [this] with t ht - let f := fun q : ℚ ↦ mLimsupIic κ a t ⌊q⌋₊ - let g := fun q : ℚ ↦ mLimsupIic κ a t ⌈q⌉₊ - have hf_le : ∀ᶠ q in atTop, f q ≤ mLimsupIic κ a t q := by + let f := fun q : ℚ ↦ mLimsupIic κ ν a t ⌊q⌋₊ + let g := fun q : ℚ ↦ mLimsupIic κ ν a t ⌈q⌉₊ + have hf_le : ∀ᶠ q in atTop, f q ≤ mLimsupIic κ ν a t q := by simp only [eventually_atTop, ge_iff_le] - exact ⟨0, fun q hq ↦ monotone_mLimsupIic κ a t (Nat.floor_le hq)⟩ - have hg_le : ∀ q, mLimsupIic κ a t q ≤ g q := fun q ↦ monotone_mLimsupIic κ a t (Nat.le_ceil _) + exact ⟨0, fun q hq ↦ monotone_mLimsupIic le_rfl a t (Nat.floor_le hq)⟩ + have hg_le : ∀ q, mLimsupIic κ ν a t q ≤ g q := + fun q ↦ monotone_mLimsupIic le_rfl a t (Nat.le_ceil _) refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ hf_le (eventually_of_forall hg_le) · exact ht.comp tendsto_nat_floor_atTop · exact ht.comp tendsto_nat_ceil_atTop @@ -1007,28 +1093,29 @@ lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] simp only [mem_iUnion, mem_Iic, mem_univ, iff_true] exact ⟨Nat.ceil x, Nat.le_ceil x⟩ have hs_meas : ∀ n, MeasurableSet (s n) := fun _ ↦ measurableSet_Iic - filter_upwards [tendsto_mLimsup_atTop_ae_of_monotone κ a s hs hs_iUnion hs_meas] + filter_upwards [tendsto_mLimsup_atTop_ae_of_monotone a s hs hs_iUnion hs_meas] with x hx using hx -lemma tendsto_atBot_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ a t q) atBot (𝓝 0) := by - suffices ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ a t (-q)) atTop (𝓝 0) by +lemma tendsto_atBot_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) : + ∀ᵐ t ∂(ν a), Tendsto (fun q ↦ mLimsupIic κ ν a t q) atBot (𝓝 0) := by + suffices ∀ᵐ t ∂(ν a), Tendsto (fun q ↦ mLimsupIic κ ν a t (-q)) atTop (𝓝 0) by filter_upwards [this] with t ht - have h_eq_neg : (fun q ↦ mLimsupIic κ a t q) = fun q ↦ mLimsupIic κ a t (- -q) := by + have h_eq_neg : (fun q ↦ mLimsupIic κ ν a t q) = fun q ↦ mLimsupIic κ ν a t (- -q) := by simp_rw [neg_neg] rw [h_eq_neg] exact ht.comp tendsto_neg_atBot_atTop - suffices ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ a t (-n)) atTop (𝓝 0) by + suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ ν a t (-n)) atTop (𝓝 0) by filter_upwards [this] with t ht - let f := fun q : ℚ ↦ mLimsupIic κ a t (-⌊q⌋₊) - let g := fun q : ℚ ↦ mLimsupIic κ a t (-⌈q⌉₊) - have hf_le : ∀ᶠ q in atTop, mLimsupIic κ a t (-q) ≤ f q := by + let f := fun q : ℚ ↦ mLimsupIic κ ν a t (-⌊q⌋₊) + let g := fun q : ℚ ↦ mLimsupIic κ ν a t (-⌈q⌉₊) + have hf_le : ∀ᶠ q in atTop, mLimsupIic κ ν a t (-q) ≤ f q := by simp only [eventually_atTop, ge_iff_le] - refine ⟨0, fun q hq ↦ monotone_mLimsupIic κ a t ?_⟩ + refine ⟨0, fun q hq ↦ monotone_mLimsupIic hκν a t ?_⟩ rw [neg_le_neg_iff] exact Nat.floor_le hq - have hg_le : ∀ q, g q ≤ mLimsupIic κ a t (-q) := by - refine fun q ↦ monotone_mLimsupIic κ a t ?_ + have hg_le : ∀ q, g q ≤ mLimsupIic κ ν a t (-q) := by + refine fun q ↦ monotone_mLimsupIic hκν a t ?_ rw [neg_le_neg_iff] exact Nat.le_ceil _ refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ (eventually_of_forall hg_le) hf_le @@ -1041,95 +1128,100 @@ lemma tendsto_atBot_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, neg_lt] exact exists_nat_gt (-x) have hs_meas : ∀ n, MeasurableSet (s n) := fun _ ↦ measurableSet_Iic - convert tendsto_mLimsup_atTop_ae_of_antitone κ a s hs hs_iInter hs_meas with x n + convert tendsto_mLimsup_atTop_ae_of_antitone hκν a s hs hs_iInter hs_meas with x n rw [mLimsupIic] simp -lemma set_integral_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] +lemma set_integral_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (q : ℚ) {A : Set ℝ} (hA : MeasurableSet A) : - ∫ t in A, mLimsupIic κ a t q ∂(kernel.fst κ a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := - set_integral_mLimsup κ a measurableSet_Iic hA + ∫ t in A, mLimsupIic κ ν a t q ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := + set_integral_mLimsup hκν a measurableSet_Iic hA -lemma integrable_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel (kernel.fst κ)] +lemma integrable_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (q : ℚ) : - Integrable (fun t ↦ mLimsupIic κ a t q) (kernel.fst κ a) := - integrable_mLimsup _ _ measurableSet_Iic + Integrable (fun t ↦ mLimsupIic κ ν a t q) (ν a) := + integrable_mLimsup hκν _ measurableSet_Iic -lemma bddBelow_range_mLimsupIic (κ : kernel α (ℝ × ℝ)) (a : α) (t : ℝ) (q : ℚ) : - BddBelow (range fun (r : Ioi q) ↦ mLimsupIic κ a t r) := by +lemma bddBelow_range_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : + BddBelow (range fun (r : Ioi q) ↦ mLimsupIic κ ν a t r) := by refine ⟨0, ?_⟩ rw [mem_lowerBounds] rintro x ⟨y, rfl⟩ - exact mLimsupIic_nonneg _ _ _ _ + exact mLimsupIic_nonneg hκν _ _ _ -lemma integrable_iInf_rat_gt_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) (q : ℚ) : - Integrable (fun t ↦ ⨅ r : Ioi q, mLimsupIic κ a t r) (kernel.fst κ a) := by +lemma integrable_iInf_rat_gt_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] + (a : α) (q : ℚ) : + Integrable (fun t ↦ ⨅ r : Ioi q, mLimsupIic κ ν a t r) (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ - · exact measurable_iInf fun i ↦ measurable_mLimsupIic_right κ a i - refine (?_ : _ ≤ (kernel.fst κ a univ : ℝ≥0∞)).trans_lt (measure_lt_top _ _) + · exact measurable_iInf fun i ↦ measurable_mLimsupIic_right κ ν a i + refine (?_ : _ ≤ (ν a univ : ℝ≥0∞)).trans_lt (measure_lt_top _ _) refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ · rw [Real.norm_eq_abs, abs_of_nonneg] · refine ciInf_le_of_le ?_ ?_ ?_ - · exact bddBelow_range_mLimsupIic κ a t _ + · exact bddBelow_range_mLimsupIic hκν a t _ · exact ⟨q + 1, by simp⟩ - · exact mLimsupIic_le_one _ _ _ _ - · exact le_ciInf fun r ↦ mLimsupIic_nonneg κ a t r + · exact mLimsupIic_le_one hκν _ _ _ + · exact le_ciInf fun r ↦ mLimsupIic_nonneg hκν a t r · simp -lemma set_integral_iInf_rat_gt_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] +lemma set_integral_iInf_rat_gt_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) (q : ℚ) {A : Set ℝ} (hA : MeasurableSet A) : - ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ a t r ∂(kernel.fst κ a) + ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ ν a t r ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by refine le_antisymm ?_ ?_ - · have h : ∀ r : Ioi q, ∫ t in A, ⨅ r' : Ioi q, mLimsupIic κ a t r' ∂(kernel.fst κ a) + · have h : ∀ r : Ioi q, ∫ t in A, ⨅ r' : Ioi q, mLimsupIic κ ν a t r' ∂(ν a) ≤ (κ a (A ×ˢ Iic (r : ℝ))).toReal := by intro r - rw [← set_integral_mLimsupIic κ a r hA] + rw [← set_integral_mLimsupIic hκν a r hA] refine set_integral_mono ?_ ?_ ?_ - · exact (integrable_iInf_rat_gt_mLimsupIic _ _ _).integrableOn - · exact (integrable_mLimsupIic _ _ _).integrableOn - · exact fun t ↦ ciInf_le (bddBelow_range_mLimsupIic _ _ _ _) r - calc ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ a t r ∂(kernel.fst κ a) + · exact (integrable_iInf_rat_gt_mLimsupIic hκν _ _).integrableOn + · exact (integrable_mLimsupIic hκν _ _).integrableOn + · exact fun t ↦ ciInf_le (bddBelow_range_mLimsupIic hκν _ _ _) r + calc ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ ν a t r ∂(ν a) ≤ ⨅ r : Ioi q, (κ a (A ×ˢ Iic (r : ℝ))).toReal := le_ciInf h _ = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by rw [← Measure.iInf_Iic_gt_prod hA q] exact (ENNReal.toReal_iInf (fun r ↦ measure_ne_top _ _)).symm - · rw [← set_integral_mLimsupIic κ a q hA] + · rw [← set_integral_mLimsupIic hκν a q hA] refine set_integral_mono ?_ ?_ ?_ - · exact (integrable_mLimsupIic _ _ _).integrableOn - · exact (integrable_iInf_rat_gt_mLimsupIic _ _ _).integrableOn - · exact fun t ↦ le_ciInf (fun r ↦ monotone_mLimsupIic _ _ _ (le_of_lt r.prop)) + · exact (integrable_mLimsupIic hκν _ _).integrableOn + · exact (integrable_iInf_rat_gt_mLimsupIic hκν _ _).integrableOn + · exact fun t ↦ le_ciInf (fun r ↦ monotone_mLimsupIic hκν _ _ (le_of_lt r.prop)) -lemma iInf_rat_gt_mLimsupIic_eq (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), ∀ q : ℚ, ⨅ r : Ioi q, mLimsupIic κ a t r = mLimsupIic κ a t q := by +lemma iInf_rat_gt_mLimsupIic_eq (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) : + ∀ᵐ t ∂(ν a), ∀ q : ℚ, ⨅ r : Ioi q, mLimsupIic κ ν a t r = mLimsupIic κ ν a t q := by rw [ae_all_iff] - refine fun q ↦ ae_eq_of_forall_set_integral_eq_of_sigmaFinite (μ := kernel.fst κ a) ?_ ?_ ?_ + refine fun q ↦ ae_eq_of_forall_set_integral_eq_of_sigmaFinite (μ := ν a) ?_ ?_ ?_ · intro s _ _ refine Integrable.integrableOn ?_ - exact integrable_iInf_rat_gt_mLimsupIic κ _ _ - · exact fun s _ _ ↦ (integrable_mLimsupIic κ a q).integrableOn + exact integrable_iInf_rat_gt_mLimsupIic hκν _ _ + · exact fun s _ _ ↦ (integrable_mLimsupIic hκν a q).integrableOn · intro s hs _ - rw [set_integral_mLimsupIic _ _ _ hs, set_integral_iInf_rat_gt_mLimsupIic _ _ _ hs] + rw [set_integral_mLimsupIic hκν _ _ hs, set_integral_iInf_rat_gt_mLimsupIic hκν _ _ hs] end Iic_Q lemma isRatStieltjesPoint_mLimsupIic_ae (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), IsRatStieltjesPoint (fun p q ↦ mLimsupIic κ p.1 p.2 q) (a, t) := by - filter_upwards [tendsto_atTop_mLimsupIic κ a, tendsto_atBot_mLimsupIic κ a, - iInf_rat_gt_mLimsupIic_eq κ a] with t ht_top ht_bot ht_iInf - exact ⟨monotone_mLimsupIic κ a t, ht_top, ht_bot, ht_iInf⟩ + ∀ᵐ t ∂(kernel.fst κ a), + IsRatStieltjesPoint (fun p q ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2 q) (a, t) := by + filter_upwards [tendsto_atTop_mLimsupIic κ a, tendsto_atBot_mLimsupIic le_rfl a, + iInf_rat_gt_mLimsupIic_eq le_rfl a] with t ht_top ht_bot ht_iInf + exact ⟨monotone_mLimsupIic le_rfl a t, ht_top, ht_bot, ht_iInf⟩ lemma isRatKernelCDF_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : - IsRatKernelCDF (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2) κ (kernel.fst κ) where - measurable := measurable_mLimsupIic κ + IsRatKernelCDF (fun p : α × ℝ ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2) κ (kernel.fst κ) where + measurable := measurable_mLimsupIic κ (kernel.fst κ) isRatStieltjesPoint_ae := isRatStieltjesPoint_mLimsupIic_ae κ - integrable := integrable_mLimsupIic κ - set_integral := fun _ _ hs _ ↦ set_integral_mLimsupIic _ _ _ hs + integrable := integrable_mLimsupIic le_rfl + set_integral := fun _ _ hs _ ↦ set_integral_mLimsupIic le_rfl _ _ hs noncomputable -def mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : α × ℝ → StieltjesFunction := - stieltjesOfMeasurableRat (fun p : α × ℝ ↦ mLimsupIic κ p.1 p.2) +def mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : + α × ℝ → StieltjesFunction := + stieltjesOfMeasurableRat (fun p : α × ℝ ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2) (isRatKernelCDF_mLimsupIic κ).measurable lemma isKernelCDF_mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : From e993e5ddfcfa647bb9c44a79e3f9376f59ee63fd Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sat, 17 Feb 2024 11:26:14 +0100 Subject: [PATCH 028/129] rename --- .../Kernel/Disintegration/AuxLemmas.lean | 18 +- .../Kernel/Disintegration/KernelCDFBorel.lean | 285 +++++++++--------- Mathlib/Topology/Instances/ENNReal.lean | 3 + 3 files changed, 151 insertions(+), 155 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean b/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean index daec77bca6d2d..c7b2952e77706 100644 --- a/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean +++ b/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean @@ -45,6 +45,7 @@ lemma integral_tendsto_of_tendsto_of_monotone {μ : Measure α} {f : ℕ → α (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Monotone fun n ↦ f n x) (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by + -- switch from the Bochner to the Lebesgue integral let f' := fun n x ↦ f n x - f 0 x have hf'_nonneg : ∀ᵐ x ∂μ, ∀ n, 0 ≤ f' n x := by filter_upwards [h_mono] with a ha n @@ -69,21 +70,12 @@ lemma integral_tendsto_of_tendsto_of_monotone {μ : Measure α} {f : ℕ → α rw [ae_all_iff] at hf'_nonneg simp_rw [integral_eq_lintegral_of_nonneg_ae (hf'_nonneg _) (hf'_meas _).1] rw [integral_eq_lintegral_of_nonneg_ae hF_ge (hF.1.sub (hf 0).1)] - have h_cont := ENNReal.continuousOn_toReal.continuousAt - (x := ∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ) ?_ + have h_cont := ENNReal.continuousAt_toReal (x := ∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ) ?_ swap - · rw [mem_nhds_iff] - refine ⟨Iio (∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ + 1), ?_, isOpen_Iio, ?_⟩ - · intro x - simp only [Pi.sub_apply, mem_Iio, ne_eq, mem_setOf_eq] - exact ne_top_of_lt - · simp only [Pi.sub_apply, mem_Iio] - refine ENNReal.lt_add_right ?_ one_ne_zero - rw [← ofReal_integral_eq_lintegral_ofReal] - · exact ENNReal.ofReal_ne_top - · exact hF.sub (hf 0) - · exact hF_ge + · rw [← ofReal_integral_eq_lintegral_ofReal (hF.sub (hf 0)) hF_ge] + exact ENNReal.ofReal_ne_top refine h_cont.tendsto.comp ?_ + -- use the result for the Lebesgue integral refine lintegral_tendsto_of_tendsto_of_monotone ?_ ?_ ?_ · exact fun n ↦ ((hf n).sub (hf 0)).aemeasurable.ennreal_ofReal · filter_upwards [h_mono] with x hx diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean index c6ee8e8bdc2a4..eb72ecd36b4cf 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean @@ -250,14 +250,15 @@ section M variable {κ : kernel α (ℝ × β)} {ν : kernel α ℝ} noncomputable -def M (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : ℝ := +def densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : + ℝ := (κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t))).toReal -lemma m_def (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) : - M κ ν a s n = fun t ↦ (κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t))).toReal := +lemma densityProcess_def (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) : + densityProcess κ ν a s n = fun t ↦ (κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t))).toReal := rfl -lemma measurable_m_aux (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) +lemma measurable_densityProcess_aux (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : Measurable (fun (p : α × ℝ) ↦ κ p.1 (I n (indexI n p.2) ×ˢ s) / ν p.1 (I n (indexI n p.2))) := by @@ -282,24 +283,24 @@ lemma measurable_m_aux (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) refine h1.comp (measurable_fst.prod_mk ?_) exact (Measurable.mono (measurable_indexI n) (ℱ.le n) le_rfl).comp measurable_snd -lemma measurable_m (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) +lemma measurable_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : - Measurable (fun (p : α × ℝ) ↦ M κ ν p.1 s n p.2) := - (measurable_m_aux κ ν hs n).ennreal_toReal + Measurable (fun (p : α × ℝ) ↦ densityProcess κ ν p.1 s n p.2) := + (measurable_densityProcess_aux κ ν hs n).ennreal_toReal -lemma measurable_m_left (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) +lemma measurable_densityProcess_left (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) {s : Set β} (hs : MeasurableSet s) (n : ℕ) (t : ℝ) : - Measurable (fun a ↦ M κ ν a s n t) := - (measurable_m κ ν hs n).comp (measurable_id.prod_mk measurable_const) + Measurable (fun a ↦ densityProcess κ ν a s n t) := + (measurable_densityProcess κ ν hs n).comp (measurable_id.prod_mk measurable_const) -lemma measurable_m_right (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) +lemma measurable_densityProcess_right (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) {s : Set β} (a : α) (hs : MeasurableSet s) (n : ℕ) : - Measurable (M κ ν a s n) := - (measurable_m κ ν hs n).comp (measurable_const.prod_mk measurable_id) + Measurable (densityProcess κ ν a s n) := + (measurable_densityProcess κ ν hs n).comp (measurable_const.prod_mk measurable_id) -lemma measurable_ℱ_m (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) : - Measurable[ℱ n] (M κ ν a s n) := by - rw [m_def] +lemma measurable_ℱ_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) : + Measurable[ℱ n] (densityProcess κ ν a s n) := by + rw [densityProcess_def] refine @Measurable.ennreal_toReal _ (ℱ n) _ ?_ refine Measurable.div ?_ ?_ · change Measurable[ℱ n] ((fun k ↦ κ a (I n k ×ˢ s)) ∘ (indexI n)) @@ -309,25 +310,22 @@ lemma measurable_ℱ_m (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α refine Measurable.comp ?_ (measurable_indexI n) exact measurable_of_countable _ -lemma stronglyMeasurable_ℱ_m (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) : - StronglyMeasurable[ℱ n] (M κ ν a s n) := - (measurable_ℱ_m κ ν a s n).stronglyMeasurable +lemma stronglyMeasurable_ℱ_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) : + StronglyMeasurable[ℱ n] (densityProcess κ ν a s n) := + (measurable_ℱ_densityProcess κ ν a s n).stronglyMeasurable -lemma adapted_m (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) : - Adapted ℱ (M κ ν a s) := - stronglyMeasurable_ℱ_m κ ν a s +lemma adapted_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) : + Adapted ℱ (densityProcess κ ν a s) := + stronglyMeasurable_ℱ_densityProcess κ ν a s -lemma m_nonneg (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : - 0 ≤ M κ ν a s n t := +lemma densityProcess_nonneg (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : + 0 ≤ densityProcess κ ν a s n t := ENNReal.toReal_nonneg -lemma m_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : - M κ ν a s n t ≤ 1 := by - rw [M] - refine ENNReal.toReal_le_of_le_ofReal zero_le_one ?_ - rw [ENNReal.ofReal_one] - refine ENNReal.div_le_of_le_mul ?_ - rw [one_mul] +lemma densityProcess_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : + densityProcess κ ν a s n t ≤ 1 := by + refine ENNReal.toReal_le_of_le_ofReal zero_le_one (ENNReal.div_le_of_le_mul ?_) + rw [ENNReal.ofReal_one, one_mul] calc κ a (I n (indexI n t) ×ˢ s) ≤ kernel.fst κ a (I n (indexI n t)) := by rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] refine measure_mono (fun x ↦ ?_) @@ -335,28 +333,29 @@ lemma m_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (n : ℕ) (t exact fun h _ ↦ h _ ≤ ν a (I n (indexI n t)) := hκν a _ (measurableSet_I _ _) -lemma snorm_m_le (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (n : ℕ) : - snorm (M κ ν a s n) 1 (ν a) ≤ ν a univ := by +lemma snorm_densityProcess_le (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (n : ℕ) : + snorm (densityProcess κ ν a s n) 1 (ν a) ≤ ν a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun x ↦ ?_))).trans ?_ - · simp only [Real.norm_eq_abs, abs_of_nonneg (m_nonneg κ ν a s n x), m_le_one hκν a s n x] + · simp only [Real.norm_eq_abs, abs_of_nonneg (densityProcess_nonneg κ ν a s n x), + densityProcess_le_one hκν a s n x] · simp -lemma integrable_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] +lemma integrable_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : - Integrable (M κ ν a s n) (ν a) := by + Integrable (densityProcess κ ν a s n) (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ - · exact measurable_m_right κ ν a hs n - · exact (snorm_m_le hκν a s n).trans_lt (measure_lt_top _ _) + · exact measurable_densityProcess_right κ ν a hs n + · exact (snorm_densityProcess_le hκν a s n).trans_lt (measure_lt_top _ _) -lemma set_integral_m_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma set_integral_densityProcess_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) (k : ℤ) : - ∫ t in I n k, M κ ν a s n t ∂(ν a) = (κ a (I n k ×ˢ s)).toReal := by - simp_rw [M] + ∫ t in I n k, densityProcess κ ν a s n t ∂(ν a) = (κ a (I n k ×ˢ s)).toReal := by + simp_rw [densityProcess] rw [integral_toReal] rotate_left · refine Measurable.aemeasurable ?_ - have h := measurable_m_aux κ ν hs n + have h := measurable_densityProcess_aux κ ν hs n change Measurable ((fun (p : α × ℝ) ↦ κ p.1 (I n (indexI n p.2) ×ˢ s) / ν p.1 (I n (indexI n p.2))) ∘ (fun t ↦ (a, t))) @@ -391,9 +390,9 @@ lemma set_integral_m_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFin rw [div_eq_mul_inv, mul_assoc, ENNReal.inv_mul_cancel h0, mul_one] exact measure_ne_top _ _ -lemma integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : - ∫ t, M κ ν a s n t ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by + ∫ t, densityProcess κ ν a s n t ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by rw [← integral_univ, ← iUnion_I n, iUnion_prod_const, measure_iUnion] rotate_left · intro i j hij @@ -401,16 +400,16 @@ lemma integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKer exact Or.inl (pairwise_disjoint_I n hij) · exact fun k ↦ (measurableSet_I n k).prod hs rw [integral_iUnion (measurableSet_I n) (pairwise_disjoint_I n) - (integrable_m hκν a hs n).integrableOn] + (integrable_densityProcess hκν a hs n).integrableOn] rw [ENNReal.tsum_toReal_eq (fun _ ↦ measure_ne_top _ _)] congr with k - rw [set_integral_m_I hκν _ hs] + rw [set_integral_densityProcess_I hκν _ hs] -lemma set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : - ∫ t in A, M κ ν a s n t ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + ∫ t in A, densityProcess κ ν a s n t ∂(ν a) = (κ a (A ×ˢ s)).toReal := by refine MeasurableSpace.induction_on_inter (m := ℱ n) (s := {s | ∃ k, s = I n k}) - (C := fun A ↦ ∫ t in A, M κ ν a s n t ∂(ν a) = (κ a (A ×ˢ s)).toReal) rfl + (C := fun A ↦ ∫ t in A, densityProcess κ ν a s n t ∂(ν a) = (κ a (A ×ˢ s)).toReal) rfl ?_ ?_ ?_ ?_ ?_ hA · rintro s ⟨i, rfl⟩ t ⟨j, rfl⟩ hst refine ⟨i, ?_⟩ @@ -422,11 +421,11 @@ lemma set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFinit rwa [Function.onFun, disjoint_iff_inter_eq_empty] at h_disj · simp · rintro _ ⟨k, rfl⟩ - rw [set_integral_m_I hκν _ hs] + rw [set_integral_densityProcess_I hκν _ hs] · intro A hA hA_eq have hA' : MeasurableSet A := ℱ.le _ _ hA - have h := integral_add_compl hA' (integrable_m hκν a hs n) - rw [hA_eq, integral_m hκν a hs] at h + have h := integral_add_compl hA' (integrable_densityProcess hκν a hs n) + rw [hA_eq, integral_densityProcess hκν a hs] at h have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by rw [prod_diff_prod, compl_eq_univ_diff] simp @@ -437,7 +436,8 @@ lemma set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFinit · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] exact measure_mono (by intro x; simp) · intro f hf_disj hf h_eq - rw [integral_iUnion (fun i ↦ ℱ.le n _ (hf i)) hf_disj (integrable_m hκν _ hs _).integrableOn] + rw [integral_iUnion (fun i ↦ ℱ.le n _ (hf i)) hf_disj + (integrable_densityProcess hκν _ hs _).integrableOn] simp_rw [h_eq] rw [iUnion_prod_const, measure_iUnion _ (fun i ↦ (ℱ.le n _ (hf i)).prod hs)] · rw [ENNReal.tsum_toReal_eq] @@ -446,34 +446,34 @@ lemma set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFinit rw [Function.onFun, Set.disjoint_prod] exact Or.inl (hf_disj hij) -lemma set_integral_m_of_le (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma set_integral_densityProcess_of_le (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {n m : ℕ} (hnm : n ≤ m) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : - ∫ t in A, M κ ν a s m t ∂(ν a) = (κ a (A ×ˢ s)).toReal := - set_integral_m hκν a hs m (ℱ.mono hnm A hA) + ∫ t in A, densityProcess κ ν a s m t ∂(ν a) = (κ a (A ×ˢ s)).toReal := + set_integral_densityProcess hκν a hs m (ℱ.mono hnm A hA) -lemma condexp_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma condexp_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {i j : ℕ} (hij : i ≤ j) : - (ν a)[M κ ν a s j | ℱ i] =ᵐ[ν a] M κ ν a s i := by + (ν a)[densityProcess κ ν a s j | ℱ i] =ᵐ[ν a] densityProcess κ ν a s i := by symm refine ae_eq_condexp_of_forall_set_integral_eq ?_ ?_ ?_ ?_ ?_ - · exact integrable_m hκν a hs j + · exact integrable_densityProcess hκν a hs j · refine fun t _ _ ↦ Integrable.integrableOn ?_ - exact integrable_m hκν _ hs _ + exact integrable_densityProcess hκν _ hs _ · intro t ht _ - rw [set_integral_m hκν a hs i ht, set_integral_m_of_le hκν a hs hij ht] - · exact StronglyMeasurable.aeStronglyMeasurable' (stronglyMeasurable_ℱ_m κ ν a s i) + rw [set_integral_densityProcess hκν a hs i ht, set_integral_densityProcess_of_le hκν a hs hij ht] + · exact StronglyMeasurable.aeStronglyMeasurable' (stronglyMeasurable_ℱ_densityProcess κ ν a s i) -lemma martingale_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma martingale_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Martingale (M κ ν a s) ℱ (ν a) := - ⟨adapted_m κ ν a s, fun _ _ ↦ condexp_m hκν a hs⟩ + Martingale (densityProcess κ ν a s) ℱ (ν a) := + ⟨adapted_densityProcess κ ν a s, fun _ _ ↦ condexp_densityProcess hκν a hs⟩ -lemma m_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) {s s' : Set β} (h : s ⊆ s') +lemma densityProcess_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) {s s' : Set β} (h : s ⊆ s') (n : ℕ) (t : ℝ) : - M κ ν a s n t ≤ M κ ν a s' n t := by - unfold M + densityProcess κ ν a s n t ≤ densityProcess κ ν a s' n t := by + unfold densityProcess by_cases h0 : ν a (I n (indexI n t)) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] simp @@ -495,10 +495,10 @@ lemma m_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) {s s' : Set β} (h : s · exact h_ne_top s · exact h_ne_top s' -lemma m_mono_kernel_left {κ' : kernel α (ℝ × β)} (hκκ' : κ ≤ κ') (hκ'ν : kernel.fst κ' ≤ ν) - (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) (t : ℝ) : - M κ ν a s n t ≤ M κ' ν a s n t := by - unfold M +lemma densityProcess_mono_kernel_left {κ' : kernel α (ℝ × β)} (hκκ' : κ ≤ κ') + (hκ'ν : kernel.fst κ' ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) (t : ℝ) : + densityProcess κ ν a s n t ≤ densityProcess κ' ν a s n t := by + unfold densityProcess by_cases h0 : ν a (I n (indexI n t)) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] simp @@ -520,10 +520,10 @@ lemma m_mono_kernel_left {κ' : kernel α (ℝ × β)} (hκκ' : κ ≤ κ') (h simp only [ne_eq, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top -lemma m_antitone_kernel_right {ν' : kernel α ℝ} (hνν' : ν ≤ ν') (hκν : kernel.fst κ ≤ ν) +lemma densityProcess_antitone_kernel_right {ν' : kernel α ℝ} (hνν' : ν ≤ ν') (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : - M κ ν' a s n t ≤ M κ ν a s n t := by - unfold M + densityProcess κ ν' a s n t ≤ densityProcess κ ν a s n t := by + unfold densityProcess have h_le : κ a (I n (indexI n t) ×ˢ s) ≤ ν a (I n (indexI n t)) := by calc κ a (I n (indexI n t) ×ˢ s) ≤ kernel.fst κ a (I n (indexI n t)) := by rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] @@ -547,16 +547,15 @@ lemma m_antitone_kernel_right {ν' : kernel α ℝ} (hνν' : ν ≤ ν') (hκν · simp only [ne_eq, ENNReal.div_eq_top, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top -lemma m_empty (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (n : ℕ) (t : ℝ) : - M κ ν a ∅ n t = 0 := by - rw [M] - simp +lemma densityProcess_empty (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (n : ℕ) (t : ℝ) : + densityProcess κ ν a ∅ n t = 0 := by + simp [densityProcess] -lemma tendsto_m_atTop_empty_of_antitone (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) [IsFiniteKernel κ] +lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) (n : ℕ) (t : ℝ) : - Tendsto (fun m ↦ M κ ν a (s m) n t) atTop (𝓝 (M κ ν a ∅ n t)) := by - simp_rw [M] + Tendsto (fun m ↦ densityProcess κ ν a (s m) n t) atTop (𝓝 (densityProcess κ ν a ∅ n t)) := by + simp_rw [densityProcess] by_cases h0 : ν a (I n (indexI n t)) = 0 · simp_rw [h0, ENNReal.toReal_div] simp @@ -576,81 +575,81 @@ lemma tendsto_m_atTop_empty_of_antitone (κ : kernel α (ℝ × β)) (ν : kerne · simp only [prod_empty, OuterMeasure.empty', ne_eq, not_true_eq_false, false_or, h0, not_false_iff] -lemma tendsto_m_atTop_of_antitone (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) [IsFiniteKernel κ] +lemma tendsto_densityProcess_atTop_of_antitone (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) (n : ℕ) (t : ℝ) : - Tendsto (fun m ↦ M κ ν a (s m) n t) atTop (𝓝 0) := by - rw [← m_empty κ ν a n t] - exact tendsto_m_atTop_empty_of_antitone κ ν a s hs hs_iInter hs_meas n t + Tendsto (fun m ↦ densityProcess κ ν a (s m) n t) atTop (𝓝 0) := by + rw [← densityProcess_empty κ ν a n t] + exact tendsto_densityProcess_atTop_empty_of_antitone κ ν a s hs hs_iInter hs_meas n t -lemma tendsto_m_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma tendsto_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - ∀ᵐ t ∂(ν a), - Tendsto (fun n ↦ M κ ν a s n t) atTop (𝓝 (ℱ.limitProcess (M κ ν a s) (ν a) t)) := by - refine Submartingale.ae_tendsto_limitProcess (martingale_m hκν a hs).submartingale + ∀ᵐ t ∂(ν a), Tendsto (fun n ↦ densityProcess κ ν a s n t) atTop + (𝓝 (ℱ.limitProcess (densityProcess κ ν a s) (ν a) t)) := by + refine Submartingale.ae_tendsto_limitProcess (martingale_densityProcess hκν a hs).submartingale (R := (ν a univ).toNNReal) (fun n ↦ ?_) - refine (snorm_m_le hκν a s n).trans_eq ?_ + refine (snorm_densityProcess_le hκν a s n).trans_eq ?_ rw [ENNReal.coe_toNNReal] exact measure_ne_top _ _ lemma limitProcess_mem_L1 (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Memℒp (ℱ.limitProcess (M κ ν a s) (ν a)) 1 (ν a) := by - refine Submartingale.memℒp_limitProcess (martingale_m hκν a hs).submartingale + Memℒp (ℱ.limitProcess (densityProcess κ ν a s) (ν a)) 1 (ν a) := by + refine Submartingale.memℒp_limitProcess (martingale_densityProcess hκν a hs).submartingale (R := (ν a univ).toNNReal) (fun n ↦ ?_) - refine (snorm_m_le hκν a s n).trans_eq ?_ + refine (snorm_densityProcess_le hκν a s n).trans_eq ?_ rw [ENNReal.coe_toNNReal] exact measure_ne_top _ _ -lemma tendsto_snorm_one_m_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] +lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : Tendsto - (fun n ↦ snorm (M κ ν a s n - ℱ.limitProcess (M κ ν a s) (ν a)) 1 (ν a)) + (fun n ↦ snorm (densityProcess κ ν a s n - ℱ.limitProcess (densityProcess κ ν a s) (ν a)) 1 (ν a)) atTop (𝓝 0) := by refine Submartingale.tendsto_snorm_one_limitProcess ?_ ?_ - · exact (martingale_m hκν a hs).submartingale + · exact (martingale_densityProcess hκν a hs).submartingale · refine uniformIntegrable_of le_rfl ENNReal.one_ne_top ?_ ?_ - · exact fun n ↦ (measurable_m_right κ ν a hs n).aestronglyMeasurable + · exact fun n ↦ (measurable_densityProcess_right κ ν a hs n).aestronglyMeasurable · intro ε _ refine ⟨2, fun n ↦ ?_⟩ refine le_of_eq_of_le ?_ (?_ : 0 ≤ ENNReal.ofReal ε) - · have : {x | 2 ≤ ‖M κ ν a s n x‖₊} = ∅ := by + · have : {x | 2 ≤ ‖densityProcess κ ν a s n x‖₊} = ∅ := by ext x simp only [mem_setOf_eq, mem_empty_iff_false, iff_false, not_le] refine (?_ : _ ≤ (1 : ℝ≥0)).trans_lt one_lt_two - rw [Real.nnnorm_of_nonneg (m_nonneg _ _ _ _ _ _)] - exact mod_cast (m_le_one hκν _ _ _ _) + rw [Real.nnnorm_of_nonneg (densityProcess_nonneg _ _ _ _ _ _)] + exact mod_cast (densityProcess_le_one hκν _ _ _ _) rw [this] simp · simp -lemma tendsto_snorm_one_restrict_m_limitProcess [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma tendsto_snorm_one_restrict_densityProcess_limitProcess [IsFiniteKernel κ] [IsFiniteKernel ν] (hκν : kernel.fst κ ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : - Tendsto (fun n ↦ snorm (M κ ν a s n - ℱ.limitProcess (M κ ν a s) (ν a)) 1 ((ν a).restrict A)) + Tendsto (fun n ↦ snorm (densityProcess κ ν a s n - ℱ.limitProcess (densityProcess κ ν a s) (ν a)) 1 ((ν a).restrict A)) atTop (𝓝 0) := - tendsto_snorm_restrict_zero (tendsto_snorm_one_m_limitProcess hκν a hs) A + tendsto_snorm_restrict_zero (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) A noncomputable def MLimsup (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (t : ℝ) : ℝ := - limsup (fun n ↦ M κ ν a s n t) atTop + limsup (fun n ↦ densityProcess κ ν a s n t) atTop lemma mLimsup_ae_eq_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - MLimsup κ ν a s =ᵐ[ν a] ℱ.limitProcess (M κ ν a s) (ν a) := by - filter_upwards [tendsto_m_limitProcess hκν a hs] with t ht using ht.limsup_eq + MLimsup κ ν a s =ᵐ[ν a] ℱ.limitProcess (densityProcess κ ν a s) (ν a) := by + filter_upwards [tendsto_densityProcess_limitProcess hκν a hs] with t ht using ht.limsup_eq lemma tendsto_m_mLimsup (hκν : kernel.fst κ ≤ ν) (a : α) [IsFiniteKernel κ] [IsFiniteKernel ν] {s : Set β} (hs : MeasurableSet s) : ∀ᵐ t ∂(ν a), - Tendsto (fun n ↦ M κ ν a s n t) atTop (𝓝 (MLimsup κ ν a s t)) := by - filter_upwards [tendsto_m_limitProcess hκν a hs, mLimsup_ae_eq_limitProcess hκν a hs] with t h1 h2 + Tendsto (fun n ↦ densityProcess κ ν a s n t) atTop (𝓝 (MLimsup κ ν a s t)) := by + filter_upwards [tendsto_densityProcess_limitProcess hκν a hs, mLimsup_ae_eq_limitProcess hκν a hs] with t h1 h2 rw [h2] exact h1 lemma measurable_mLimsup (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) {s : Set β} (hs : MeasurableSet s) : Measurable (fun (p : α × ℝ) ↦ MLimsup κ ν p.1 s p.2) := - measurable_limsup (fun n ↦ measurable_m κ ν hs n) + measurable_limsup (fun n ↦ measurable_densityProcess κ ν hs n) lemma measurable_mLimsup_left (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) {s : Set β} (hs : MeasurableSet s) (t : ℝ) : @@ -668,21 +667,21 @@ lemma mLimsup_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) {s s' : Set β} ( MLimsup κ ν a s t ≤ MLimsup κ ν a s' t := by rw [MLimsup, MLimsup] refine limsup_le_limsup ?_ ?_ ?_ - · exact eventually_of_forall (fun n ↦ m_mono_set hκν a h n t) - · exact isCoboundedUnder_le_of_le atTop (fun i ↦ m_nonneg _ _ _ _ _ _) - · exact isBoundedUnder_of ⟨1, fun n ↦ m_le_one hκν _ _ _ _⟩ + · exact eventually_of_forall (fun n ↦ densityProcess_mono_set hκν a h n t) + · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) + · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ lemma mLimsup_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (t : ℝ) : 0 ≤ MLimsup κ ν a s t := by refine le_limsup_of_frequently_le ?_ ?_ - · exact frequently_of_forall (fun n ↦ m_nonneg _ _ _ _ _ _) - · exact isBoundedUnder_of ⟨1, fun n ↦ m_le_one hκν _ _ _ _⟩ + · exact frequently_of_forall (fun n ↦ densityProcess_nonneg _ _ _ _ _ _) + · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ lemma mLimsup_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (t : ℝ) : MLimsup κ ν a s t ≤ 1 := by refine limsup_le_of_le ?_ ?_ - · exact isCoboundedUnder_le_of_le atTop (fun i ↦ m_nonneg _ _ _ _ _ _) - · exact eventually_of_forall (fun n ↦ m_le_one hκν _ _ _ _) + · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) + · exact eventually_of_forall (fun n ↦ densityProcess_le_one hκν _ _ _ _) lemma snorm_mLimsup_le (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) : snorm (fun t ↦ MLimsup κ ν a s t) 1 (ν a) ≤ ν a univ := by @@ -701,12 +700,12 @@ lemma integrable_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] lemma tendsto_set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : - Tendsto (fun i ↦ ∫ x in A, M κ ν a s i x ∂(ν a)) atTop + Tendsto (fun i ↦ ∫ x in A, densityProcess κ ν a s i x ∂(ν a)) atTop (𝓝 (∫ x in A, MLimsup κ ν a s x ∂(ν a))) := by refine tendsto_set_integral_of_L1' (μ := ν a) (MLimsup κ ν a s) - (integrable_mLimsup hκν a hs) (F := fun i t ↦ M κ ν a s i t) (l := atTop) - (eventually_of_forall (fun n ↦ integrable_m hκν _ hs _)) ?_ A - refine (tendsto_congr fun n ↦ ?_).mp (tendsto_snorm_one_m_limitProcess hκν a hs) + (integrable_mLimsup hκν a hs) (F := fun i t ↦ densityProcess κ ν a s i t) (l := atTop) + (eventually_of_forall (fun n ↦ integrable_densityProcess hκν _ hs _)) ?_ A + refine (tendsto_congr fun n ↦ ?_).mp (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) refine snorm_congr_ae ?_ refine EventuallyEq.sub EventuallyEq.rfl ?_ exact (mLimsup_ae_eq_limitProcess hκν a hs).symm @@ -715,16 +714,17 @@ lemma set_integral_mLimsup_of_measurableSet (hκν : kernel.fst κ ≤ ν) [IsFi [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : ∫ t in A, MLimsup κ ν a s t ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - suffices ∫ t in A, MLimsup κ ν a s t ∂(ν a) = ∫ t in A, M κ ν a s n t ∂(ν a) by + suffices ∫ t in A, MLimsup κ ν a s t ∂(ν a) = ∫ t in A, densityProcess κ ν a s n t ∂(ν a) by rw [this] - exact set_integral_m hκν _ hs _ hA + exact set_integral_densityProcess hκν _ hs _ hA suffices ∫ t in A, MLimsup κ ν a s t ∂(ν a) - = limsup (fun i ↦ ∫ t in A, M κ ν a s i t ∂(ν a)) atTop by - rw [this, ← limsup_const (α := ℕ) (f := atTop) (∫ t in A, M κ ν a s n t ∂(ν a)), + = limsup (fun i ↦ ∫ t in A, densityProcess κ ν a s i t ∂(ν a)) atTop by + rw [this, ← limsup_const (α := ℕ) (f := atTop) (∫ t in A, densityProcess κ ν a s n t ∂(ν a)), limsup_congr] simp only [eventually_atTop, ge_iff_le] refine ⟨n, fun m hnm ↦ ?_⟩ - rw [set_integral_m_of_le hκν _ hs hnm hA, set_integral_m hκν _ hs _ hA] + rw [set_integral_densityProcess_of_le hκν _ hs hnm hA, + set_integral_densityProcess hκν _ hs _ hA] -- use L1 convergence have h := tendsto_set_integral_m hκν a hs A rw [h.limsup_eq] @@ -880,9 +880,9 @@ lemma tendsto_mLimsup_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin section UnivFst -lemma m_univ [IsFiniteKernel κ] (a : α) (n : ℕ) (t : ℝ) : - M κ (kernel.fst κ) a univ n t = if kernel.fst κ a (I n (indexI n t)) = 0 then 0 else 1 := by - rw [M] +lemma densityProcess_univ [IsFiniteKernel κ] (a : α) (n : ℕ) (t : ℝ) : + densityProcess κ (kernel.fst κ) a univ n t = if kernel.fst κ a (I n (indexI n t)) = 0 then 0 else 1 := by + rw [densityProcess] by_cases h : kernel.fst κ a (I n (indexI n t)) = 0 · simp [h] by_cases h' : κ a (I n (indexI n t) ×ˢ univ) = 0 @@ -899,14 +899,14 @@ lemma m_univ [IsFiniteKernel κ] (a : α) (n : ℕ) (t : ℝ) : · rwa [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h · exact measure_ne_top _ _ -lemma m_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ) : - ∀ᵐ t ∂(kernel.fst κ a), M κ (kernel.fst κ) a univ n t = 1 := by +lemma densityProcess_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ) : + ∀ᵐ t ∂(kernel.fst κ a), densityProcess κ (kernel.fst κ) a univ n t = 1 := by rw [ae_iff] - have : {t | ¬ M κ (kernel.fst κ) a univ n t = 1} + have : {t | ¬ densityProcess κ (kernel.fst κ) a univ n t = 1} ⊆ {t | kernel.fst κ a (I n (indexI n t)) = 0} := by intro t ht simp only [mem_setOf_eq] at ht ⊢ - rw [m_univ] at ht + rw [densityProcess_univ] at ht simpa using ht refine measure_mono_null this ?_ have : {t | kernel.fst κ a (I n (indexI n t)) = 0} @@ -920,11 +920,11 @@ lemma m_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : intro i simp -lemma tendsto_m_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) +lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) (t : ℝ) : - Tendsto (fun m ↦ M κ (kernel.fst κ) a (s m) n t) atTop - (𝓝 (M κ (kernel.fst κ) a univ n t)) := by - simp_rw [M] + Tendsto (fun m ↦ densityProcess κ (kernel.fst κ) a (s m) n t) atTop + (𝓝 (densityProcess κ (kernel.fst κ) a univ n t)) := by + simp_rw [densityProcess] refine (ENNReal.tendsto_toReal ?_).comp ?_ · rw [ne_eq, ENNReal.div_eq_top] push_neg @@ -957,16 +957,17 @@ lemma tendsto_m_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) rw [← prod_iUnion, hs_iUnion] · exact Or.inr h0 -lemma tendsto_m_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma tendsto_densityProcess_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) : - ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ M κ (kernel.fst κ) a (s m) n t) atTop (𝓝 1) := by - filter_upwards [m_univ_ae κ a n] with t ht + ∀ᵐ t ∂(kernel.fst κ a), + Tendsto (fun m ↦ densityProcess κ (kernel.fst κ) a (s m) n t) atTop (𝓝 1) := by + filter_upwards [densityProcess_univ_ae κ a n] with t ht rw [← ht] - exact tendsto_m_atTop_univ_of_monotone κ a s hs hs_iUnion n t + exact tendsto_densityProcess_atTop_univ_of_monotone κ a s hs hs_iUnion n t lemma mLimsup_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), MLimsup κ (kernel.fst κ) a Set.univ t = 1 := by - have h := m_univ_ae κ a + have h := densityProcess_univ_ae κ a rw [← ae_all_iff] at h filter_upwards [h] with t ht rw [MLimsup] diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index e9613a35e1f75..4d01df9284f3f 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -139,6 +139,9 @@ theorem tendsto_toReal {a : ℝ≥0∞} (ha : a ≠ ∞) : Tendsto ENNReal.toRea lemma continuousOn_toReal : ContinuousOn ENNReal.toReal { a | a ≠ ∞ } := NNReal.continuous_coe.comp_continuousOn continuousOn_toNNReal +lemma continuousAt_toReal (hx : x ≠ ∞) : ContinuousAt ENNReal.toReal x := + continuousOn_toReal.continuousAt (isOpen_ne_top.mem_nhds_iff.mpr hx) + /-- The set of finite `ℝ≥0∞` numbers is homeomorphic to `ℝ≥0`. -/ def neTopHomeomorphNNReal : { a | a ≠ ∞ } ≃ₜ ℝ≥0 where toEquiv := neTopEquivNNReal From b9d443b90668be9763afe981786adb627f87a23b Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sat, 17 Feb 2024 12:32:13 +0100 Subject: [PATCH 029/129] reorder arguments --- .../Kernel/Disintegration/KernelCDFBorel.lean | 412 +++++++++--------- 1 file changed, 216 insertions(+), 196 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean index eb72ecd36b4cf..c5013b8b21609 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean @@ -8,7 +8,22 @@ import Mathlib.Probability.Martingale.Convergence import Mathlib.Probability.Kernel.Disintegration.CDFKernel /-! -# KernelCDFBorel +# Kernel density + +Let `κ : kernel α (ℝ × β)` and `ν : kernel α ℝ` be two finite kernels with `kernel.fst κ ≤ ν`. +We build a function `f : α → ℝ → Set β → ℝ` jointly measurable in the first two arguments such that +for all `a : α` and all measurable sets `s : Set β` and `A : Set ℝ`, +`∫ t in A, f a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal`. + +If we were interested only in a fixed `a : α`, then we could use the Radon-Nikodym derivative to +build the function `f`, as follows +``` +def f_a (κ : kernel α (ℝ × β)) (ν : kernel a ℝ) (a : α) (t : ℝ) (s : Set β) : ℝ := + (((κ a).restrict (univ ×ˢ s)).fst.rnDeriv (ν a) t).toReal +``` +However, we can't turn those functions for each `a` into a measurable function of the pair `(a, t)`. + +TODO: what do we do then? ## Main definitions @@ -27,13 +42,11 @@ import Mathlib.Probability.Kernel.Disintegration.CDFKernel issue with the following: joint measurability -def M' (κ : kernel α (ℝ × β)) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : ℝ≥0∞ := - (((κ a).restrict (univ ×ˢ s)).fst.trim (ℱ.le n)).rnDeriv (((kernel.fst κ a)).trim (ℱ.le n)) t - ## References -* [F. Bar, *Quuxes*][bibkey] +The construction of the density process in this file follows the proof of Theorem 9.27 in +[O. Kallenberg, Foundations of modern probability][kallenberg2021]. ## Tags @@ -250,16 +263,17 @@ section M variable {κ : kernel α (ℝ × β)} {ν : kernel α ℝ} noncomputable -def densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : +def densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) (a : α) (t : ℝ) (s : Set β) : ℝ := (κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t))).toReal -lemma densityProcess_def (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) : - densityProcess κ ν a s n = fun t ↦ (κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t))).toReal := +lemma densityProcess_def (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) (a : α) (s : Set β) : + (fun t ↦ densityProcess κ ν n a t s) + = fun t ↦ (κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t))).toReal := rfl -lemma measurable_densityProcess_aux (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) - {s : Set β} (hs : MeasurableSet s) (n : ℕ) : +lemma measurable_densityProcess_aux (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) + {s : Set β} (hs : MeasurableSet s) : Measurable (fun (p : α × ℝ) ↦ κ p.1 (I n (indexI n p.2) ×ˢ s) / ν p.1 (I n (indexI n p.2))) := by change Measurable ((fun (p : α × ℤ) ↦ @@ -283,23 +297,24 @@ lemma measurable_densityProcess_aux (κ : kernel α (ℝ × β)) (ν : kernel α refine h1.comp (measurable_fst.prod_mk ?_) exact (Measurable.mono (measurable_indexI n) (ℱ.le n) le_rfl).comp measurable_snd -lemma measurable_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) - {s : Set β} (hs : MeasurableSet s) (n : ℕ) : - Measurable (fun (p : α × ℝ) ↦ densityProcess κ ν p.1 s n p.2) := - (measurable_densityProcess_aux κ ν hs n).ennreal_toReal - -lemma measurable_densityProcess_left (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) - {s : Set β} (hs : MeasurableSet s) (n : ℕ) (t : ℝ) : - Measurable (fun a ↦ densityProcess κ ν a s n t) := - (measurable_densityProcess κ ν hs n).comp (measurable_id.prod_mk measurable_const) - -lemma measurable_densityProcess_right (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) - {s : Set β} (a : α) (hs : MeasurableSet s) (n : ℕ) : - Measurable (densityProcess κ ν a s n) := - (measurable_densityProcess κ ν hs n).comp (measurable_const.prod_mk measurable_id) - -lemma measurable_ℱ_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) : - Measurable[ℱ n] (densityProcess κ ν a s n) := by +lemma measurable_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) + {s : Set β} (hs : MeasurableSet s) : + Measurable (fun (p : α × ℝ) ↦ densityProcess κ ν n p.1 p.2 s) := + (measurable_densityProcess_aux κ ν n hs).ennreal_toReal + +lemma measurable_densityProcess_left (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) + (t : ℝ) {s : Set β} (hs : MeasurableSet s) : + Measurable (fun a ↦ densityProcess κ ν n a t s) := + (measurable_densityProcess κ ν n hs).comp (measurable_id.prod_mk measurable_const) + +lemma measurable_densityProcess_right (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) + {s : Set β} (a : α) (hs : MeasurableSet s) : + Measurable (fun t ↦ densityProcess κ ν n a t s) := + (measurable_densityProcess κ ν n hs).comp (measurable_const.prod_mk measurable_id) + +lemma measurable_ℱ_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) + (a : α) (s : Set β) : + Measurable[ℱ n] (fun t ↦ densityProcess κ ν n a t s) := by rw [densityProcess_def] refine @Measurable.ennreal_toReal _ (ℱ n) _ ?_ refine Measurable.div ?_ ?_ @@ -310,20 +325,22 @@ lemma measurable_ℱ_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α refine Measurable.comp ?_ (measurable_indexI n) exact measurable_of_countable _ -lemma stronglyMeasurable_ℱ_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) : - StronglyMeasurable[ℱ n] (densityProcess κ ν a s n) := - (measurable_ℱ_densityProcess κ ν a s n).stronglyMeasurable +lemma stronglyMeasurable_ℱ_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) + (a : α) (s : Set β) : + StronglyMeasurable[ℱ n] (fun t ↦ densityProcess κ ν n a t s) := + (measurable_ℱ_densityProcess κ ν n a s).stronglyMeasurable lemma adapted_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) : - Adapted ℱ (densityProcess κ ν a s) := - stronglyMeasurable_ℱ_densityProcess κ ν a s + Adapted ℱ (fun n t ↦ densityProcess κ ν n a t s) := + fun n ↦ stronglyMeasurable_ℱ_densityProcess κ ν n a s -lemma densityProcess_nonneg (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : - 0 ≤ densityProcess κ ν a s n t := +lemma densityProcess_nonneg (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) + (a : α) (t : ℝ) (s : Set β) : + 0 ≤ densityProcess κ ν n a t s := ENNReal.toReal_nonneg -lemma densityProcess_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (n : ℕ) (t : ℝ) : - densityProcess κ ν a s n t ≤ 1 := by +lemma densityProcess_le_one (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : ℝ) (s : Set β) : + densityProcess κ ν n a t s ≤ 1 := by refine ENNReal.toReal_le_of_le_ofReal zero_le_one (ENNReal.div_le_of_le_mul ?_) rw [ENNReal.ofReal_one, one_mul] calc κ a (I n (indexI n t) ×ˢ s) ≤ kernel.fst κ a (I n (indexI n t)) := by @@ -333,29 +350,29 @@ lemma densityProcess_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) exact fun h _ ↦ h _ ≤ ν a (I n (indexI n t)) := hκν a _ (measurableSet_I _ _) -lemma snorm_densityProcess_le (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (n : ℕ) : - snorm (densityProcess κ ν a s n) 1 (ν a) ≤ ν a univ := by +lemma snorm_densityProcess_le (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (s : Set β) : + snorm (fun t ↦ densityProcess κ ν n a t s) 1 (ν a) ≤ ν a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun x ↦ ?_))).trans ?_ - · simp only [Real.norm_eq_abs, abs_of_nonneg (densityProcess_nonneg κ ν a s n x), - densityProcess_le_one hκν a s n x] + · simp only [Real.norm_eq_abs, abs_of_nonneg (densityProcess_nonneg κ ν n a x s), + densityProcess_le_one hκν n a x s] · simp -lemma integrable_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] - (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : - Integrable (densityProcess κ ν a s n) (ν a) := by +lemma integrable_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] (n : ℕ) + (a : α) {s : Set β} (hs : MeasurableSet s) : + Integrable (fun t ↦ densityProcess κ ν n a t s) (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ - · exact measurable_densityProcess_right κ ν a hs n - · exact (snorm_densityProcess_le hκν a s n).trans_lt (measure_lt_top _ _) + · exact measurable_densityProcess_right κ ν n a hs + · exact (snorm_densityProcess_le hκν n a s).trans_lt (measure_lt_top _ _) lemma set_integral_densityProcess_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) (k : ℤ) : - ∫ t in I n k, densityProcess κ ν a s n t ∂(ν a) = (κ a (I n k ×ˢ s)).toReal := by + (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) (k : ℤ) : + ∫ t in I n k, densityProcess κ ν n a t s ∂(ν a) = (κ a (I n k ×ˢ s)).toReal := by simp_rw [densityProcess] rw [integral_toReal] rotate_left · refine Measurable.aemeasurable ?_ - have h := measurable_densityProcess_aux κ ν hs n + have h := measurable_densityProcess_aux κ ν n hs change Measurable ((fun (p : α × ℝ) ↦ κ p.1 (I n (indexI n p.2) ×ˢ s) / ν p.1 (I n (indexI n p.2))) ∘ (fun t ↦ (a, t))) @@ -391,8 +408,8 @@ lemma set_integral_densityProcess_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKern exact measure_ne_top _ _ lemma integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) : - ∫ t, densityProcess κ ν a s n t ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by + (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : + ∫ t, densityProcess κ ν n a t s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by rw [← integral_univ, ← iUnion_I n, iUnion_prod_const, measure_iUnion] rotate_left · intro i j hij @@ -400,16 +417,16 @@ lemma integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] exact Or.inl (pairwise_disjoint_I n hij) · exact fun k ↦ (measurableSet_I n k).prod hs rw [integral_iUnion (measurableSet_I n) (pairwise_disjoint_I n) - (integrable_densityProcess hκν a hs n).integrableOn] + (integrable_densityProcess hκν n a hs).integrableOn] rw [ENNReal.tsum_toReal_eq (fun _ ↦ measure_ne_top _ _)] congr with k - rw [set_integral_densityProcess_I hκν _ hs] + rw [set_integral_densityProcess_I hκν _ _ hs] lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : - ∫ t in A, densityProcess κ ν a s n t ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : + ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by refine MeasurableSpace.induction_on_inter (m := ℱ n) (s := {s | ∃ k, s = I n k}) - (C := fun A ↦ ∫ t in A, densityProcess κ ν a s n t ∂(ν a) = (κ a (A ×ˢ s)).toReal) rfl + (C := fun A ↦ ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) rfl ?_ ?_ ?_ ?_ ?_ hA · rintro s ⟨i, rfl⟩ t ⟨j, rfl⟩ hst refine ⟨i, ?_⟩ @@ -421,11 +438,11 @@ lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel rwa [Function.onFun, disjoint_iff_inter_eq_empty] at h_disj · simp · rintro _ ⟨k, rfl⟩ - rw [set_integral_densityProcess_I hκν _ hs] + rw [set_integral_densityProcess_I hκν _ _ hs] · intro A hA hA_eq have hA' : MeasurableSet A := ℱ.le _ _ hA - have h := integral_add_compl hA' (integrable_densityProcess hκν a hs n) - rw [hA_eq, integral_densityProcess hκν a hs] at h + have h := integral_add_compl hA' (integrable_densityProcess hκν n a hs) + rw [hA_eq, integral_densityProcess hκν n a hs] at h have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by rw [prod_diff_prod, compl_eq_univ_diff] simp @@ -437,7 +454,7 @@ lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel exact measure_mono (by intro x; simp) · intro f hf_disj hf h_eq rw [integral_iUnion (fun i ↦ ℱ.le n _ (hf i)) hf_disj - (integrable_densityProcess hκν _ hs _).integrableOn] + (integrable_densityProcess hκν _ _ hs).integrableOn] simp_rw [h_eq] rw [iUnion_prod_const, measure_iUnion _ (fun i ↦ (ℱ.le n _ (hf i)).prod hs)] · rw [ENNReal.tsum_toReal_eq] @@ -446,33 +463,34 @@ lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel rw [Function.onFun, Set.disjoint_prod] exact Or.inl (hf_disj hij) -lemma set_integral_densityProcess_of_le (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) {s : Set β} (hs : MeasurableSet s) {n m : ℕ} (hnm : n ≤ m) +lemma set_integral_densityProcess_of_le (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] {n m : ℕ} (hnm : n ≤ m) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : - ∫ t in A, densityProcess κ ν a s m t ∂(ν a) = (κ a (A ×ˢ s)).toReal := - set_integral_densityProcess hκν a hs m (ℱ.mono hnm A hA) + ∫ t in A, densityProcess κ ν m a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := + set_integral_densityProcess hκν m a hs (ℱ.mono hnm A hA) lemma condexp_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) {s : Set β} - (hs : MeasurableSet s) {i j : ℕ} (hij : i ≤ j) : - (ν a)[densityProcess κ ν a s j | ℱ i] =ᵐ[ν a] densityProcess κ ν a s i := by + {i j : ℕ} (hij : i ≤ j) (a : α) {s : Set β} (hs : MeasurableSet s) : + (ν a)[fun t ↦ densityProcess κ ν j a t s | ℱ i] + =ᵐ[ν a] fun t ↦ densityProcess κ ν i a t s := by symm refine ae_eq_condexp_of_forall_set_integral_eq ?_ ?_ ?_ ?_ ?_ - · exact integrable_densityProcess hκν a hs j + · exact integrable_densityProcess hκν j a hs · refine fun t _ _ ↦ Integrable.integrableOn ?_ - exact integrable_densityProcess hκν _ hs _ + exact integrable_densityProcess hκν _ _ hs · intro t ht _ - rw [set_integral_densityProcess hκν a hs i ht, set_integral_densityProcess_of_le hκν a hs hij ht] - · exact StronglyMeasurable.aeStronglyMeasurable' (stronglyMeasurable_ℱ_densityProcess κ ν a s i) + rw [set_integral_densityProcess hκν i a hs ht, + set_integral_densityProcess_of_le hκν hij a hs ht] + · exact StronglyMeasurable.aeStronglyMeasurable' (stronglyMeasurable_ℱ_densityProcess κ ν i a s) lemma martingale_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Martingale (densityProcess κ ν a s) ℱ (ν a) := - ⟨adapted_densityProcess κ ν a s, fun _ _ ↦ condexp_densityProcess hκν a hs⟩ + Martingale (fun n t ↦ densityProcess κ ν n a t s) ℱ (ν a) := + ⟨adapted_densityProcess κ ν a s, fun _ _ h ↦ condexp_densityProcess hκν h a hs⟩ -lemma densityProcess_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) {s s' : Set β} (h : s ⊆ s') - (n : ℕ) (t : ℝ) : - densityProcess κ ν a s n t ≤ densityProcess κ ν a s' n t := by +lemma densityProcess_mono_set (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : ℝ) + {s s' : Set β} (h : s ⊆ s') : + densityProcess κ ν n a t s ≤ densityProcess κ ν n a t s' := by unfold densityProcess by_cases h0 : ν a (I n (indexI n t)) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] @@ -496,8 +514,8 @@ lemma densityProcess_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) {s s' : Se · exact h_ne_top s' lemma densityProcess_mono_kernel_left {κ' : kernel α (ℝ × β)} (hκκ' : κ ≤ κ') - (hκ'ν : kernel.fst κ' ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) (t : ℝ) : - densityProcess κ ν a s n t ≤ densityProcess κ' ν a s n t := by + (hκ'ν : kernel.fst κ' ≤ ν) (n : ℕ) (a : α) (t : ℝ) {s : Set β} (hs : MeasurableSet s) : + densityProcess κ ν n a t s ≤ densityProcess κ' ν n a t s := by unfold densityProcess by_cases h0 : ν a (I n (indexI n t)) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] @@ -520,9 +538,9 @@ lemma densityProcess_mono_kernel_left {κ' : kernel α (ℝ × β)} (hκκ' : κ simp only [ne_eq, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top -lemma densityProcess_antitone_kernel_right {ν' : kernel α ℝ} (hνν' : ν ≤ ν') (hκν : kernel.fst κ ≤ ν) - (a : α) (s : Set β) (n : ℕ) (t : ℝ) : - densityProcess κ ν' a s n t ≤ densityProcess κ ν a s n t := by +lemma densityProcess_antitone_kernel_right {ν' : kernel α ℝ} + (hνν' : ν ≤ ν') (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : ℝ) (s : Set β) : + densityProcess κ ν' n a t s ≤ densityProcess κ ν n a t s := by unfold densityProcess have h_le : κ a (I n (indexI n t) ×ˢ s) ≤ ν a (I n (indexI n t)) := by calc κ a (I n (indexI n t) ×ˢ s) ≤ kernel.fst κ a (I n (indexI n t)) := by @@ -547,14 +565,15 @@ lemma densityProcess_antitone_kernel_right {ν' : kernel α ℝ} (hνν' : ν · simp only [ne_eq, ENNReal.div_eq_top, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top -lemma densityProcess_empty (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (n : ℕ) (t : ℝ) : - densityProcess κ ν a ∅ n t = 0 := by +lemma densityProcess_empty (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) (a : α) (t : ℝ) : + densityProcess κ ν n a t ∅ = 0 := by simp [densityProcess] -lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) [IsFiniteKernel κ] - (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) - (hs_meas : ∀ n, MeasurableSet (s n)) (n : ℕ) (t : ℝ) : - Tendsto (fun m ↦ densityProcess κ ν a (s m) n t) atTop (𝓝 (densityProcess κ ν a ∅ n t)) := by +lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) + [IsFiniteKernel κ] (n : ℕ) (a : α) (t : ℝ) + (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) + (hs_meas : ∀ n, MeasurableSet (s n)) : + Tendsto (fun m ↦ densityProcess κ ν n a t (s m)) atTop (𝓝 (densityProcess κ ν n a t ∅)) := by simp_rw [densityProcess] by_cases h0 : ν a (I n (indexI n t)) = 0 · simp_rw [h0, ENNReal.toReal_div] @@ -575,45 +594,46 @@ lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (ℝ × β) · simp only [prod_empty, OuterMeasure.empty', ne_eq, not_true_eq_false, false_or, h0, not_false_iff] -lemma tendsto_densityProcess_atTop_of_antitone (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) [IsFiniteKernel κ] - (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) - (hs_meas : ∀ n, MeasurableSet (s n)) (n : ℕ) (t : ℝ) : - Tendsto (fun m ↦ densityProcess κ ν a (s m) n t) atTop (𝓝 0) := by - rw [← densityProcess_empty κ ν a n t] - exact tendsto_densityProcess_atTop_empty_of_antitone κ ν a s hs hs_iInter hs_meas n t +lemma tendsto_densityProcess_atTop_of_antitone (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) + [IsFiniteKernel κ] (n : ℕ) (a : α) (t : ℝ) + (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) + (hs_meas : ∀ n, MeasurableSet (s n)) : + Tendsto (fun m ↦ densityProcess κ ν n a t (s m)) atTop (𝓝 0) := by + rw [← densityProcess_empty κ ν n a t] + exact tendsto_densityProcess_atTop_empty_of_antitone κ ν n a t s hs hs_iInter hs_meas -lemma tendsto_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) {s : Set β} (hs : MeasurableSet s) : - ∀ᵐ t ∂(ν a), Tendsto (fun n ↦ densityProcess κ ν a s n t) atTop - (𝓝 (ℱ.limitProcess (densityProcess κ ν a s) (ν a) t)) := by +lemma tendsto_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : + ∀ᵐ t ∂(ν a), Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop + (𝓝 (ℱ.limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a) t)) := by refine Submartingale.ae_tendsto_limitProcess (martingale_densityProcess hκν a hs).submartingale (R := (ν a univ).toNNReal) (fun n ↦ ?_) - refine (snorm_densityProcess_le hκν a s n).trans_eq ?_ + refine (snorm_densityProcess_le hκν n a s).trans_eq ?_ rw [ENNReal.coe_toNNReal] exact measure_ne_top _ _ lemma limitProcess_mem_L1 (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Memℒp (ℱ.limitProcess (densityProcess κ ν a s) (ν a)) 1 (ν a) := by + Memℒp (ℱ.limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 (ν a) := by refine Submartingale.memℒp_limitProcess (martingale_densityProcess hκν a hs).submartingale (R := (ν a univ).toNNReal) (fun n ↦ ?_) - refine (snorm_densityProcess_le hκν a s n).trans_eq ?_ + refine (snorm_densityProcess_le hκν n a s).trans_eq ?_ rw [ENNReal.coe_toNNReal] exact measure_ne_top _ _ lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Tendsto - (fun n ↦ snorm (densityProcess κ ν a s n - ℱ.limitProcess (densityProcess κ ν a s) (ν a)) 1 (ν a)) + Tendsto (fun n ↦ snorm ((fun t ↦ densityProcess κ ν n a t s) + - ℱ.limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 (ν a)) atTop (𝓝 0) := by refine Submartingale.tendsto_snorm_one_limitProcess ?_ ?_ · exact (martingale_densityProcess hκν a hs).submartingale · refine uniformIntegrable_of le_rfl ENNReal.one_ne_top ?_ ?_ - · exact fun n ↦ (measurable_densityProcess_right κ ν a hs n).aestronglyMeasurable + · exact fun n ↦ (measurable_densityProcess_right κ ν n a hs).aestronglyMeasurable · intro ε _ refine ⟨2, fun n ↦ ?_⟩ refine le_of_eq_of_le ?_ (?_ : 0 ≤ ENNReal.ofReal ε) - · have : {x | 2 ≤ ‖densityProcess κ ν a s n x‖₊} = ∅ := by + · have : {x | 2 ≤ ‖densityProcess κ ν n a x s‖₊} = ∅ := by ext x simp only [mem_setOf_eq, mem_empty_iff_false, iff_false, not_le] refine (?_ : _ ≤ (1 : ℝ≥0)).trans_lt one_lt_two @@ -625,74 +645,77 @@ lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : kernel.fst κ ≤ lemma tendsto_snorm_one_restrict_densityProcess_limitProcess [IsFiniteKernel κ] [IsFiniteKernel ν] (hκν : kernel.fst κ ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : - Tendsto (fun n ↦ snorm (densityProcess κ ν a s n - ℱ.limitProcess (densityProcess κ ν a s) (ν a)) 1 ((ν a).restrict A)) + Tendsto (fun n ↦ snorm ((fun t ↦ densityProcess κ ν n a t s) + - ℱ.limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 ((ν a).restrict A)) atTop (𝓝 0) := tendsto_snorm_restrict_zero (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) A noncomputable -def MLimsup (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) (t : ℝ) : ℝ := - limsup (fun n ↦ densityProcess κ ν a s n t) atTop +def MLimsup (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (t : ℝ) (s : Set β) : ℝ := + limsup (fun n ↦ densityProcess κ ν n a t s) atTop lemma mLimsup_ae_eq_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - MLimsup κ ν a s =ᵐ[ν a] ℱ.limitProcess (densityProcess κ ν a s) (ν a) := by + (fun t ↦ MLimsup κ ν a t s) + =ᵐ[ν a] ℱ.limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a) := by filter_upwards [tendsto_densityProcess_limitProcess hκν a hs] with t ht using ht.limsup_eq lemma tendsto_m_mLimsup (hκν : kernel.fst κ ≤ ν) (a : α) [IsFiniteKernel κ] [IsFiniteKernel ν] {s : Set β} (hs : MeasurableSet s) : ∀ᵐ t ∂(ν a), - Tendsto (fun n ↦ densityProcess κ ν a s n t) atTop (𝓝 (MLimsup κ ν a s t)) := by - filter_upwards [tendsto_densityProcess_limitProcess hκν a hs, mLimsup_ae_eq_limitProcess hκν a hs] with t h1 h2 + Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop (𝓝 (MLimsup κ ν a t s)) := by + filter_upwards [tendsto_densityProcess_limitProcess hκν a hs, mLimsup_ae_eq_limitProcess hκν a hs] + with t h1 h2 rw [h2] exact h1 lemma measurable_mLimsup (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) {s : Set β} (hs : MeasurableSet s) : - Measurable (fun (p : α × ℝ) ↦ MLimsup κ ν p.1 s p.2) := - measurable_limsup (fun n ↦ measurable_densityProcess κ ν hs n) + Measurable (fun (p : α × ℝ) ↦ MLimsup κ ν p.1 p.2 s) := + measurable_limsup (fun n ↦ measurable_densityProcess κ ν n hs) -lemma measurable_mLimsup_left (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) - {s : Set β} (hs : MeasurableSet s) (t : ℝ) : - Measurable (fun a ↦ MLimsup κ ν a s t) := by - change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ ν p.1 s p.2) ∘ (fun a ↦ (a, t))) +lemma measurable_mLimsup_left (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (t : ℝ) + {s : Set β} (hs : MeasurableSet s) : + Measurable (fun a ↦ MLimsup κ ν a t s) := by + change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ ν p.1 p.2 s) ∘ (fun a ↦ (a, t))) exact (measurable_mLimsup κ ν hs).comp (measurable_id.prod_mk measurable_const) lemma measurable_mLimsup_right (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) {s : Set β} (hs : MeasurableSet s) (a : α) : - Measurable (MLimsup κ ν a s) := by - change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ ν p.1 s p.2) ∘ (fun t ↦ (a, t))) + Measurable (fun t ↦ MLimsup κ ν a t s) := by + change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ ν p.1 p.2 s) ∘ (fun t ↦ (a, t))) exact (measurable_mLimsup κ ν hs).comp (measurable_const.prod_mk measurable_id) -lemma mLimsup_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) {s s' : Set β} (h : s ⊆ s') (t : ℝ) : - MLimsup κ ν a s t ≤ MLimsup κ ν a s' t := by +lemma mLimsup_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) {s s' : Set β} (h : s ⊆ s') : + MLimsup κ ν a t s ≤ MLimsup κ ν a t s' := by rw [MLimsup, MLimsup] refine limsup_le_limsup ?_ ?_ ?_ - · exact eventually_of_forall (fun n ↦ densityProcess_mono_set hκν a h n t) + · exact eventually_of_forall (fun n ↦ densityProcess_mono_set hκν n a t h) · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ -lemma mLimsup_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (t : ℝ) : - 0 ≤ MLimsup κ ν a s t := by +lemma mLimsup_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (s : Set β) : + 0 ≤ MLimsup κ ν a t s := by refine le_limsup_of_frequently_le ?_ ?_ · exact frequently_of_forall (fun n ↦ densityProcess_nonneg _ _ _ _ _ _) · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ -lemma mLimsup_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) (t : ℝ) : - MLimsup κ ν a s t ≤ 1 := by +lemma mLimsup_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (s : Set β) : + MLimsup κ ν a t s ≤ 1 := by refine limsup_le_of_le ?_ ?_ · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) · exact eventually_of_forall (fun n ↦ densityProcess_le_one hκν _ _ _ _) lemma snorm_mLimsup_le (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) : - snorm (fun t ↦ MLimsup κ ν a s t) 1 (ν a) ≤ ν a univ := by + snorm (fun t ↦ MLimsup κ ν a t s) 1 (ν a) ≤ ν a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ - · simp only [Real.norm_eq_abs, abs_of_nonneg (mLimsup_nonneg hκν a s t), - mLimsup_le_one hκν a s t] + · simp only [Real.norm_eq_abs, abs_of_nonneg (mLimsup_nonneg hκν a t s), + mLimsup_le_one hκν a t s] · simp lemma integrable_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Integrable (fun t ↦ MLimsup κ ν a s t) (ν a) := by + Integrable (fun t ↦ MLimsup κ ν a t s) (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ · exact measurable_mLimsup_right κ ν hs a @@ -700,11 +723,11 @@ lemma integrable_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] lemma tendsto_set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : - Tendsto (fun i ↦ ∫ x in A, densityProcess κ ν a s i x ∂(ν a)) atTop - (𝓝 (∫ x in A, MLimsup κ ν a s x ∂(ν a))) := by - refine tendsto_set_integral_of_L1' (μ := ν a) (MLimsup κ ν a s) - (integrable_mLimsup hκν a hs) (F := fun i t ↦ densityProcess κ ν a s i t) (l := atTop) - (eventually_of_forall (fun n ↦ integrable_densityProcess hκν _ hs _)) ?_ A + Tendsto (fun i ↦ ∫ x in A, densityProcess κ ν i a x s ∂(ν a)) atTop + (𝓝 (∫ x in A, MLimsup κ ν a x s ∂(ν a))) := by + refine tendsto_set_integral_of_L1' (μ := ν a) (fun t ↦ MLimsup κ ν a t s) + (integrable_mLimsup hκν a hs) (F := fun i t ↦ densityProcess κ ν i a t s) (l := atTop) + (eventually_of_forall (fun n ↦ integrable_densityProcess hκν _ _ hs)) ?_ A refine (tendsto_congr fun n ↦ ?_).mp (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) refine snorm_congr_ae ?_ refine EventuallyEq.sub EventuallyEq.rfl ?_ @@ -712,40 +735,40 @@ lemma tendsto_set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] lemma set_integral_mLimsup_of_measurableSet (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) {s : Set β} (hs : MeasurableSet s) (n : ℕ) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : - ∫ t in A, MLimsup κ ν a s t ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - suffices ∫ t in A, MLimsup κ ν a s t ∂(ν a) = ∫ t in A, densityProcess κ ν a s n t ∂(ν a) by + (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : + ∫ t in A, MLimsup κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + suffices ∫ t in A, MLimsup κ ν a t s ∂(ν a) = ∫ t in A, densityProcess κ ν n a t s ∂(ν a) by rw [this] - exact set_integral_densityProcess hκν _ hs _ hA - suffices ∫ t in A, MLimsup κ ν a s t ∂(ν a) - = limsup (fun i ↦ ∫ t in A, densityProcess κ ν a s i t ∂(ν a)) atTop by - rw [this, ← limsup_const (α := ℕ) (f := atTop) (∫ t in A, densityProcess κ ν a s n t ∂(ν a)), + exact set_integral_densityProcess hκν _ _ hs hA + suffices ∫ t in A, MLimsup κ ν a t s ∂(ν a) + = limsup (fun i ↦ ∫ t in A, densityProcess κ ν i a t s ∂(ν a)) atTop by + rw [this, ← limsup_const (α := ℕ) (f := atTop) (∫ t in A, densityProcess κ ν n a t s ∂(ν a)), limsup_congr] simp only [eventually_atTop, ge_iff_le] refine ⟨n, fun m hnm ↦ ?_⟩ - rw [set_integral_densityProcess_of_le hκν _ hs hnm hA, - set_integral_densityProcess hκν _ hs _ hA] + rw [set_integral_densityProcess_of_le hκν hnm _ hs hA, + set_integral_densityProcess hκν _ _ hs hA] -- use L1 convergence have h := tendsto_set_integral_m hκν a hs A rw [h.limsup_eq] lemma integral_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - ∫ t, MLimsup κ ν a s t ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by - rw [← integral_univ, set_integral_mLimsup_of_measurableSet hκν a hs 0 MeasurableSet.univ] + ∫ t, MLimsup κ ν a t s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by + rw [← integral_univ, set_integral_mLimsup_of_measurableSet hκν 0 a hs MeasurableSet.univ] lemma set_integral_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set ℝ} (hA : MeasurableSet A) : - ∫ t in A, MLimsup κ ν a s t ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + ∫ t in A, MLimsup κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by have hA' : MeasurableSet[⨆ n, ℱ n] A := by rwa [iSup_ℱ] refine MeasurableSpace.induction_on_inter (m := ⨆ n, ℱ n) - (C := fun A ↦ ∫ t in A, MLimsup κ ν a s t ∂(ν a) = (κ a (A ×ˢ s)).toReal) + (C := fun A ↦ ∫ t in A, MLimsup κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) (MeasurableSpace.measurableSpace_iSup_eq ℱ) ?_ ?_ ?_ ?_ ?_ hA' · rintro s ⟨n, hs⟩ t ⟨m, ht⟩ _ exact ⟨max n m, (ℱ.mono (le_max_left n m) _ hs).inter (ℱ.mono (le_max_right n m) _ ht)⟩ · simp · intro A ⟨n, hA⟩ - exact set_integral_mLimsup_of_measurableSet hκν a hs n hA + exact set_integral_mLimsup_of_measurableSet hκν n a hs hA · intro A hA hA_eq rw [iSup_ℱ] at hA have h := integral_add_compl hA (integrable_mLimsup hκν a hs) @@ -776,7 +799,7 @@ lemma tendsto_integral_mLimsup_of_monotone (hκν : kernel.fst κ ≤ ν) [IsFin [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : - Tendsto (fun m ↦ ∫ t, MLimsup κ ν a (s m) t ∂(ν a)) atTop (𝓝 (κ a univ).toReal) := by + Tendsto (fun m ↦ ∫ t, MLimsup κ ν a t (s m) ∂(ν a)) atTop (𝓝 (κ a univ).toReal) := by simp_rw [integral_mLimsup hκν a (hs_meas _)] have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := κ a univ) ?_ swap @@ -796,7 +819,7 @@ lemma tendsto_integral_mLimsup_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : - Tendsto (fun m ↦ ∫ t, MLimsup κ ν a (s m) t ∂(ν a)) atTop (𝓝 0) := by + Tendsto (fun m ↦ ∫ t, MLimsup κ ν a t (s m) ∂(ν a)) atTop (𝓝 0) := by simp_rw [integral_mLimsup hκν a (hs_meas _)] rw [← ENNReal.zero_toReal] have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := 0) ?_ @@ -816,31 +839,30 @@ lemma tendsto_integral_mLimsup_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin simp only [ne_eq, prod_empty, OuterMeasure.empty', forall_exists_index] lemma tendsto_mLimsup_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] - [IsFiniteKernel ν] - (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) + [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : - ∀ᵐ t ∂(ν a), Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 0) := by - have h_anti : ∀ t, Antitone (fun m ↦ MLimsup κ ν a (s m) t) := - fun t n m hnm ↦ mLimsup_mono_set hκν a (hs hnm) t - have h_le_one : ∀ m t, MLimsup κ ν a (s m) t ≤ 1 := fun m t ↦ mLimsup_le_one hκν a (s m) t + ∀ᵐ t ∂(ν a), Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 0) := by + have h_anti : ∀ t, Antitone (fun m ↦ MLimsup κ ν a t (s m)) := + fun t n m hnm ↦ mLimsup_mono_set hκν a t (hs hnm) + have h_le_one : ∀ m t, MLimsup κ ν a t (s m) ≤ 1 := fun m t ↦ mLimsup_le_one hκν a t (s m) -- for all `t`, `fun m ↦ MLimsup κ a (s m) t` has a limit - have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 l) := by + have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 l) := by intro t - have h_tendsto : Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop atBot ∨ - ∃ l, Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 l) := + have h_tendsto : Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop atBot ∨ + ∃ l, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 l) := tendsto_of_antitone (h_anti t) cases' h_tendsto with h_absurd h_tendsto · rw [tendsto_atTop_atBot_iff_of_antitone (h_anti t)] at h_absurd obtain ⟨r, hr⟩ := h_absurd (-1) - have h_nonneg := mLimsup_nonneg hκν a (s r) t + have h_nonneg := mLimsup_nonneg hκν a t (s r) linarith · exact h_tendsto -- let `F` be the pointwise limit of `fun m ↦ MLimsup κ a (s m) t` for all `t` let F : ℝ → ℝ := fun t ↦ (h_exists t).choose - have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 (F t)) := + have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 (F t)) := fun t ↦ (h_exists t).choose_spec have hF_nonneg : ∀ t, 0 ≤ F t := - fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ mLimsup_nonneg hκν a (s m) t) + fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ mLimsup_nonneg hκν a t (s m)) have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) have hF_int : Integrable F (ν a) := by rw [← memℒp_one_iff_integrable] @@ -866,22 +888,22 @@ lemma tendsto_mLimsup_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin suffices ∀ᵐ (t : ℝ) ∂(ν a), 0 = F t by filter_upwards [this] with a ha; simp [ha] refine ae_eq_of_integral_eq_of_ae_le (integrable_const _) hF_int (ae_of_all _ hF_nonneg) ?_ have h_integral : - Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a (s m) t ∂(ν a)) atTop (𝓝 (∫ t, F t ∂(ν a))) := by + Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ t, F t ∂(ν a))) := by refine integral_tendsto_of_tendsto_of_antitone ?_ hF_int ?_ ?_ · exact fun n ↦ integrable_mLimsup hκν _ (hs_meas n) · exact ae_of_all _ h_anti · exact ae_of_all _ hF_tendsto have h_integral' : - Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a (s m) t ∂(ν a)) atTop - (𝓝 (∫ _, 0 ∂(ν a))) := by + Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ _, 0 ∂(ν a))) := by simp only [integral_zero] exact tendsto_integral_mLimsup_of_antitone hκν a s hs hs_iInter hs_meas exact (tendsto_nhds_unique h_integral h_integral').symm section UnivFst -lemma densityProcess_univ [IsFiniteKernel κ] (a : α) (n : ℕ) (t : ℝ) : - densityProcess κ (kernel.fst κ) a univ n t = if kernel.fst κ a (I n (indexI n t)) = 0 then 0 else 1 := by +lemma densityProcess_univ [IsFiniteKernel κ] (n : ℕ) (a : α) (t : ℝ) : + densityProcess κ (kernel.fst κ) n a t univ + = if kernel.fst κ a (I n (indexI n t)) = 0 then 0 else 1 := by rw [densityProcess] by_cases h : kernel.fst κ a (I n (indexI n t)) = 0 · simp [h] @@ -899,10 +921,10 @@ lemma densityProcess_univ [IsFiniteKernel κ] (a : α) (n : ℕ) (t : ℝ) : · rwa [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h · exact measure_ne_top _ _ -lemma densityProcess_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) (n : ℕ) : - ∀ᵐ t ∂(kernel.fst κ a), densityProcess κ (kernel.fst κ) a univ n t = 1 := by +lemma densityProcess_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (n : ℕ) (a : α) : + ∀ᵐ t ∂(kernel.fst κ a), densityProcess κ (kernel.fst κ) n a t univ = 1 := by rw [ae_iff] - have : {t | ¬ densityProcess κ (kernel.fst κ) a univ n t = 1} + have : {t | ¬ densityProcess κ (kernel.fst κ) n a t univ = 1} ⊆ {t | kernel.fst κ a (I n (indexI n t)) = 0} := by intro t ht simp only [mem_setOf_eq] at ht ⊢ @@ -921,9 +943,9 @@ lemma densityProcess_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a simp lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) - (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) (t : ℝ) : - Tendsto (fun m ↦ densityProcess κ (kernel.fst κ) a (s m) n t) atTop - (𝓝 (densityProcess κ (kernel.fst κ) a univ n t)) := by + (n : ℕ) (a : α) (t : ℝ) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) : + Tendsto (fun m ↦ densityProcess κ (kernel.fst κ) n a t (s m)) atTop + (𝓝 (densityProcess κ (kernel.fst κ) n a t univ)) := by simp_rw [densityProcess] refine (ENNReal.tendsto_toReal ?_).comp ?_ · rw [ne_eq, ENNReal.div_eq_top] @@ -958,16 +980,16 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) · exact Or.inr h0 lemma tendsto_densityProcess_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] - (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (n : ℕ) : + (n : ℕ) (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) : ∀ᵐ t ∂(kernel.fst κ a), - Tendsto (fun m ↦ densityProcess κ (kernel.fst κ) a (s m) n t) atTop (𝓝 1) := by - filter_upwards [densityProcess_univ_ae κ a n] with t ht + Tendsto (fun m ↦ densityProcess κ (kernel.fst κ) n a t (s m)) atTop (𝓝 1) := by + filter_upwards [densityProcess_univ_ae κ n a] with t ht rw [← ht] - exact tendsto_densityProcess_atTop_univ_of_monotone κ a s hs hs_iUnion n t + exact tendsto_densityProcess_atTop_univ_of_monotone κ n a t s hs hs_iUnion lemma mLimsup_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), MLimsup κ (kernel.fst κ) a Set.univ t = 1 := by - have h := densityProcess_univ_ae κ a + ∀ᵐ t ∂(kernel.fst κ a), MLimsup κ (kernel.fst κ) a t univ = 1 := by + have h := fun n ↦ densityProcess_univ_ae κ n a rw [← ae_all_iff] at h filter_upwards [h] with t ht rw [MLimsup] @@ -976,16 +998,16 @@ lemma mLimsup_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) : lemma tendsto_mLimsup_atTop_ae_of_monotone [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : - ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ (kernel.fst κ) a (s m) t) atTop (𝓝 1) := by + ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ (kernel.fst κ) a t (s m)) atTop (𝓝 1) := by let ν := kernel.fst κ - have h_mono : ∀ t, Monotone (fun m ↦ MLimsup κ (kernel.fst κ) a (s m) t) := - fun t n m hnm ↦ mLimsup_mono_set le_rfl a (hs hnm) t - have h_le_one : ∀ m t, MLimsup κ ν a (s m) t ≤ 1 := fun m t ↦ mLimsup_le_one le_rfl a (s m) t + have h_mono : ∀ t, Monotone (fun m ↦ MLimsup κ (kernel.fst κ) a t (s m)) := + fun t n m hnm ↦ mLimsup_mono_set le_rfl a t (hs hnm) + have h_le_one : ∀ m t, MLimsup κ ν a t (s m) ≤ 1 := fun m t ↦ mLimsup_le_one le_rfl a t (s m) -- for all `t`, `fun m ↦ MLimsup κ a (s m) t` has a limit - have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 l) := by + have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 l) := by intro t - have h_tendsto : Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop atTop ∨ - ∃ l, Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 l) := + have h_tendsto : Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop atTop ∨ + ∃ l, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 l) := tendsto_of_monotone (h_mono t) cases' h_tendsto with h_absurd h_tendsto · rw [tendsto_atTop_atTop_iff_of_monotone (h_mono t)] at h_absurd @@ -994,14 +1016,14 @@ lemma tendsto_mLimsup_atTop_ae_of_monotone [IsFiniteKernel κ] · exact h_tendsto -- let `F` be the pointwise limit of `fun m ↦ MLimsup κ a (s m) t` for all `t` let F : ℝ → ℝ := fun t ↦ (h_exists t).choose - have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ ν a (s m) t) atTop (𝓝 (F t)) := + have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 (F t)) := fun t ↦ (h_exists t).choose_spec have hF_nonneg : ∀ t, 0 ≤ F t := - fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ mLimsup_nonneg le_rfl a (s m) t) + fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ mLimsup_nonneg le_rfl a t (s m)) have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) have hF_int : Integrable F (ν a) := by rw [← memℒp_one_iff_integrable] - refine ⟨?_, ?_⟩ + constructor · refine aestronglyMeasurable_of_tendsto_ae atTop (fun n ↦ ?_) (ae_of_all _ hF_tendsto) exact (measurable_mLimsup_right κ ν (hs_meas _) a).aestronglyMeasurable · rw [snorm_one_eq_lintegral_nnnorm] @@ -1020,15 +1042,13 @@ lemma tendsto_mLimsup_atTop_ae_of_monotone [IsFiniteKernel κ] -- us that `F` is 1 a.e. refine ae_eq_of_integral_eq_of_ae_le hF_int (integrable_const _) (ae_of_all _ hF_le_one) ?_ have h_integral : - Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a (s m) t ∂(ν a)) atTop - (𝓝 (∫ t, F t ∂(ν a))) := by + Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ t, F t ∂(ν a))) := by refine integral_tendsto_of_tendsto_of_monotone ?_ hF_int ?_ ?_ · exact fun n ↦ integrable_mLimsup le_rfl _ (hs_meas n) · exact ae_of_all _ h_mono · exact ae_of_all _ hF_tendsto have h_integral' : - Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a (s m) t ∂(ν a)) atTop - (𝓝 (∫ _, 1 ∂(ν a))) := by + Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ _, 1 ∂(ν a))) := by rw [integral_const] simp only [smul_eq_mul, mul_one] rw [kernel.fst_apply' _ _ MeasurableSet.univ] @@ -1045,7 +1065,7 @@ section Iic_Q noncomputable def mLimsupIic (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) (a : α) (t : ℝ) (q : ℚ) : ℝ := - MLimsup κ ν a (Set.Iic q) t + MLimsup κ ν a t (Set.Iic q) lemma measurable_mLimsupIic (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) : Measurable (fun p : α × ℝ ↦ mLimsupIic κ ν p.1 p.2) := by @@ -1060,17 +1080,17 @@ lemma monotone_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) : Monotone (mLimsupIic κ ν a t) := by intro q r hqr rw [mLimsupIic, mLimsupIic] - refine mLimsup_mono_set hκν a ?_ t + refine mLimsup_mono_set hκν a t ?_ refine Iic_subset_Iic.mpr ?_ exact_mod_cast hqr lemma mLimsupIic_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : 0 ≤ mLimsupIic κ ν a t q := - mLimsup_nonneg hκν a _ t + mLimsup_nonneg hκν a t _ lemma mLimsupIic_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : mLimsupIic κ ν a t q ≤ 1 := - mLimsup_le_one hκν a _ t + mLimsup_le_one hκν a t _ lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ (kernel.fst κ) a t q) atTop (𝓝 1) := by From 3460efd3b4e8c055797f243705e38c74804c3953 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 18 Feb 2024 16:11:42 +0100 Subject: [PATCH 030/129] split file --- .../{KernelCDFBorel.lean => Density.lean} | 247 +++--------------- .../Kernel/Disintegration/KernelCDFReal.lean | 220 ++++++++++++++++ 2 files changed, 251 insertions(+), 216 deletions(-) rename Mathlib/Probability/Kernel/Disintegration/{KernelCDFBorel.lean => Density.lean} (83%) create mode 100644 Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean similarity index 83% rename from Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean rename to Mathlib/Probability/Kernel/Disintegration/Density.lean index c5013b8b21609..5a9c48828e7bb 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -3,9 +3,10 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Analysis.SpecialFunctions.Log.Base +import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Martingale.Convergence -import Mathlib.Probability.Kernel.Disintegration.CDFKernel +import Mathlib.Analysis.SpecialFunctions.Log.Base +import Mathlib.Probability.Kernel.Disintegration.AuxLemmas /-! # Kernel density @@ -23,7 +24,20 @@ def f_a (κ : kernel α (ℝ × β)) (ν : kernel a ℝ) (a : α) (t : ℝ) (s : ``` However, we can't turn those functions for each `a` into a measurable function of the pair `(a, t)`. -TODO: what do we do then? +In order to obtain measurability through countability, we discretize the real line. +For each `n : ℕ`, we define the intervals `I n k = [k * 2^-n, (k + 1) * 2^-n)` for `k : ℤ`. +For `t : ℝ`, let `indexI n t = ⌊t * 2^n⌋` be the integer such that `t ∈ I n (indexI n t)`. + +For a given `n`, the function `densityProcess κ ν n : α → ℝ → Set β → ℝ` defined by +`fun a t s ↦ (κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t))).toReal` has the desired +property that `∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal` for all `A` +in the σ-algebra generated by the intervals `I n k` for `k : ℤ` and is measurable in `(a, t)`. + +Let `ℱ` be the filtration of those σ-algebras for all `n : ℕ`. The functions `densityProcess κ ν n` +described here are a bounded `ν`-martingale for the filtration `ℱ`. By Doob's L1 martingale +convergence theorem, that martingale converges to a limit, which has a product-measurable version +and satisfies the integral equality for all `A` in `⨆ n, ℱ n = borel ℝ`. We have obtained the +desired density function. ## Main definitions @@ -33,24 +47,20 @@ TODO: what do we do then? * `fooBar_unique` -## Notation - - - ## Implementation details -issue with the following: joint measurability - - ## References The construction of the density process in this file follows the proof of Theorem 9.27 in [O. Kallenberg, Foundations of modern probability][kallenberg2021]. -## Tags +## TODO + +The construction in this file can be used to obtain a kernel Radon-Nikodym defivative of +`κ : kernel α Ω` with `Ω` standard Borel with respect to `ν : kernel α Ω`. +See Theorem 1.28 in O. Kallenberg, Random Measures, Theory and Applications. -Foobars, barfoos -/ @@ -258,7 +268,7 @@ end dissection_system variable [MeasurableSpace β] -section M +section DensityProcess variable {κ : kernel α (ℝ × β)} {ν : kernel α ℝ} @@ -665,9 +675,7 @@ lemma tendsto_m_mLimsup (hκν : kernel.fst κ ≤ ν) (a : α) [IsFiniteKernel ∀ᵐ t ∂(ν a), Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop (𝓝 (MLimsup κ ν a t s)) := by filter_upwards [tendsto_densityProcess_limitProcess hκν a hs, mLimsup_ae_eq_limitProcess hκν a hs] - with t h1 h2 - rw [h2] - exact h1 + with t h1 h2 using h2 ▸ h1 lemma measurable_mLimsup (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) {s : Set β} (hs : MeasurableSet s) : @@ -678,17 +686,16 @@ lemma measurable_mLimsup_left (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) {s : Set β} (hs : MeasurableSet s) : Measurable (fun a ↦ MLimsup κ ν a t s) := by change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ ν p.1 p.2 s) ∘ (fun a ↦ (a, t))) - exact (measurable_mLimsup κ ν hs).comp (measurable_id.prod_mk measurable_const) + exact (measurable_mLimsup κ ν hs).comp measurable_prod_mk_right lemma measurable_mLimsup_right (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) {s : Set β} (hs : MeasurableSet s) (a : α) : Measurable (fun t ↦ MLimsup κ ν a t s) := by change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ ν p.1 p.2 s) ∘ (fun t ↦ (a, t))) - exact (measurable_mLimsup κ ν hs).comp (measurable_const.prod_mk measurable_id) + exact (measurable_mLimsup κ ν hs).comp measurable_prod_mk_left lemma mLimsup_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) {s s' : Set β} (h : s ⊆ s') : MLimsup κ ν a t s ≤ MLimsup κ ν a t s' := by - rw [MLimsup, MLimsup] refine limsup_le_limsup ?_ ?_ ?_ · exact eventually_of_forall (fun n ↦ densityProcess_mono_set hκν n a t h) · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) @@ -730,9 +737,9 @@ lemma tendsto_set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] (eventually_of_forall (fun n ↦ integrable_densityProcess hκν _ _ hs)) ?_ A refine (tendsto_congr fun n ↦ ?_).mp (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) refine snorm_congr_ae ?_ - refine EventuallyEq.sub EventuallyEq.rfl ?_ - exact (mLimsup_ae_eq_limitProcess hκν a hs).symm + exact EventuallyEq.rfl.sub (mLimsup_ae_eq_limitProcess hκν a hs).symm +/-- Auxiliary lemma for `set_integral_mLimsup`. -/ lemma set_integral_mLimsup_of_measurableSet (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : @@ -933,13 +940,11 @@ lemma densityProcess_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (n refine measure_mono_null this ?_ have : {t | kernel.fst κ a (I n (indexI n t)) = 0} ⊆ ⋃ (k) (_ : kernel.fst κ a (I n k) = 0), I n k := by - intro t - simp only [mem_setOf_eq, mem_iUnion, exists_prop] - intro ht + intro t ht + simp only [mem_setOf_eq, mem_iUnion, exists_prop] at ht ⊢ exact ⟨indexI n t, ht, mem_I_indexI _ _⟩ refine measure_mono_null this ?_ rw [measure_iUnion_null] - intro i simp lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) @@ -1057,196 +1062,6 @@ lemma tendsto_mLimsup_atTop_ae_of_monotone [IsFiniteKernel κ] end UnivFst -end M - -variable {κ : kernel α (ℝ × ℝ)} {ν : kernel α ℝ} - -section Iic_Q - -noncomputable -def mLimsupIic (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) (a : α) (t : ℝ) (q : ℚ) : ℝ := - MLimsup κ ν a t (Set.Iic q) - -lemma measurable_mLimsupIic (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) : - Measurable (fun p : α × ℝ ↦ mLimsupIic κ ν p.1 p.2) := by - rw [measurable_pi_iff] - exact fun _ ↦ measurable_mLimsup κ ν measurableSet_Iic - -lemma measurable_mLimsupIic_right (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) (a : α) (q : ℚ) : - Measurable (fun t ↦ mLimsupIic κ ν a t q) := - measurable_mLimsup_right _ _ measurableSet_Iic _ - -lemma monotone_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) : - Monotone (mLimsupIic κ ν a t) := by - intro q r hqr - rw [mLimsupIic, mLimsupIic] - refine mLimsup_mono_set hκν a t ?_ - refine Iic_subset_Iic.mpr ?_ - exact_mod_cast hqr - -lemma mLimsupIic_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : - 0 ≤ mLimsupIic κ ν a t q := - mLimsup_nonneg hκν a t _ - -lemma mLimsupIic_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : - mLimsupIic κ ν a t q ≤ 1 := - mLimsup_le_one hκν a t _ - -lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ (kernel.fst κ) a t q) atTop (𝓝 1) := by - let ν := kernel.fst κ - suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ ν a t n) atTop (𝓝 1) by - filter_upwards [this] with t ht - let f := fun q : ℚ ↦ mLimsupIic κ ν a t ⌊q⌋₊ - let g := fun q : ℚ ↦ mLimsupIic κ ν a t ⌈q⌉₊ - have hf_le : ∀ᶠ q in atTop, f q ≤ mLimsupIic κ ν a t q := by - simp only [eventually_atTop, ge_iff_le] - exact ⟨0, fun q hq ↦ monotone_mLimsupIic le_rfl a t (Nat.floor_le hq)⟩ - have hg_le : ∀ q, mLimsupIic κ ν a t q ≤ g q := - fun q ↦ monotone_mLimsupIic le_rfl a t (Nat.le_ceil _) - refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ hf_le (eventually_of_forall hg_le) - · exact ht.comp tendsto_nat_floor_atTop - · exact ht.comp tendsto_nat_ceil_atTop - let s : ℕ → Set ℝ := fun n ↦ Iic n - have hs : Monotone s := fun i j hij ↦ Iic_subset_Iic.mpr (by exact mod_cast hij) - have hs_iUnion : ⋃ i, s i = univ := by - ext x - simp only [mem_iUnion, mem_Iic, mem_univ, iff_true] - exact ⟨Nat.ceil x, Nat.le_ceil x⟩ - have hs_meas : ∀ n, MeasurableSet (s n) := fun _ ↦ measurableSet_Iic - filter_upwards [tendsto_mLimsup_atTop_ae_of_monotone a s hs hs_iUnion hs_meas] - with x hx using hx - -lemma tendsto_atBot_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) : - ∀ᵐ t ∂(ν a), Tendsto (fun q ↦ mLimsupIic κ ν a t q) atBot (𝓝 0) := by - suffices ∀ᵐ t ∂(ν a), Tendsto (fun q ↦ mLimsupIic κ ν a t (-q)) atTop (𝓝 0) by - filter_upwards [this] with t ht - have h_eq_neg : (fun q ↦ mLimsupIic κ ν a t q) = fun q ↦ mLimsupIic κ ν a t (- -q) := by - simp_rw [neg_neg] - rw [h_eq_neg] - exact ht.comp tendsto_neg_atBot_atTop - suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ ν a t (-n)) atTop (𝓝 0) by - filter_upwards [this] with t ht - let f := fun q : ℚ ↦ mLimsupIic κ ν a t (-⌊q⌋₊) - let g := fun q : ℚ ↦ mLimsupIic κ ν a t (-⌈q⌉₊) - have hf_le : ∀ᶠ q in atTop, mLimsupIic κ ν a t (-q) ≤ f q := by - simp only [eventually_atTop, ge_iff_le] - refine ⟨0, fun q hq ↦ monotone_mLimsupIic hκν a t ?_⟩ - rw [neg_le_neg_iff] - exact Nat.floor_le hq - have hg_le : ∀ q, g q ≤ mLimsupIic κ ν a t (-q) := by - refine fun q ↦ monotone_mLimsupIic hκν a t ?_ - rw [neg_le_neg_iff] - exact Nat.le_ceil _ - refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ (eventually_of_forall hg_le) hf_le - · exact ht.comp tendsto_nat_ceil_atTop - · exact ht.comp tendsto_nat_floor_atTop - let s : ℕ → Set ℝ := fun n ↦ Iic (-n) - have hs : Antitone s := fun i j hij ↦ Iic_subset_Iic.mpr (neg_le_neg (by exact mod_cast hij)) - have hs_iInter : ⋂ i, s i = ∅ := by - ext x - simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, neg_lt] - exact exists_nat_gt (-x) - have hs_meas : ∀ n, MeasurableSet (s n) := fun _ ↦ measurableSet_Iic - convert tendsto_mLimsup_atTop_ae_of_antitone hκν a s hs hs_iInter hs_meas with x n - rw [mLimsupIic] - simp - -lemma set_integral_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) (q : ℚ) {A : Set ℝ} (hA : MeasurableSet A) : - ∫ t in A, mLimsupIic κ ν a t q ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := - set_integral_mLimsup hκν a measurableSet_Iic hA - -lemma integrable_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) (q : ℚ) : - Integrable (fun t ↦ mLimsupIic κ ν a t q) (ν a) := - integrable_mLimsup hκν _ measurableSet_Iic - -lemma bddBelow_range_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : - BddBelow (range fun (r : Ioi q) ↦ mLimsupIic κ ν a t r) := by - refine ⟨0, ?_⟩ - rw [mem_lowerBounds] - rintro x ⟨y, rfl⟩ - exact mLimsupIic_nonneg hκν _ _ _ - -lemma integrable_iInf_rat_gt_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] - (a : α) (q : ℚ) : - Integrable (fun t ↦ ⨅ r : Ioi q, mLimsupIic κ ν a t r) (ν a) := by - rw [← memℒp_one_iff_integrable] - refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ - · exact measurable_iInf fun i ↦ measurable_mLimsupIic_right κ ν a i - refine (?_ : _ ≤ (ν a univ : ℝ≥0∞)).trans_lt (measure_lt_top _ _) - refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ - · rw [Real.norm_eq_abs, abs_of_nonneg] - · refine ciInf_le_of_le ?_ ?_ ?_ - · exact bddBelow_range_mLimsupIic hκν a t _ - · exact ⟨q + 1, by simp⟩ - · exact mLimsupIic_le_one hκν _ _ _ - · exact le_ciInf fun r ↦ mLimsupIic_nonneg hκν a t r - · simp - -lemma set_integral_iInf_rat_gt_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] - [IsFiniteKernel ν] - (a : α) (q : ℚ) {A : Set ℝ} (hA : MeasurableSet A) : - ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ ν a t r ∂(ν a) - = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by - refine le_antisymm ?_ ?_ - · have h : ∀ r : Ioi q, ∫ t in A, ⨅ r' : Ioi q, mLimsupIic κ ν a t r' ∂(ν a) - ≤ (κ a (A ×ˢ Iic (r : ℝ))).toReal := by - intro r - rw [← set_integral_mLimsupIic hκν a r hA] - refine set_integral_mono ?_ ?_ ?_ - · exact (integrable_iInf_rat_gt_mLimsupIic hκν _ _).integrableOn - · exact (integrable_mLimsupIic hκν _ _).integrableOn - · exact fun t ↦ ciInf_le (bddBelow_range_mLimsupIic hκν _ _ _) r - calc ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ ν a t r ∂(ν a) - ≤ ⨅ r : Ioi q, (κ a (A ×ˢ Iic (r : ℝ))).toReal := le_ciInf h - _ = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by - rw [← Measure.iInf_Iic_gt_prod hA q] - exact (ENNReal.toReal_iInf (fun r ↦ measure_ne_top _ _)).symm - · rw [← set_integral_mLimsupIic hκν a q hA] - refine set_integral_mono ?_ ?_ ?_ - · exact (integrable_mLimsupIic hκν _ _).integrableOn - · exact (integrable_iInf_rat_gt_mLimsupIic hκν _ _).integrableOn - · exact fun t ↦ le_ciInf (fun r ↦ monotone_mLimsupIic hκν _ _ (le_of_lt r.prop)) - -lemma iInf_rat_gt_mLimsupIic_eq (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) : - ∀ᵐ t ∂(ν a), ∀ q : ℚ, ⨅ r : Ioi q, mLimsupIic κ ν a t r = mLimsupIic κ ν a t q := by - rw [ae_all_iff] - refine fun q ↦ ae_eq_of_forall_set_integral_eq_of_sigmaFinite (μ := ν a) ?_ ?_ ?_ - · intro s _ _ - refine Integrable.integrableOn ?_ - exact integrable_iInf_rat_gt_mLimsupIic hκν _ _ - · exact fun s _ _ ↦ (integrable_mLimsupIic hκν a q).integrableOn - · intro s hs _ - rw [set_integral_mLimsupIic hκν _ _ hs, set_integral_iInf_rat_gt_mLimsupIic hκν _ _ hs] - -end Iic_Q - -lemma isRatStieltjesPoint_mLimsupIic_ae (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), - IsRatStieltjesPoint (fun p q ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2 q) (a, t) := by - filter_upwards [tendsto_atTop_mLimsupIic κ a, tendsto_atBot_mLimsupIic le_rfl a, - iInf_rat_gt_mLimsupIic_eq le_rfl a] with t ht_top ht_bot ht_iInf - exact ⟨monotone_mLimsupIic le_rfl a t, ht_top, ht_bot, ht_iInf⟩ - -lemma isRatKernelCDF_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : - IsRatKernelCDF (fun p : α × ℝ ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2) κ (kernel.fst κ) where - measurable := measurable_mLimsupIic κ (kernel.fst κ) - isRatStieltjesPoint_ae := isRatStieltjesPoint_mLimsupIic_ae κ - integrable := integrable_mLimsupIic le_rfl - set_integral := fun _ _ hs _ ↦ set_integral_mLimsupIic le_rfl _ _ hs - -noncomputable -def mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : - α × ℝ → StieltjesFunction := - stieltjesOfMeasurableRat (fun p : α × ℝ ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2) - (isRatKernelCDF_mLimsupIic κ).measurable - -lemma isKernelCDF_mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : - IsKernelCDF (mLimsupCDF κ) κ (kernel.fst κ) := - isKernelCDF_stieltjesOfMeasurableRat (isRatKernelCDF_mLimsupIic κ) +end DensityProcess end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean new file mode 100644 index 0000000000000..9613594348142 --- /dev/null +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean @@ -0,0 +1,220 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Probability.Kernel.Disintegration.CDFKernel +import Mathlib.Probability.Kernel.Disintegration.Density + +/-! +# Kernel CDF + +## Main definitions + +* `FooBar` + +## Main statements + +* `fooBar_unique` + +## Implementation details + + +## References + +-/ + + +open MeasureTheory Set Filter + +open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory + +namespace ProbabilityTheory + +variable {α : Type*} {mα : MeasurableSpace α} {κ : kernel α (ℝ × ℝ)} {ν : kernel α ℝ} + +noncomputable +def mLimsupIic (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) (a : α) (t : ℝ) (q : ℚ) : ℝ := + MLimsup κ ν a t (Set.Iic q) + +lemma measurable_mLimsupIic (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) : + Measurable (fun p : α × ℝ ↦ mLimsupIic κ ν p.1 p.2) := by + rw [measurable_pi_iff] + exact fun _ ↦ measurable_mLimsup κ ν measurableSet_Iic + +lemma measurable_mLimsupIic_right (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) (a : α) (q : ℚ) : + Measurable (fun t ↦ mLimsupIic κ ν a t q) := + measurable_mLimsup_right _ _ measurableSet_Iic _ + +lemma monotone_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) : + Monotone (mLimsupIic κ ν a t) := by + intro q r hqr + rw [mLimsupIic, mLimsupIic] + refine mLimsup_mono_set hκν a t ?_ + refine Iic_subset_Iic.mpr ?_ + exact_mod_cast hqr + +lemma mLimsupIic_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : + 0 ≤ mLimsupIic κ ν a t q := + mLimsup_nonneg hκν a t _ + +lemma mLimsupIic_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : + mLimsupIic κ ν a t q ≤ 1 := + mLimsup_le_one hκν a t _ + +lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : + ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ (kernel.fst κ) a t q) atTop (𝓝 1) := by + let ν := kernel.fst κ + suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ ν a t n) atTop (𝓝 1) by + filter_upwards [this] with t ht + let f := fun q : ℚ ↦ mLimsupIic κ ν a t ⌊q⌋₊ + let g := fun q : ℚ ↦ mLimsupIic κ ν a t ⌈q⌉₊ + have hf_le : ∀ᶠ q in atTop, f q ≤ mLimsupIic κ ν a t q := by + simp only [eventually_atTop, ge_iff_le] + exact ⟨0, fun q hq ↦ monotone_mLimsupIic le_rfl a t (Nat.floor_le hq)⟩ + have hg_le : ∀ q, mLimsupIic κ ν a t q ≤ g q := + fun q ↦ monotone_mLimsupIic le_rfl a t (Nat.le_ceil _) + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ hf_le (eventually_of_forall hg_le) + · exact ht.comp tendsto_nat_floor_atTop + · exact ht.comp tendsto_nat_ceil_atTop + let s : ℕ → Set ℝ := fun n ↦ Iic n + have hs : Monotone s := fun i j hij ↦ Iic_subset_Iic.mpr (by exact mod_cast hij) + have hs_iUnion : ⋃ i, s i = univ := by + ext x + simp only [mem_iUnion, mem_Iic, mem_univ, iff_true] + exact ⟨Nat.ceil x, Nat.le_ceil x⟩ + have hs_meas : ∀ n, MeasurableSet (s n) := fun _ ↦ measurableSet_Iic + filter_upwards [tendsto_mLimsup_atTop_ae_of_monotone a s hs hs_iUnion hs_meas] + with x hx using hx + +lemma tendsto_atBot_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) : + ∀ᵐ t ∂(ν a), Tendsto (fun q ↦ mLimsupIic κ ν a t q) atBot (𝓝 0) := by + suffices ∀ᵐ t ∂(ν a), Tendsto (fun q ↦ mLimsupIic κ ν a t (-q)) atTop (𝓝 0) by + filter_upwards [this] with t ht + have h_eq_neg : (fun q ↦ mLimsupIic κ ν a t q) = fun q ↦ mLimsupIic κ ν a t (- -q) := by + simp_rw [neg_neg] + rw [h_eq_neg] + exact ht.comp tendsto_neg_atBot_atTop + suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ ν a t (-n)) atTop (𝓝 0) by + filter_upwards [this] with t ht + let f := fun q : ℚ ↦ mLimsupIic κ ν a t (-⌊q⌋₊) + let g := fun q : ℚ ↦ mLimsupIic κ ν a t (-⌈q⌉₊) + have hf_le : ∀ᶠ q in atTop, mLimsupIic κ ν a t (-q) ≤ f q := by + simp only [eventually_atTop, ge_iff_le] + refine ⟨0, fun q hq ↦ monotone_mLimsupIic hκν a t ?_⟩ + rw [neg_le_neg_iff] + exact Nat.floor_le hq + have hg_le : ∀ q, g q ≤ mLimsupIic κ ν a t (-q) := by + refine fun q ↦ monotone_mLimsupIic hκν a t ?_ + rw [neg_le_neg_iff] + exact Nat.le_ceil _ + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ (eventually_of_forall hg_le) hf_le + · exact ht.comp tendsto_nat_ceil_atTop + · exact ht.comp tendsto_nat_floor_atTop + let s : ℕ → Set ℝ := fun n ↦ Iic (-n) + have hs : Antitone s := fun i j hij ↦ Iic_subset_Iic.mpr (neg_le_neg (by exact mod_cast hij)) + have hs_iInter : ⋂ i, s i = ∅ := by + ext x + simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, neg_lt] + exact exists_nat_gt (-x) + have hs_meas : ∀ n, MeasurableSet (s n) := fun _ ↦ measurableSet_Iic + convert tendsto_mLimsup_atTop_ae_of_antitone hκν a s hs hs_iInter hs_meas with x n + rw [mLimsupIic] + simp + +lemma set_integral_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) (q : ℚ) {A : Set ℝ} (hA : MeasurableSet A) : + ∫ t in A, mLimsupIic κ ν a t q ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := + set_integral_mLimsup hκν a measurableSet_Iic hA + +lemma integrable_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] + (a : α) (q : ℚ) : + Integrable (fun t ↦ mLimsupIic κ ν a t q) (ν a) := + integrable_mLimsup hκν _ measurableSet_Iic + +lemma bddBelow_range_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : + BddBelow (range fun (r : Ioi q) ↦ mLimsupIic κ ν a t r) := by + refine ⟨0, ?_⟩ + rw [mem_lowerBounds] + rintro x ⟨y, rfl⟩ + exact mLimsupIic_nonneg hκν _ _ _ + +lemma integrable_iInf_rat_gt_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] + (a : α) (q : ℚ) : + Integrable (fun t ↦ ⨅ r : Ioi q, mLimsupIic κ ν a t r) (ν a) := by + rw [← memℒp_one_iff_integrable] + refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ + · exact measurable_iInf fun i ↦ measurable_mLimsupIic_right κ ν a i + refine (?_ : _ ≤ (ν a univ : ℝ≥0∞)).trans_lt (measure_lt_top _ _) + refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ + · rw [Real.norm_eq_abs, abs_of_nonneg] + · refine ciInf_le_of_le ?_ ?_ ?_ + · exact bddBelow_range_mLimsupIic hκν a t _ + · exact ⟨q + 1, by simp⟩ + · exact mLimsupIic_le_one hκν _ _ _ + · exact le_ciInf fun r ↦ mLimsupIic_nonneg hκν a t r + · simp + +lemma set_integral_iInf_rat_gt_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] + (a : α) (q : ℚ) {A : Set ℝ} (hA : MeasurableSet A) : + ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ ν a t r ∂(ν a) + = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by + refine le_antisymm ?_ ?_ + · have h : ∀ r : Ioi q, ∫ t in A, ⨅ r' : Ioi q, mLimsupIic κ ν a t r' ∂(ν a) + ≤ (κ a (A ×ˢ Iic (r : ℝ))).toReal := by + intro r + rw [← set_integral_mLimsupIic hκν a r hA] + refine set_integral_mono ?_ ?_ ?_ + · exact (integrable_iInf_rat_gt_mLimsupIic hκν _ _).integrableOn + · exact (integrable_mLimsupIic hκν _ _).integrableOn + · exact fun t ↦ ciInf_le (bddBelow_range_mLimsupIic hκν _ _ _) r + calc ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ ν a t r ∂(ν a) + ≤ ⨅ r : Ioi q, (κ a (A ×ˢ Iic (r : ℝ))).toReal := le_ciInf h + _ = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by + rw [← Measure.iInf_Iic_gt_prod hA q] + exact (ENNReal.toReal_iInf (fun r ↦ measure_ne_top _ _)).symm + · rw [← set_integral_mLimsupIic hκν a q hA] + refine set_integral_mono ?_ ?_ ?_ + · exact (integrable_mLimsupIic hκν _ _).integrableOn + · exact (integrable_iInf_rat_gt_mLimsupIic hκν _ _).integrableOn + · exact fun t ↦ le_ciInf (fun r ↦ monotone_mLimsupIic hκν _ _ (le_of_lt r.prop)) + +lemma iInf_rat_gt_mLimsupIic_eq (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) : + ∀ᵐ t ∂(ν a), ∀ q : ℚ, ⨅ r : Ioi q, mLimsupIic κ ν a t r = mLimsupIic κ ν a t q := by + rw [ae_all_iff] + refine fun q ↦ ae_eq_of_forall_set_integral_eq_of_sigmaFinite (μ := ν a) ?_ ?_ ?_ + · intro s _ _ + refine Integrable.integrableOn ?_ + exact integrable_iInf_rat_gt_mLimsupIic hκν _ _ + · exact fun s _ _ ↦ (integrable_mLimsupIic hκν a q).integrableOn + · intro s hs _ + rw [set_integral_mLimsupIic hκν _ _ hs, set_integral_iInf_rat_gt_mLimsupIic hκν _ _ hs] + +lemma isRatStieltjesPoint_mLimsupIic_ae (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : + ∀ᵐ t ∂(kernel.fst κ a), + IsRatStieltjesPoint (fun p q ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2 q) (a, t) := by + filter_upwards [tendsto_atTop_mLimsupIic κ a, tendsto_atBot_mLimsupIic le_rfl a, + iInf_rat_gt_mLimsupIic_eq le_rfl a] with t ht_top ht_bot ht_iInf + exact ⟨monotone_mLimsupIic le_rfl a t, ht_top, ht_bot, ht_iInf⟩ + +lemma isRatKernelCDF_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : + IsRatKernelCDF (fun p : α × ℝ ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2) κ (kernel.fst κ) where + measurable := measurable_mLimsupIic κ (kernel.fst κ) + isRatStieltjesPoint_ae := isRatStieltjesPoint_mLimsupIic_ae κ + integrable := integrable_mLimsupIic le_rfl + set_integral := fun _ _ hs _ ↦ set_integral_mLimsupIic le_rfl _ _ hs + +noncomputable +def mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : + α × ℝ → StieltjesFunction := + stieltjesOfMeasurableRat (fun p : α × ℝ ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2) + (isRatKernelCDF_mLimsupIic κ).measurable + +lemma isKernelCDF_mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : + IsKernelCDF (mLimsupCDF κ) κ (kernel.fst κ) := + isKernelCDF_stieltjesOfMeasurableRat (isRatKernelCDF_mLimsupIic κ) + +end ProbabilityTheory From 3bb8792d15f55c05e089e4f7a0746042eee772a8 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 18 Feb 2024 16:18:12 +0100 Subject: [PATCH 031/129] fix --- .../Probability/Kernel/Disintegration/CondCdf.lean | 11 +++++------ 1 file changed, 5 insertions(+), 6 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index 6cc2bdac58370..a7c121a79def1 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -294,7 +294,7 @@ theorem tendsto_preCDF_atTop_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] -- us that `F` is 1 a.e. have h_lintegral_eq : ∫⁻ a, F a ∂ρ.fst = ∫⁻ _, 1 ∂ρ.fst := by have h_lintegral : - Tendsto (fun r : ℕ => ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (∫⁻ a, F a ∂ρ.fst)) := by + Tendsto (fun r : ℕ => ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (∫⁻ a, F a ∂ρ.fst)) := by refine' lintegral_tendsto_of_tendsto_of_monotone (-- does this exist only for ℕ? @@ -304,7 +304,7 @@ theorem tendsto_preCDF_atTop_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] refine' fun n m hnm => ha _ exact mod_cast hnm have h_lintegral' : - Tendsto (fun r : ℕ => ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (∫⁻ _, 1 ∂ρ.fst)) := by + Tendsto (fun r : ℕ => ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (∫⁻ _, 1 ∂ρ.fst)) := by rw [lintegral_one, Measure.fst_univ] exact (tendsto_lintegral_preCDF_atTop ρ).comp tendsto_nat_cast_atTop_atTop exact tendsto_nhds_unique h_lintegral h_lintegral' @@ -351,7 +351,7 @@ theorem tendsto_preCDF_atBot_zero (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ filter_upwards [h_tendsto, h_lintegral_eq] with a ha_tendsto ha_eq rwa [ha_eq] at ha_tendsto have h_lintegral : - Tendsto (fun r => ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) atTop (𝓝 (∫⁻ a, F a ∂ρ.fst)) := by + Tendsto (fun r => ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) atTop (𝓝 (∫⁻ a, F a ∂ρ.fst)) := by refine' tendsto_lintegral_filter_of_dominated_convergence (fun _ => 1) (eventually_of_forall fun _ => measurable_preCDF) (eventually_of_forall fun _ => _) _ @@ -361,7 +361,7 @@ theorem tendsto_preCDF_atBot_zero (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ exact measure_ne_top _ _ have h_lintegral' : Tendsto (fun r => ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) atTop (𝓝 0) := by have h_lintegral_eq : - (fun r => ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) = fun r : ℚ => ρ (univ ×ˢ Iic (-r : ℝ)) := by + (fun r => ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) = fun r : ℚ => ρ (univ ×ˢ Iic (-r : ℝ)) := by ext1 n rw [← set_lintegral_univ, set_lintegral_preCDF_fst ρ _ MeasurableSet.univ, Measure.IicSnd_univ] @@ -426,8 +426,7 @@ lemma isRatStieltjesPoint_ae (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : tendsto_preCDF_atBot_zero ρ, inf_gt_preCDF ρ] with a h1 h2 h3 h4 h5 constructor · intro r r' hrr' - have h_ne_top : ∀ r, preCDF ρ r a ≠ ∞ := fun r ↦ - ((h2 r).trans_lt ENNReal.one_lt_top).ne + have h_ne_top : ∀ r, preCDF ρ r a ≠ ∞ := fun r ↦ ((h2 r).trans_lt ENNReal.one_lt_top).ne rw [ENNReal.toReal_le_toReal (h_ne_top _) (h_ne_top _)] exact h1 hrr' · rw [← ENNReal.one_toReal, ENNReal.tendsto_toReal_iff] From 237699e37d306541df504ccdb1a903f9e640fec1 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 18 Feb 2024 16:19:21 +0100 Subject: [PATCH 032/129] delete again --- .../Probability/Kernel/Disintegration.lean | 657 ------------------ 1 file changed, 657 deletions(-) delete mode 100644 Mathlib/Probability/Kernel/Disintegration.lean diff --git a/Mathlib/Probability/Kernel/Disintegration.lean b/Mathlib/Probability/Kernel/Disintegration.lean deleted file mode 100644 index 957a452c6d63e..0000000000000 --- a/Mathlib/Probability/Kernel/Disintegration.lean +++ /dev/null @@ -1,657 +0,0 @@ -/- -Copyright (c) 2023 Rémy Degenne. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Rémy Degenne, Kexing Ying --/ -import Mathlib.Probability.Kernel.CondCdf -import Mathlib.MeasureTheory.Constructions.Polish -import Mathlib.Probability.Kernel.MeasureCompProd - -#align_import probability.kernel.disintegration from "leanprover-community/mathlib"@"6315581f5650ffa2fbdbbbedc41243c8d7070981" - -/-! -# Disintegration of measures on product spaces - -Let `ρ` be a finite measure on `α × Ω`, where `Ω` is a standard Borel space. In mathlib terms, `Ω` -verifies `[Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω]`. -Then there exists a kernel `ρ.condKernel : kernel α Ω` such that for any measurable set -`s : Set (α × Ω)`, `ρ s = ∫⁻ a, ρ.condKernel a {x | (a, x) ∈ s} ∂ρ.fst`. - -In terms of kernels, `ρ.condKernel` is such that for any measurable space `γ`, we -have a disintegration of the constant kernel from `γ` with value `ρ`: -`kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ (condKernel ρ))`, -where `ρ.fst` is the marginal measure of `ρ` on `α`. In particular, `ρ = ρ.fst ⊗ₘ ρ.condKernel`. - -In order to obtain a disintegration for any standard Borel space, we use that these spaces embed -measurably into `ℝ`: it then suffices to define a suitable kernel for `Ω = ℝ`. In the real case, -we define a conditional kernel by taking for each `a : α` the measure associated to the Stieltjes -function `condCDF ρ a` (the conditional cumulative distribution function). - -## Main definitions - -* `MeasureTheory.Measure.condKernel ρ : kernel α Ω`: conditional kernel described above. - -## Main statements - -* `ProbabilityTheory.lintegral_condKernel`: - `∫⁻ a, ∫⁻ ω, f (a, ω) ∂(ρ.condKernel a) ∂ρ.fst = ∫⁻ x, f x ∂ρ` -* `ProbabilityTheory.lintegral_condKernel_mem`: - `∫⁻ a, ρ.condKernel a {x | (a, x) ∈ s} ∂ρ.fst = ρ s` -* `ProbabilityTheory.kernel.const_eq_compProd`: - `kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ ρ.condKernel)` -* `ProbabilityTheory.measure_eq_compProd`: `ρ = ρ.fst ⊗ₘ ρ.condKernel` -* `ProbabilityTheory.eq_condKernel_of_measure_eq_compProd`: a.e. uniqueness of `condKernel` - --/ - - -open MeasureTheory Set Filter - -open scoped ENNReal MeasureTheory Topology ProbabilityTheory - -namespace ProbabilityTheory - -variable {α : Type*} {mα : MeasurableSpace α} - -section Real - -/-! ### Disintegration of measures on `α × ℝ` -/ - - -/-- Conditional measure on the second space of the product given the value on the first, as a -kernel. Use the more general `condKernel`. -/ -noncomputable def condKernelReal (ρ : Measure (α × ℝ)) : kernel α ℝ where - val a := (condCDF ρ a).measure - property := measurable_measure_condCDF ρ -#align probability_theory.cond_kernel_real ProbabilityTheory.condKernelReal - -instance (ρ : Measure (α × ℝ)) : IsMarkovKernel (condKernelReal ρ) := - ⟨fun a => by rw [condKernelReal]; exact instIsProbabilityMeasure ρ a⟩ - -theorem condKernelReal_Iic (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : - condKernelReal ρ a (Iic x) = ENNReal.ofReal (condCDF ρ a x) := - measure_condCDF_Iic ρ a x -#align probability_theory.cond_kernel_real_Iic ProbabilityTheory.condKernelReal_Iic - -theorem set_lintegral_condKernelReal_Iic (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) - {s : Set α} (hs : MeasurableSet s) : - ∫⁻ a in s, condKernelReal ρ a (Iic x) ∂ρ.fst = ρ (s ×ˢ Iic x) := by - simp_rw [condKernelReal_Iic]; exact set_lintegral_condCDF ρ x hs -#align probability_theory.set_lintegral_cond_kernel_real_Iic ProbabilityTheory.set_lintegral_condKernelReal_Iic - -theorem set_lintegral_condKernelReal_univ (ρ : Measure (α × ℝ)) {s : Set α} (hs : MeasurableSet s) : - ∫⁻ a in s, condKernelReal ρ a univ ∂ρ.fst = ρ (s ×ˢ univ) := by - simp only [measure_univ, lintegral_const, Measure.restrict_apply, MeasurableSet.univ, univ_inter, - one_mul, Measure.fst_apply hs, ← prod_univ] -#align probability_theory.set_lintegral_cond_kernel_real_univ ProbabilityTheory.set_lintegral_condKernelReal_univ - -theorem lintegral_condKernelReal_univ (ρ : Measure (α × ℝ)) : - ∫⁻ a, condKernelReal ρ a univ ∂ρ.fst = ρ univ := by - rw [← set_lintegral_univ, set_lintegral_condKernelReal_univ ρ MeasurableSet.univ, - univ_prod_univ] -#align probability_theory.lintegral_cond_kernel_real_univ ProbabilityTheory.lintegral_condKernelReal_univ - -variable (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] - -theorem set_lintegral_condKernelReal_prod {s : Set α} (hs : MeasurableSet s) {t : Set ℝ} - (ht : MeasurableSet t) : ∫⁻ a in s, condKernelReal ρ a t ∂ρ.fst = ρ (s ×ˢ t) := by - -- `set_lintegral_condKernelReal_Iic` gives the result for `t = Iic x`. These sets form a - -- π-system that generates the Borel σ-algebra, hence we can get the same equality for any - -- measurable set `t`. - apply MeasurableSpace.induction_on_inter (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic _ _ _ _ ht - · simp only [measure_empty, lintegral_const, zero_mul, prod_empty] - · rintro t ⟨q, rfl⟩ - exact set_lintegral_condKernelReal_Iic ρ q hs - · intro t ht ht_lintegral - calc - ∫⁻ a in s, condKernelReal ρ a tᶜ ∂ρ.fst = - ∫⁻ a in s, condKernelReal ρ a univ - condKernelReal ρ a t ∂ρ.fst := by - congr with a; rw [measure_compl ht (measure_ne_top (condKernelReal ρ a) _)] - _ = ∫⁻ a in s, condKernelReal ρ a univ ∂ρ.fst - ∫⁻ a in s, condKernelReal ρ a t ∂ρ.fst := by - rw [lintegral_sub (kernel.measurable_coe (condKernelReal ρ) ht)] - · rw [ht_lintegral] - exact measure_ne_top ρ _ - · exact eventually_of_forall fun a => measure_mono (subset_univ _) - _ = ρ (s ×ˢ univ) - ρ (s ×ˢ t) := by - rw [set_lintegral_condKernelReal_univ ρ hs, ht_lintegral] - _ = ρ (s ×ˢ tᶜ) := by - rw [← measure_diff _ (hs.prod ht) (measure_ne_top ρ _)] - · rw [prod_diff_prod, compl_eq_univ_diff] - simp only [diff_self, empty_prod, union_empty] - · rw [prod_subset_prod_iff] - exact Or.inl ⟨subset_rfl, subset_univ t⟩ - · intro f hf_disj hf_meas hf_eq - simp_rw [measure_iUnion hf_disj hf_meas] - rw [lintegral_tsum fun i => (kernel.measurable_coe _ (hf_meas i)).aemeasurable.restrict, - prod_iUnion, measure_iUnion] - · simp_rw [hf_eq] - · intro i j hij - rw [Function.onFun, Set.disjoint_prod] - exact Or.inr (hf_disj hij) - · exact fun i => MeasurableSet.prod hs (hf_meas i) -#align probability_theory.set_lintegral_cond_kernel_real_prod ProbabilityTheory.set_lintegral_condKernelReal_prod - -theorem lintegral_condKernelReal_mem {s : Set (α × ℝ)} (hs : MeasurableSet s) : - ∫⁻ a, condKernelReal ρ a {x | (a, x) ∈ s} ∂ρ.fst = ρ s := by - -- `set_lintegral_condKernelReal_prod` gives the result for sets of the form `t₁ × t₂`. These - -- sets form a π-system that generates the product σ-algebra, hence we can get the same equality - -- for any measurable set `s`. - apply MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod _ _ _ _ hs - · simp only [mem_empty_iff_false, setOf_false, measure_empty, lintegral_const, - zero_mul] - · rintro _ ⟨t₁, ht₁, t₂, ht₂, rfl⟩ - have h_prod_eq_snd : ∀ a ∈ t₁, {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} = t₂ := by - intro a ha - simp only [ha, prod_mk_mem_set_prod_eq, true_and_iff, setOf_mem_eq] - rw [← lintegral_add_compl _ ht₁] - have h_eq1 : ∫⁻ a in t₁, condKernelReal ρ a {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} ∂ρ.fst = - ∫⁻ a in t₁, condKernelReal ρ a t₂ ∂ρ.fst := by - refine' set_lintegral_congr_fun ht₁ (eventually_of_forall fun a ha => _) - rw [h_prod_eq_snd a ha] - have h_eq2 : ∫⁻ a in t₁ᶜ, condKernelReal ρ a {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} ∂ρ.fst = 0 := by - suffices h_eq_zero : ∀ a ∈ t₁ᶜ, condKernelReal ρ a {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} = 0 by - rw [set_lintegral_congr_fun ht₁.compl (eventually_of_forall h_eq_zero)] - simp only [lintegral_const, zero_mul] - intro a hat₁ - rw [mem_compl_iff] at hat₁ - simp only [hat₁, prod_mk_mem_set_prod_eq, false_and_iff, setOf_false, measure_empty] - rw [h_eq1, h_eq2, add_zero] - exact set_lintegral_condKernelReal_prod ρ ht₁ ht₂ - · intro t ht ht_eq - calc - ∫⁻ a, condKernelReal ρ a {x : ℝ | (a, x) ∈ tᶜ} ∂ρ.fst = - ∫⁻ a, condKernelReal ρ a {x : ℝ | (a, x) ∈ t}ᶜ ∂ρ.fst := rfl - _ = ∫⁻ a, condKernelReal ρ a univ - condKernelReal ρ a {x : ℝ | (a, x) ∈ t} ∂ρ.fst := by - congr with a : 1 - exact measure_compl (measurable_prod_mk_left ht) (measure_ne_top (condKernelReal ρ a) _) - _ = ∫⁻ a, condKernelReal ρ a univ ∂ρ.fst - - ∫⁻ a, condKernelReal ρ a {x : ℝ | (a, x) ∈ t} ∂ρ.fst := by - have h_le : (fun a => condKernelReal ρ a {x : ℝ | (a, x) ∈ t}) ≤ᵐ[ρ.fst] fun a => - condKernelReal ρ a univ := eventually_of_forall fun a => measure_mono (subset_univ _) - rw [lintegral_sub _ _ h_le] - · exact kernel.measurable_kernel_prod_mk_left ht - refine' ((lintegral_mono_ae h_le).trans_lt _).ne - rw [lintegral_condKernelReal_univ] - exact measure_lt_top ρ univ - _ = ρ univ - ρ t := by rw [ht_eq, lintegral_condKernelReal_univ] - _ = ρ tᶜ := (measure_compl ht (measure_ne_top _ _)).symm - · intro f hf_disj hf_meas hf_eq - have h_eq : ∀ a, {x | (a, x) ∈ ⋃ i, f i} = ⋃ i, {x | (a, x) ∈ f i} := by - intro a - ext1 x - simp only [mem_iUnion, mem_setOf_eq] - simp_rw [h_eq] - have h_disj : ∀ a, Pairwise (Disjoint on fun i => {x | (a, x) ∈ f i}) := by - intro a i j hij - have h_disj := hf_disj hij - rw [Function.onFun, disjoint_iff_inter_eq_empty] at h_disj ⊢ - ext1 x - simp only [mem_inter_iff, mem_setOf_eq, mem_empty_iff_false, iff_false_iff] - intro h_mem_both - suffices (a, x) ∈ ∅ by rwa [mem_empty_iff_false] at this - rwa [← h_disj, mem_inter_iff] - calc - ∫⁻ a, condKernelReal ρ a (⋃ i, {x | (a, x) ∈ f i}) ∂ρ.fst = - ∫⁻ a, ∑' i, condKernelReal ρ a {x | (a, x) ∈ f i} ∂ρ.fst := by - congr with a : 1 - rw [measure_iUnion (h_disj a) fun i => measurable_prod_mk_left (hf_meas i)] - _ = ∑' i, ∫⁻ a, condKernelReal ρ a {x | (a, x) ∈ f i} ∂ρ.fst := - (lintegral_tsum fun i => (kernel.measurable_kernel_prod_mk_left (hf_meas i)).aemeasurable) - _ = ∑' i, ρ (f i) := by simp_rw [hf_eq] - _ = ρ (iUnion f) := (measure_iUnion hf_disj hf_meas).symm -#align probability_theory.lintegral_cond_kernel_real_mem ProbabilityTheory.lintegral_condKernelReal_mem - -theorem kernel.const_eq_compProd_real (γ : Type*) [MeasurableSpace γ] (ρ : Measure (α × ℝ)) - [IsFiniteMeasure ρ] : - kernel.const γ ρ = kernel.const γ ρ.fst ⊗ₖ kernel.prodMkLeft γ (condKernelReal ρ) := by - ext a s hs : 2 - rw [kernel.compProd_apply _ _ _ hs, kernel.const_apply, kernel.const_apply] - simp_rw [kernel.prodMkLeft_apply] - rw [lintegral_condKernelReal_mem ρ hs] -#align probability_theory.kernel.const_eq_comp_prod_real ProbabilityTheory.kernel.const_eq_compProd_real - -theorem measure_eq_compProd_real : ρ = ρ.fst ⊗ₘ condKernelReal ρ := by - rw [Measure.compProd, ← kernel.const_eq_compProd_real Unit ρ, kernel.const_apply] -#align probability_theory.measure_eq_comp_prod_real ProbabilityTheory.measure_eq_compProd_real - -theorem lintegral_condKernelReal {f : α × ℝ → ℝ≥0∞} (hf : Measurable f) : - ∫⁻ a, ∫⁻ y, f (a, y) ∂condKernelReal ρ a ∂ρ.fst = ∫⁻ x, f x ∂ρ := by - nth_rw 3 [measure_eq_compProd_real ρ] - rw [Measure.lintegral_compProd hf] -#align probability_theory.lintegral_cond_kernel_real ProbabilityTheory.lintegral_condKernelReal - -theorem ae_condKernelReal_eq_one {s : Set ℝ} (hs : MeasurableSet s) (hρ : ρ {x | x.snd ∈ sᶜ} = 0) : - ∀ᵐ a ∂ρ.fst, condKernelReal ρ a s = 1 := by - have h : ρ {x | x.snd ∈ sᶜ} = (ρ.fst ⊗ₘ condKernelReal ρ) {x | x.snd ∈ sᶜ} := by - rw [← measure_eq_compProd_real] - rw [hρ, Measure.compProd_apply] at h - swap; · exact measurable_snd hs.compl - rw [eq_comm, lintegral_eq_zero_iff] at h - swap - · simp only [mem_compl_iff, mem_setOf_eq] - exact kernel.measurable_coe _ hs.compl - simp only [mem_compl_iff, mem_setOf_eq, kernel.prodMkLeft_apply'] at h - filter_upwards [h] with a ha - change condKernelReal ρ a sᶜ = 0 at ha - rwa [prob_compl_eq_zero_iff hs] at ha -#align probability_theory.ae_cond_kernel_real_eq_one ProbabilityTheory.ae_condKernelReal_eq_one - -end Real - -section StandardBorel - -/-! ### Disintegration of measures on standard Borel spaces - -Since every standard Borel space embeds measurably into `ℝ`, we can generalize the disintegration -property on `ℝ` to all these spaces. -/ - - -variable {Ω : Type*} [MeasurableSpace Ω] [StandardBorelSpace Ω] - [Nonempty Ω] (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] - -/-- Existence of a conditional kernel. Use the definition `condKernel` to get that kernel. -/ -theorem exists_cond_kernel (γ : Type*) [MeasurableSpace γ] : - ∃ (η : kernel α Ω) (_h : IsMarkovKernel η), kernel.const γ ρ = - (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ η) := by - obtain ⟨f, hf⟩ := exists_measurableEmbedding_real Ω - let ρ' : Measure (α × ℝ) := ρ.map (Prod.map id f) - -- The general idea is to define `η = kernel.comapRight (condKernelReal ρ') hf`. There is - -- however an issue: that `η` may not be a Markov kernel since its value is only a - -- probability distribution almost everywhere with respect to `ρ.fst`, not everywhere. - -- We modify it to obtain a Markov kernel which is almost everywhere equal. - let ρ_set := (toMeasurable ρ.fst {a | condKernelReal ρ' a (range f) = 1}ᶜ)ᶜ - have hm : MeasurableSet ρ_set := (measurableSet_toMeasurable _ _).compl - have h_eq_one_of_mem : ∀ a ∈ ρ_set, condKernelReal ρ' a (range f) = 1 := by - intro a ha - rw [mem_compl_iff] at ha - have h_ss := subset_toMeasurable ρ.fst {a : α | condKernelReal ρ' a (range f) = 1}ᶜ - suffices ha' : a ∉ {a : α | condKernelReal ρ' a (range f) = 1}ᶜ by - rwa [not_mem_compl_iff] at ha' - exact not_mem_subset h_ss ha - have h_prod_embed : MeasurableEmbedding (Prod.map (id : α → α) f) := - MeasurableEmbedding.id.prod_mk hf - have h_fst : ρ'.fst = ρ.fst := by - ext1 u hu - rw [Measure.fst_apply hu, Measure.fst_apply hu, - Measure.map_apply h_prod_embed.measurable (measurable_fst hu)] - rfl - have h_ae : ∀ᵐ a ∂ρ.fst, a ∈ ρ_set := by - rw [ae_iff] - simp only [not_mem_compl_iff, setOf_mem_eq, measure_toMeasurable] - change ρ.fst {a : α | a ∉ {a' : α | condKernelReal ρ' a' (range f) = 1}} = 0 - rw [← ae_iff, ← h_fst] - refine' ae_condKernelReal_eq_one ρ' hf.measurableSet_range _ - rw [Measure.map_apply h_prod_embed.measurable] - swap; · exact measurable_snd hf.measurableSet_range.compl - convert measure_empty (α := α × Ω) - ext1 x - simp only [mem_compl_iff, mem_range, preimage_setOf_eq, Prod_map, mem_setOf_eq, - mem_empty_iff_false, iff_false_iff, Classical.not_not, exists_apply_eq_apply] - classical - obtain ⟨x₀, hx₀⟩ : ∃ x, x ∈ range f := range_nonempty _ - let η' := - kernel.piecewise hm (condKernelReal ρ') (kernel.deterministic (fun _ => x₀) measurable_const) - -- We show that `kernel.comapRight η' hf` is a suitable Markov kernel. - refine' ⟨kernel.comapRight η' hf, _, _⟩ - · refine' kernel.IsMarkovKernel.comapRight _ _ fun a => _ - rw [kernel.piecewise_apply'] - split_ifs with h_mem - · exact h_eq_one_of_mem _ h_mem - · rw [kernel.deterministic_apply' _ _ hf.measurableSet_range, Set.indicator_apply, if_pos hx₀] - have : kernel.const γ ρ = kernel.comapRight (kernel.const γ ρ') h_prod_embed := by - ext c t ht : 2 - rw [kernel.const_apply, kernel.comapRight_apply' _ _ _ ht, kernel.const_apply, - Measure.map_apply h_prod_embed.measurable (h_prod_embed.measurableSet_image.mpr ht)] - congr with x : 1 - rw [← @Prod.mk.eta _ _ x] - simp only [id.def, mem_preimage, Prod.map_mk, mem_image, Prod.mk.inj_iff, Prod.exists] - refine' ⟨fun h => ⟨x.1, x.2, h, rfl, rfl⟩, _⟩ - rintro ⟨a, b, h_mem, rfl, hf_eq⟩ - rwa [hf.injective hf_eq] at h_mem - rw [this, kernel.const_eq_compProd_real _ ρ'] - ext c t ht : 2 - rw [kernel.comapRight_apply' _ _ _ ht, - kernel.compProd_apply _ _ _ (h_prod_embed.measurableSet_image.mpr ht), kernel.const_apply, - h_fst, kernel.compProd_apply _ _ _ ht, kernel.const_apply] - refine' lintegral_congr_ae _ - filter_upwards [h_ae] with a ha - rw [kernel.prodMkLeft_apply', kernel.prodMkLeft_apply', kernel.comapRight_apply'] - swap - · exact measurable_prod_mk_left ht - have h1 : {c : ℝ | (a, c) ∈ Prod.map id f '' t} = f '' {c : Ω | (a, c) ∈ t} := by - ext1 x - simp only [Prod_map, id.def, mem_image, Prod.mk.inj_iff, Prod.exists, mem_setOf_eq] - constructor - · rintro ⟨a', b, h_mem, rfl, hf_eq⟩ - exact ⟨b, h_mem, hf_eq⟩ - · rintro ⟨b, h_mem, hf_eq⟩ - exact ⟨a, b, h_mem, rfl, hf_eq⟩ - have h2 : condKernelReal ρ' (c, a).snd = η' (c, a).snd := by - rw [kernel.piecewise_apply, if_pos ha] - rw [h1, h2] -#align probability_theory.exists_cond_kernel ProbabilityTheory.exists_cond_kernel - -/-- Conditional kernel of a measure on a product space: a Markov kernel such that -`ρ = ρ.fst ⊗ₘ ρ.condKernel` (see `ProbabilityTheory.measure_eq_compProd`). -/ -noncomputable irreducible_def _root_.MeasureTheory.Measure.condKernel : kernel α Ω := - (exists_cond_kernel ρ Unit).choose -#align measure_theory.measure.cond_kernel MeasureTheory.Measure.condKernel - -theorem condKernel_def : ρ.condKernel = (exists_cond_kernel ρ Unit).choose := by - rw [MeasureTheory.Measure.condKernel] -#align probability_theory.cond_kernel_def ProbabilityTheory.condKernel_def - -instance : IsMarkovKernel ρ.condKernel := by - rw [condKernel_def]; exact (exists_cond_kernel ρ Unit).choose_spec.choose - -theorem kernel.const_unit_eq_compProd : - kernel.const Unit ρ = kernel.const Unit ρ.fst ⊗ₖ kernel.prodMkLeft Unit ρ.condKernel := by - simp_rw [condKernel_def]; exact (exists_cond_kernel ρ Unit).choose_spec.choose_spec -#align probability_theory.kernel.const_unit_eq_comp_prod ProbabilityTheory.kernel.const_unit_eq_compProd - -/-- **Disintegration** of finite product measures on `α × Ω`, where `Ω` is standard Borel. Such a -measure can be written as the composition-product of the constant kernel with value `ρ.fst` -(marginal measure over `α`) and a Markov kernel from `α` to `Ω`. We call that Markov kernel -`ProbabilityTheory.condKernel ρ`. -/ -theorem measure_eq_compProd : ρ = ρ.fst ⊗ₘ ρ.condKernel := by - rw [Measure.compProd, ← kernel.const_unit_eq_compProd, kernel.const_apply] -#align probability_theory.measure_eq_comp_prod ProbabilityTheory.measure_eq_compProd - -/-- **Disintegration** of constant kernels. A constant kernel on a product space `α × Ω`, where `Ω` -is standard Borel, can be written as the composition-product of the constant kernel with -value `ρ.fst` (marginal measure over `α`) and a Markov kernel from `α` to `Ω`. We call that -Markov kernel `ProbabilityTheory.condKernel ρ`. -/ -theorem kernel.const_eq_compProd (γ : Type*) [MeasurableSpace γ] (ρ : Measure (α × Ω)) - [IsFiniteMeasure ρ] : - kernel.const γ ρ = kernel.const γ ρ.fst ⊗ₖ kernel.prodMkLeft γ ρ.condKernel := by - ext a s hs : 2 - simpa only [kernel.const_apply, kernel.compProd_apply _ _ _ hs, kernel.prodMkLeft_apply'] using - kernel.ext_iff'.mp (kernel.const_unit_eq_compProd ρ) () s hs -#align probability_theory.kernel.const_eq_comp_prod ProbabilityTheory.kernel.const_eq_compProd - -/-- Auxiliary lemma for `condKernel_apply_of_ne_zero`. -/ -lemma condKernel_apply_of_ne_zero_of_measurableSet [MeasurableSingletonClass α] - {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] - {x : α} (hx : ρ.fst {x} ≠ 0) {s : Set Ω} (hs : MeasurableSet s) : - ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s) := by - nth_rewrite 3 [measure_eq_compProd ρ] - rw [Measure.compProd_apply (measurableSet_prod.mpr (Or.inl ⟨measurableSet_singleton x, hs⟩))] - classical - have : ∀ a, ρ.condKernel a (Prod.mk a ⁻¹' {x} ×ˢ s) - = ({x} : Set α).indicator (fun a ↦ ρ.condKernel a s) a := by - intro a - by_cases hax : a = x - · simp only [hax, Set.singleton_prod, Set.mem_singleton_iff, Set.indicator_of_mem] - congr with y - simp - · simp only [Set.singleton_prod, Set.mem_singleton_iff, hax, not_false_eq_true, - Set.indicator_of_not_mem] - have : Prod.mk a ⁻¹' (Prod.mk x '' s) = ∅ := by - ext y - simp [Ne.symm hax] - simp only [this, measure_empty] - simp_rw [this] - rw [MeasureTheory.lintegral_indicator _ (measurableSet_singleton x)] - simp only [Measure.restrict_singleton, lintegral_smul_measure, lintegral_dirac] - rw [← mul_assoc, ENNReal.inv_mul_cancel hx (measure_ne_top ρ.fst _), one_mul] - -/-- If the singleton `{x}` has non-zero mass for `ρ.fst`, then for all `s : Set Ω`, -`ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s)` . -/ -lemma condKernel_apply_of_ne_zero [MeasurableSingletonClass α] - {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {x : α} (hx : ρ.fst {x} ≠ 0) - (s : Set Ω) : - ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s) := by - have : ρ.condKernel x s = ((ρ.fst {x})⁻¹ • ρ).comap (fun (y : Ω) ↦ (x, y)) s := by - congr 2 with s hs - simp [condKernel_apply_of_ne_zero_of_measurableSet hx hs, - (measurableEmbedding_prod_mk_left x).comap_apply] - simp [this, (measurableEmbedding_prod_mk_left x).comap_apply, hx] - -theorem lintegral_condKernel_mem {s : Set (α × Ω)} (hs : MeasurableSet s) : - ∫⁻ a, ρ.condKernel a {x | (a, x) ∈ s} ∂ρ.fst = ρ s := by - conv_rhs => rw [measure_eq_compProd ρ] - simp_rw [Measure.compProd_apply hs] - rfl -#align probability_theory.lintegral_cond_kernel_mem ProbabilityTheory.lintegral_condKernel_mem - -theorem set_lintegral_condKernel_eq_measure_prod {s : Set α} (hs : MeasurableSet s) {t : Set Ω} - (ht : MeasurableSet t) : ∫⁻ a in s, ρ.condKernel a t ∂ρ.fst = ρ (s ×ˢ t) := by - have : ρ (s ×ˢ t) = (ρ.fst ⊗ₘ ρ.condKernel) (s ×ˢ t) := by - congr; exact measure_eq_compProd ρ - rw [this, Measure.compProd_apply (hs.prod ht)] - classical - have : ∀ a, ρ.condKernel a (Prod.mk a ⁻¹' s ×ˢ t) - = s.indicator (fun a ↦ ρ.condKernel a t) a := by - intro a - by_cases ha : a ∈ s <;> simp [ha] - simp_rw [this] - rw [lintegral_indicator _ hs] -#align probability_theory.set_lintegral_cond_kernel_eq_measure_prod ProbabilityTheory.set_lintegral_condKernel_eq_measure_prod - -theorem lintegral_condKernel {f : α × Ω → ℝ≥0∞} (hf : Measurable f) : - ∫⁻ a, ∫⁻ ω, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫⁻ x, f x ∂ρ := by - conv_rhs => rw [measure_eq_compProd ρ] - rw [Measure.lintegral_compProd hf] -#align probability_theory.lintegral_cond_kernel ProbabilityTheory.lintegral_condKernel - -theorem set_lintegral_condKernel {f : α × Ω → ℝ≥0∞} (hf : Measurable f) {s : Set α} - (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) : - ∫⁻ a in s, ∫⁻ ω in t, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫⁻ x in s ×ˢ t, f x ∂ρ := by - conv_rhs => rw [measure_eq_compProd ρ] - rw [Measure.set_lintegral_compProd hf hs ht] -#align probability_theory.set_lintegral_cond_kernel ProbabilityTheory.set_lintegral_condKernel - -theorem set_lintegral_condKernel_univ_right {f : α × Ω → ℝ≥0∞} (hf : Measurable f) {s : Set α} - (hs : MeasurableSet s) : - ∫⁻ a in s, ∫⁻ ω, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫⁻ x in s ×ˢ univ, f x ∂ρ := by - rw [← set_lintegral_condKernel ρ hf hs MeasurableSet.univ]; simp_rw [Measure.restrict_univ] -#align probability_theory.set_lintegral_cond_kernel_univ_right ProbabilityTheory.set_lintegral_condKernel_univ_right - -theorem set_lintegral_condKernel_univ_left {f : α × Ω → ℝ≥0∞} (hf : Measurable f) {t : Set Ω} - (ht : MeasurableSet t) : - ∫⁻ a, ∫⁻ ω in t, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫⁻ x in univ ×ˢ t, f x ∂ρ := by - rw [← set_lintegral_condKernel ρ hf MeasurableSet.univ ht]; simp_rw [Measure.restrict_univ] -#align probability_theory.set_lintegral_cond_kernel_univ_left ProbabilityTheory.set_lintegral_condKernel_univ_left - -section IntegralCondKernel - -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] - -theorem _root_.MeasureTheory.AEStronglyMeasurable.integral_condKernel {ρ : Measure (α × Ω)} - [IsFiniteMeasure ρ] {f : α × Ω → E} (hf : AEStronglyMeasurable f ρ) : - AEStronglyMeasurable (fun x => ∫ y, f (x, y) ∂ρ.condKernel x) ρ.fst := by - rw [measure_eq_compProd ρ] at hf - exact AEStronglyMeasurable.integral_kernel_compProd hf -#align measure_theory.ae_strongly_measurable.integral_cond_kernel MeasureTheory.AEStronglyMeasurable.integral_condKernel - -theorem integral_condKernel {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} - (hf : Integrable f ρ) : ∫ a, ∫ x, f (a, x) ∂ρ.condKernel a ∂ρ.fst = ∫ ω, f ω ∂ρ := by - conv_rhs => rw [measure_eq_compProd ρ] - have hf': Integrable f (ρ.fst ⊗ₘ ρ.condKernel) := by rwa [measure_eq_compProd ρ] at hf - rw [Measure.integral_compProd hf'] -#align probability_theory.integral_cond_kernel ProbabilityTheory.integral_condKernel - -theorem set_integral_condKernel {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} - {s : Set α} (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) - (hf : IntegrableOn f (s ×ˢ t) ρ) : - ∫ a in s, ∫ ω in t, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫ x in s ×ˢ t, f x ∂ρ := by - conv_rhs => rw [measure_eq_compProd ρ] - rw [Measure.set_integral_compProd hs ht] - rwa [measure_eq_compProd ρ] at hf -#align probability_theory.set_integral_cond_kernel ProbabilityTheory.set_integral_condKernel - -theorem set_integral_condKernel_univ_right {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} - {s : Set α} (hs : MeasurableSet s) (hf : IntegrableOn f (s ×ˢ univ) ρ) : - ∫ a in s, ∫ ω, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫ x in s ×ˢ univ, f x ∂ρ := by - rw [← set_integral_condKernel hs MeasurableSet.univ hf]; simp_rw [Measure.restrict_univ] -#align probability_theory.set_integral_cond_kernel_univ_right ProbabilityTheory.set_integral_condKernel_univ_right - -theorem set_integral_condKernel_univ_left {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] {f : α × Ω → E} - {t : Set Ω} (ht : MeasurableSet t) (hf : IntegrableOn f (univ ×ˢ t) ρ) : - ∫ a, ∫ ω in t, f (a, ω) ∂ρ.condKernel a ∂ρ.fst = ∫ x in univ ×ˢ t, f x ∂ρ := by - rw [← set_integral_condKernel MeasurableSet.univ ht hf]; simp_rw [Measure.restrict_univ] -#align probability_theory.set_integral_cond_kernel_univ_left ProbabilityTheory.set_integral_condKernel_univ_left - -end IntegralCondKernel - -section Unique - -/-! ### Uniqueness - -This section will prove that the conditional kernel is unique almost everywhere. -/ - -/-- A s-finite kernel which satisfy the disintegration property of the given measure `ρ` is almost -everywhere equal to the disintegration kernel of `ρ` when evaluated on a measurable set. - -This theorem in the case of finite kernels is weaker than `eq_condKernel_of_measure_eq_compProd` -which asserts that the kernels are equal almost everywhere and not just on a given measurable -set. -/ -theorem eq_condKernel_of_measure_eq_compProd' (κ : kernel α Ω) [IsSFiniteKernel κ] - (hκ : ρ = ρ.fst ⊗ₘ κ) {s : Set Ω} (hs : MeasurableSet s) : - ∀ᵐ x ∂ρ.fst, κ x s = ρ.condKernel x s := by - refine' ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite - (kernel.measurable_coe κ hs) (kernel.measurable_coe ρ.condKernel hs) _ - intros t ht _ - conv_rhs => rw [set_lintegral_condKernel_eq_measure_prod _ ht hs, hκ] - simp only [Measure.compProd_apply (ht.prod hs), Set.mem_prod, ← lintegral_indicator _ ht] - congr; ext x - by_cases hx : x ∈ t - all_goals simp [hx] - --- This lemma establishes uniqueness of the disintegration kernel on ℝ -lemma eq_condKernel_of_measure_eq_compProd_real (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] - (κ : kernel α ℝ) [IsFiniteKernel κ] (hκ : ρ = ρ.fst ⊗ₘ κ) : - ∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by - have huniv : ∀ᵐ x ∂ρ.fst, κ x Set.univ = ρ.condKernel x Set.univ := - eq_condKernel_of_measure_eq_compProd' ρ κ hκ MeasurableSet.univ - suffices ∀ᵐ x ∂ρ.fst, ∀ ⦃t⦄, MeasurableSet t → κ x t = ρ.condKernel x t by - filter_upwards [this] with x hx - ext t ht; exact hx ht - apply MeasurableSpace.ae_induction_on_inter Real.borel_eq_generateFrom_Iic_rat - Real.isPiSystem_Iic_rat - · simp only [OuterMeasure.empty', Filter.eventually_true] - · simp only [iUnion_singleton_eq_range, mem_range, forall_exists_index, forall_apply_eq_imp_iff] - exact ae_all_iff.2 fun q => eq_condKernel_of_measure_eq_compProd' ρ κ hκ measurableSet_Iic - · filter_upwards [huniv] with x hxuniv t ht heq - rw [measure_compl ht <| measure_ne_top _ _, heq, hxuniv, measure_compl ht <| measure_ne_top _ _] - · refine' ae_of_all _ (fun x f hdisj hf heq => _) - rw [measure_iUnion hdisj hf, measure_iUnion hdisj hf] - exact tsum_congr heq - -/-- A finite kernel which satisfies the disintegration property is almost everywhere equal to the -disintegration kernel. -/ -theorem eq_condKernel_of_measure_eq_compProd (κ : kernel α Ω) [IsFiniteKernel κ] - (hκ : ρ = ρ.fst ⊗ₘ κ) : - ∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by --- The idea is to transporting the question to `ℝ` from `Ω` using `exists_measurableEmbedding_real` --- and then constructing a measure on `α × ℝ` using the obtained measurable embedding - obtain ⟨f, hf⟩ := exists_measurableEmbedding_real Ω - set ρ' : Measure (α × ℝ) := ρ.map (Prod.map id f) with hρ'def - have hρ' : ρ'.fst = ρ.fst := by - ext s hs - rw [hρ'def, Measure.fst_apply, Measure.fst_apply, Measure.map_apply] - exacts [rfl, Measurable.prod measurable_fst <| hf.measurable.comp measurable_snd, - measurable_fst hs, hs, hs] - have hρ'' : ∀ᵐ x ∂ρ'.fst, kernel.map κ f hf.measurable x = ρ'.condKernel x := by - refine' eq_condKernel_of_measure_eq_compProd_real ρ' (kernel.map κ f hf.measurable) _ - ext s hs - simp only [Measure.map_apply (measurable_id.prod_map hf.measurable) hs] - conv_lhs => congr; rw [hκ] - rw [Measure.compProd_apply hs, Measure.compProd_apply - (measurable_id.prod_map hf.measurable hs), (_ : (ρ.map (Prod.map id f)).fst = ρ.fst)] - · congr - ext x - simp only [Set.mem_preimage, Prod_map, id_eq, kernel.prodMkLeft_apply, kernel.map_apply] - rw [Measure.map_apply hf.measurable] - · rfl - · exact measurable_prod_mk_left hs - · rw [Measure.fst_map_prod_mk] - · simp only [Prod_map, id_eq]; rfl - · exact (hf.measurable.comp measurable_snd) - rw [hρ'] at hρ'' - suffices ∀ᵐ x ∂ρ.fst, ∀ s, MeasurableSet s → - ((ρ.map (Prod.map id f)).condKernel x) s = (ρ.condKernel x) (f ⁻¹' s) by - filter_upwards [hρ'', this] with x hx h - rw [kernel.map_apply] at hx - ext s hs - rw [← Set.preimage_image_eq s hf.injective, - ← Measure.map_apply hf.measurable <| hf.measurableSet_image.2 hs, hx, - h _ <| hf.measurableSet_image.2 hs] - have hprod : (ρ.map (Prod.map id f)).fst = ρ.fst := by - ext s hs - rw [Measure.fst_apply hs, - Measure.map_apply (measurable_id.prod_map hf.measurable) (measurable_fst hs), - ← Set.preimage_comp, Measure.fst_apply hs] - rfl - suffices ρ.map (Prod.map id f) = - ((ρ.map (Prod.map id f)).fst ⊗ₘ (kernel.map (Measure.condKernel ρ) f hf.measurable)) by - have heq := eq_condKernel_of_measure_eq_compProd_real _ _ this - rw [hprod] at heq - filter_upwards [heq] with x hx s hs - rw [← hx, kernel.map_apply, Measure.map_apply hf.measurable hs] - ext s hs - simp only [hprod, Measure.compProd_apply hs, kernel.map_apply _ hf.measurable] - rw [Measure.map_apply (measurable_id.prod_map hf.measurable) hs, ← lintegral_condKernel_mem] - · simp only [mem_preimage, Prod_map, id_eq] - congr with a - rw [Measure.map_apply hf.measurable (measurable_prod_mk_left hs)] - rfl - · exact measurable_id.prod_map hf.measurable hs - -end Unique - -end StandardBorel - -end ProbabilityTheory - -namespace MeasureTheory - -/-! ### Integrability - -We place these lemmas in the `MeasureTheory` namespace to enable dot notation. -/ - - -open ProbabilityTheory - -variable {α Ω E F : Type*} {mα : MeasurableSpace α} [MeasurableSpace Ω] - [StandardBorelSpace Ω] [Nonempty Ω] [NormedAddCommGroup E] [NormedSpace ℝ E] - [CompleteSpace E] [NormedAddCommGroup F] {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] - -theorem AEStronglyMeasurable.ae_integrable_condKernel_iff {f : α × Ω → F} - (hf : AEStronglyMeasurable f ρ) : - (∀ᵐ a ∂ρ.fst, Integrable (fun ω => f (a, ω)) (ρ.condKernel a)) ∧ - Integrable (fun a => ∫ ω, ‖f (a, ω)‖ ∂ρ.condKernel a) ρ.fst ↔ Integrable f ρ := by - rw [measure_eq_compProd ρ] at hf - conv_rhs => rw [measure_eq_compProd ρ] - rw [Measure.integrable_compProd_iff hf] -#align measure_theory.ae_strongly_measurable.ae_integrable_cond_kernel_iff MeasureTheory.AEStronglyMeasurable.ae_integrable_condKernel_iff - -theorem Integrable.condKernel_ae {f : α × Ω → F} (hf_int : Integrable f ρ) : - ∀ᵐ a ∂ρ.fst, Integrable (fun ω => f (a, ω)) (ρ.condKernel a) := by - have hf_ae : AEStronglyMeasurable f ρ := hf_int.1 - rw [← hf_ae.ae_integrable_condKernel_iff] at hf_int - exact hf_int.1 -#align measure_theory.integrable.cond_kernel_ae MeasureTheory.Integrable.condKernel_ae - -theorem Integrable.integral_norm_condKernel {f : α × Ω → F} (hf_int : Integrable f ρ) : - Integrable (fun x => ∫ y, ‖f (x, y)‖ ∂ρ.condKernel x) ρ.fst := by - have hf_ae : AEStronglyMeasurable f ρ := hf_int.1 - rw [← hf_ae.ae_integrable_condKernel_iff] at hf_int - exact hf_int.2 -#align measure_theory.integrable.integral_norm_cond_kernel MeasureTheory.Integrable.integral_norm_condKernel - -theorem Integrable.norm_integral_condKernel {f : α × Ω → E} (hf_int : Integrable f ρ) : - Integrable (fun x => ‖∫ y, f (x, y) ∂ρ.condKernel x‖) ρ.fst := by - refine' hf_int.integral_norm_condKernel.mono hf_int.1.integral_condKernel.norm _ - refine' eventually_of_forall fun x => _ - rw [norm_norm] - refine' (norm_integral_le_integral_norm _).trans_eq (Real.norm_of_nonneg _).symm - exact integral_nonneg_of_ae (eventually_of_forall fun y => norm_nonneg _) -#align measure_theory.integrable.norm_integral_cond_kernel MeasureTheory.Integrable.norm_integral_condKernel - -theorem Integrable.integral_condKernel {f : α × Ω → E} (hf_int : Integrable f ρ) : - Integrable (fun x => ∫ y, f (x, y) ∂ρ.condKernel x) ρ.fst := - (integrable_norm_iff hf_int.1.integral_condKernel).mp hf_int.norm_integral_condKernel -#align measure_theory.integrable.integral_cond_kernel MeasureTheory.Integrable.integral_condKernel - -end MeasureTheory From ac529500460ac6d12dbc169b5266351bb55d5f45 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 18 Feb 2024 16:20:17 +0100 Subject: [PATCH 033/129] Mathlib.lean --- Mathlib.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib.lean b/Mathlib.lean index 2d7b682ca5bbd..9b1baa8860d80 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3117,8 +3117,9 @@ import Mathlib.Probability.Kernel.Condexp import Mathlib.Probability.Kernel.Disintegration.Basic import Mathlib.Probability.Kernel.Disintegration.CDFKernel import Mathlib.Probability.Kernel.Disintegration.CondCdf +import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.Disintegration.Integral -import Mathlib.Probability.Kernel.Disintegration.KernelCDFBorel +import Mathlib.Probability.Kernel.Disintegration.KernelCDFReal import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes import Mathlib.Probability.Kernel.Disintegration.Unique import Mathlib.Probability.Kernel.IntegralCompProd From 3a77c06da60fb95757f5bf593df0207fd4d46e44 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 18 Feb 2024 16:22:27 +0100 Subject: [PATCH 034/129] fix name --- Mathlib/Probability/Kernel/Disintegration/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index b0fc4351b8789..3605950f91482 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.Disintegration.CondCdf -import Mathlib.Probability.Kernel.Disintegration.KernelCDFBorel +import Mathlib.Probability.Kernel.Disintegration.KernelCDFReal /-! # Disintegration of kernels and measures From 3f9519b688a292ae87c75dac4ce85a51ed9c1241 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 18 Feb 2024 17:14:52 +0100 Subject: [PATCH 035/129] rename --- .../Kernel/Disintegration/Density.lean | 279 +++++++++--------- 1 file changed, 141 insertions(+), 138 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 5a9c48828e7bb..f14504d9f0c29 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -74,18 +74,19 @@ variable {α β : Type*} {mα : MeasurableSpace α} section dissection_system -def I (n : ℕ) (k : ℤ) : Set ℝ := Set.Ico (k * (2⁻¹ : ℝ) ^ n) ((k + 1) * ((2 : ℝ) ^ n)⁻¹) +def IcoPowTwo (n : ℕ) (k : ℤ) : Set ℝ := Set.Ico (k * (2⁻¹ : ℝ) ^ n) ((k + 1) * ((2 : ℝ) ^ n)⁻¹) -lemma mem_I_iff_mul {n : ℕ} {k : ℤ} (x : ℝ) : x ∈ I n k ↔ k ≤ x * 2 ^ n ∧ x * 2 ^ n < k + 1 := by - simp only [I, inv_pow, mem_Ico] +lemma mem_IcoPowTwo_iff_mul {n : ℕ} {k : ℤ} (x : ℝ) : + x ∈ IcoPowTwo n k ↔ k ≤ x * 2 ^ n ∧ x * 2 ^ n < k + 1 := by + simp only [IcoPowTwo, inv_pow, mem_Ico] rw [← div_eq_mul_inv, div_le_iff, ← div_eq_mul_inv, lt_div_iff] · positivity · positivity -lemma mem_I_iff_floor {n : ℕ} {k : ℤ} (x : ℝ) : x ∈ I n k ↔ ⌊x * 2 ^ n⌋ = k := by - simp [mem_I_iff_mul, Int.floor_eq_iff] +lemma mem_IcoPowTwo_iff_floor {n : ℕ} {k : ℤ} (x : ℝ) : x ∈ IcoPowTwo n k ↔ ⌊x * 2 ^ n⌋ = k := by + simp [mem_IcoPowTwo_iff_mul, Int.floor_eq_iff] -lemma measurableSet_I (n : ℕ) (k : ℤ) : MeasurableSet (I n k) := measurableSet_Ico +lemma measurableSet_IcoPowTwo (n : ℕ) (k : ℤ) : MeasurableSet (IcoPowTwo n k) := measurableSet_Ico lemma Measure.iInf_Iic_gt_prod {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] {s : Set α} (hs : MeasurableSet s) (t : ℚ) : @@ -104,20 +105,21 @@ lemma Measure.iInf_Iic_gt_prod {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] exact mod_cast hrr' · exact ⟨⟨t + 1, lt_add_one _⟩, measure_ne_top ρ _⟩ -lemma pairwise_disjoint_I (n : ℕ) : Pairwise (Disjoint on fun k ↦ I n k) := by +lemma pairwise_disjoint_IcoPowTwo (n : ℕ) : Pairwise (Disjoint on fun k ↦ IcoPowTwo n k) := by intro i j hij rw [Function.onFun, Set.disjoint_iff] intro x - simp only [mem_inter_iff, mem_I_iff_floor, mem_empty_iff_false, and_imp, imp_false] + simp only [mem_inter_iff, mem_IcoPowTwo_iff_floor, mem_empty_iff_false, and_imp, imp_false] intro hi hj rw [hi] at hj exact hij hj -lemma I_succ_union (n : ℕ) (k : ℤ) : I (n+1) (2 * k) ∪ I (n+1) (2 * k + 1) = I n k := by +lemma IcoPowTwo_succ_union (n : ℕ) (k : ℤ) : + IcoPowTwo (n+1) (2 * k) ∪ IcoPowTwo (n+1) (2 * k + 1) = IcoPowTwo n k := by ext x cases lt_or_le x ((2 * k + 1) * ((2 : ℝ) ^ (n+1))⁻¹) with | inl h => - simp only [I, inv_pow, mem_Ico, Int.cast_mul, Int.int_cast_ofNat, Int.cast_add, + simp only [IcoPowTwo, inv_pow, mem_Ico, Int.cast_mul, Int.int_cast_ofNat, Int.cast_add, Int.cast_one, mem_union, h, and_true, not_le.mpr h, false_and, or_false] have : x < (k + 1) * (2 ^ n)⁻¹ := by refine h.trans_le ?_ @@ -130,7 +132,7 @@ lemma I_succ_union (n : ℕ) (k : ℤ) : I (n+1) (2 * k) ∪ I (n+1) (2 * k + 1) rw [pow_add, pow_one, mul_inv, mul_comm _ 2⁻¹, ← mul_assoc, mul_comm _ 2⁻¹, ← mul_assoc, inv_mul_cancel two_ne_zero, one_mul] | inr h => - simp only [I, inv_pow, mem_Ico, Int.cast_mul, Int.int_cast_ofNat, Int.cast_add, + simp only [IcoPowTwo, inv_pow, mem_Ico, Int.cast_mul, Int.int_cast_ofNat, Int.cast_add, Int.cast_one, mem_union, not_lt.mpr h, and_false, h, true_and, false_or] have : k * (2 ^ n)⁻¹ ≤ x := by refine le_trans ?_ h @@ -144,10 +146,10 @@ lemma I_succ_union (n : ℕ) (k : ℤ) : I (n+1) (2 * k) ∪ I (n+1) (2 * k + 1) norm_num rw [one_div, mul_add, ← mul_assoc, inv_mul_cancel two_ne_zero, one_mul] -noncomputable def indexI (n : ℕ) (t : ℝ) : ℤ := Int.floor (t * 2 ^ n) +noncomputable def indexIcoPowTwo (n : ℕ) (t : ℝ) : ℤ := Int.floor (t * 2 ^ n) -lemma mem_I_indexI (n : ℕ) (t : ℝ) : t ∈ I n (indexI n t) := by - rw [indexI, I] +lemma mem_IcoPowTwo_indexIcoPowTwo (n : ℕ) (t : ℝ) : t ∈ IcoPowTwo n (indexIcoPowTwo n t) := by + rw [indexIcoPowTwo, IcoPowTwo] simp only [inv_pow, mem_Ico] constructor · rw [← div_eq_mul_inv, div_le_iff] @@ -157,37 +159,40 @@ lemma mem_I_indexI (n : ℕ) (t : ℝ) : t ∈ I n (indexI n t) := by · exact Int.lt_floor_add_one (t * 2 ^ n) · positivity -lemma indexI_of_mem (n : ℕ) (k : ℤ) (t : ℝ) (ht : t ∈ I n k) : indexI n t = k := by - rw [indexI] - simp only [I, inv_pow, mem_Ico, ← div_eq_mul_inv] at ht +lemma indexIcoPowTwo_of_mem (n : ℕ) (k : ℤ) (t : ℝ) (ht : t ∈ IcoPowTwo n k) : + indexIcoPowTwo n t = k := by + rw [indexIcoPowTwo] + simp only [IcoPowTwo, inv_pow, mem_Ico, ← div_eq_mul_inv] at ht rw [div_le_iff, lt_div_iff] at ht · rw [Int.floor_eq_iff] exact ht · positivity · positivity -lemma mem_I_iff_indexI (n : ℕ) (k : ℤ) (t : ℝ) : t ∈ I n k ↔ indexI n t = k := - ⟨fun h ↦ indexI_of_mem n k t h, fun h ↦ h ▸ mem_I_indexI n t⟩ +lemma mem_IcoPowTwo_iff_indexIcoPowTwo (n : ℕ) (k : ℤ) (t : ℝ) : + t ∈ IcoPowTwo n k ↔ indexIcoPowTwo n t = k := + ⟨fun h ↦ indexIcoPowTwo_of_mem n k t h, fun h ↦ h ▸ mem_IcoPowTwo_indexIcoPowTwo n t⟩ -lemma iUnion_I (n : ℕ) : ⋃ k, I n k = univ := by +lemma iUnion_IcoPowTwo (n : ℕ) : ⋃ k, IcoPowTwo n k = univ := by ext x simp only [mem_iUnion, mem_univ, iff_true] - exact ⟨indexI n x, mem_I_indexI n x⟩ + exact ⟨indexIcoPowTwo n x, mem_IcoPowTwo_indexIcoPowTwo n x⟩ -lemma indexI_le_indexI_iff (n : ℕ) (t x : ℝ) : - indexI n t ≤ indexI n x ↔ ⌊t * 2 ^ n⌋ * (2 ^ n)⁻¹ ≤ x := by - simp only [indexI._eq_1] +lemma indexIcoPowTwo_le_indexIcoPowTwo_iff (n : ℕ) (t x : ℝ) : + indexIcoPowTwo n t ≤ indexIcoPowTwo n x ↔ ⌊t * 2 ^ n⌋ * (2 ^ n)⁻¹ ≤ x := by + simp only [indexIcoPowTwo] rw [← div_eq_mul_inv, div_le_iff, Int.le_floor] positivity -lemma iUnion_ge_I (n : ℕ) (t : ℝ) : - ⋃ (k) (_ : indexI n t ≤ k), I n k = Ici (⌊t * 2 ^ n⌋ * (2 ^ n)⁻¹ : ℝ) := by +lemma iUnion_ge_IcoPowTwo (n : ℕ) (t : ℝ) : + ⋃ (k) (_ : indexIcoPowTwo n t ≤ k), IcoPowTwo n k = Ici (⌊t * 2 ^ n⌋ * (2 ^ n)⁻¹ : ℝ) := by ext x - simp [mem_I_iff_indexI, indexI_le_indexI_iff] + simp [mem_IcoPowTwo_iff_indexIcoPowTwo, indexIcoPowTwo_le_indexIcoPowTwo_iff] -lemma iInter_biUnion_I (x : ℝ) : ⋂ n, ⋃ (k) (_ : indexI n x ≤ k), I n k = Ici x := by +lemma iInter_biUnion_I (x : ℝ) : + ⋂ n, ⋃ (k) (_ : indexIcoPowTwo n x ≤ k), IcoPowTwo n k = Ici x := by ext t - simp [iUnion_ge_I] + simp [iUnion_ge_IcoPowTwo] refine ⟨fun h ↦ ?_, fun h n ↦ le_trans ?_ h⟩ · by_contra h_lt push_neg at h_lt @@ -226,31 +231,31 @@ lemma iInter_biUnion_I (x : ℝ) : ⋂ n, ⋃ (k) (_ : indexI n x ≤ k), I n k -- todo : `Filtration` should be renamed to `filtration` def ℱ : Filtration ℕ (borel ℝ) where - seq := fun n ↦ MeasurableSpace.generateFrom {s | ∃ k, s = I n k} + seq := fun n ↦ MeasurableSpace.generateFrom {s | ∃ k, s = IcoPowTwo n k} mono' := by refine monotone_nat_of_le_succ ?_ intro n refine MeasurableSpace.generateFrom_le fun s ⟨k, hs⟩ ↦ ?_ - rw [hs, ← I_succ_union n k] + rw [hs, ← IcoPowTwo_succ_union n k] refine MeasurableSet.union ?_ ?_ · exact MeasurableSpace.measurableSet_generateFrom ⟨2 * k, rfl⟩ · exact MeasurableSpace.measurableSet_generateFrom ⟨2 * k + 1, rfl⟩ le' := fun n ↦ by refine MeasurableSpace.generateFrom_le fun s ⟨k, hs⟩ ↦ ?_ rw [hs] - exact measurableSet_I n k + exact measurableSet_IcoPowTwo n k -lemma measurableSet_ℱ_I (n : ℕ) (k : ℤ) : MeasurableSet[ℱ n] (I n k) := +lemma measurableSet_ℱ_IcoPowTwo (n : ℕ) (k : ℤ) : MeasurableSet[ℱ n] (IcoPowTwo n k) := MeasurableSpace.measurableSet_generateFrom ⟨k, rfl⟩ -lemma measurable_indexI (n : ℕ) : Measurable[ℱ n] (indexI n) := by - unfold indexI +lemma measurable_indexIcoPowTwo (n : ℕ) : Measurable[ℱ n] (indexIcoPowTwo n) := by refine @measurable_to_countable' ℤ ℝ _ _ (ℱ n) _ (fun k ↦ ?_) - have : (fun t ↦ ⌊t * (2 : ℝ) ^ n⌋) ⁻¹' {k} = I n k := by + have : (fun t ↦ ⌊t * (2 : ℝ) ^ n⌋) ⁻¹' {k} = IcoPowTwo n k := by ext t - simp only [mem_I_iff_floor, mem_preimage, mem_singleton_iff] + simp only [mem_IcoPowTwo_iff_floor, mem_preimage, mem_singleton_iff] + unfold indexIcoPowTwo rw [this] - exact measurableSet_ℱ_I n k + exact measurableSet_ℱ_IcoPowTwo n k lemma iSup_ℱ : ⨆ n, ℱ n = borel ℝ := by refine le_antisymm ?_ ?_ @@ -262,7 +267,7 @@ lemma iSup_ℱ : ⨆ n, ℱ n = borel ℝ := by refine MeasurableSet.iInter (fun n ↦ ?_) refine MeasurableSet.biUnion ?_ (fun k _ ↦ ?_) · exact to_countable _ - · exact le_iSup ℱ n _ (measurableSet_ℱ_I n k) + · exact le_iSup ℱ n _ (measurableSet_ℱ_IcoPowTwo n k) end dissection_system @@ -275,37 +280,39 @@ variable {κ : kernel α (ℝ × β)} {ν : kernel α ℝ} noncomputable def densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) (a : α) (t : ℝ) (s : Set β) : ℝ := - (κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t))).toReal + (κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) / ν a (IcoPowTwo n (indexIcoPowTwo n t))).toReal lemma densityProcess_def (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) (a : α) (s : Set β) : (fun t ↦ densityProcess κ ν n a t s) - = fun t ↦ (κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t))).toReal := + = fun t ↦ (κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) + / ν a (IcoPowTwo n (indexIcoPowTwo n t))).toReal := rfl lemma measurable_densityProcess_aux (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) {s : Set β} (hs : MeasurableSet s) : - Measurable (fun (p : α × ℝ) ↦ - κ p.1 (I n (indexI n p.2) ×ˢ s) / ν p.1 (I n (indexI n p.2))) := by - change Measurable ((fun (p : α × ℤ) ↦ - κ p.1 (I n p.2 ×ˢ s) / ν p.1 (I n p.2)) ∘ (fun (p : α × ℝ) ↦ (p.1, indexI n p.2))) - have h1 : Measurable (fun (p : α × ℤ) ↦ κ p.1 (I n p.2 ×ˢ s) / ν p.1 (I n p.2)) := by + Measurable (fun (p : α × ℝ) ↦ κ p.1 (IcoPowTwo n (indexIcoPowTwo n p.2) ×ˢ s) + / ν p.1 (IcoPowTwo n (indexIcoPowTwo n p.2))) := by + change Measurable ((fun (p : α × ℤ) ↦ κ p.1 (IcoPowTwo n p.2 ×ˢ s) / ν p.1 (IcoPowTwo n p.2)) + ∘ (fun (p : α × ℝ) ↦ (p.1, indexIcoPowTwo n p.2))) + have h1 : + Measurable (fun (p : α × ℤ) ↦ κ p.1 (IcoPowTwo n p.2 ×ˢ s) / ν p.1 (IcoPowTwo n p.2)) := by refine Measurable.div ?_ ?_ - · have h_swap : Measurable fun (p : ℤ × α) ↦ κ p.2 (I n p.1 ×ˢ s) := by + · have h_swap : Measurable fun (p : ℤ × α) ↦ κ p.2 (IcoPowTwo n p.1 ×ˢ s) := by refine measurable_uncurry_of_continuous_of_measurable - (u := fun k a ↦ κ a (I n k ×ˢ s)) ?_ ?_ + (u := fun k a ↦ κ a (IcoPowTwo n k ×ˢ s)) ?_ ?_ · exact fun _ ↦ continuous_bot - · exact fun _ ↦ kernel.measurable_coe _ ((measurableSet_I _ _).prod hs) - change Measurable ((fun (p : ℤ × α) ↦ κ p.2 (I n p.1 ×ˢ s)) ∘ Prod.swap) + · exact fun _ ↦ kernel.measurable_coe _ ((measurableSet_IcoPowTwo _ _).prod hs) + change Measurable ((fun (p : ℤ × α) ↦ κ p.2 (IcoPowTwo n p.1 ×ˢ s)) ∘ Prod.swap) exact h_swap.comp measurable_swap - · have h_swap : Measurable fun (p : ℤ × α) ↦ ν p.2 (I n p.1) := by + · have h_swap : Measurable fun (p : ℤ × α) ↦ ν p.2 (IcoPowTwo n p.1) := by refine measurable_uncurry_of_continuous_of_measurable - (u := fun k a ↦ ν a (I n k)) ?_ ?_ + (u := fun k a ↦ ν a (IcoPowTwo n k)) ?_ ?_ · exact fun _ ↦ continuous_bot - · exact fun _ ↦ kernel.measurable_coe _ (measurableSet_I _ _) - change Measurable ((fun (p : ℤ × α) ↦ ν p.2 (I n p.1)) ∘ Prod.swap) + · exact fun _ ↦ kernel.measurable_coe _ (measurableSet_IcoPowTwo _ _) + change Measurable ((fun (p : ℤ × α) ↦ ν p.2 (IcoPowTwo n p.1)) ∘ Prod.swap) exact h_swap.comp measurable_swap refine h1.comp (measurable_fst.prod_mk ?_) - exact (Measurable.mono (measurable_indexI n) (ℱ.le n) le_rfl).comp measurable_snd + exact (Measurable.mono (measurable_indexIcoPowTwo n) (ℱ.le n) le_rfl).comp measurable_snd lemma measurable_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) {s : Set β} (hs : MeasurableSet s) : @@ -328,11 +335,11 @@ lemma measurable_ℱ_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α rw [densityProcess_def] refine @Measurable.ennreal_toReal _ (ℱ n) _ ?_ refine Measurable.div ?_ ?_ - · change Measurable[ℱ n] ((fun k ↦ κ a (I n k ×ˢ s)) ∘ (indexI n)) - refine Measurable.comp ?_ (measurable_indexI n) + · change Measurable[ℱ n] ((fun k ↦ κ a (IcoPowTwo n k ×ˢ s)) ∘ (indexIcoPowTwo n)) + refine Measurable.comp ?_ (measurable_indexIcoPowTwo n) exact measurable_of_countable _ - · change Measurable[ℱ n] ((fun k ↦ ν a (I n k)) ∘ (indexI n)) - refine Measurable.comp ?_ (measurable_indexI n) + · change Measurable[ℱ n] ((fun k ↦ ν a (IcoPowTwo n k)) ∘ (indexIcoPowTwo n)) + refine Measurable.comp ?_ (measurable_indexIcoPowTwo n) exact measurable_of_countable _ lemma stronglyMeasurable_ℱ_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) @@ -349,16 +356,21 @@ lemma densityProcess_nonneg (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n 0 ≤ densityProcess κ ν n a t s := ENNReal.toReal_nonneg +lemma apply_IcoPowTwo_le_of_fst_le (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : ℝ) (s : Set β) : + κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) ≤ ν a (IcoPowTwo n (indexIcoPowTwo n t)) := by + calc κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) + ≤ kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) := by + rw [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] + refine measure_mono (fun x ↦ ?_) + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + _ ≤ ν a (IcoPowTwo n (indexIcoPowTwo n t)) := hκν a _ (measurableSet_IcoPowTwo _ _) + lemma densityProcess_le_one (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : ℝ) (s : Set β) : densityProcess κ ν n a t s ≤ 1 := by refine ENNReal.toReal_le_of_le_ofReal zero_le_one (ENNReal.div_le_of_le_mul ?_) rw [ENNReal.ofReal_one, one_mul] - calc κ a (I n (indexI n t) ×ˢ s) ≤ kernel.fst κ a (I n (indexI n t)) := by - rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] - refine measure_mono (fun x ↦ ?_) - simp only [mem_prod, mem_setOf_eq, and_imp] - exact fun h _ ↦ h - _ ≤ ν a (I n (indexI n t)) := hκν a _ (measurableSet_I _ _) + exact apply_IcoPowTwo_le_of_fst_le hκν n a t s lemma snorm_densityProcess_le (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (s : Set β) : snorm (fun t ↦ densityProcess κ ν n a t s) 1 (ν a) ≤ ν a univ := by @@ -377,39 +389,40 @@ lemma integrable_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel lemma set_integral_densityProcess_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) (k : ℤ) : - ∫ t in I n k, densityProcess κ ν n a t s ∂(ν a) = (κ a (I n k ×ˢ s)).toReal := by + ∫ t in IcoPowTwo n k, densityProcess κ ν n a t s ∂(ν a) + = (κ a (IcoPowTwo n k ×ˢ s)).toReal := by simp_rw [densityProcess] rw [integral_toReal] rotate_left · refine Measurable.aemeasurable ?_ have h := measurable_densityProcess_aux κ ν n hs - change Measurable - ((fun (p : α × ℝ) ↦ κ p.1 (I n (indexI n p.2) ×ˢ s) / ν p.1 (I n (indexI n p.2))) - ∘ (fun t ↦ (a, t))) + change Measurable ((fun (p : α × ℝ) ↦ κ p.1 (IcoPowTwo n (indexIcoPowTwo n p.2) ×ˢ s) + / ν p.1 (IcoPowTwo n (indexIcoPowTwo n p.2))) ∘ (fun t ↦ (a, t))) exact h.comp measurable_prod_mk_left · refine ae_of_all _ (fun t ↦ ?_) - by_cases h0 : ν a (I n (indexI n t)) = 0 - · suffices κ a (I n (indexI n t) ×ˢ s) = 0 by simp [h0, this] - have h0' : kernel.fst κ a (I n (indexI n t)) = 0 := - le_antisymm ((hκν a _ (measurableSet_I _ _)).trans h0.le) zero_le' - rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0' + by_cases h0 : ν a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 + · suffices κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) = 0 by simp [h0, this] + have h0' : kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 := + le_antisymm ((hκν a _ (measurableSet_IcoPowTwo _ _)).trans h0.le) zero_le' + rw [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] at h0' refine measure_mono_null (fun x ↦ ?_) h0' simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h · refine ENNReal.div_lt_top ?_ h0 exact measure_ne_top _ _ congr - have : ∫⁻ t in I n k, κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t)) ∂(ν a) - = ∫⁻ _ in I n k, κ a (I n k ×ˢ s) / ν a (I n k) ∂(ν a) := by - refine set_lintegral_congr_fun (measurableSet_I _ _) (ae_of_all _ (fun t ht ↦ ?_)) - rw [indexI_of_mem _ _ _ ht] + have : ∫⁻ t in IcoPowTwo n k, + κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) / ν a (IcoPowTwo n (indexIcoPowTwo n t)) ∂(ν a) + = ∫⁻ _ in IcoPowTwo n k, κ a (IcoPowTwo n k ×ˢ s) / ν a (IcoPowTwo n k) ∂(ν a) := by + refine set_lintegral_congr_fun (measurableSet_IcoPowTwo _ _) (ae_of_all _ (fun t ht ↦ ?_)) + rw [indexIcoPowTwo_of_mem _ _ _ ht] rw [this] simp only [lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter] - by_cases h0 : ν a (I n k) = 0 + by_cases h0 : ν a (IcoPowTwo n k) = 0 · simp only [h0, mul_zero] - have h0' : kernel.fst κ a (I n k) = 0 := - le_antisymm ((hκν a _ (measurableSet_I _ _)).trans h0.le) zero_le' - rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0' + have h0' : kernel.fst κ a (IcoPowTwo n k) = 0 := + le_antisymm ((hκν a _ (measurableSet_IcoPowTwo _ _)).trans h0.le) zero_le' + rw [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] at h0' refine (measure_mono_null ?_ h0').symm intro p simp only [mem_prod, mem_setOf_eq, and_imp] @@ -420,13 +433,13 @@ lemma set_integral_densityProcess_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKern lemma integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : ∫ t, densityProcess κ ν n a t s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by - rw [← integral_univ, ← iUnion_I n, iUnion_prod_const, measure_iUnion] + rw [← integral_univ, ← iUnion_IcoPowTwo n, iUnion_prod_const, measure_iUnion] rotate_left · intro i j hij simp only [Set.disjoint_prod, disjoint_self, bot_eq_empty] - exact Or.inl (pairwise_disjoint_I n hij) - · exact fun k ↦ (measurableSet_I n k).prod hs - rw [integral_iUnion (measurableSet_I n) (pairwise_disjoint_I n) + exact Or.inl (pairwise_disjoint_IcoPowTwo n hij) + · exact fun k ↦ (measurableSet_IcoPowTwo n k).prod hs + rw [integral_iUnion (measurableSet_IcoPowTwo n) (pairwise_disjoint_IcoPowTwo n) (integrable_densityProcess hκν n a hs).integrableOn] rw [ENNReal.tsum_toReal_eq (fun _ ↦ measure_ne_top _ _)] congr with k @@ -435,14 +448,14 @@ lemma integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - refine MeasurableSpace.induction_on_inter (m := ℱ n) (s := {s | ∃ k, s = I n k}) + refine MeasurableSpace.induction_on_inter (m := ℱ n) (s := {s | ∃ k, s = IcoPowTwo n k}) (C := fun A ↦ ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) rfl ?_ ?_ ?_ ?_ ?_ hA · rintro s ⟨i, rfl⟩ t ⟨j, rfl⟩ hst refine ⟨i, ?_⟩ suffices i = j by rw [this, inter_self] by_contra h_ne - have h_disj := pairwise_disjoint_I n h_ne + have h_disj := pairwise_disjoint_IcoPowTwo n h_ne rw [nonempty_iff_ne_empty] at hst refine hst ?_ rwa [Function.onFun, disjoint_iff_inter_eq_empty] at h_disj @@ -502,20 +515,16 @@ lemma densityProcess_mono_set (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) {s s' : Set β} (h : s ⊆ s') : densityProcess κ ν n a t s ≤ densityProcess κ ν n a t s' := by unfold densityProcess - by_cases h0 : ν a (I n (indexI n t)) = 0 + by_cases h0 : ν a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] simp - have h_ne_top : ∀ s, κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t)) ≠ ⊤ := by + have h_ne_top : ∀ s, κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) + / ν a (IcoPowTwo n (indexIcoPowTwo n t)) ≠ ⊤ := by intro s rw [ne_eq, ENNReal.div_eq_top] simp only [ne_eq, h0, and_false, false_or, not_and, not_not] refine fun h_top ↦ eq_top_mono ?_ h_top - calc κ a (I n (indexI n t) ×ˢ s) ≤ kernel.fst κ a (I n (indexI n t)) := by - rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] - refine measure_mono (fun x ↦ ?_) - simp only [mem_prod, mem_setOf_eq, and_imp] - exact fun h _ ↦ h - _ ≤ ν a (I n (indexI n t)) := hκν a _ (measurableSet_I _ _) + exact apply_IcoPowTwo_le_of_fst_le hκν n a t s rw [ENNReal.toReal_le_toReal] · gcongr rw [prod_subset_prod_iff] @@ -527,23 +536,18 @@ lemma densityProcess_mono_kernel_left {κ' : kernel α (ℝ × β)} (hκκ' : κ (hκ'ν : kernel.fst κ' ≤ ν) (n : ℕ) (a : α) (t : ℝ) {s : Set β} (hs : MeasurableSet s) : densityProcess κ ν n a t s ≤ densityProcess κ' ν n a t s := by unfold densityProcess - by_cases h0 : ν a (I n (indexI n t)) = 0 + by_cases h0 : ν a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] simp - have h_le : κ' a (I n (indexI n t) ×ˢ s) ≤ ν a (I n (indexI n t)) := by - calc κ' a (I n (indexI n t) ×ˢ s) ≤ kernel.fst κ' a (I n (indexI n t)) := by - rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] - refine measure_mono (fun x ↦ ?_) - simp only [mem_prod, mem_setOf_eq, and_imp] - exact fun h _ ↦ h - _ ≤ ν a (I n (indexI n t)) := hκ'ν a _ (measurableSet_I _ _) + have h_le : κ' a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) + ≤ ν a (IcoPowTwo n (indexIcoPowTwo n t)) := apply_IcoPowTwo_le_of_fst_le hκ'ν n a t s rw [ENNReal.toReal_le_toReal] · gcongr - exact hκκ' _ _ ((measurableSet_I _ _).prod hs) + exact hκκ' _ _ ((measurableSet_IcoPowTwo _ _).prod hs) · rw [ne_eq, ENNReal.div_eq_top] simp only [ne_eq, h0, and_false, false_or, not_and, not_not] refine fun h_top ↦ eq_top_mono ?_ h_top - exact (hκκ' _ _ ((measurableSet_I _ _).prod hs)).trans h_le + exact (hκκ' _ _ ((measurableSet_IcoPowTwo _ _).prod hs)).trans h_le · rw [ne_eq, ENNReal.div_eq_top] simp only [ne_eq, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top @@ -552,26 +556,21 @@ lemma densityProcess_antitone_kernel_right {ν' : kernel α ℝ} (hνν' : ν ≤ ν') (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : ℝ) (s : Set β) : densityProcess κ ν' n a t s ≤ densityProcess κ ν n a t s := by unfold densityProcess - have h_le : κ a (I n (indexI n t) ×ˢ s) ≤ ν a (I n (indexI n t)) := by - calc κ a (I n (indexI n t) ×ˢ s) ≤ kernel.fst κ a (I n (indexI n t)) := by - rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] - refine measure_mono (fun x ↦ ?_) - simp only [mem_prod, mem_setOf_eq, and_imp] - exact fun h _ ↦ h - _ ≤ ν a (I n (indexI n t)) := hκν a _ (measurableSet_I _ _) - by_cases h0 : ν a (I n (indexI n t)) = 0 - · suffices κ a (I n (indexI n t) ×ˢ s) = 0 by + have h_le : κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) + ≤ ν a (IcoPowTwo n (indexIcoPowTwo n t)) := apply_IcoPowTwo_le_of_fst_le hκν n a t s + by_cases h0 : ν a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 + · suffices κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) = 0 by simp only [this, ENNReal.zero_div, ENNReal.zero_toReal, h0, le_refl] exact le_antisymm (h_le.trans h0.le) zero_le' - have h0' : ν' a (I n (indexI n t)) ≠ 0 := by + have h0' : ν' a (IcoPowTwo n (indexIcoPowTwo n t)) ≠ 0 := by refine fun h ↦ h0 (le_antisymm (le_trans ?_ h.le) zero_le') - exact hνν' _ _ (measurableSet_I _ _) + exact hνν' _ _ (measurableSet_IcoPowTwo _ _) rw [ENNReal.toReal_le_toReal] · gcongr - exact hνν' _ _ (measurableSet_I _ _) + exact hνν' _ _ (measurableSet_IcoPowTwo _ _) · simp only [ne_eq, ENNReal.div_eq_top, h0', and_false, false_or, not_and, not_not] refine fun h_top ↦ eq_top_mono ?_ h_top - exact h_le.trans (hνν' _ _ (measurableSet_I _ _)) + exact h_le.trans (hνν' _ _ (measurableSet_IcoPowTwo _ _)) · simp only [ne_eq, ENNReal.div_eq_top, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top @@ -585,7 +584,7 @@ lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (ℝ × β) (hs_meas : ∀ n, MeasurableSet (s n)) : Tendsto (fun m ↦ densityProcess κ ν n a t (s m)) atTop (𝓝 (densityProcess κ ν n a t ∅)) := by simp_rw [densityProcess] - by_cases h0 : ν a (I n (indexI n t)) = 0 + by_cases h0 : ν a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 · simp_rw [h0, ENNReal.toReal_div] simp refine (ENNReal.tendsto_toReal ?_).comp ?_ @@ -593,10 +592,11 @@ lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (ℝ × β) push_neg simp refine ENNReal.Tendsto.div_const ?_ ?_ - · have h := tendsto_measure_iInter (μ := κ a) (s := fun m ↦ I n (indexI n t) ×ˢ s m) ?_ ?_ ?_ + · have h := tendsto_measure_iInter (μ := κ a) + (s := fun m ↦ IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s m) ?_ ?_ ?_ · convert h rw [← prod_iInter, hs_iInter] - · exact fun n ↦ MeasurableSet.prod (measurableSet_I _ _) (hs_meas n) + · exact fun n ↦ MeasurableSet.prod (measurableSet_IcoPowTwo _ _) (hs_meas n) · intro m m' hmm' simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] exact Or.inl <| hs hmm' @@ -910,39 +910,40 @@ section UnivFst lemma densityProcess_univ [IsFiniteKernel κ] (n : ℕ) (a : α) (t : ℝ) : densityProcess κ (kernel.fst κ) n a t univ - = if kernel.fst κ a (I n (indexI n t)) = 0 then 0 else 1 := by + = if kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 then 0 else 1 := by rw [densityProcess] - by_cases h : kernel.fst κ a (I n (indexI n t)) = 0 + by_cases h : kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 · simp [h] - by_cases h' : κ a (I n (indexI n t) ×ˢ univ) = 0 + by_cases h' : κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ univ) = 0 · simp [h'] · rw [ENNReal.div_zero h'] simp · simp only [h, ite_false] - rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] - have : I n (indexI n t) ×ˢ univ = {p : ℝ × β | p.1 ∈ I n (indexI n t)} := by + rw [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] + have : IcoPowTwo n (indexIcoPowTwo n t) ×ˢ univ + = {p : ℝ × β | p.1 ∈ IcoPowTwo n (indexIcoPowTwo n t)} := by ext x simp rw [this, ENNReal.div_self] · simp - · rwa [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h + · rwa [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] at h · exact measure_ne_top _ _ lemma densityProcess_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (n : ℕ) (a : α) : ∀ᵐ t ∂(kernel.fst κ a), densityProcess κ (kernel.fst κ) n a t univ = 1 := by rw [ae_iff] have : {t | ¬ densityProcess κ (kernel.fst κ) n a t univ = 1} - ⊆ {t | kernel.fst κ a (I n (indexI n t)) = 0} := by + ⊆ {t | kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) = 0} := by intro t ht simp only [mem_setOf_eq] at ht ⊢ rw [densityProcess_univ] at ht simpa using ht refine measure_mono_null this ?_ - have : {t | kernel.fst κ a (I n (indexI n t)) = 0} - ⊆ ⋃ (k) (_ : kernel.fst κ a (I n k) = 0), I n k := by + have : {t | kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) = 0} + ⊆ ⋃ (k) (_ : kernel.fst κ a (IcoPowTwo n k) = 0), IcoPowTwo n k := by intro t ht simp only [mem_setOf_eq, mem_iUnion, exists_prop] at ht ⊢ - exact ⟨indexI n t, ht, mem_I_indexI _ _⟩ + exact ⟨indexIcoPowTwo n t, ht, mem_IcoPowTwo_indexIcoPowTwo _ _⟩ refine measure_mono_null this ?_ rw [measure_iUnion_null] simp @@ -955,7 +956,7 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) refine (ENNReal.tendsto_toReal ?_).comp ?_ · rw [ne_eq, ENNReal.div_eq_top] push_neg - simp_rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] + simp_rw [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] constructor · refine fun h h0 ↦ h (measure_mono_null (fun x ↦ ?_) h0) simp only [mem_prod, mem_setOf_eq, and_imp] @@ -963,11 +964,12 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) · refine fun h_top ↦ eq_top_mono (measure_mono (fun x ↦ ?_)) h_top simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h - by_cases h0 : kernel.fst κ a (I n (indexI n t)) = 0 - · rw [kernel.fst_apply' _ _ (measurableSet_I _ _)] at h0 ⊢ - suffices ∀ m, κ a (I n (indexI n t) ×ˢ s m) = 0 by + by_cases h0 : kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 + · rw [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] at h0 ⊢ + suffices ∀ m, κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s m) = 0 by simp only [this, h0, ENNReal.zero_div, tendsto_const_nhds_iff] - suffices κ a (I n (indexI n t) ×ˢ univ) = 0 by simp only [this, ENNReal.zero_div] + suffices κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ univ) = 0 by + simp only [this, ENNReal.zero_div] convert h0 ext x simp only [mem_prod, mem_univ, and_true, mem_setOf_eq] @@ -975,7 +977,8 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h refine ENNReal.Tendsto.div_const ?_ ?_ - · have h := tendsto_measure_iUnion (μ := κ a) (s := fun m ↦ I n (indexI n t) ×ˢ s m) ?_ + · have h := tendsto_measure_iUnion (μ := κ a) + (s := fun m ↦ IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s m) ?_ swap · intro m m' hmm' simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] From e6145fd20da5f08e988deeb44ec08be00f1ca4f4 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 19 Feb 2024 16:24:44 +0100 Subject: [PATCH 036/129] add RadonNikodym --- .../Kernel/Disintegration/RadonNikodym.lean | 412 ++++++++++++++++++ 1 file changed, 412 insertions(+) create mode 100644 Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean diff --git a/Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean b/Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean new file mode 100644 index 0000000000000..266139c1047a3 --- /dev/null +++ b/Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean @@ -0,0 +1,412 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Probability.Kernel.Disintegration.KernelCDFReal +import Mathlib.Probability.Kernel.WithDensity + +/-! +# Density + +## Main definitions + +* `FooBar` + +## Main statements + +* `fooBar_unique` + +## Notation + + + +## Implementation details + + + +## References + +* [F. Bar, *Quuxes*][bibkey] + +## Tags + +Foobars, barfoos +-/ + +open MeasureTheory Set Filter + +open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory + +namespace ProbabilityTheory + +variable {α : Type*} {mα : MeasurableSpace α} + +section aux + +lemma MeasurableEmbedding.comap_add {β : Type*} {mβ : MeasurableSpace β} + {f : α → β} (hf : MeasurableEmbedding f) {μ ν : Measure β} : + (μ + ν).comap f = μ.comap f + ν.comap f := by + ext s hs + rw [← Measure.comapₗ_eq_comap _ hf.injective (fun s hs ↦ hf.measurableSet_image.mpr hs) _ hs, + map_add] + simp only [Measure.add_toOuterMeasure, OuterMeasure.coe_add, Pi.add_apply] + rw [← Measure.comapₗ_eq_comap _ hf.injective (fun s hs ↦ hf.measurableSet_image.mpr hs) _ hs, + ← Measure.comapₗ_eq_comap _ hf.injective (fun s hs ↦ hf.measurableSet_image.mpr hs) _ hs] + +@[simp] +lemma kernel.map_id {β : Type*} {_ : MeasurableSpace β} {κ : kernel α β} : + kernel.map κ id measurable_id = κ := by ext a; rw [kernel.map_apply]; simp + +@[simp] +lemma kernel.map_id' {β : Type*} {_ : MeasurableSpace β} {κ : kernel α β} : + kernel.map κ (fun a ↦ a) measurable_id = κ := kernel.map_id + +end aux + +section withDensity + +variable {β : Type*} {mβ : MeasurableSpace β} {κ : kernel α β} + +@[simp] +lemma Measure.absolutelyContinuous_zero (μ : Measure α) : 0 ≪ μ := + Measure.absolutelyContinuous_of_le (Measure.zero_le μ) + +nonrec lemma kernel.withDensity_absolutelyContinuous [IsSFiniteKernel κ] + (f : α → β → ℝ≥0∞) (a : α) : + kernel.withDensity κ f a ≪ κ a := by + by_cases hf : Measurable (Function.uncurry f) + · rw [kernel.withDensity_apply _ hf] + exact withDensity_absolutelyContinuous _ _ + · rw [withDensity_of_not_measurable _ hf] + simp + +@[simp] +lemma kernel.withDensity_one [IsSFiniteKernel κ] : + kernel.withDensity κ 1 = κ := by + ext a + rw [kernel.withDensity_apply _ measurable_const] + simp + +@[simp] +lemma kernel.withDensity_one' [IsSFiniteKernel κ] : + kernel.withDensity κ (fun _ _ ↦ 1) = κ := kernel.withDensity_one + +lemma kernel.withDensity_sub_add [IsSFiniteKernel κ] {f g : α → β → ℝ≥0∞} + (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) + (hg_int : ∀ a, ∫⁻ x, g a x ∂(κ a) ≠ ∞) (hfg : ∀ a, g a ≤ᵐ[κ a] f a) : + kernel.withDensity κ (fun a x ↦ f a x - g a x) + kernel.withDensity κ g + = kernel.withDensity κ f := by + ext a s + simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add] + rw [kernel.withDensity_apply' _ hf, kernel.withDensity_apply' _ hg, kernel.withDensity_apply'] + swap; · exact hf.sub hg + rw [lintegral_sub] + · rw [tsub_add_cancel_iff_le] + exact lintegral_mono_ae (ae_restrict_of_ae (hfg a)) + · exact hg.comp measurable_prod_mk_left + · exact ((set_lintegral_le_lintegral _ _).trans_lt (hg_int a).lt_top).ne + · exact ae_restrict_of_ae (hfg a) + +nonrec lemma kernel.withDensity_mul [IsSFiniteKernel κ] {f : α → β → ℝ≥0} {g : α → β → ℝ≥0∞} + (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) : + kernel.withDensity κ (fun a x ↦ f a x * g a x) + = kernel.withDensity (kernel.withDensity κ fun a x ↦ f a x) fun a x ↦ g a x := by + ext a : 1 + rw [kernel.withDensity_apply] + swap; · exact (measurable_coe_nnreal_ennreal.comp hf).mul hg + suffices (Measure.withDensity (κ a) ((fun x ↦ (f a x : ℝ≥0∞)) * (fun x ↦ (g a x : ℝ≥0∞)))) = + (withDensity (withDensity κ fun a x ↦ f a x) fun a x ↦ g a x) a by + convert this + rw [withDensity_mul] + · rw [kernel.withDensity_apply _ hg, kernel.withDensity_apply] + exact measurable_coe_nnreal_ennreal.comp hf + · rw [measurable_coe_nnreal_ennreal_iff] + exact hf.comp measurable_prod_mk_left + · exact hg.comp measurable_prod_mk_left + +end withDensity + +section Real + +variable {κ ν : kernel α ℝ} + +lemma fst_map_le_of_le (hκν : κ ≤ ν) : + kernel.fst (kernel.map κ (fun a ↦ (a, ())) + (@measurable_prod_mk_right ℝ Unit _ inferInstance _)) ≤ ν := by + rwa [kernel.fst_map_prod _ measurable_id' measurable_const, kernel.map_id'] + +lemma todo (κ : kernel α ℝ) (ν : kernel α ℝ) : κ ≤ κ + ν := by + -- todo improve this: `le_add_of_nonneg_right` should work on kernels + intro a + simp only [AddSubmonoid.coe_add, Pi.add_apply] + exact le_add_of_nonneg_right (Measure.zero_le (ν a)) + +noncomputable +def g (κ ν : kernel α ℝ) (a : α) (x : ℝ) : ℝ := + MLimsup (kernel.map κ (fun a ↦ (a, ())) + (@measurable_prod_mk_right ℝ Unit _ inferInstance _)) ν a x univ + +lemma g_nonneg (hκν : κ ≤ ν) {a : α} {x : ℝ} : 0 ≤ g κ ν a x := + mLimsup_nonneg (fst_map_le_of_le hκν) _ _ _ + +lemma g_le_one (hκν : κ ≤ ν) {a : α} {x : ℝ} : g κ ν a x ≤ 1 := + mLimsup_le_one (fst_map_le_of_le hκν) _ _ _ + +lemma measurable_g (κ : kernel α ℝ) (ν : kernel α ℝ) : + Measurable (fun p : α × ℝ ↦ g κ ν p.1 p.2) := + measurable_mLimsup _ ν MeasurableSet.univ + +lemma measurable_g_right (κ : kernel α ℝ) (ν : kernel α ℝ) (a : α) : + Measurable (fun x : ℝ ↦ g κ ν a x) := by + change Measurable ((fun p : α × ℝ ↦ g κ ν p.1 p.2) ∘ (fun x ↦ (a, x))) + exact (measurable_g _ _).comp measurable_prod_mk_left + +noncomputable +def kernel.rnDerivReal (κ ν : kernel α ℝ) (a : α) (x : ℝ) : ℝ≥0∞ := + ENNReal.ofReal (g κ (κ + ν) a x) / ENNReal.ofReal (1 - g κ (κ + ν) a x) + +lemma kernel.rnDerivReal_def (κ ν : kernel α ℝ) : + kernel.rnDerivReal κ ν + = fun a x ↦ ENNReal.ofReal (g κ (κ + ν) a x) / ENNReal.ofReal (1 - g κ (κ + ν) a x) := rfl + +lemma measurable_rnDerivReal : Measurable (fun p : α × ℝ ↦ kernel.rnDerivReal κ ν p.1 p.2) := + (measurable_g κ _).ennreal_ofReal.div (measurable_const.sub (measurable_g κ _)).ennreal_ofReal + +lemma measurable_rnDerivReal_right (κ ν : kernel α ℝ) (a : α) : + Measurable (fun x : ℝ ↦ kernel.rnDerivReal κ ν a x) := by + change Measurable ((fun p : α × ℝ ↦ kernel.rnDerivReal κ ν p.1 p.2) ∘ (fun x ↦ (a, x))) + exact measurable_rnDerivReal.comp measurable_prod_mk_left + +lemma withDensity_g (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFiniteKernel ν] : + kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (g κ (κ + ν) a x)) = κ := by + have h_le : κ ≤ κ + ν := todo κ ν + ext a s hs + rw [kernel.withDensity_apply'] + swap; exact (measurable_g _ _).ennreal_ofReal + have : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl + simp_rw [this] + rw [← ofReal_integral_eq_lintegral_ofReal] + · unfold g + rw [set_integral_mLimsup (fst_map_le_of_le h_le) a MeasurableSet.univ hs, + ENNReal.ofReal_toReal, kernel.map_apply' _ _ _ (hs.prod MeasurableSet.univ)] + · congr with x + simp + · exact measure_ne_top _ _ + · exact (integrable_mLimsup (fst_map_le_of_le h_le) a MeasurableSet.univ).restrict + · exact ae_of_all _ (fun x ↦ mLimsup_nonneg (fst_map_le_of_le h_le) _ _ _) + +lemma withDensity_one_sub_g (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFiniteKernel ν] : + kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (1 - g κ (κ + ν) a x)) = ν := by + have h_le : κ ≤ κ + ν := todo κ ν + suffices kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (1 - g κ (κ + ν) a x)) + + kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (g κ (κ + ν) a x)) = κ + ν by + ext a s + have h : (kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (1 - g κ (κ + ν) a x)) + + kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (g κ (κ + ν) a x))) a s + = κ a s + ν a s := by + rw [this] + simp + simp only [kernel.coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add] + at h + rwa [withDensity_g, add_comm, ENNReal.add_right_inj (measure_ne_top _ _)] at h + have h_nonneg : ∀ a x, 0 ≤ g κ (κ + ν) a x := + fun _ _ ↦ mLimsup_nonneg (fst_map_le_of_le h_le) _ _ _ + have : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl + simp_rw [this, ENNReal.ofReal_sub _ (h_nonneg _ _), ENNReal.ofReal_one] + rw [kernel.withDensity_sub_add] + · rw [kernel.withDensity_one'] + · exact measurable_const + · exact (measurable_g _ _).ennreal_ofReal + · refine fun a ↦ ne_of_lt ?_ + calc ∫⁻ x, ENNReal.ofReal (g κ (κ + ν) a x) ∂(κ + ν) a + ≤ ∫⁻ _, 1 ∂(κ + ν) a := by + refine lintegral_mono (fun x ↦ ?_) + simp [g_le_one (todo κ ν)] + _ < ⊤ := by rw [MeasureTheory.lintegral_const, one_mul]; exact measure_lt_top _ _ + · refine fun a ↦ ae_of_all _ (fun x ↦ ?_) + simp only [ENNReal.ofReal_le_one] + exact mLimsup_le_one (fst_map_le_of_le h_le) _ _ _ + +noncomputable +def kernel.singularPartReal (κ ν : kernel α ℝ) [IsSFiniteKernel κ] [IsSFiniteKernel ν] : + kernel α ℝ := + kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (g κ (κ + ν) a x) + - Real.toNNReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x) + +lemma kernel.mutuallySingular_singularPartReal (κ ν : kernel α ℝ) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + kernel.singularPartReal κ ν a ⟂ₘ ν a := by + symm + have h_coe : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl + refine ⟨{x | g κ (κ + ν) a x = 1}, measurable_g_right _ _ a (measurableSet_singleton _), ?_, ?_⟩ + · suffices kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (1 - g κ (κ + ν) a x)) a + {x | g κ (κ + ν) a x = 1} = 0 by + rwa [withDensity_one_sub_g κ ν] at this + simp_rw [h_coe] + rw [kernel.withDensity_apply', lintegral_eq_zero_iff, EventuallyEq, ae_restrict_iff] + rotate_left + · exact (measurable_const.sub ((measurable_g _ _).comp measurable_prod_mk_left)).ennreal_ofReal + (measurableSet_singleton _) + · exact (measurable_const.sub ((measurable_g _ _).comp measurable_prod_mk_left)).ennreal_ofReal + · exact (measurable_const.sub (measurable_g _ _)).ennreal_ofReal + refine ae_of_all _ (fun x hx ↦ ?_) + simp only [mem_setOf_eq] at hx + simp [hx] + · rw [kernel.singularPartReal, kernel.withDensity_apply', lintegral_eq_zero_iff, EventuallyEq, + ae_restrict_iff] + all_goals simp_rw [h_coe] + rotate_left + · refine measurableSet_preimage (Measurable.sub ?_ ?_) (measurableSet_singleton _) + · exact (measurable_g_right _ _ _).ennreal_ofReal + · refine Measurable.mul ?_ ?_ + · exact (measurable_const.sub (measurable_g_right _ _ _)).ennreal_ofReal + · exact measurable_rnDerivReal_right _ _ _ + · refine Measurable.sub ?_ ?_ + · exact (measurable_g_right _ _ _).ennreal_ofReal + · refine Measurable.mul ?_ ?_ + · exact (measurable_const.sub (measurable_g_right _ _ _)).ennreal_ofReal + · exact measurable_rnDerivReal_right _ _ _ + · refine Measurable.sub ?_ ?_ + · exact (measurable_g _ _).ennreal_ofReal + · refine Measurable.mul ?_ ?_ + · exact (measurable_const.sub (measurable_g _ _)).ennreal_ofReal + · exact measurable_rnDerivReal + refine ae_of_all _ (fun x hx ↦ ?_) + simp only [mem_compl_iff, mem_setOf_eq] at hx + simp_rw [rnDerivReal] + rw [← ENNReal.ofReal_div_of_pos, div_eq_inv_mul, ← ENNReal.ofReal_mul, ← mul_assoc, + mul_inv_cancel, one_mul, tsub_self] + · rfl + · rw [ne_eq, sub_eq_zero] + exact Ne.symm hx + · simp [g_le_one (todo κ ν)] + · simp [lt_of_le_of_ne (g_le_one (todo κ ν)) hx] + +lemma kernel.rnDerivReal_add_singularPartReal (κ ν : kernel α ℝ) + [IsFiniteKernel κ] [IsFiniteKernel ν] : + kernel.withDensity ν (kernel.rnDerivReal κ ν) + kernel.singularPartReal κ ν = κ := by + have : kernel.withDensity ν (kernel.rnDerivReal κ ν) + = kernel.withDensity (kernel.withDensity (κ + ν) + (fun a x ↦ Real.toNNReal (1 - g κ (κ + ν) a x))) (kernel.rnDerivReal κ ν) := by + rw [kernel.rnDerivReal_def] + congr + exact (withDensity_one_sub_g κ ν).symm + rw [this, kernel.singularPartReal, add_comm, ← kernel.withDensity_mul] + rotate_left + · exact (measurable_const.sub (measurable_g _ _)).real_toNNReal + · exact measurable_rnDerivReal + have h_coe : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl + have h_le : ∀ a x, ENNReal.ofReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x + ≤ ENNReal.ofReal (g κ (κ + ν) a x) := by + intro a x + unfold kernel.rnDerivReal + by_cases hg_one : g κ (κ + ν) a x = 1 + · simp [hg_one] + rw [← ENNReal.ofReal_div_of_pos, div_eq_inv_mul, ← ENNReal.ofReal_mul, ← mul_assoc, + mul_inv_cancel, one_mul] + · rw [ne_eq, sub_eq_zero] + exact Ne.symm hg_one + · simp [g_le_one (todo κ ν)] + · simp [lt_of_le_of_ne (g_le_one (todo κ ν)) hg_one] + rw [kernel.withDensity_sub_add, withDensity_g] + all_goals simp_rw [h_coe] + · exact (measurable_g _ _).ennreal_ofReal + · exact Measurable.mul (measurable_const.sub (measurable_g _ _)).ennreal_ofReal + measurable_rnDerivReal + · intro a + have : ∀ x, ENNReal.ofReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x ≤ 1 := by + refine fun x ↦ (h_le a x).trans ?_ + simp only [ENNReal.ofReal_le_one, g_le_one (todo κ ν)] + refine ne_of_lt ?_ + calc ∫⁻ x, ENNReal.ofReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x ∂(κ + ν) a + ≤ ∫⁻ _, 1 ∂(κ + ν) a := lintegral_mono this + _ < ⊤ := by rw [MeasureTheory.lintegral_const, one_mul]; exact measure_lt_top _ _ + · exact fun a ↦ ae_of_all _ (h_le a) + +end Real + +section StandardBorel + +variable {Ω : Type*} [MeasurableSpace Ω] [StandardBorelSpace Ω] {κ ν : kernel α Ω} + +protected noncomputable def kernel.rnDeriv (κ ν : kernel α Ω) (a : α) (ω : Ω) : ℝ≥0∞ := + let f := measurableEmbedding_real Ω + let hf := measurableEmbedding_measurableEmbedding_real Ω + kernel.rnDerivReal (kernel.map κ f hf.measurable) (kernel.map ν f hf.measurable) a (f ω) + +protected noncomputable def kernel.singularPart (κ ν : kernel α Ω) + [IsSFiniteKernel κ] [IsSFiniteKernel ν] : kernel α Ω := + let f := measurableEmbedding_real Ω + let hf := measurableEmbedding_measurableEmbedding_real Ω + kernel.comapRight + (kernel.singularPartReal (kernel.map κ f hf.measurable) (kernel.map ν f hf.measurable)) + hf + +lemma kernel.mutuallySingular_singularPart (κ ν : kernel α Ω) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) : + kernel.singularPart κ ν a ⟂ₘ ν a := by + let f := measurableEmbedding_real Ω + let hf := measurableEmbedding_measurableEmbedding_real Ω + change kernel.comapRight + (kernel.singularPartReal (kernel.map κ f hf.measurable) (kernel.map ν f hf.measurable)) + hf a ⟂ₘ ν a + let s := (kernel.mutuallySingular_singularPartReal (kernel.map κ f hf.measurable) + (kernel.map ν f hf.measurable) a).nullSet + have hs : MeasurableSet s := Measure.MutuallySingular.measurableSet_nullSet _ + have hνs := Measure.MutuallySingular.measure_compl_nullSet + (kernel.mutuallySingular_singularPartReal (kernel.map κ f hf.measurable) + (kernel.map ν f hf.measurable) a) + have hsing := Measure.MutuallySingular.measure_nullSet + (kernel.mutuallySingular_singularPartReal (kernel.map κ f hf.measurable) + (kernel.map ν f hf.measurable) a) + refine ⟨f ⁻¹' s, hf.measurable hs, ?_, ?_⟩ + · rw [kernel.comapRight_apply' _ _ _ (hf.measurable hs)] + refine measure_mono_null (fun x hx ↦ ?_) hsing + exact image_preimage_subset _ _ hx + · have : ν = kernel.comapRight (kernel.map ν f hf.measurable) hf := by + ext a + rw [kernel.comapRight_apply _ hf, kernel.map_apply, MeasurableEmbedding.comap_map] + exact hf + rw [this, kernel.comapRight_apply' _ _ _ (hf.measurable hs).compl] + refine measure_mono_null (fun x ↦ ?_) hνs + rw [image_compl_preimage, mem_diff] + exact fun h ↦ h.2 + +lemma kernel.rnDeriv_add_singularPart (κ ν : kernel α Ω) [IsFiniteKernel κ] [IsFiniteKernel ν] : + kernel.withDensity ν (kernel.rnDeriv κ ν) + kernel.singularPart κ ν = κ := by + let f := measurableEmbedding_real Ω + let hf := measurableEmbedding_measurableEmbedding_real Ω + ext a : 1 + simp only [coeFn_add, Pi.add_apply] + change kernel.withDensity ν + (fun a ω ↦ kernel.rnDerivReal (kernel.map κ f hf.measurable) + (kernel.map ν f hf.measurable) a (f ω)) a + kernel.comapRight + (kernel.singularPartReal (kernel.map κ f hf.measurable) (kernel.map ν f hf.measurable)) + hf a = κ a + have h := kernel.rnDerivReal_add_singularPartReal (kernel.map κ f hf.measurable) + (kernel.map ν f hf.measurable) + have : κ = kernel.comapRight (kernel.map κ f hf.measurable) hf := by + ext a + rw [kernel.comapRight_apply _ hf, kernel.map_apply, MeasurableEmbedding.comap_map] + exact hf + conv_rhs => rw [this, ← h] + rw [comapRight_apply, comapRight_apply] + simp only [coeFn_add, Pi.add_apply] + rw [MeasurableEmbedding.comap_add hf] + congr + ext s hs + rw [Measure.comap_apply _ hf.injective (fun s hs ↦ hf.measurableSet_image.mpr hs) _ hs, + kernel.withDensity_apply', kernel.withDensity_apply', kernel.map_apply, + Measure.restrict_map hf.measurable (hf.measurableSet_image.mpr hs), + MeasureTheory.lintegral_map _ hf.measurable, preimage_image_eq _ hf.injective] + · exact measurable_rnDerivReal_right _ _ _ + · exact measurable_rnDerivReal + · change Measurable ((Function.uncurry fun a x ↦ + rnDerivReal (map κ f hf.measurable) (map ν f hf.measurable) a x) + ∘ (fun p : α × Ω ↦ (p.1, f p.2))) + exact measurable_rnDerivReal.comp (measurable_fst.prod_mk (hf.measurable.comp measurable_snd)) + +end StandardBorel + +end ProbabilityTheory From 746ef138bbcc3fba7885580d2d1acc4c46ba941c Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 19 Feb 2024 16:25:20 +0100 Subject: [PATCH 037/129] Mathlib.lean --- Mathlib.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib.lean b/Mathlib.lean index 9b1baa8860d80..aca67f25b225d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3121,6 +3121,7 @@ import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.Disintegration.Integral import Mathlib.Probability.Kernel.Disintegration.KernelCDFReal import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes +import Mathlib.Probability.Kernel.Disintegration.RadonNikodym import Mathlib.Probability.Kernel.Disintegration.Unique import Mathlib.Probability.Kernel.IntegralCompProd import Mathlib.Probability.Kernel.Invariance From 857e3a23c8b93515f1890b47afbbe8eba7b0a09f Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 19 Feb 2024 17:06:10 +0100 Subject: [PATCH 038/129] move lemmas --- Mathlib/Analysis/SpecificLimits/Basic.lean | 5 ++ Mathlib/MeasureTheory/Integral/Bochner.lean | 59 ++++++++++++++ Mathlib/Order/LiminfLimsup.lean | 10 +++ .../Kernel/Disintegration/AuxLemmas.lean | 79 ------------------- 4 files changed, 74 insertions(+), 79 deletions(-) diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index 26412ecbb362d..66986a544eab6 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -624,6 +624,11 @@ theorem tendsto_nat_floor_atTop {α : Type*} [LinearOrderedSemiring α] [FloorSe Nat.floor_mono.tendsto_atTop_atTop fun x ↦ ⟨max 0 (x + 1), by simp [Nat.le_floor_iff]⟩ #align tendsto_nat_floor_at_top tendsto_nat_floor_atTop +lemma tendsto_nat_ceil_atTop {α : Type*} [LinearOrderedSemiring α] [FloorSemiring α] : + Tendsto (fun x : α ↦ ⌈x⌉₊) atTop atTop := by + refine Nat.ceil_mono.tendsto_atTop_atTop (fun x ↦ ⟨x, ?_⟩) + simp only [Nat.ceil_natCast, le_refl] + lemma tendsto_nat_floor_mul_atTop {α : Type _} [LinearOrderedSemifield α] [FloorSemiring α] [Archimedean α] (a : α) (ha : 0 < a) : Tendsto (fun (x:ℕ) => ⌊a * x⌋₊) atTop atTop := Tendsto.comp tendsto_nat_floor_atTop diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 74aad611159d8..7f70cfb338e19 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1348,6 +1348,65 @@ lemma integral_exp_pos {μ : Measure α} {f : α → ℝ} [hμ : NeZero μ] ext1 x simp only [Function.mem_support, ne_eq, (Real.exp_pos _).ne', not_false_eq_true, Set.mem_univ] +/-- Monotone convergence theorem for real-valued functions and Bochner integrals -/ +lemma integral_tendsto_of_tendsto_of_monotone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} + (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Monotone fun n ↦ f n x) + (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : + Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by + -- switch from the Bochner to the Lebesgue integral + let f' := fun n x ↦ f n x - f 0 x + have hf'_nonneg : ∀ᵐ x ∂μ, ∀ n, 0 ≤ f' n x := by + filter_upwards [h_mono] with a ha n + simp [ha (zero_le n)] + have hf'_meas : ∀ n, Integrable (f' n) μ := fun n ↦ (hf n).sub (hf 0) + suffices Tendsto (fun n ↦ ∫ x, f' n x ∂μ) atTop (𝓝 (∫ x, (F - f 0) x ∂μ)) by + rw [integral_sub' hF (hf 0)] at this + have h_sub : ∀ n, ∫ x, f' n x ∂μ = ∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ := by + intro n + rw [integral_sub (hf n) (hf 0)] + simp_rw [h_sub] at this + have h1 : (fun n ↦ ∫ x, f n x ∂μ) + = fun n ↦ (∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by ext n; abel + have h2 : ∫ x, F x ∂μ = (∫ x, F x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by abel + rw [h1, h2] + exact this.add tendsto_const_nhds + have hF_ge : 0 ≤ᵐ[μ] fun x ↦ (F - f 0) x := by + filter_upwards [h_tendsto, h_mono] with x hx_tendsto hx_mono + simp only [Pi.zero_apply, Pi.sub_apply, sub_nonneg] + exact ge_of_tendsto' hx_tendsto (fun n ↦ hx_mono (zero_le _)) + rw [ae_all_iff] at hf'_nonneg + simp_rw [integral_eq_lintegral_of_nonneg_ae (hf'_nonneg _) (hf'_meas _).1] + rw [integral_eq_lintegral_of_nonneg_ae hF_ge (hF.1.sub (hf 0).1)] + have h_cont := ENNReal.continuousAt_toReal (x := ∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ) ?_ + swap + · rw [← ofReal_integral_eq_lintegral_ofReal (hF.sub (hf 0)) hF_ge] + exact ENNReal.ofReal_ne_top + refine h_cont.tendsto.comp ?_ + -- use the result for the Lebesgue integral + refine lintegral_tendsto_of_tendsto_of_monotone ?_ ?_ ?_ + · exact fun n ↦ ((hf n).sub (hf 0)).aemeasurable.ennreal_ofReal + · filter_upwards [h_mono] with x hx n m hnm + refine ENNReal.ofReal_le_ofReal ?_ + simp only [tsub_le_iff_right, sub_add_cancel] + exact hx hnm + · filter_upwards [h_tendsto] with x hx + refine (ENNReal.continuous_ofReal.tendsto _).comp ?_ + simp only [Pi.sub_apply] + exact Tendsto.sub hx tendsto_const_nhds + +/-- Monotone convergence theorem for real-valued functions and Bochner integrals -/ +lemma integral_tendsto_of_tendsto_of_antitone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} + (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Antitone fun n ↦ f n x) + (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : + Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by + suffices Tendsto (fun n ↦ ∫ x, -f n x ∂μ) atTop (𝓝 (∫ x, -F x ∂μ)) by + suffices Tendsto (fun n ↦ ∫ x, - -f n x ∂μ) atTop (𝓝 (∫ x, - -F x ∂μ)) by + simpa [neg_neg] using this + convert this.neg <;> rw [integral_neg] + refine integral_tendsto_of_tendsto_of_monotone (fun n ↦ (hf n).neg) hF.neg ?_ ?_ + · filter_upwards [h_mono] with x hx n m hnm using neg_le_neg_iff.mpr <| hx hnm + · filter_upwards [h_tendsto] with x hx using hx.neg + section NormedAddCommGroup variable {H : Type*} [NormedAddCommGroup H] diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index f5c16fff032a7..dbda631677d78 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -242,6 +242,16 @@ theorem IsBoundedUnder.isCoboundedUnder_le {u : γ → α} {l : Filter γ} [Preo (h : l.IsBoundedUnder (· ≥ ·) u) : l.IsCoboundedUnder (· ≤ ·) u := h.isCoboundedUnder_flip +lemma isCoboundedUnder_le_of_eventually_le [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} {x : α} + (hf : ∀ᶠ i in l, x ≤ f i) : + IsCoboundedUnder (· ≤ ·) l f := + IsBoundedUnder.isCoboundedUnder_le ⟨x, hf⟩ + +lemma isCoboundedUnder_le_of_le [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} {x : α} + (hf : ∀ i, x ≤ f i) : + IsCoboundedUnder (· ≤ ·) l f := + isCoboundedUnder_le_of_eventually_le l (eventually_of_forall hf) + theorem IsBoundedUnder.isCoboundedUnder_ge {u : γ → α} {l : Filter γ} [Preorder α] [NeBot l] (h : l.IsBoundedUnder (· ≤ ·) u) : l.IsCoboundedUnder (· ≥ ·) u := h.isCoboundedUnder_flip diff --git a/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean b/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean index c7b2952e77706..6960ec563a1d5 100644 --- a/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean +++ b/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean @@ -40,82 +40,3 @@ lemma measurableSet_tendsto_fun {β γ ι : Type*} [MeasurableSpace β] rw [tendsto_iff_dist_tendsto_zero] rw [this] exact measurableSet_tendsto_nhds (fun n ↦ (hf n).dist hg) 0 - -lemma integral_tendsto_of_tendsto_of_monotone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} - (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Monotone fun n ↦ f n x) - (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : - Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by - -- switch from the Bochner to the Lebesgue integral - let f' := fun n x ↦ f n x - f 0 x - have hf'_nonneg : ∀ᵐ x ∂μ, ∀ n, 0 ≤ f' n x := by - filter_upwards [h_mono] with a ha n - simp [ha (zero_le n)] - have hf'_meas : ∀ n, Integrable (f' n) μ := fun n ↦ (hf n).sub (hf 0) - suffices Tendsto (fun n ↦ ∫ x, f' n x ∂μ) atTop (𝓝 (∫ x, (F - f 0) x ∂μ)) by - rw [integral_sub' hF (hf 0)] at this - have h_sub : ∀ n, ∫ x, f' n x ∂μ = ∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ := by - intro n - simp only - rw [integral_sub (hf n) (hf 0)] - simp_rw [h_sub] at this - have h1 : (fun n ↦ ∫ x, f n x ∂μ) - = fun n ↦ (∫ x, f n x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by ext n; abel - have h2 : ∫ x, F x ∂μ = (∫ x, F x ∂μ - ∫ x, f 0 x ∂μ) + ∫ x, f 0 x ∂μ := by abel - rw [h1, h2] - exact this.add tendsto_const_nhds - have hF_ge : 0 ≤ᵐ[μ] fun x ↦ (F - f 0) x := by - filter_upwards [h_tendsto, h_mono] with x hx_tendsto hx_mono - simp only [Pi.zero_apply, Pi.sub_apply, sub_nonneg] - exact ge_of_tendsto' hx_tendsto (fun n ↦ hx_mono (zero_le _)) - rw [ae_all_iff] at hf'_nonneg - simp_rw [integral_eq_lintegral_of_nonneg_ae (hf'_nonneg _) (hf'_meas _).1] - rw [integral_eq_lintegral_of_nonneg_ae hF_ge (hF.1.sub (hf 0).1)] - have h_cont := ENNReal.continuousAt_toReal (x := ∫⁻ a, ENNReal.ofReal ((F - f 0) a) ∂μ) ?_ - swap - · rw [← ofReal_integral_eq_lintegral_ofReal (hF.sub (hf 0)) hF_ge] - exact ENNReal.ofReal_ne_top - refine h_cont.tendsto.comp ?_ - -- use the result for the Lebesgue integral - refine lintegral_tendsto_of_tendsto_of_monotone ?_ ?_ ?_ - · exact fun n ↦ ((hf n).sub (hf 0)).aemeasurable.ennreal_ofReal - · filter_upwards [h_mono] with x hx - intro n m hnm - refine ENNReal.ofReal_le_ofReal ?_ - simp only [tsub_le_iff_right, sub_add_cancel] - exact hx hnm - · filter_upwards [h_tendsto] with x hx - refine (ENNReal.continuous_ofReal.tendsto _).comp ?_ - simp only [Pi.sub_apply] - exact Tendsto.sub hx tendsto_const_nhds - -lemma integral_tendsto_of_tendsto_of_antitone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} - (hf : ∀ n, Integrable (f n) μ) (hF : Integrable F μ) (h_mono : ∀ᵐ x ∂μ, Antitone fun n ↦ f n x) - (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : - Tendsto (fun n ↦ ∫ x, f n x ∂μ) atTop (𝓝 (∫ x, F x ∂μ)) := by - suffices Tendsto (fun n ↦ ∫ x, -f n x ∂μ) atTop (𝓝 (∫ x, -F x ∂μ)) by - suffices Tendsto (fun n ↦ ∫ x, - -f n x ∂μ) atTop (𝓝 (∫ x, - -F x ∂μ)) by - simp_rw [neg_neg] at this - exact this - convert this.neg <;> rw [integral_neg] - refine integral_tendsto_of_tendsto_of_monotone (fun n ↦ (hf n).neg) hF.neg ?_ ?_ - · filter_upwards [h_mono] with x hx - intro n m hnm - simp only [neg_le_neg_iff] - exact hx hnm - · filter_upwards [h_tendsto] with x hx - exact hx.neg - -lemma tendsto_nat_ceil_atTop {α : Type*} [LinearOrderedSemiring α] [FloorSemiring α] : - Tendsto (fun x : α ↦ ⌈x⌉₊) atTop atTop := by - refine Nat.ceil_mono.tendsto_atTop_atTop (fun x ↦ ⟨x, ?_⟩) - simp only [Nat.ceil_natCast, le_refl] - -lemma isCoboundedUnder_le_of_eventually_le {ι α : Type*} [Preorder α] (l : Filter ι) [NeBot l] - {f : ι → α} {x : α} (hf : ∀ᶠ i in l, x ≤ f i) : - IsCoboundedUnder (· ≤ ·) l f := - IsBoundedUnder.isCoboundedUnder_le ⟨x, hf⟩ - -lemma isCoboundedUnder_le_of_le {ι α : Type*} [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} - {x : α} (hf : ∀ i, x ≤ f i) : - IsCoboundedUnder (· ≤ ·) l f := - isCoboundedUnder_le_of_eventually_le l (eventually_of_forall hf) From 29ecce1b33594150ded4eb07788a459e46300b4b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Mon, 19 Feb 2024 17:13:57 +0100 Subject: [PATCH 039/129] Update Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean b/Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean index 266139c1047a3..587a81e7a7677 100644 --- a/Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean @@ -230,7 +230,7 @@ lemma withDensity_one_sub_g (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFinit noncomputable def kernel.singularPartReal (κ ν : kernel α ℝ) [IsSFiniteKernel κ] [IsSFiniteKernel ν] : - kernel α ℝ := + kernel α ℝ := kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (g κ (κ + ν) a x) - Real.toNNReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x) From 4ec8637c6d58c4368d5f4b339a571fda4d15b443 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 19 Feb 2024 18:49:01 +0100 Subject: [PATCH 040/129] move lemmas --- .../MeasureTheory/Measure/MeasureSpace.lean | 10 ++ Mathlib/Probability/Kernel/Basic.lean | 10 ++ Mathlib/Probability/Kernel/Composition.lean | 6 + .../Kernel/Disintegration/Density.lean | 7 - .../Kernel/Disintegration/RadonNikodym.lean | 133 +++--------------- .../Kernel/Disintegration/Unique.lean | 2 +- Mathlib/Probability/Kernel/WithDensity.lean | 59 ++++++++ docs/references.bib | 13 +- 8 files changed, 121 insertions(+), 119 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 156cb146c91f8..81ea883afac39 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -1617,6 +1617,7 @@ instance instIsRefl [MeasurableSpace α] : IsRefl (Measure α) (· ≪ ·) := ⟨fun _ => AbsolutelyContinuous.rfl⟩ #align measure_theory.measure.absolutely_continuous.is_refl MeasureTheory.Measure.AbsolutelyContinuous.instIsRefl +@[simp] protected lemma zero (μ : Measure α) : 0 ≪ μ := fun s _ ↦ by simp @[trans] @@ -2123,6 +2124,15 @@ nonrec theorem map_apply (μ : Measure α) (s : Set β) : μ.map f s = μ (f ⁻ _ = μ (f ⁻¹' s) := by rw [map_apply hf.measurable htm, hft, measure_toMeasurable] #align measurable_embedding.map_apply MeasurableEmbedding.map_apply +lemma comap_add (μ ν : Measure β) : + (μ + ν).comap f = μ.comap f + ν.comap f := by + ext s hs + rw [← Measure.comapₗ_eq_comap _ hf.injective (fun s hs ↦ hf.measurableSet_image.mpr hs) _ hs, + _root_.map_add] + simp only [Measure.add_toOuterMeasure, OuterMeasure.coe_add, Pi.add_apply] + rw [← Measure.comapₗ_eq_comap _ hf.injective (fun s hs ↦ hf.measurableSet_image.mpr hs) _ hs, + ← Measure.comapₗ_eq_comap _ hf.injective (fun s hs ↦ hf.measurableSet_image.mpr hs) _ hs] + end MeasurableEmbedding namespace MeasurableEquiv diff --git a/Mathlib/Probability/Kernel/Basic.lean b/Mathlib/Probability/Kernel/Basic.lean index c05dad0306633..372efea2c8c39 100644 --- a/Mathlib/Probability/Kernel/Basic.lean +++ b/Mathlib/Probability/Kernel/Basic.lean @@ -70,6 +70,16 @@ instance {α β : Type*} [MeasurableSpace α] [MeasurableSpace β] : coe := Subtype.val coe_injective' := Subtype.val_injective +instance kernel.covariantAddLE {α β : Type*} [MeasurableSpace α] [MeasurableSpace β] : + CovariantClass (kernel α β) (kernel α β) (· + ·) (· ≤ ·) := + ⟨fun _ _ _ hμ a ↦ add_le_add_left (hμ a) _⟩ + +noncomputable +instance kernel.instOrderBot {α β : Type*} [MeasurableSpace α] [MeasurableSpace β] : + OrderBot (kernel α β) where + bot := 0 + bot_le κ a := by simp only [ZeroMemClass.coe_zero, Pi.zero_apply, Measure.zero_le] + variable {α β ι : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} namespace kernel diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index 03cfeed68e94b..21eb691126e50 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -588,6 +588,12 @@ theorem map_apply' (κ : kernel α β) (hf : Measurable f) (a : α) {s : Set γ} lemma map_zero (hf : Measurable f) : kernel.map (0 : kernel α β) f hf = 0 := by ext; rw [kernel.map_apply]; simp +@[simp] +lemma map_id (κ : kernel α β) : map κ id measurable_id = κ := by ext a; rw [map_apply]; simp + +@[simp] +lemma map_id' (κ : kernel α β) : map κ (fun a ↦ a) measurable_id = κ := map_id κ + nonrec theorem lintegral_map (κ : kernel α β) (hf : Measurable f) (a : α) {g' : γ → ℝ≥0∞} (hg : Measurable g') : ∫⁻ b, g' b ∂map κ f hf a = ∫⁻ a, g' (f a) ∂κ a := by rw [map_apply _ hf, lintegral_map hg hf] diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index f14504d9f0c29..7d24ed07895b7 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -54,13 +54,6 @@ desired density function. The construction of the density process in this file follows the proof of Theorem 9.27 in [O. Kallenberg, Foundations of modern probability][kallenberg2021]. - -## TODO - -The construction in this file can be used to obtain a kernel Radon-Nikodym defivative of -`κ : kernel α Ω` with `Ω` standard Borel with respect to `ν : kernel α Ω`. -See Theorem 1.28 in O. Kallenberg, Random Measures, Theory and Applications. - -/ diff --git a/Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean b/Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean index 587a81e7a7677..cd1efd595cea6 100644 --- a/Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean @@ -7,11 +7,14 @@ import Mathlib.Probability.Kernel.Disintegration.KernelCDFReal import Mathlib.Probability.Kernel.WithDensity /-! -# Density +# Radon-Nikodym derivative and Lebesgue decompusition for kernels + +Let `Ω` be a standard Borel space and `κ ν : kernel α Ω`. Then there exists... TODO ## Main definitions -* `FooBar` +* `ProbabilityTheory.kernel.rnDeriv` +* `ProbabilityTheory.kernel.singularPart` ## Main statements @@ -27,7 +30,7 @@ import Mathlib.Probability.Kernel.WithDensity ## References -* [F. Bar, *Quuxes*][bibkey] +Theorem 1.28 in [O. Kallenberg, Random Measures, Theory and Applications][kallenberg2017]. ## Tags @@ -42,91 +45,6 @@ namespace ProbabilityTheory variable {α : Type*} {mα : MeasurableSpace α} -section aux - -lemma MeasurableEmbedding.comap_add {β : Type*} {mβ : MeasurableSpace β} - {f : α → β} (hf : MeasurableEmbedding f) {μ ν : Measure β} : - (μ + ν).comap f = μ.comap f + ν.comap f := by - ext s hs - rw [← Measure.comapₗ_eq_comap _ hf.injective (fun s hs ↦ hf.measurableSet_image.mpr hs) _ hs, - map_add] - simp only [Measure.add_toOuterMeasure, OuterMeasure.coe_add, Pi.add_apply] - rw [← Measure.comapₗ_eq_comap _ hf.injective (fun s hs ↦ hf.measurableSet_image.mpr hs) _ hs, - ← Measure.comapₗ_eq_comap _ hf.injective (fun s hs ↦ hf.measurableSet_image.mpr hs) _ hs] - -@[simp] -lemma kernel.map_id {β : Type*} {_ : MeasurableSpace β} {κ : kernel α β} : - kernel.map κ id measurable_id = κ := by ext a; rw [kernel.map_apply]; simp - -@[simp] -lemma kernel.map_id' {β : Type*} {_ : MeasurableSpace β} {κ : kernel α β} : - kernel.map κ (fun a ↦ a) measurable_id = κ := kernel.map_id - -end aux - -section withDensity - -variable {β : Type*} {mβ : MeasurableSpace β} {κ : kernel α β} - -@[simp] -lemma Measure.absolutelyContinuous_zero (μ : Measure α) : 0 ≪ μ := - Measure.absolutelyContinuous_of_le (Measure.zero_le μ) - -nonrec lemma kernel.withDensity_absolutelyContinuous [IsSFiniteKernel κ] - (f : α → β → ℝ≥0∞) (a : α) : - kernel.withDensity κ f a ≪ κ a := by - by_cases hf : Measurable (Function.uncurry f) - · rw [kernel.withDensity_apply _ hf] - exact withDensity_absolutelyContinuous _ _ - · rw [withDensity_of_not_measurable _ hf] - simp - -@[simp] -lemma kernel.withDensity_one [IsSFiniteKernel κ] : - kernel.withDensity κ 1 = κ := by - ext a - rw [kernel.withDensity_apply _ measurable_const] - simp - -@[simp] -lemma kernel.withDensity_one' [IsSFiniteKernel κ] : - kernel.withDensity κ (fun _ _ ↦ 1) = κ := kernel.withDensity_one - -lemma kernel.withDensity_sub_add [IsSFiniteKernel κ] {f g : α → β → ℝ≥0∞} - (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) - (hg_int : ∀ a, ∫⁻ x, g a x ∂(κ a) ≠ ∞) (hfg : ∀ a, g a ≤ᵐ[κ a] f a) : - kernel.withDensity κ (fun a x ↦ f a x - g a x) + kernel.withDensity κ g - = kernel.withDensity κ f := by - ext a s - simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add] - rw [kernel.withDensity_apply' _ hf, kernel.withDensity_apply' _ hg, kernel.withDensity_apply'] - swap; · exact hf.sub hg - rw [lintegral_sub] - · rw [tsub_add_cancel_iff_le] - exact lintegral_mono_ae (ae_restrict_of_ae (hfg a)) - · exact hg.comp measurable_prod_mk_left - · exact ((set_lintegral_le_lintegral _ _).trans_lt (hg_int a).lt_top).ne - · exact ae_restrict_of_ae (hfg a) - -nonrec lemma kernel.withDensity_mul [IsSFiniteKernel κ] {f : α → β → ℝ≥0} {g : α → β → ℝ≥0∞} - (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) : - kernel.withDensity κ (fun a x ↦ f a x * g a x) - = kernel.withDensity (kernel.withDensity κ fun a x ↦ f a x) fun a x ↦ g a x := by - ext a : 1 - rw [kernel.withDensity_apply] - swap; · exact (measurable_coe_nnreal_ennreal.comp hf).mul hg - suffices (Measure.withDensity (κ a) ((fun x ↦ (f a x : ℝ≥0∞)) * (fun x ↦ (g a x : ℝ≥0∞)))) = - (withDensity (withDensity κ fun a x ↦ f a x) fun a x ↦ g a x) a by - convert this - rw [withDensity_mul] - · rw [kernel.withDensity_apply _ hg, kernel.withDensity_apply] - exact measurable_coe_nnreal_ennreal.comp hf - · rw [measurable_coe_nnreal_ennreal_iff] - exact hf.comp measurable_prod_mk_left - · exact hg.comp measurable_prod_mk_left - -end withDensity - section Real variable {κ ν : kernel α ℝ} @@ -136,12 +54,6 @@ lemma fst_map_le_of_le (hκν : κ ≤ ν) : (@measurable_prod_mk_right ℝ Unit _ inferInstance _)) ≤ ν := by rwa [kernel.fst_map_prod _ measurable_id' measurable_const, kernel.map_id'] -lemma todo (κ : kernel α ℝ) (ν : kernel α ℝ) : κ ≤ κ + ν := by - -- todo improve this: `le_add_of_nonneg_right` should work on kernels - intro a - simp only [AddSubmonoid.coe_add, Pi.add_apply] - exact le_add_of_nonneg_right (Measure.zero_le (ν a)) - noncomputable def g (κ ν : kernel α ℝ) (a : α) (x : ℝ) : ℝ := MLimsup (kernel.map κ (fun a ↦ (a, ())) @@ -170,17 +82,18 @@ lemma kernel.rnDerivReal_def (κ ν : kernel α ℝ) : kernel.rnDerivReal κ ν = fun a x ↦ ENNReal.ofReal (g κ (κ + ν) a x) / ENNReal.ofReal (1 - g κ (κ + ν) a x) := rfl -lemma measurable_rnDerivReal : Measurable (fun p : α × ℝ ↦ kernel.rnDerivReal κ ν p.1 p.2) := +lemma measurable_rnDerivReal (κ ν : kernel α ℝ) : + Measurable (fun p : α × ℝ ↦ kernel.rnDerivReal κ ν p.1 p.2) := (measurable_g κ _).ennreal_ofReal.div (measurable_const.sub (measurable_g κ _)).ennreal_ofReal lemma measurable_rnDerivReal_right (κ ν : kernel α ℝ) (a : α) : Measurable (fun x : ℝ ↦ kernel.rnDerivReal κ ν a x) := by change Measurable ((fun p : α × ℝ ↦ kernel.rnDerivReal κ ν p.1 p.2) ∘ (fun x ↦ (a, x))) - exact measurable_rnDerivReal.comp measurable_prod_mk_left + exact (measurable_rnDerivReal _ _).comp measurable_prod_mk_left lemma withDensity_g (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFiniteKernel ν] : kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (g κ (κ + ν) a x)) = κ := by - have h_le : κ ≤ κ + ν := todo κ ν + have h_le : κ ≤ κ + ν := le_add_of_nonneg_right bot_le ext a s hs rw [kernel.withDensity_apply'] swap; exact (measurable_g _ _).ennreal_ofReal @@ -198,7 +111,7 @@ lemma withDensity_g (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFiniteKernel lemma withDensity_one_sub_g (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFiniteKernel ν] : kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (1 - g κ (κ + ν) a x)) = ν := by - have h_le : κ ≤ κ + ν := todo κ ν + have h_le : κ ≤ κ + ν := le_add_of_nonneg_right bot_le suffices kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (1 - g κ (κ + ν) a x)) + kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (g κ (κ + ν) a x)) = κ + ν by ext a s @@ -222,7 +135,7 @@ lemma withDensity_one_sub_g (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFinit calc ∫⁻ x, ENNReal.ofReal (g κ (κ + ν) a x) ∂(κ + ν) a ≤ ∫⁻ _, 1 ∂(κ + ν) a := by refine lintegral_mono (fun x ↦ ?_) - simp [g_le_one (todo κ ν)] + simp [g_le_one (le_add_of_nonneg_right bot_le)] _ < ⊤ := by rw [MeasureTheory.lintegral_const, one_mul]; exact measure_lt_top _ _ · refine fun a ↦ ae_of_all _ (fun x ↦ ?_) simp only [ENNReal.ofReal_le_one] @@ -271,7 +184,7 @@ lemma kernel.mutuallySingular_singularPartReal (κ ν : kernel α ℝ) · exact (measurable_g _ _).ennreal_ofReal · refine Measurable.mul ?_ ?_ · exact (measurable_const.sub (measurable_g _ _)).ennreal_ofReal - · exact measurable_rnDerivReal + · exact measurable_rnDerivReal _ _ refine ae_of_all _ (fun x hx ↦ ?_) simp only [mem_compl_iff, mem_setOf_eq] at hx simp_rw [rnDerivReal] @@ -280,8 +193,8 @@ lemma kernel.mutuallySingular_singularPartReal (κ ν : kernel α ℝ) · rfl · rw [ne_eq, sub_eq_zero] exact Ne.symm hx - · simp [g_le_one (todo κ ν)] - · simp [lt_of_le_of_ne (g_le_one (todo κ ν)) hx] + · simp [g_le_one (le_add_of_nonneg_right bot_le)] + · simp [lt_of_le_of_ne (g_le_one (le_add_of_nonneg_right bot_le)) hx] lemma kernel.rnDerivReal_add_singularPartReal (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFiniteKernel ν] : @@ -295,7 +208,7 @@ lemma kernel.rnDerivReal_add_singularPartReal (κ ν : kernel α ℝ) rw [this, kernel.singularPartReal, add_comm, ← kernel.withDensity_mul] rotate_left · exact (measurable_const.sub (measurable_g _ _)).real_toNNReal - · exact measurable_rnDerivReal + · exact measurable_rnDerivReal _ _ have h_coe : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl have h_le : ∀ a x, ENNReal.ofReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x ≤ ENNReal.ofReal (g κ (κ + ν) a x) := by @@ -307,17 +220,16 @@ lemma kernel.rnDerivReal_add_singularPartReal (κ ν : kernel α ℝ) mul_inv_cancel, one_mul] · rw [ne_eq, sub_eq_zero] exact Ne.symm hg_one - · simp [g_le_one (todo κ ν)] - · simp [lt_of_le_of_ne (g_le_one (todo κ ν)) hg_one] + · simp [g_le_one (le_add_of_nonneg_right bot_le)] + · simp [lt_of_le_of_ne (g_le_one (le_add_of_nonneg_right bot_le)) hg_one] rw [kernel.withDensity_sub_add, withDensity_g] all_goals simp_rw [h_coe] · exact (measurable_g _ _).ennreal_ofReal - · exact Measurable.mul (measurable_const.sub (measurable_g _ _)).ennreal_ofReal - measurable_rnDerivReal + · exact (measurable_const.sub (measurable_g _ _)).ennreal_ofReal.mul (measurable_rnDerivReal _ _) · intro a have : ∀ x, ENNReal.ofReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x ≤ 1 := by refine fun x ↦ (h_le a x).trans ?_ - simp only [ENNReal.ofReal_le_one, g_le_one (todo κ ν)] + simp only [ENNReal.ofReal_le_one, g_le_one (le_add_of_nonneg_right bot_le)] refine ne_of_lt ?_ calc ∫⁻ x, ENNReal.ofReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x ∂(κ + ν) a ≤ ∫⁻ _, 1 ∂(κ + ν) a := lintegral_mono this @@ -401,11 +313,12 @@ lemma kernel.rnDeriv_add_singularPart (κ ν : kernel α Ω) [IsFiniteKernel κ] Measure.restrict_map hf.measurable (hf.measurableSet_image.mpr hs), MeasureTheory.lintegral_map _ hf.measurable, preimage_image_eq _ hf.injective] · exact measurable_rnDerivReal_right _ _ _ - · exact measurable_rnDerivReal + · exact measurable_rnDerivReal _ _ · change Measurable ((Function.uncurry fun a x ↦ rnDerivReal (map κ f hf.measurable) (map ν f hf.measurable) a x) ∘ (fun p : α × Ω ↦ (p.1, f p.2))) - exact measurable_rnDerivReal.comp (measurable_fst.prod_mk (hf.measurable.comp measurable_snd)) + exact (measurable_rnDerivReal _ _).comp + (measurable_fst.prod_mk (hf.measurable.comp measurable_snd)) end StandardBorel diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index d096646caa25e..40635d7202994 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kexing Ying +Authors: Kexing Ying, Rémy Degenne -/ import Mathlib.Probability.Kernel.Disintegration.Integral diff --git a/Mathlib/Probability/Kernel/WithDensity.lean b/Mathlib/Probability/Kernel/WithDensity.lean index 4c4820fa08c24..efbe335f6bc44 100644 --- a/Mathlib/Probability/Kernel/WithDensity.lean +++ b/Mathlib/Probability/Kernel/WithDensity.lean @@ -72,6 +72,33 @@ protected theorem withDensity_apply' (κ : kernel α β) [IsSFiniteKernel κ] rw [kernel.withDensity_apply κ hf, withDensity_apply' _ s] #align probability_theory.kernel.with_density_apply' ProbabilityTheory.kernel.withDensity_apply' +nonrec lemma kernel.withDensity_absolutelyContinuous [IsSFiniteKernel κ] + (f : α → β → ℝ≥0∞) (a : α) : + kernel.withDensity κ f a ≪ κ a := by + by_cases hf : Measurable (Function.uncurry f) + · rw [kernel.withDensity_apply _ hf] + exact withDensity_absolutelyContinuous _ _ + · rw [withDensity_of_not_measurable _ hf] + simp + +@[simp] +lemma withDensity_one (κ : kernel α β) [IsSFiniteKernel κ] : + kernel.withDensity κ 1 = κ := by + ext a; rw [kernel.withDensity_apply _ measurable_const]; simp + +@[simp] +lemma withDensity_one' (κ : kernel α β) [IsSFiniteKernel κ] : + kernel.withDensity κ (fun _ _ ↦ 1) = κ := kernel.withDensity_one _ + +@[simp] +lemma withDensity_zero (κ : kernel α β) [IsSFiniteKernel κ] : + kernel.withDensity κ 0 = 0 := by + ext a; rw [kernel.withDensity_apply _ measurable_const]; simp + +@[simp] +lemma withDensity_zero' (κ : kernel α β) [IsSFiniteKernel κ] : + kernel.withDensity κ (fun _ _ ↦ 0) = 0 := kernel.withDensity_zero _ + theorem lintegral_withDensity (κ : kernel α β) [IsSFiniteKernel κ] (hf : Measurable (Function.uncurry f)) (a : α) {g : β → ℝ≥0∞} (hg : Measurable g) : ∫⁻ b, g b ∂withDensity κ f a = ∫⁻ b, f a b * g b ∂κ a := by @@ -134,6 +161,21 @@ theorem withDensity_tsum [Countable ι] (κ : kernel α β) [IsSFiniteKernel κ] rw [kernel.withDensity_apply' _ (hf n) a s] #align probability_theory.kernel.with_density_tsum ProbabilityTheory.kernel.withDensity_tsum +lemma withDensity_sub_add [IsSFiniteKernel κ] {f g : α → β → ℝ≥0∞} + (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) + (hg_int : ∀ a, ∫⁻ x, g a x ∂(κ a) ≠ ∞) (hfg : ∀ a, g a ≤ᵐ[κ a] f a) : + withDensity κ (fun a x ↦ f a x - g a x) + withDensity κ g = withDensity κ f := by + ext a s + simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add] + rw [kernel.withDensity_apply' _ hf, kernel.withDensity_apply' _ hg, kernel.withDensity_apply'] + swap; · exact hf.sub hg + rw [lintegral_sub] + · rw [tsub_add_cancel_iff_le] + exact lintegral_mono_ae (ae_restrict_of_ae (hfg a)) + · exact hg.comp measurable_prod_mk_left + · exact ((set_lintegral_le_lintegral _ _).trans_lt (hg_int a).lt_top).ne + · exact ae_restrict_of_ae (hfg a) + /-- If a kernel `κ` is finite and a function `f : α → β → ℝ≥0∞` is bounded, then `withDensity κ f` is finite. -/ theorem isFiniteKernel_withDensity_of_bounded (κ : kernel α β) [IsFiniteKernel κ] {B : ℝ≥0∞} @@ -230,4 +272,21 @@ instance (κ : kernel α β) [IsSFiniteKernel κ] (f : α → β → ℝ≥0) : IsSFiniteKernel (withDensity κ fun a b => f a b) := IsSFiniteKernel.withDensity κ fun _ _ => ENNReal.coe_ne_top +nonrec lemma withDensity_mul [IsSFiniteKernel κ] {f : α → β → ℝ≥0} {g : α → β → ℝ≥0∞} + (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) : + withDensity κ (fun a x ↦ f a x * g a x) + = withDensity (withDensity κ fun a x ↦ f a x) fun a x ↦ g a x := by + ext a : 1 + rw [kernel.withDensity_apply] + swap; · exact (measurable_coe_nnreal_ennreal.comp hf).mul hg + suffices (Measure.withDensity (κ a) ((fun x ↦ (f a x : ℝ≥0∞)) * (fun x ↦ (g a x : ℝ≥0∞)))) = + (withDensity (withDensity κ fun a x ↦ f a x) fun a x ↦ g a x) a by + convert this + rw [withDensity_mul] + · rw [kernel.withDensity_apply _ hg, kernel.withDensity_apply] + exact measurable_coe_nnreal_ennreal.comp hf + · rw [measurable_coe_nnreal_ennreal_iff] + exact hf.comp measurable_prod_mk_left + · exact hg.comp measurable_prod_mk_left + end ProbabilityTheory.kernel diff --git a/docs/references.bib b/docs/references.bib index 0b53c6728569a..6e12855aa6ece 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1769,8 +1769,19 @@ @Article{ KahnMaltsiniotis2008 url = {https://doi.org/10.1016/j.aim.2008.03.010} } +@Book{ kallenberg2017, + title = {Random Measures, Theory and Applications}, + author = {Kallenberg, Olav}, + series = {Probability Theory and Stochastic Modelling}, + volume = {77}, + year = {2017}, + publisher = {Springer}, + doi = {10.1007/978-3-319-41598-7}, + url = {https://doi.org/10.1007/978-3-319-41598-7} +} + @Book{ kallenberg2021, - author = {Olav Kallenberg}, + author = {Kallenberg, Olav}, title = {Foundations of modern probability}, series = {Probability Theory and Stochastic Modelling}, volume = {99}, From 6cc7107d8c902230997ee78313938cf46213eb1e Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Tue, 20 Feb 2024 20:36:33 +0100 Subject: [PATCH 041/129] use CountablyGenerated --- Mathlib.lean | 2 +- Mathlib/Data/Set/Countable.lean | 8 + .../MeasureTheory/MeasurableSpace/Basic.lean | 176 +++++ .../MeasureTheory/MeasurableSpace/Defs.lean | 4 + .../Kernel/Disintegration/Basic.lean | 120 +--- .../Kernel/Disintegration/Density.lean | 677 ++++++++---------- .../Kernel/Disintegration/Integral.lean | 6 +- .../Kernel/Disintegration/KernelCDFReal.lean | 39 +- .../Kernel/Disintegration/Unique.lean | 6 +- .../{Disintegration => }/RadonNikodym.lean | 49 +- 10 files changed, 573 insertions(+), 514 deletions(-) rename Mathlib/Probability/Kernel/{Disintegration => }/RadonNikodym.lean (88%) diff --git a/Mathlib.lean b/Mathlib.lean index aca67f25b225d..d1de337e5bfce 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3121,12 +3121,12 @@ import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.Disintegration.Integral import Mathlib.Probability.Kernel.Disintegration.KernelCDFReal import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes -import Mathlib.Probability.Kernel.Disintegration.RadonNikodym import Mathlib.Probability.Kernel.Disintegration.Unique import Mathlib.Probability.Kernel.IntegralCompProd import Mathlib.Probability.Kernel.Invariance import Mathlib.Probability.Kernel.MeasurableIntegral import Mathlib.Probability.Kernel.MeasureCompProd +import Mathlib.Probability.Kernel.RadonNikodym import Mathlib.Probability.Kernel.WithDensity import Mathlib.Probability.Martingale.Basic import Mathlib.Probability.Martingale.BorelCantelli diff --git a/Mathlib/Data/Set/Countable.lean b/Mathlib/Data/Set/Countable.lean index ab146e88767fd..c52d235209afb 100644 --- a/Mathlib/Data/Set/Countable.lean +++ b/Mathlib/Data/Set/Countable.lean @@ -95,6 +95,14 @@ theorem subset_range_enumerate {s : Set α} (h : s.Countable) (default : α) : simp [enumerateCountable, Encodable.encodek]⟩ #align set.subset_range_enumerate Set.subset_range_enumerate +lemma enumerateCountable_mem {s : Set α} (h : s.Countable) {default : α} (h_mem : default ∈ s) + (n : ℕ) : + enumerateCountable h default n ∈ s := by + rw [enumerateCountable] + match @decode s (Countable.toEncodable h) n with + | none => exact h_mem + | some val => simp only [Subtype.coe_prop] + end Enumerate theorem Countable.mono {s₁ s₂ : Set α} (h : s₁ ⊆ s₂) (hs : s₂.Countable) : s₁.Countable := diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 568cfec34e02a..0ca1cb8f854ee 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -1994,6 +1994,182 @@ instance [MeasurableSpace α] {s : Set α} [h : CountablyGenerated s] [Measurabl rw [← forall_generateFrom_mem_iff_mem_iff, ← hb] at h simpa using h {y} +def partitioned (f : ℕ → Set α) : ℕ → Set (Set α) + | 0 => {f 0, (f 0)ᶜ} + | n + 1 => ⋃ u ∈ partitioned f n, {u ∩ f (n + 1), u \ f (n + 1)} + +@[simp] +lemma partitioned_zero (f : ℕ → Set α) : partitioned f 0 = {f 0, (f 0)ᶜ} := rfl + +lemma partitioned_succ (f : ℕ → Set α) (n : ℕ) : + partitioned f (n + 1) = ⋃ u ∈ partitioned f n, {u ∩ f (n + 1), u \ f (n + 1)} := rfl + +lemma disjoint_partitioned (f : ℕ → Set α) (n : ℕ) : + ∀ u ∈ partitioned f n, ∀ v ∈ partitioned f n, u ≠ v → Disjoint u v := by + induction n with + | zero => + simp only [Nat.zero_eq, partitioned_zero, union_singleton, mem_insert_iff, + mem_singleton_iff, ne_eq, forall_eq_or_imp, forall_eq, not_true_eq_false, disjoint_self, + bot_eq_empty, compl_empty_iff, IsEmpty.forall_iff, true_and, and_true] + exact ⟨fun _ ↦ disjoint_compl_right, fun _ ↦ disjoint_compl_left⟩ + | succ n ih => + intro u hu v hv huv + rw [partitioned_succ] at hu hv + simp only [union_singleton, mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop] at hu hv + obtain ⟨u', hu', hu'_eq⟩ := hu + obtain ⟨v', hv', hv'_eq⟩ := hv + rcases hu'_eq with rfl | rfl <;> rcases hv'_eq with rfl | rfl + · refine Disjoint.mono (inter_subset_left _ _) (inter_subset_left _ _) (ih u' hu' v' hv' ?_) + exact fun huv' ↦ huv (huv' ▸ rfl) + · by_cases huv' : u' = v' + · rw [huv'] + exact Disjoint.mono_left (inter_subset_right _ _) Set.disjoint_sdiff_right + · exact Disjoint.mono (inter_subset_left _ _) (diff_subset _ _) (ih u' hu' v' hv' huv') + · by_cases huv' : u' = v' + · rw [huv'] + exact Disjoint.mono_right (inter_subset_right _ _) Set.disjoint_sdiff_left + · exact Disjoint.mono (diff_subset _ _) (inter_subset_left _ _) (ih u' hu' v' hv' huv') + · refine Disjoint.mono (diff_subset _ _) (diff_subset _ _) (ih u' hu' v' hv' ?_) + refine fun huv' ↦ huv (huv' ▸ rfl) + +lemma sUnion_partitioned (f : ℕ → Set α) (n : ℕ) : ⋃₀ partitioned f n = univ := by + induction n with + | zero => simp + | succ n ih => + rw [partitioned_succ] + ext x + have : x ∈ ⋃₀ partitioned f n := by simp [ih] + simp only [mem_sUnion, mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop, mem_univ, + iff_true] at this ⊢ + obtain ⟨t, ht, hxt⟩ := this + by_cases hxf : x ∈ f (n + 1) + · exact ⟨t ∩ f (n + 1), ⟨t, ht, Or.inl rfl⟩, hxt, hxf⟩ + · exact ⟨t \ f (n + 1), ⟨t, ht, Or.inr rfl⟩, hxt, hxf⟩ + +lemma partitioned_finite (f : ℕ → Set α) (n : ℕ) : Set.Finite (partitioned f n) := by + induction n with + | zero => simp + | succ n ih => + rw [partitioned_succ] + have : Finite (partitioned f n) := Set.finite_coe_iff.mp ih + rw [← Set.finite_coe_iff] + refine Finite.Set.finite_biUnion (partitioned f n) _ (fun u _ ↦ ?_) + rw [Set.finite_coe_iff] + simp + +lemma todo [m : MeasurableSpace α] [h : CountablyGenerated α] : + ∃ s : ℕ → Set (Set α), (∀ n, Set.Finite (s n)) + ∧ (∀ n, ∀ u ∈ s n, ∀ v ∈ s n, u ≠ v → Disjoint u v) + ∧ (∀ n, ⋃₀ s n = univ) + ∧ (∀ n, ∀ t ∈ s n, MeasurableSet[generateFrom (s (n + 1))] t) + ∧ m = generateFrom (⋃ n, s n) := by + obtain ⟨b, hb_count, hb_gen⟩ := h.isCountablyGenerated + cases Set.eq_empty_or_nonempty b with + | inl h_empty => + refine ⟨fun _ ↦ {univ}, by simp, by simp, by simp, ?_⟩ + simp only [iUnion_singleton_eq_range, range_const, generateFrom_singleton, mem_univ, + comap_const] + rw [hb_gen] + simp [h_empty] + | inr h_nonempty => + let s₀ := h_nonempty.choose + let t : ℕ → Set α := enumerateCountable hb_count s₀ + have ht_mem : ∀ n, t n ∈ b := enumerateCountable_mem hb_count h_nonempty.choose_spec + refine ⟨partitioned t, partitioned_finite t, disjoint_partitioned t, sUnion_partitioned t, ?_, ?_⟩ + · intro n u hun + rw [← diff_union_inter u (t (n + 1))] + refine MeasurableSet.union ?_ ?_ <;> + · refine measurableSet_generateFrom ?_ + rw [partitioned_succ] + simp_rw [mem_iUnion] + exact ⟨u, hun, by simp⟩ + · rw [hb_gen] + refine le_antisymm (generateFrom_le fun u hu ↦ ?_) (generateFrom_le fun u hu ↦ ?_) + · obtain ⟨n, rfl⟩ : ∃ n, u = t n := by + have h := subset_range_enumerate hb_count s₀ hu + obtain ⟨n, hn⟩ := mem_range.mpr h + exact ⟨n, hn.symm⟩ + suffices MeasurableSet[generateFrom (partitioned t n)] (t n) by + suffices h_le : generateFrom (partitioned t n) ≤ generateFrom (⋃ n, partitioned t n) by + exact h_le _ this + exact generateFrom_mono (subset_iUnion _ _) + cases n with + | zero => + simp only [Nat.zero_eq, partitioned_zero] + refine measurableSet_generateFrom ?_ + simp + | succ n => + have : t (n + 1) = ⋃ u ∈ partitioned t n, u ∩ t (n + 1) := by + simp_rw [← iUnion_inter] + suffices ⋃ u ∈ partitioned t n, u = univ by simp only [this, univ_inter] + rw [← sUnion_eq_biUnion] + exact sUnion_partitioned _ _ + rw [this] + refine MeasurableSet.biUnion (Finite.countable (partitioned_finite _ _)) (fun v hv ↦ ?_) + refine measurableSet_generateFrom ?_ + rw [partitioned_succ] + simp only [mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop] + exact ⟨v, hv, Or.inl rfl⟩ + · simp only [mem_iUnion] at hu + obtain ⟨n, hun⟩ := hu + induction n generalizing u with + | zero => + simp only [Nat.zero_eq, partitioned_zero, mem_insert_iff, mem_singleton_iff] at hun + rcases hun with rfl | rfl + · exact measurableSet_generateFrom (ht_mem _) + · exact (measurableSet_generateFrom (ht_mem _)).compl + | succ n ih => + simp only [partitioned_succ, mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop] + at hun + obtain ⟨v, hv, huv⟩ := hun + rcases huv with rfl | rfl + · exact MeasurableSet.inter (ih v hv) (measurableSet_generateFrom (ht_mem _)) + · exact MeasurableSet.diff (ih v hv) (measurableSet_generateFrom (ht_mem _)) + +def countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] : ℕ → Set (Set α) := + todo.choose + +lemma finite_countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : + Set.Finite (countablePartition α n) := + todo.choose_spec.1 n + +lemma disjoint_countablePartition [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) {s t : Set α} + (hs : s ∈ countablePartition α n) (ht : t ∈ countablePartition α n) (hst : s ≠ t) : + Disjoint s t := + todo.choose_spec.2.1 n s hs t ht hst + +lemma sUnion_countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : + ⋃₀ countablePartition α n = univ := + todo.choose_spec.2.2.1 n + +lemma measurableSet_succ_countablePartition [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) + {s : Set α} (hs : s ∈ countablePartition α n) : + MeasurableSet[generateFrom (countablePartition α (n + 1))] s := + todo.choose_spec.2.2.2.1 n s hs + +lemma generateFrom_iUnion_countablePartition (α : Type*) [m : MeasurableSpace α] + [CountablyGenerated α] : + generateFrom (⋃ n, countablePartition α n) = m := + todo.choose_spec.2.2.2.2.symm + +lemma generateFrom_countablePartition_le_succ (α : Type*) [MeasurableSpace α] [CountablyGenerated α] + (n : ℕ) : + generateFrom (countablePartition α n) ≤ generateFrom (countablePartition α (n + 1)) := + generateFrom_le (fun _ hs ↦ measurableSet_succ_countablePartition n hs) + +lemma generateFrom_countablePartition_le (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] + (n : ℕ) : + generateFrom (countablePartition α n) ≤ m := by + conv_rhs => rw [← generateFrom_iUnion_countablePartition α] + exact generateFrom_mono (subset_iUnion _ _) + +lemma measurableSet_countablePartition [m : MeasurableSpace α] [CountablyGenerated α] (n : ℕ) + {s : Set α} (hs : s ∈ countablePartition α n) : + MeasurableSet s := by + have : MeasurableSet[generateFrom (countablePartition α n)] s := + measurableSet_generateFrom hs + exact generateFrom_countablePartition_le α n _ this + variable (α) open Classical diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean b/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean index 31b1db9d8045f..e1a665fc570a7 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean @@ -443,6 +443,10 @@ instance : CompleteLattice (MeasurableSpace α) := instance : Inhabited (MeasurableSpace α) := ⟨⊤⟩ +@[simp] +lemma generateFrom_empty : generateFrom (∅ : Set (Set α)) = ⊥ := + le_bot_iff.mp (generateFrom_le (by simp)) + @[mono] theorem generateFrom_mono {s t : Set (Set α)} (h : s ⊆ t) : generateFrom s ≤ generateFrom t := giGenerateFrom.gc.monotone_l h diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 3605950f91482..7bde52b4639f6 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -34,11 +34,12 @@ The first step (building the measurable function on `ℚ`) is done differently d put those together into a `kernel (α × β) ℝ`. The construction of that `f` is done in the `CondCDF.lean` file. * If `α` is not countable, we can't proceed separately for each `a : α` and have to build a function - `f : α × β → ℚ → ℝ` which is measurable on the product. We are able to do so if `β` is standard - borel by reducing to the case `β = ℝ`. See the file `KernelCDFBorel.lean`. + `f : α × β → ℚ → ℝ` which is measurable on the product. We are able to do so if `β` has a + countably generated σ-algebra (this is the case in particular for standard Borel spaces). + See the file `KernelCDFBorel.lean`. -We define a class `CountableOrStandardBorel α β` which encodes the property -`(Countable α ∧ MeasurableSingletonClass α) ∨ StandardBorelSpace β`. The conditional kernel is +We define a class `CountableOrCountablyGenerated α β` which encodes the property +`(Countable α ∧ MeasurableSingletonClass α) ∨ CountablyGenerated β`. The conditional kernel is defined under that assumption. Properties of integrals involving `kernel.condKernel` are collated in the file `Integral.lean`. @@ -83,7 +84,8 @@ open scoped ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory -variable {α β Ω Ω': Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} +variable {α β γ Ω Ω': Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + {mγ : MeasurableSpace γ} [MeasurableSpace.CountablyGenerated γ] [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] [MeasurableSpace Ω'] [StandardBorelSpace Ω'] [Nonempty Ω'] @@ -208,91 +210,31 @@ end BorelSnd section StandardBorel noncomputable -def kernel.condKernelReal (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : kernel (α × ℝ) ℝ := +def kernel.condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : kernel (α × γ) ℝ := cdfKernel _ (isKernelCDF_mLimsupCDF κ) -instance (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : IsMarkovKernel (kernel.condKernelReal κ) := by +instance (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : IsMarkovKernel (kernel.condKernelReal κ) := by unfold kernel.condKernelReal; infer_instance -lemma kernel.eq_compProd_condKernelReal (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : +lemma kernel.eq_compProd_condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : κ = kernel.fst κ ⊗ₖ kernel.condKernelReal κ := kernel.eq_compProd_cdfKernel (isKernelCDF_mLimsupCDF κ) noncomputable -def condKernelBorelAux (κ : kernel α (ℝ × Ω)) [IsFiniteKernel κ] : kernel (α × ℝ) Ω := +def condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : kernel (α × γ) Ω := let f := measurableEmbedding_real Ω let hf := measurableEmbedding_measurableEmbedding_real Ω - let κ' := kernel.map κ (Prod.map (id : ℝ → ℝ) f) (measurable_id.prod_map hf.measurable) + let κ' := kernel.map κ (Prod.map (id : γ → γ) f) (measurable_id.prod_map hf.measurable) condKernelBorelSnd κ (isKernelCDF_mLimsupCDF κ') -instance instIsMarkovKernel_condKernelBorelAux (κ : kernel α (ℝ × Ω)) [IsFiniteKernel κ] : - IsMarkovKernel (condKernelBorelAux κ) := by - rw [condKernelBorelAux] +instance instIsMarkovKernel_condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : + IsMarkovKernel (condKernelBorel κ) := by + rw [condKernelBorel] infer_instance -lemma compProd_fst_condKernelBorelAux (κ : kernel α (ℝ × Ω)) [IsFiniteKernel κ] : - kernel.fst κ ⊗ₖ condKernelBorelAux κ = κ := by - rw [condKernelBorelAux, compProd_fst_condKernelBorelSnd] - -noncomputable -def condKernelBorel (κ : kernel α (Ω × Ω')) [IsFiniteKernel κ] : kernel (α × Ω) Ω' := - let f := measurableEmbedding_real Ω - let hf := measurableEmbedding_measurableEmbedding_real Ω - let κ' := kernel.map κ (Prod.map f (id : Ω' → Ω')) (hf.measurable.prod_map measurable_id) - kernel.comap (condKernelBorelAux κ') (Prod.map (id : α → α) f) - (measurable_id.prod_map hf.measurable) - -instance instIsMarkovKernel_condKernelBorel (κ : kernel α (Ω × Ω')) [IsFiniteKernel κ] : - IsMarkovKernel (condKernelBorel κ) := by rw [condKernelBorel]; infer_instance - -lemma compProd_fst_condKernelBorel (κ : kernel α (Ω × Ω')) [IsFiniteKernel κ] : +lemma compProd_fst_condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : kernel.fst κ ⊗ₖ condKernelBorel κ = κ := by - let f := measurableEmbedding_real Ω - let hf := measurableEmbedding_measurableEmbedding_real Ω - let κ' : kernel α (ℝ × Ω') := - kernel.map κ (Prod.map f (id : Ω' → Ω')) (hf.measurable.prod_map measurable_id) - have h_condKernel : condKernelBorel κ - = kernel.comap (condKernelBorelAux κ') (Prod.map (id : α → α) f) - (measurable_id.prod_map hf.measurable) := rfl - have h_compProd := compProd_fst_condKernelBorelAux κ' - have h_prod_embed : MeasurableEmbedding (Prod.map f (id : Ω' → Ω')) := - hf.prod_mk MeasurableEmbedding.id - have : κ = kernel.comapRight κ' h_prod_embed := by - ext c t ht : 2 - unfold_let κ' - rw [kernel.comapRight_apply' _ _ _ ht, kernel.map_apply' κ h_prod_embed.measurable - _ (h_prod_embed.measurableSet_image.mpr ht)] - congr with x : 1 - rw [← @Prod.mk.eta _ _ x] - simp only [Prod.mk.eta, Prod_map, id_eq, mem_preimage, mem_image, Prod.mk.injEq, Prod.exists, - exists_eq_right_right] - refine ⟨fun h ↦ ⟨x.1, h, rfl⟩, ?_⟩ - rintro ⟨ω, h_mem, h⟩ - rwa [hf.injective h] at h_mem - have h_fst : kernel.fst κ' = kernel.map (kernel.fst κ) f hf.measurable := by - ext a s hs - unfold_let κ' - rw [kernel.map_apply' _ _ _ hs, kernel.fst_apply', kernel.fst_apply', kernel.map_apply'] - · congr - · exact measurable_fst hs - · exact hf.measurable hs - · exact hs - conv_rhs => rw [this, ← h_compProd] - ext a s hs - rw [h_condKernel, h_fst] - rw [kernel.comapRight_apply' _ _ _ hs, kernel.compProd_apply _ _ _ hs, kernel.compProd_apply] - swap; · exact h_prod_embed.measurableSet_image.mpr hs - rw [kernel.map_apply, lintegral_map] - congr with ω - rw [kernel.comap_apply'] - · congr with ω' - simp only [mem_setOf_eq, Prod_map, id_eq, mem_image, Prod.mk.injEq, Prod.exists, - exists_eq_right_right] - refine ⟨fun h ↦ ⟨ω, h, rfl⟩, ?_⟩ - rintro ⟨a, h_mem, h⟩ - rwa [hf.injective h] at h_mem - · exact kernel.measurable_kernel_prod_mk_left' (h_prod_embed.measurableSet_image.mpr hs) _ - · exact hf.measurable + rw [condKernelBorel, compProd_fst_condKernelBorelSnd] end StandardBorel @@ -442,35 +384,37 @@ lemma compProd_fst_condKernelCountable (κ : kernel α (β × Ω)) [IsFiniteKern end Countable -section CountableOrStandardBorel +section CountableOrCountablyGenerated -class CountableOrStandardBorel (α β : Type*) [MeasurableSpace α] [MeasurableSpace β] : Prop := - (countableOrStandardBorel : (Countable α ∧ MeasurableSingletonClass α) ∨ StandardBorelSpace β) +class CountableOrCountablyGenerated (α β : Type*) [MeasurableSpace α] [MeasurableSpace β] : Prop := + (countableOrCountablyGenerated : + (Countable α ∧ MeasurableSingletonClass α) ∨ MeasurableSpace.CountablyGenerated β) -instance instCountableOrStandardBorel_of_countable +instance instCountableOrCountablyGenerated_of_countable [h1 : Countable α] [h2 : MeasurableSingletonClass α] : - CountableOrStandardBorel α β := ⟨Or.inl ⟨h1, h2⟩⟩ + CountableOrCountablyGenerated α β := ⟨Or.inl ⟨h1, h2⟩⟩ -instance instCountableOrStandardBorel_of_standardBorelSpace [h : StandardBorelSpace β] : - CountableOrStandardBorel α β := ⟨Or.inr h⟩ +instance instCountableOrCountablyGenerated_of_standardBorelSpace + [h : MeasurableSpace.CountablyGenerated β] : + CountableOrCountablyGenerated α β := ⟨Or.inr h⟩ open Classical in noncomputable -def kernel.condKernel [h : CountableOrStandardBorel α β] +def kernel.condKernel [h : CountableOrCountablyGenerated α β] (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : kernel (α × β) Ω' := if hα : Countable α ∧ MeasurableSingletonClass α then letI := hα.1; letI := hα.2; condKernelCountable κ else - letI := h.countableOrStandardBorel.resolve_left hα; condKernelBorel κ + letI := h.countableOrCountablyGenerated.resolve_left hα; condKernelBorel κ -instance kernel.instIsMarkovKernel_condKernel [CountableOrStandardBorel α β] +instance kernel.instIsMarkovKernel_condKernel [CountableOrCountablyGenerated α β] (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : IsMarkovKernel (kernel.condKernel κ) := by unfold kernel.condKernel split_ifs <;> infer_instance -lemma kernel.compProd_fst_condKernel [hαβ : CountableOrStandardBorel α β] +lemma kernel.compProd_fst_condKernel [hαβ : CountableOrCountablyGenerated α β] (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : kernel.fst κ ⊗ₖ kernel.condKernel κ = κ := by unfold kernel.condKernel @@ -478,9 +422,9 @@ lemma kernel.compProd_fst_condKernel [hαβ : CountableOrStandardBorel α β] · have := h.1 have := h.2 exact compProd_fst_condKernelCountable κ - · have := hαβ.countableOrStandardBorel.resolve_left h + · have := hαβ.countableOrCountablyGenerated.resolve_left h exact compProd_fst_condKernelBorel κ -end CountableOrStandardBorel +end CountableOrCountablyGenerated end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 7d24ed07895b7..f97954d067c75 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -54,8 +54,11 @@ desired density function. The construction of the density process in this file follows the proof of Theorem 9.27 in [O. Kallenberg, Foundations of modern probability][kallenberg2021]. + +TODO: adapted to use countably generated sigma-algebras. -/ +variable {α β : Type*} {mα : MeasurableSpace α} open MeasureTheory Set Filter @@ -63,24 +66,6 @@ open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory -variable {α β : Type*} {mα : MeasurableSpace α} - -section dissection_system - -def IcoPowTwo (n : ℕ) (k : ℤ) : Set ℝ := Set.Ico (k * (2⁻¹ : ℝ) ^ n) ((k + 1) * ((2 : ℝ) ^ n)⁻¹) - -lemma mem_IcoPowTwo_iff_mul {n : ℕ} {k : ℤ} (x : ℝ) : - x ∈ IcoPowTwo n k ↔ k ≤ x * 2 ^ n ∧ x * 2 ^ n < k + 1 := by - simp only [IcoPowTwo, inv_pow, mem_Ico] - rw [← div_eq_mul_inv, div_le_iff, ← div_eq_mul_inv, lt_div_iff] - · positivity - · positivity - -lemma mem_IcoPowTwo_iff_floor {n : ℕ} {k : ℤ} (x : ℝ) : x ∈ IcoPowTwo n k ↔ ⌊x * 2 ^ n⌋ = k := by - simp [mem_IcoPowTwo_iff_mul, Int.floor_eq_iff] - -lemma measurableSet_IcoPowTwo (n : ℕ) (k : ℤ) : MeasurableSet (IcoPowTwo n k) := measurableSet_Ico - lemma Measure.iInf_Iic_gt_prod {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] {s : Set α} (hs : MeasurableSet s) (t : ℚ) : ⨅ r : { r' : ℚ // t < r' }, ρ (s ×ˢ Iic (r : ℝ)) = ρ (s ×ˢ Iic (t : ℝ)) := by @@ -98,272 +83,203 @@ lemma Measure.iInf_Iic_gt_prod {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] exact mod_cast hrr' · exact ⟨⟨t + 1, lt_add_one _⟩, measure_ne_top ρ _⟩ -lemma pairwise_disjoint_IcoPowTwo (n : ℕ) : Pairwise (Disjoint on fun k ↦ IcoPowTwo n k) := by - intro i j hij - rw [Function.onFun, Set.disjoint_iff] - intro x - simp only [mem_inter_iff, mem_IcoPowTwo_iff_floor, mem_empty_iff_false, and_imp, imp_false] - intro hi hj - rw [hi] at hj - exact hij hj - -lemma IcoPowTwo_succ_union (n : ℕ) (k : ℤ) : - IcoPowTwo (n+1) (2 * k) ∪ IcoPowTwo (n+1) (2 * k + 1) = IcoPowTwo n k := by - ext x - cases lt_or_le x ((2 * k + 1) * ((2 : ℝ) ^ (n+1))⁻¹) with - | inl h => - simp only [IcoPowTwo, inv_pow, mem_Ico, Int.cast_mul, Int.int_cast_ofNat, Int.cast_add, - Int.cast_one, mem_union, h, and_true, not_le.mpr h, false_and, or_false] - have : x < (k + 1) * (2 ^ n)⁻¹ := by - refine h.trans_le ?_ - rw [pow_add, pow_one, mul_inv, mul_comm _ 2⁻¹, ← mul_assoc] - gcongr - rw [add_mul, one_mul, mul_comm, ← mul_assoc, inv_mul_cancel two_ne_zero, one_mul] - gcongr - norm_num - simp only [this, and_true] - rw [pow_add, pow_one, mul_inv, mul_comm _ 2⁻¹, ← mul_assoc, mul_comm _ 2⁻¹, ← mul_assoc, - inv_mul_cancel two_ne_zero, one_mul] - | inr h => - simp only [IcoPowTwo, inv_pow, mem_Ico, Int.cast_mul, Int.int_cast_ofNat, Int.cast_add, - Int.cast_one, mem_union, not_lt.mpr h, and_false, h, true_and, false_or] - have : k * (2 ^ n)⁻¹ ≤ x := by - refine le_trans ?_ h - rw [pow_add, pow_one, mul_inv, mul_comm _ 2⁻¹, ← mul_assoc, mul_comm _ 2⁻¹, mul_add, - ← mul_assoc, inv_mul_cancel two_ne_zero, mul_one, one_mul, add_mul] - simp only [le_add_iff_nonneg_right, gt_iff_lt, inv_pos, zero_lt_two, - mul_nonneg_iff_of_pos_left, inv_nonneg] - positivity - simp only [this, true_and] - rw [pow_add, pow_one, mul_inv, mul_comm _ 2⁻¹, ← mul_assoc, mul_comm _ 2⁻¹, add_assoc] - norm_num - rw [one_div, mul_add, ← mul_assoc, inv_mul_cancel two_ne_zero, one_mul] - -noncomputable def indexIcoPowTwo (n : ℕ) (t : ℝ) : ℤ := Int.floor (t * 2 ^ n) - -lemma mem_IcoPowTwo_indexIcoPowTwo (n : ℕ) (t : ℝ) : t ∈ IcoPowTwo n (indexIcoPowTwo n t) := by - rw [indexIcoPowTwo, IcoPowTwo] - simp only [inv_pow, mem_Ico] - constructor - · rw [← div_eq_mul_inv, div_le_iff] - · exact Int.floor_le (t * 2 ^ n) - · positivity - · rw [← div_eq_mul_inv, lt_div_iff] - · exact Int.lt_floor_add_one (t * 2 ^ n) - · positivity - -lemma indexIcoPowTwo_of_mem (n : ℕ) (k : ℤ) (t : ℝ) (ht : t ∈ IcoPowTwo n k) : - indexIcoPowTwo n t = k := by - rw [indexIcoPowTwo] - simp only [IcoPowTwo, inv_pow, mem_Ico, ← div_eq_mul_inv] at ht - rw [div_le_iff, lt_div_iff] at ht - · rw [Int.floor_eq_iff] - exact ht - · positivity - · positivity - -lemma mem_IcoPowTwo_iff_indexIcoPowTwo (n : ℕ) (k : ℤ) (t : ℝ) : - t ∈ IcoPowTwo n k ↔ indexIcoPowTwo n t = k := - ⟨fun h ↦ indexIcoPowTwo_of_mem n k t h, fun h ↦ h ▸ mem_IcoPowTwo_indexIcoPowTwo n t⟩ - -lemma iUnion_IcoPowTwo (n : ℕ) : ⋃ k, IcoPowTwo n k = univ := by - ext x - simp only [mem_iUnion, mem_univ, iff_true] - exact ⟨indexIcoPowTwo n x, mem_IcoPowTwo_indexIcoPowTwo n x⟩ - -lemma indexIcoPowTwo_le_indexIcoPowTwo_iff (n : ℕ) (t x : ℝ) : - indexIcoPowTwo n t ≤ indexIcoPowTwo n x ↔ ⌊t * 2 ^ n⌋ * (2 ^ n)⁻¹ ≤ x := by - simp only [indexIcoPowTwo] - rw [← div_eq_mul_inv, div_le_iff, Int.le_floor] - positivity - -lemma iUnion_ge_IcoPowTwo (n : ℕ) (t : ℝ) : - ⋃ (k) (_ : indexIcoPowTwo n t ≤ k), IcoPowTwo n k = Ici (⌊t * 2 ^ n⌋ * (2 ^ n)⁻¹ : ℝ) := by - ext x - simp [mem_IcoPowTwo_iff_indexIcoPowTwo, indexIcoPowTwo_le_indexIcoPowTwo_iff] - -lemma iInter_biUnion_I (x : ℝ) : - ⋂ n, ⋃ (k) (_ : indexIcoPowTwo n x ≤ k), IcoPowTwo n k = Ici x := by - ext t - simp [iUnion_ge_IcoPowTwo] - refine ⟨fun h ↦ ?_, fun h n ↦ le_trans ?_ h⟩ - · by_contra h_lt - push_neg at h_lt - have h_pos : ∀ i, 0 < (2 : ℝ) ^ i := fun i ↦ by positivity - simp_rw [← div_eq_mul_inv, div_le_iff (h_pos _)] at h - obtain ⟨i, hi⟩ : ∃ i, 1 < (x - t) * 2 ^ i := by - suffices ∃ i : ℝ, 1 ≤ (x - t) * 2 ^ i by - obtain ⟨i, hi⟩ := this - refine ⟨⌈i⌉₊ + 1, hi.trans_lt ?_⟩ - gcongr - · simp [h_lt] - · refine ((Real.rpow_lt_rpow_left_iff one_lt_two).mpr (?_ : i < ⌈i⌉₊ + 1)).trans_eq ?_ - · refine (Nat.le_ceil _).trans_lt ?_ - norm_num - · norm_cast - use Real.logb 2 ((x - t)⁻¹) - rw [Real.rpow_logb] - · rw [mul_inv_cancel] - rw [sub_ne_zero] - exact h_lt.ne' - · exact zero_lt_two - · simp - · simp [h_lt] - specialize h i - rw [mul_comm, mul_sub, lt_sub_iff_add_lt', mul_comm] at hi - have h' : ⌈x * 2 ^ i⌉ ≤ t * 2 ^ i + 1 := by - calc (⌈x * 2 ^ i⌉ : ℝ) ≤ ⌊x * 2 ^ i⌋ + 1 := by - exact mod_cast (Int.ceil_le_floor_add_one (x * 2 ^ i)) - _ ≤ t * 2 ^ i + 1 := by gcongr - have h'' : ↑⌈x * 2 ^ i⌉ < 2 ^ i * x := h'.trans_lt hi - rw [← not_le, mul_comm] at h'' - exact h'' (Int.le_ceil _) - · rw [← div_eq_mul_inv, div_le_iff] - · exact Int.floor_le (x * 2 ^ n) - · positivity +variable {α β γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} + [MeasurableSpace.CountablyGenerated γ] -- todo : `Filtration` should be renamed to `filtration` -def ℱ : Filtration ℕ (borel ℝ) where - seq := fun n ↦ MeasurableSpace.generateFrom {s | ∃ k, s = IcoPowTwo n k} - mono' := by - refine monotone_nat_of_le_succ ?_ - intro n - refine MeasurableSpace.generateFrom_le fun s ⟨k, hs⟩ ↦ ?_ - rw [hs, ← IcoPowTwo_succ_union n k] - refine MeasurableSet.union ?_ ?_ - · exact MeasurableSpace.measurableSet_generateFrom ⟨2 * k, rfl⟩ - · exact MeasurableSpace.measurableSet_generateFrom ⟨2 * k + 1, rfl⟩ - le' := fun n ↦ by - refine MeasurableSpace.generateFrom_le fun s ⟨k, hs⟩ ↦ ?_ - rw [hs] - exact measurableSet_IcoPowTwo n k - -lemma measurableSet_ℱ_IcoPowTwo (n : ℕ) (k : ℤ) : MeasurableSet[ℱ n] (IcoPowTwo n k) := - MeasurableSpace.measurableSet_generateFrom ⟨k, rfl⟩ - -lemma measurable_indexIcoPowTwo (n : ℕ) : Measurable[ℱ n] (indexIcoPowTwo n) := by - refine @measurable_to_countable' ℤ ℝ _ _ (ℱ n) _ (fun k ↦ ?_) - have : (fun t ↦ ⌊t * (2 : ℝ) ^ n⌋) ⁻¹' {k} = IcoPowTwo n k := by - ext t - simp only [mem_IcoPowTwo_iff_floor, mem_preimage, mem_singleton_iff] - unfold indexIcoPowTwo +def ℱ (γ : Type*) [m : MeasurableSpace γ] [MeasurableSpace.CountablyGenerated γ] : + Filtration ℕ m where + seq := fun n ↦ MeasurableSpace.generateFrom <| MeasurableSpace.countablePartition γ n + mono' := monotone_nat_of_le_succ (MeasurableSpace.generateFrom_countablePartition_le_succ _) + le' := MeasurableSpace.generateFrom_countablePartition_le γ + +lemma measurableSet_ℱ_countablePartition (n : ℕ) {s : Set γ} + (hs : s ∈ MeasurableSpace.countablePartition γ n) : + MeasurableSet[ℱ γ n] s := + MeasurableSpace.measurableSet_generateFrom hs + +lemma existsPartitionSet_mem (n : ℕ) (t : γ) : + ∃ s, s ∈ MeasurableSpace.countablePartition γ n ∧ t ∈ s := by + have h_univ := MeasurableSpace.sUnion_countablePartition γ n + have h_mem_univ := mem_univ t + rw [← h_univ] at h_mem_univ + simpa only [mem_sUnion] using h_mem_univ + +def partitionSet (n : ℕ) (t : γ) : Set γ := + (existsPartitionSet_mem n t).choose + +lemma partitionSet_mem (n : ℕ) (t : γ) : + partitionSet n t ∈ MeasurableSpace.countablePartition γ n := + (existsPartitionSet_mem n t).choose_spec.1 + +lemma mem_partitionSet (n : ℕ) (t : γ) : t ∈ partitionSet n t := + (existsPartitionSet_mem n t).choose_spec.2 + +lemma mem_countablePartition_iff (n : ℕ) (t : γ) {s : Set γ} + (hs : s ∈ MeasurableSpace.countablePartition γ n) : + t ∈ s ↔ partitionSet n t = s := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · by_contra h_ne + have h_disj : Disjoint s (partitionSet n t) := + MeasurableSpace.disjoint_countablePartition n hs (partitionSet_mem n t) (Ne.symm h_ne) + refine absurd h_disj ?_ + rw [not_disjoint_iff_nonempty_inter] + exact ⟨t, h, mem_partitionSet n t⟩ + · rw [← h] + exact mem_partitionSet n t + +lemma partitionSet_of_mem {n : ℕ} {t : γ} {s : Set γ} + (hs : s ∈ MeasurableSpace.countablePartition γ n) (ht : t ∈ s) : + partitionSet n t = s := by + rwa [← mem_countablePartition_iff n t hs] + +lemma measurableSet_ℱ_partitionSet (n : ℕ) (t : γ) : + MeasurableSet[ℱ γ n] (partitionSet n t) := + measurableSet_ℱ_countablePartition n (partitionSet_mem n t) + +lemma measurableSet_partitionSet (n : ℕ) (t : γ) : + MeasurableSet (partitionSet n t) := + (ℱ γ).le n _ (measurableSet_ℱ_partitionSet n t) + +instance todo_move (n : ℕ) : Countable (MeasurableSpace.countablePartition γ n) := + Set.Finite.countable (MeasurableSpace.finite_countablePartition _ _) + +lemma measurable_aux (n : ℕ) (m : MeasurableSpace (MeasurableSpace.countablePartition γ n)) : + @Measurable γ (MeasurableSpace.countablePartition γ n) (ℱ γ n) m + (fun c : γ ↦ ⟨partitionSet n c, partitionSet_mem n c⟩) := by + refine @measurable_to_countable' (MeasurableSpace.countablePartition γ n) γ m _ + (ℱ γ n) _ (fun t ↦ ?_) + rcases t with ⟨t, ht⟩ + suffices MeasurableSet[ℱ γ n] {x | partitionSet n x = t} by + convert this + ext x + simp + have : {x | partitionSet n x = t} = t := by + ext x + rw [mem_setOf_eq, mem_countablePartition_iff n x ht] rw [this] - exact measurableSet_ℱ_IcoPowTwo n k - -lemma iSup_ℱ : ⨆ n, ℱ n = borel ℝ := by - refine le_antisymm ?_ ?_ - · rw [iSup_le_iff] - exact ℱ.le - · conv_lhs => rw [borel_eq_generateFrom_Ici ℝ] - refine MeasurableSpace.generateFrom_le (fun s ⟨x, hx⟩ ↦ ?_) - rw [← hx, ← iInter_biUnion_I x] - refine MeasurableSet.iInter (fun n ↦ ?_) - refine MeasurableSet.biUnion ?_ (fun k _ ↦ ?_) - · exact to_countable _ - · exact le_iSup ℱ n _ (measurableSet_ℱ_IcoPowTwo n k) - -end dissection_system + exact measurableSet_ℱ_countablePartition _ ht + +lemma measurable_partitionSet (n : ℕ) : Measurable[ℱ γ n] (partitionSet n) := by + have : partitionSet n = (fun s : MeasurableSpace.countablePartition γ n ↦ (s : Set γ)) + ∘ (fun t ↦ ⟨partitionSet n t, partitionSet_mem n t⟩) := rfl + rw [this] + refine Measurable.comp + (?_ : Measurable (fun s : MeasurableSpace.countablePartition γ n ↦ (s : Set γ))) ?_ + · measurability + exact measurable_aux _ _ + +lemma MeasurableSpace.generateFrom_iSup {ι : Type*} (s : ι → Set (Set α)) : + MeasurableSpace.generateFrom (⋃ i, s i) = ⨆ i, MeasurableSpace.generateFrom (s i) := + (@MeasurableSpace.giGenerateFrom α).gc.l_iSup + +lemma iSup_ℱ (β : Type*) [m : MeasurableSpace β] [MeasurableSpace.CountablyGenerated β] : + ⨆ n, ℱ β n = m := by + conv_rhs => rw [← MeasurableSpace.generateFrom_iUnion_countablePartition β] + rw [MeasurableSpace.generateFrom_iSup] + rfl variable [MeasurableSpace β] section DensityProcess -variable {κ : kernel α (ℝ × β)} {ν : kernel α ℝ} +variable {κ : kernel α (γ × β)} {ν : kernel α γ} noncomputable -def densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) (a : α) (t : ℝ) (s : Set β) : +def densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (t : γ) (s : Set β) : ℝ := - (κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) / ν a (IcoPowTwo n (indexIcoPowTwo n t))).toReal + (κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t)).toReal -lemma densityProcess_def (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) (a : α) (s : Set β) : +lemma densityProcess_def (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (s : Set β) : (fun t ↦ densityProcess κ ν n a t s) - = fun t ↦ (κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) - / ν a (IcoPowTwo n (indexIcoPowTwo n t))).toReal := + = fun t ↦ (κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t)).toReal := rfl -lemma measurable_densityProcess_aux (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) +lemma measurable_densityProcess_ℱ_aux (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) {s : Set β} (hs : MeasurableSet s) : - Measurable (fun (p : α × ℝ) ↦ κ p.1 (IcoPowTwo n (indexIcoPowTwo n p.2) ×ˢ s) - / ν p.1 (IcoPowTwo n (indexIcoPowTwo n p.2))) := by - change Measurable ((fun (p : α × ℤ) ↦ κ p.1 (IcoPowTwo n p.2 ×ˢ s) / ν p.1 (IcoPowTwo n p.2)) - ∘ (fun (p : α × ℝ) ↦ (p.1, indexIcoPowTwo n p.2))) - have h1 : - Measurable (fun (p : α × ℤ) ↦ κ p.1 (IcoPowTwo n p.2 ×ˢ s) / ν p.1 (IcoPowTwo n p.2)) := by + Measurable[MeasurableSpace.prod mα (ℱ γ n)] (fun (p : α × γ) ↦ + κ p.1 (partitionSet n p.2 ×ˢ s) / ν p.1 (partitionSet n p.2)) := by + change Measurable[MeasurableSpace.prod mα (ℱ γ n)] + ((fun (p : α × MeasurableSpace.countablePartition γ n) ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) + ∘ (fun (p : α × γ) ↦ (p.1, ⟨partitionSet n p.2, partitionSet_mem n p.2⟩))) + have h1 : @Measurable _ _ (MeasurableSpace.prod mα ⊤) _ + (fun (p : α × MeasurableSpace.countablePartition γ n) ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) := by refine Measurable.div ?_ ?_ - · have h_swap : Measurable fun (p : ℤ × α) ↦ κ p.2 (IcoPowTwo n p.1 ×ˢ s) := by - refine measurable_uncurry_of_continuous_of_measurable - (u := fun k a ↦ κ a (IcoPowTwo n k ×ˢ s)) ?_ ?_ - · exact fun _ ↦ continuous_bot - · exact fun _ ↦ kernel.measurable_coe _ ((measurableSet_IcoPowTwo _ _).prod hs) - change Measurable ((fun (p : ℤ × α) ↦ κ p.2 (IcoPowTwo n p.1 ×ˢ s)) ∘ Prod.swap) - exact h_swap.comp measurable_swap - · have h_swap : Measurable fun (p : ℤ × α) ↦ ν p.2 (IcoPowTwo n p.1) := by - refine measurable_uncurry_of_continuous_of_measurable - (u := fun k a ↦ ν a (IcoPowTwo n k)) ?_ ?_ - · exact fun _ ↦ continuous_bot - · exact fun _ ↦ kernel.measurable_coe _ (measurableSet_IcoPowTwo _ _) - change Measurable ((fun (p : ℤ × α) ↦ ν p.2 (IcoPowTwo n p.1)) ∘ Prod.swap) - exact h_swap.comp measurable_swap - refine h1.comp (measurable_fst.prod_mk ?_) - exact (Measurable.mono (measurable_indexIcoPowTwo n) (ℱ.le n) le_rfl).comp measurable_snd - -lemma measurable_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) + · refine measurable_from_prod_countable ?_ + rintro ⟨t, ht⟩ + simp only + refine kernel.measurable_coe _ (MeasurableSet.prod ?_ hs) + exact MeasurableSpace.measurableSet_countablePartition _ ht + · refine measurable_from_prod_countable ?_ + rintro ⟨t, ht⟩ + simp only + exact kernel.measurable_coe _ (MeasurableSpace.measurableSet_countablePartition _ ht) + refine h1.comp ?_ + refine measurable_fst.prod_mk ?_ + change @Measurable (α × γ) (MeasurableSpace.countablePartition γ n) + (MeasurableSpace.prod mα (ℱ γ n)) ⊤ + ((fun c : γ ↦ ⟨partitionSet n c, partitionSet_mem n c⟩) ∘ (fun p : α × γ ↦ p.2)) + refine Measurable.comp ?_ measurable_snd + exact measurable_aux n ⊤ + +lemma measurable_densityProcess_aux (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) + {s : Set β} (hs : MeasurableSet s) : + Measurable (fun (p : α × γ) ↦ + κ p.1 (partitionSet n p.2 ×ˢ s) / ν p.1 (partitionSet n p.2)) := by + refine Measurable.mono (measurable_densityProcess_ℱ_aux κ ν n hs) ?_ le_rfl + exact sup_le_sup le_rfl (MeasurableSpace.comap_mono ((ℱ γ).le _)) + +lemma measurable_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) {s : Set β} (hs : MeasurableSet s) : - Measurable (fun (p : α × ℝ) ↦ densityProcess κ ν n p.1 p.2 s) := + Measurable (fun (p : α × γ) ↦ densityProcess κ ν n p.1 p.2 s) := (measurable_densityProcess_aux κ ν n hs).ennreal_toReal -lemma measurable_densityProcess_left (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) - (t : ℝ) {s : Set β} (hs : MeasurableSet s) : +lemma measurable_densityProcess_left (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) + (t : γ) {s : Set β} (hs : MeasurableSet s) : Measurable (fun a ↦ densityProcess κ ν n a t s) := (measurable_densityProcess κ ν n hs).comp (measurable_id.prod_mk measurable_const) -lemma measurable_densityProcess_right (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) +lemma measurable_densityProcess_right (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) {s : Set β} (a : α) (hs : MeasurableSet s) : Measurable (fun t ↦ densityProcess κ ν n a t s) := (measurable_densityProcess κ ν n hs).comp (measurable_const.prod_mk measurable_id) -lemma measurable_ℱ_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) - (a : α) (s : Set β) : - Measurable[ℱ n] (fun t ↦ densityProcess κ ν n a t s) := by - rw [densityProcess_def] - refine @Measurable.ennreal_toReal _ (ℱ n) _ ?_ - refine Measurable.div ?_ ?_ - · change Measurable[ℱ n] ((fun k ↦ κ a (IcoPowTwo n k ×ˢ s)) ∘ (indexIcoPowTwo n)) - refine Measurable.comp ?_ (measurable_indexIcoPowTwo n) - exact measurable_of_countable _ - · change Measurable[ℱ n] ((fun k ↦ ν a (IcoPowTwo n k)) ∘ (indexIcoPowTwo n)) - refine Measurable.comp ?_ (measurable_indexIcoPowTwo n) - exact measurable_of_countable _ - -lemma stronglyMeasurable_ℱ_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) - (a : α) (s : Set β) : - StronglyMeasurable[ℱ n] (fun t ↦ densityProcess κ ν n a t s) := - (measurable_ℱ_densityProcess κ ν n a s).stronglyMeasurable - -lemma adapted_densityProcess (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (s : Set β) : - Adapted ℱ (fun n t ↦ densityProcess κ ν n a t s) := - fun n ↦ stronglyMeasurable_ℱ_densityProcess κ ν n a s - -lemma densityProcess_nonneg (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) - (a : α) (t : ℝ) (s : Set β) : +lemma measurable_ℱ_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) + (a : α) {s : Set β} (hs : MeasurableSet s) : + Measurable[ℱ γ n] (fun t ↦ densityProcess κ ν n a t s) := by + refine @Measurable.ennreal_toReal _ (ℱ γ n) _ ?_ + exact (measurable_densityProcess_ℱ_aux κ ν n hs).comp measurable_prod_mk_left + +lemma stronglyMeasurable_ℱ_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) + (a : α) {s : Set β} (hs : MeasurableSet s) : + StronglyMeasurable[ℱ γ n] (fun t ↦ densityProcess κ ν n a t s) := + (measurable_ℱ_densityProcess κ ν n a hs).stronglyMeasurable + +lemma adapted_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) + {s : Set β} (hs : MeasurableSet s) : + Adapted (ℱ γ) (fun n t ↦ densityProcess κ ν n a t s) := + fun n ↦ stronglyMeasurable_ℱ_densityProcess κ ν n a hs + +lemma densityProcess_nonneg (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) + (a : α) (t : γ) (s : Set β) : 0 ≤ densityProcess κ ν n a t s := ENNReal.toReal_nonneg -lemma apply_IcoPowTwo_le_of_fst_le (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : ℝ) (s : Set β) : - κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) ≤ ν a (IcoPowTwo n (indexIcoPowTwo n t)) := by - calc κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) - ≤ kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) := by - rw [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] +lemma apply_partitionSet_le_of_fst_le (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : + κ a (partitionSet n t ×ˢ s) ≤ ν a (partitionSet n t) := by + calc κ a (partitionSet n t ×ˢ s) + ≤ kernel.fst κ a (partitionSet n t) := by + rw [kernel.fst_apply' _ _ (measurableSet_partitionSet _ _)] refine measure_mono (fun x ↦ ?_) simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h - _ ≤ ν a (IcoPowTwo n (indexIcoPowTwo n t)) := hκν a _ (measurableSet_IcoPowTwo _ _) + _ ≤ ν a (partitionSet n t) := hκν a _ (measurableSet_partitionSet _ _) -lemma densityProcess_le_one (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : ℝ) (s : Set β) : +lemma densityProcess_le_one (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : densityProcess κ ν n a t s ≤ 1 := by refine ENNReal.toReal_le_of_le_ofReal zero_le_one (ENNReal.div_le_of_le_mul ?_) rw [ENNReal.ofReal_one, one_mul] - exact apply_IcoPowTwo_le_of_fst_le hκν n a t s + exact apply_partitionSet_le_of_fst_le hκν n a t s lemma snorm_densityProcess_le (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (s : Set β) : snorm (fun t ↦ densityProcess κ ν n a t s) 1 (ν a) ≤ ν a univ := by @@ -381,41 +297,41 @@ lemma integrable_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel · exact (snorm_densityProcess_le hκν n a s).trans_lt (measure_lt_top _ _) lemma set_integral_densityProcess_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) (k : ℤ) : - ∫ t in IcoPowTwo n k, densityProcess κ ν n a t s ∂(ν a) - = (κ a (IcoPowTwo n k ×ˢ s)).toReal := by + (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {u : Set γ} + (hu : u ∈ MeasurableSpace.countablePartition γ n) : + ∫ t in u, densityProcess κ ν n a t s ∂(ν a) = (κ a (u ×ˢ s)).toReal := by + have hu_meas : MeasurableSet u := MeasurableSpace.measurableSet_countablePartition n hu simp_rw [densityProcess] rw [integral_toReal] rotate_left · refine Measurable.aemeasurable ?_ have h := measurable_densityProcess_aux κ ν n hs - change Measurable ((fun (p : α × ℝ) ↦ κ p.1 (IcoPowTwo n (indexIcoPowTwo n p.2) ×ˢ s) - / ν p.1 (IcoPowTwo n (indexIcoPowTwo n p.2))) ∘ (fun t ↦ (a, t))) + change Measurable ((fun (p : α × _) ↦ κ p.1 (partitionSet n p.2 ×ˢ s) + / ν p.1 (partitionSet n p.2)) ∘ (fun t ↦ (a, t))) exact h.comp measurable_prod_mk_left · refine ae_of_all _ (fun t ↦ ?_) - by_cases h0 : ν a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 - · suffices κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) = 0 by simp [h0, this] - have h0' : kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 := - le_antisymm ((hκν a _ (measurableSet_IcoPowTwo _ _)).trans h0.le) zero_le' - rw [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] at h0' + by_cases h0 : ν a (partitionSet n t) = 0 + · suffices κ a (partitionSet n t ×ˢ s) = 0 by simp [h0, this] + have h0' : kernel.fst κ a (partitionSet n t) = 0 := + le_antisymm ((hκν a _ (measurableSet_partitionSet _ _)).trans h0.le) zero_le' + rw [kernel.fst_apply' _ _ (measurableSet_partitionSet _ _)] at h0' refine measure_mono_null (fun x ↦ ?_) h0' simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h · refine ENNReal.div_lt_top ?_ h0 exact measure_ne_top _ _ congr - have : ∫⁻ t in IcoPowTwo n k, - κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) / ν a (IcoPowTwo n (indexIcoPowTwo n t)) ∂(ν a) - = ∫⁻ _ in IcoPowTwo n k, κ a (IcoPowTwo n k ×ˢ s) / ν a (IcoPowTwo n k) ∂(ν a) := by - refine set_lintegral_congr_fun (measurableSet_IcoPowTwo _ _) (ae_of_all _ (fun t ht ↦ ?_)) - rw [indexIcoPowTwo_of_mem _ _ _ ht] + have : ∫⁻ t in u, κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t) ∂(ν a) + = ∫⁻ _ in u, κ a (u ×ˢ s) / ν a u ∂(ν a) := by + refine set_lintegral_congr_fun hu_meas (ae_of_all _ (fun t ht ↦ ?_)) + rw [partitionSet_of_mem hu ht] rw [this] simp only [lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter] - by_cases h0 : ν a (IcoPowTwo n k) = 0 + by_cases h0 : ν a u = 0 · simp only [h0, mul_zero] - have h0' : kernel.fst κ a (IcoPowTwo n k) = 0 := - le_antisymm ((hκν a _ (measurableSet_IcoPowTwo _ _)).trans h0.le) zero_le' - rw [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] at h0' + have h0' : kernel.fst κ a u = 0 := + le_antisymm ((hκν a _ hu_meas).trans h0.le) zero_le' + rw [kernel.fst_apply' _ _ hu_meas] at h0' refine (measure_mono_null ?_ h0').symm intro p simp only [mem_prod, mem_setOf_eq, and_imp] @@ -426,37 +342,44 @@ lemma set_integral_densityProcess_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKern lemma integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : ∫ t, densityProcess κ ν n a t s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by - rw [← integral_univ, ← iUnion_IcoPowTwo n, iUnion_prod_const, measure_iUnion] + rw [← integral_univ, ← MeasurableSpace.sUnion_countablePartition γ n, sUnion_eq_iUnion, + iUnion_prod_const, measure_iUnion] rotate_left · intro i j hij simp only [Set.disjoint_prod, disjoint_self, bot_eq_empty] - exact Or.inl (pairwise_disjoint_IcoPowTwo n hij) - · exact fun k ↦ (measurableSet_IcoPowTwo n k).prod hs - rw [integral_iUnion (measurableSet_IcoPowTwo n) (pairwise_disjoint_IcoPowTwo n) - (integrable_densityProcess hκν n a hs).integrableOn] + refine Or.inl (MeasurableSpace.disjoint_countablePartition n i.prop j.prop ?_) + rw [ne_eq, Subtype.coe_inj] + exact hij + · exact fun k ↦ (MeasurableSpace.measurableSet_countablePartition n k.prop).prod hs + rw [integral_iUnion] + rotate_left + · exact fun k ↦ MeasurableSpace.measurableSet_countablePartition n k.prop + · intro i j hij + refine MeasurableSpace.disjoint_countablePartition n i.prop j.prop ?_ + rw [ne_eq, Subtype.coe_inj] + exact hij + · exact (integrable_densityProcess hκν n a hs).integrableOn rw [ENNReal.tsum_toReal_eq (fun _ ↦ measure_ne_top _ _)] congr with k - rw [set_integral_densityProcess_I hκν _ _ hs] + rw [set_integral_densityProcess_I hκν _ _ hs k.prop] lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : + (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet[ℱ γ n] A) : ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - refine MeasurableSpace.induction_on_inter (m := ℱ n) (s := {s | ∃ k, s = IcoPowTwo n k}) + refine MeasurableSpace.induction_on_inter (m := ℱ γ n) + (s := MeasurableSpace.countablePartition γ n) (C := fun A ↦ ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) rfl ?_ ?_ ?_ ?_ ?_ hA - · rintro s ⟨i, rfl⟩ t ⟨j, rfl⟩ hst - refine ⟨i, ?_⟩ - suffices i = j by rw [this, inter_self] + · rintro s hs t ht hst + suffices s = t by rwa [this, inter_self] by_contra h_ne - have h_disj := pairwise_disjoint_IcoPowTwo n h_ne - rw [nonempty_iff_ne_empty] at hst - refine hst ?_ - rwa [Function.onFun, disjoint_iff_inter_eq_empty] at h_disj + rw [← not_disjoint_iff_nonempty_inter] at hst + exact hst <| MeasurableSpace.disjoint_countablePartition n hs ht h_ne · simp - · rintro _ ⟨k, rfl⟩ - rw [set_integral_densityProcess_I hκν _ _ hs] + · rintro u hu + rw [set_integral_densityProcess_I hκν _ _ hs hu] · intro A hA hA_eq - have hA' : MeasurableSet A := ℱ.le _ _ hA + have hA' : MeasurableSet A := (ℱ γ).le _ _ hA have h := integral_add_compl hA' (integrable_densityProcess hκν n a hs) rw [hA_eq, integral_densityProcess hκν n a hs] at h have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by @@ -469,10 +392,10 @@ lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] exact measure_mono (by intro x; simp) · intro f hf_disj hf h_eq - rw [integral_iUnion (fun i ↦ ℱ.le n _ (hf i)) hf_disj + rw [integral_iUnion (fun i ↦ (ℱ γ).le n _ (hf i)) hf_disj (integrable_densityProcess hκν _ _ hs).integrableOn] simp_rw [h_eq] - rw [iUnion_prod_const, measure_iUnion _ (fun i ↦ (ℱ.le n _ (hf i)).prod hs)] + rw [iUnion_prod_const, measure_iUnion _ (fun i ↦ ((ℱ γ).le n _ (hf i)).prod hs)] · rw [ENNReal.tsum_toReal_eq] exact fun _ ↦ measure_ne_top _ _ · intro i j hij @@ -481,13 +404,13 @@ lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel lemma set_integral_densityProcess_of_le (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] {n m : ℕ} (hnm : n ≤ m) (a : α) {s : Set β} (hs : MeasurableSet s) - {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : + {A : Set γ} (hA : MeasurableSet[ℱ γ n] A) : ∫ t in A, densityProcess κ ν m a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := - set_integral_densityProcess hκν m a hs (ℱ.mono hnm A hA) + set_integral_densityProcess hκν m a hs ((ℱ γ).mono hnm A hA) lemma condexp_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] {i j : ℕ} (hij : i ≤ j) (a : α) {s : Set β} (hs : MeasurableSet s) : - (ν a)[fun t ↦ densityProcess κ ν j a t s | ℱ i] + (ν a)[fun t ↦ densityProcess κ ν j a t s | ℱ γ i] =ᵐ[ν a] fun t ↦ densityProcess κ ν i a t s := by symm refine ae_eq_condexp_of_forall_set_integral_eq ?_ ?_ ?_ ?_ ?_ @@ -497,27 +420,26 @@ lemma condexp_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] · intro t ht _ rw [set_integral_densityProcess hκν i a hs ht, set_integral_densityProcess_of_le hκν hij a hs ht] - · exact StronglyMeasurable.aeStronglyMeasurable' (stronglyMeasurable_ℱ_densityProcess κ ν i a s) + · exact StronglyMeasurable.aeStronglyMeasurable' (stronglyMeasurable_ℱ_densityProcess κ ν i a hs) lemma martingale_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Martingale (fun n t ↦ densityProcess κ ν n a t s) ℱ (ν a) := - ⟨adapted_densityProcess κ ν a s, fun _ _ h ↦ condexp_densityProcess hκν h a hs⟩ + Martingale (fun n t ↦ densityProcess κ ν n a t s) (ℱ γ) (ν a) := + ⟨adapted_densityProcess κ ν a hs, fun _ _ h ↦ condexp_densityProcess hκν h a hs⟩ -lemma densityProcess_mono_set (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : ℝ) +lemma densityProcess_mono_set (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) {s s' : Set β} (h : s ⊆ s') : densityProcess κ ν n a t s ≤ densityProcess κ ν n a t s' := by unfold densityProcess - by_cases h0 : ν a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 + by_cases h0 : ν a (partitionSet n t) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] simp - have h_ne_top : ∀ s, κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) - / ν a (IcoPowTwo n (indexIcoPowTwo n t)) ≠ ⊤ := by + have h_ne_top : ∀ s, κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t) ≠ ⊤ := by intro s rw [ne_eq, ENNReal.div_eq_top] simp only [ne_eq, h0, and_false, false_or, not_and, not_not] refine fun h_top ↦ eq_top_mono ?_ h_top - exact apply_IcoPowTwo_le_of_fst_le hκν n a t s + exact apply_partitionSet_le_of_fst_le hκν n a t s rw [ENNReal.toReal_le_toReal] · gcongr rw [prod_subset_prod_iff] @@ -525,59 +447,59 @@ lemma densityProcess_mono_set (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) · exact h_ne_top s · exact h_ne_top s' -lemma densityProcess_mono_kernel_left {κ' : kernel α (ℝ × β)} (hκκ' : κ ≤ κ') - (hκ'ν : kernel.fst κ' ≤ ν) (n : ℕ) (a : α) (t : ℝ) {s : Set β} (hs : MeasurableSet s) : +lemma densityProcess_mono_kernel_left {κ' : kernel α (γ × β)} (hκκ' : κ ≤ κ') + (hκ'ν : kernel.fst κ' ≤ ν) (n : ℕ) (a : α) (t : γ) {s : Set β} (hs : MeasurableSet s) : densityProcess κ ν n a t s ≤ densityProcess κ' ν n a t s := by unfold densityProcess - by_cases h0 : ν a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 + by_cases h0 : ν a (partitionSet n t) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] simp - have h_le : κ' a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) - ≤ ν a (IcoPowTwo n (indexIcoPowTwo n t)) := apply_IcoPowTwo_le_of_fst_le hκ'ν n a t s + have h_le : κ' a (partitionSet n t ×ˢ s) + ≤ ν a (partitionSet n t) := apply_partitionSet_le_of_fst_le hκ'ν n a t s rw [ENNReal.toReal_le_toReal] · gcongr - exact hκκ' _ _ ((measurableSet_IcoPowTwo _ _).prod hs) + exact hκκ' _ _ ((measurableSet_partitionSet _ _).prod hs) · rw [ne_eq, ENNReal.div_eq_top] simp only [ne_eq, h0, and_false, false_or, not_and, not_not] refine fun h_top ↦ eq_top_mono ?_ h_top - exact (hκκ' _ _ ((measurableSet_IcoPowTwo _ _).prod hs)).trans h_le + exact (hκκ' _ _ ((measurableSet_partitionSet _ _).prod hs)).trans h_le · rw [ne_eq, ENNReal.div_eq_top] simp only [ne_eq, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top -lemma densityProcess_antitone_kernel_right {ν' : kernel α ℝ} - (hνν' : ν ≤ ν') (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : ℝ) (s : Set β) : +lemma densityProcess_antitone_kernel_right {ν' : kernel α γ} + (hνν' : ν ≤ ν') (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : densityProcess κ ν' n a t s ≤ densityProcess κ ν n a t s := by unfold densityProcess - have h_le : κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) - ≤ ν a (IcoPowTwo n (indexIcoPowTwo n t)) := apply_IcoPowTwo_le_of_fst_le hκν n a t s - by_cases h0 : ν a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 - · suffices κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s) = 0 by + have h_le : κ a (partitionSet n t ×ˢ s) + ≤ ν a (partitionSet n t) := apply_partitionSet_le_of_fst_le hκν n a t s + by_cases h0 : ν a (partitionSet n t) = 0 + · suffices κ a (partitionSet n t ×ˢ s) = 0 by simp only [this, ENNReal.zero_div, ENNReal.zero_toReal, h0, le_refl] exact le_antisymm (h_le.trans h0.le) zero_le' - have h0' : ν' a (IcoPowTwo n (indexIcoPowTwo n t)) ≠ 0 := by + have h0' : ν' a (partitionSet n t) ≠ 0 := by refine fun h ↦ h0 (le_antisymm (le_trans ?_ h.le) zero_le') - exact hνν' _ _ (measurableSet_IcoPowTwo _ _) + exact hνν' _ _ (measurableSet_partitionSet _ _) rw [ENNReal.toReal_le_toReal] · gcongr - exact hνν' _ _ (measurableSet_IcoPowTwo _ _) + exact hνν' _ _ (measurableSet_partitionSet _ _) · simp only [ne_eq, ENNReal.div_eq_top, h0', and_false, false_or, not_and, not_not] refine fun h_top ↦ eq_top_mono ?_ h_top - exact h_le.trans (hνν' _ _ (measurableSet_IcoPowTwo _ _)) + exact h_le.trans (hνν' _ _ (measurableSet_partitionSet _ _)) · simp only [ne_eq, ENNReal.div_eq_top, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top -lemma densityProcess_empty (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (n : ℕ) (a : α) (t : ℝ) : +lemma densityProcess_empty (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (t : γ) : densityProcess κ ν n a t ∅ = 0 := by simp [densityProcess] -lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) - [IsFiniteKernel κ] (n : ℕ) (a : α) (t : ℝ) +lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (γ × β)) (ν : kernel α γ) + [IsFiniteKernel κ] (n : ℕ) (a : α) (t : γ) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : Tendsto (fun m ↦ densityProcess κ ν n a t (s m)) atTop (𝓝 (densityProcess κ ν n a t ∅)) := by simp_rw [densityProcess] - by_cases h0 : ν a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 + by_cases h0 : ν a (partitionSet n t) = 0 · simp_rw [h0, ENNReal.toReal_div] simp refine (ENNReal.tendsto_toReal ?_).comp ?_ @@ -586,10 +508,10 @@ lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (ℝ × β) simp refine ENNReal.Tendsto.div_const ?_ ?_ · have h := tendsto_measure_iInter (μ := κ a) - (s := fun m ↦ IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s m) ?_ ?_ ?_ + (s := fun m ↦ partitionSet n t ×ˢ s m) ?_ ?_ ?_ · convert h rw [← prod_iInter, hs_iInter] - · exact fun n ↦ MeasurableSet.prod (measurableSet_IcoPowTwo _ _) (hs_meas n) + · exact fun n ↦ MeasurableSet.prod (measurableSet_partitionSet _ _) (hs_meas n) · intro m m' hmm' simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] exact Or.inl <| hs hmm' @@ -597,8 +519,8 @@ lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (ℝ × β) · simp only [prod_empty, OuterMeasure.empty', ne_eq, not_true_eq_false, false_or, h0, not_false_iff] -lemma tendsto_densityProcess_atTop_of_antitone (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) - [IsFiniteKernel κ] (n : ℕ) (a : α) (t : ℝ) +lemma tendsto_densityProcess_atTop_of_antitone (κ : kernel α (γ × β)) (ν : kernel α γ) + [IsFiniteKernel κ] (n : ℕ) (a : α) (t : γ) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : Tendsto (fun m ↦ densityProcess κ ν n a t (s m)) atTop (𝓝 0) := by @@ -608,7 +530,7 @@ lemma tendsto_densityProcess_atTop_of_antitone (κ : kernel α (ℝ × β)) (ν lemma tendsto_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : ∀ᵐ t ∂(ν a), Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop - (𝓝 (ℱ.limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a) t)) := by + (𝓝 ((ℱ γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a) t)) := by refine Submartingale.ae_tendsto_limitProcess (martingale_densityProcess hκν a hs).submartingale (R := (ν a univ).toNNReal) (fun n ↦ ?_) refine (snorm_densityProcess_le hκν n a s).trans_eq ?_ @@ -617,7 +539,7 @@ lemma tendsto_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFini lemma limitProcess_mem_L1 (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Memℒp (ℱ.limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 (ν a) := by + Memℒp ((ℱ γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 (ν a) := by refine Submartingale.memℒp_limitProcess (martingale_densityProcess hκν a hs).submartingale (R := (ν a univ).toNNReal) (fun n ↦ ?_) refine (snorm_densityProcess_le hκν n a s).trans_eq ?_ @@ -627,7 +549,7 @@ lemma limitProcess_mem_L1 (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [Is lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : Tendsto (fun n ↦ snorm ((fun t ↦ densityProcess κ ν n a t s) - - ℱ.limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 (ν a)) + - (ℱ γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 (ν a)) atTop (𝓝 0) := by refine Submartingale.tendsto_snorm_one_limitProcess ?_ ?_ · exact (martingale_densityProcess hκν a hs).submartingale @@ -647,20 +569,20 @@ lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : kernel.fst κ ≤ · simp lemma tendsto_snorm_one_restrict_densityProcess_limitProcess [IsFiniteKernel κ] [IsFiniteKernel ν] - (hκν : kernel.fst κ ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : + (hκν : kernel.fst κ ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : Tendsto (fun n ↦ snorm ((fun t ↦ densityProcess κ ν n a t s) - - ℱ.limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 ((ν a).restrict A)) + - (ℱ γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 ((ν a).restrict A)) atTop (𝓝 0) := tendsto_snorm_restrict_zero (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) A noncomputable -def MLimsup (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (a : α) (t : ℝ) (s : Set β) : ℝ := +def MLimsup (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) (t : γ) (s : Set β) : ℝ := limsup (fun n ↦ densityProcess κ ν n a t s) atTop lemma mLimsup_ae_eq_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : (fun t ↦ MLimsup κ ν a t s) - =ᵐ[ν a] ℱ.limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a) := by + =ᵐ[ν a] (ℱ γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a) := by filter_upwards [tendsto_densityProcess_limitProcess hκν a hs] with t ht using ht.limsup_eq lemma tendsto_m_mLimsup (hκν : kernel.fst κ ≤ ν) (a : α) [IsFiniteKernel κ] [IsFiniteKernel ν] @@ -670,37 +592,37 @@ lemma tendsto_m_mLimsup (hκν : kernel.fst κ ≤ ν) (a : α) [IsFiniteKernel filter_upwards [tendsto_densityProcess_limitProcess hκν a hs, mLimsup_ae_eq_limitProcess hκν a hs] with t h1 h2 using h2 ▸ h1 -lemma measurable_mLimsup (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) +lemma measurable_mLimsup (κ : kernel α (γ × β)) (ν : kernel α γ) {s : Set β} (hs : MeasurableSet s) : - Measurable (fun (p : α × ℝ) ↦ MLimsup κ ν p.1 p.2 s) := + Measurable (fun (p : α × γ) ↦ MLimsup κ ν p.1 p.2 s) := measurable_limsup (fun n ↦ measurable_densityProcess κ ν n hs) -lemma measurable_mLimsup_left (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) (t : ℝ) +lemma measurable_mLimsup_left (κ : kernel α (γ × β)) (ν : kernel α γ) (t : γ) {s : Set β} (hs : MeasurableSet s) : Measurable (fun a ↦ MLimsup κ ν a t s) := by - change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ ν p.1 p.2 s) ∘ (fun a ↦ (a, t))) + change Measurable ((fun (p : α × γ) ↦ MLimsup κ ν p.1 p.2 s) ∘ (fun a ↦ (a, t))) exact (measurable_mLimsup κ ν hs).comp measurable_prod_mk_right -lemma measurable_mLimsup_right (κ : kernel α (ℝ × β)) (ν : kernel α ℝ) +lemma measurable_mLimsup_right (κ : kernel α (γ × β)) (ν : kernel α γ) {s : Set β} (hs : MeasurableSet s) (a : α) : Measurable (fun t ↦ MLimsup κ ν a t s) := by - change Measurable ((fun (p : α × ℝ) ↦ MLimsup κ ν p.1 p.2 s) ∘ (fun t ↦ (a, t))) + change Measurable ((fun (p : α × γ) ↦ MLimsup κ ν p.1 p.2 s) ∘ (fun t ↦ (a, t))) exact (measurable_mLimsup κ ν hs).comp measurable_prod_mk_left -lemma mLimsup_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) {s s' : Set β} (h : s ⊆ s') : +lemma mLimsup_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) {s s' : Set β} (h : s ⊆ s') : MLimsup κ ν a t s ≤ MLimsup κ ν a t s' := by refine limsup_le_limsup ?_ ?_ ?_ · exact eventually_of_forall (fun n ↦ densityProcess_mono_set hκν n a t h) · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ -lemma mLimsup_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (s : Set β) : +lemma mLimsup_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : 0 ≤ MLimsup κ ν a t s := by refine le_limsup_of_frequently_le ?_ ?_ · exact frequently_of_forall (fun n ↦ densityProcess_nonneg _ _ _ _ _ _) · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ -lemma mLimsup_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (s : Set β) : +lemma mLimsup_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : MLimsup κ ν a t s ≤ 1 := by refine limsup_le_of_le ?_ ?_ · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) @@ -722,7 +644,7 @@ lemma integrable_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] · exact (snorm_mLimsup_le hκν a s).trans_lt (measure_lt_top _ _) lemma tendsto_set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set ℝ) : + (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : Tendsto (fun i ↦ ∫ x in A, densityProcess κ ν i a x s ∂(ν a)) atTop (𝓝 (∫ x in A, MLimsup κ ν a x s ∂(ν a))) := by refine tendsto_set_integral_of_L1' (μ := ν a) (fun t ↦ MLimsup κ ν a t s) @@ -735,7 +657,7 @@ lemma tendsto_set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] /-- Auxiliary lemma for `set_integral_mLimsup`. -/ lemma set_integral_mLimsup_of_measurableSet (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set ℝ} (hA : MeasurableSet[ℱ n] A) : + (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet[ℱ γ n] A) : ∫ t in A, MLimsup κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by suffices ∫ t in A, MLimsup κ ν a t s ∂(ν a) = ∫ t in A, densityProcess κ ν n a t s ∂(ν a) by rw [this] @@ -758,14 +680,14 @@ lemma integral_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFin rw [← integral_univ, set_integral_mLimsup_of_measurableSet hκν 0 a hs MeasurableSet.univ] lemma set_integral_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set ℝ} (hA : MeasurableSet A) : + (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet A) : ∫ t in A, MLimsup κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - have hA' : MeasurableSet[⨆ n, ℱ n] A := by rwa [iSup_ℱ] - refine MeasurableSpace.induction_on_inter (m := ⨆ n, ℱ n) + have hA' : MeasurableSet[⨆ n, ℱ γ n] A := by rwa [iSup_ℱ] + refine MeasurableSpace.induction_on_inter (m := ⨆ n, ℱ γ n) (C := fun A ↦ ∫ t in A, MLimsup κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) - (MeasurableSpace.measurableSpace_iSup_eq ℱ) ?_ ?_ ?_ ?_ ?_ hA' + (MeasurableSpace.measurableSpace_iSup_eq (ℱ γ)) ?_ ?_ ?_ ?_ ?_ hA' · rintro s ⟨n, hs⟩ t ⟨m, ht⟩ _ - exact ⟨max n m, (ℱ.mono (le_max_left n m) _ hs).inter (ℱ.mono (le_max_right n m) _ ht)⟩ + exact ⟨max n m, ((ℱ γ).mono (le_max_left n m) _ hs).inter ((ℱ γ).mono (le_max_right n m) _ ht)⟩ · simp · intro A ⟨n, hA⟩ exact set_integral_mLimsup_of_measurableSet hκν n a hs hA @@ -858,7 +780,7 @@ lemma tendsto_mLimsup_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin linarith · exact h_tendsto -- let `F` be the pointwise limit of `fun m ↦ MLimsup κ a (s m) t` for all `t` - let F : ℝ → ℝ := fun t ↦ (h_exists t).choose + let F : γ → ℝ := fun t ↦ (h_exists t).choose have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 (F t)) := fun t ↦ (h_exists t).choose_spec have hF_nonneg : ∀ t, 0 ≤ F t := @@ -885,7 +807,7 @@ lemma tendsto_mLimsup_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin exact hF_tendsto t -- since `F` is nonnegative, proving that its integral is 0 is sufficient to get that -- `F` is 0 a.e. - suffices ∀ᵐ (t : ℝ) ∂(ν a), 0 = F t by filter_upwards [this] with a ha; simp [ha] + suffices ∀ᵐ (t : γ) ∂(ν a), 0 = F t by filter_upwards [this] with a ha; simp [ha] refine ae_eq_of_integral_eq_of_ae_le (integrable_const _) hF_int (ae_of_all _ hF_nonneg) ?_ have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ t, F t ∂(ν a))) := by @@ -901,55 +823,62 @@ lemma tendsto_mLimsup_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin section UnivFst -lemma densityProcess_univ [IsFiniteKernel κ] (n : ℕ) (a : α) (t : ℝ) : +lemma densityProcess_univ [IsFiniteKernel κ] (n : ℕ) (a : α) (t : γ) : densityProcess κ (kernel.fst κ) n a t univ - = if kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 then 0 else 1 := by + = if kernel.fst κ a (partitionSet n t) = 0 then 0 else 1 := by rw [densityProcess] - by_cases h : kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 + by_cases h : kernel.fst κ a (partitionSet n t) = 0 · simp [h] - by_cases h' : κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ univ) = 0 + by_cases h' : κ a (partitionSet n t ×ˢ univ) = 0 · simp [h'] · rw [ENNReal.div_zero h'] simp · simp only [h, ite_false] - rw [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] - have : IcoPowTwo n (indexIcoPowTwo n t) ×ˢ univ - = {p : ℝ × β | p.1 ∈ IcoPowTwo n (indexIcoPowTwo n t)} := by + rw [kernel.fst_apply' _ _ (measurableSet_partitionSet _ _)] + have : partitionSet n t ×ˢ univ = {p : γ × β | p.1 ∈ partitionSet n t} := by ext x simp rw [this, ENNReal.div_self] · simp - · rwa [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] at h + · rwa [kernel.fst_apply' _ _ (measurableSet_partitionSet _ _)] at h · exact measure_ne_top _ _ -lemma densityProcess_univ_ae (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (n : ℕ) (a : α) : +lemma densityProcess_univ_ae (κ : kernel α (γ × β)) [IsFiniteKernel κ] (n : ℕ) (a : α) : ∀ᵐ t ∂(kernel.fst κ a), densityProcess κ (kernel.fst κ) n a t univ = 1 := by rw [ae_iff] have : {t | ¬ densityProcess κ (kernel.fst κ) n a t univ = 1} - ⊆ {t | kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) = 0} := by + ⊆ {t | kernel.fst κ a (partitionSet n t) = 0} := by intro t ht simp only [mem_setOf_eq] at ht ⊢ rw [densityProcess_univ] at ht simpa using ht refine measure_mono_null this ?_ - have : {t | kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) = 0} - ⊆ ⋃ (k) (_ : kernel.fst κ a (IcoPowTwo n k) = 0), IcoPowTwo n k := by + have : {t | kernel.fst κ a (partitionSet n t) = 0} + ⊆ ⋃ (u) (_ : u ∈ MeasurableSpace.countablePartition γ n) (_ : kernel.fst κ a u = 0), u := by intro t ht simp only [mem_setOf_eq, mem_iUnion, exists_prop] at ht ⊢ - exact ⟨indexIcoPowTwo n t, ht, mem_IcoPowTwo_indexIcoPowTwo _ _⟩ + exact ⟨partitionSet n t, partitionSet_mem _ _, ht, mem_partitionSet _ _⟩ refine measure_mono_null this ?_ - rw [measure_iUnion_null] - simp - -lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) - (n : ℕ) (a : α) (t : ℝ) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) : + rw [measure_biUnion] + · simp + · exact (MeasurableSpace.finite_countablePartition _ _).countable + · intro s hs t ht hst + simp only [disjoint_iUnion_right, disjoint_iUnion_left] + exact fun _ _ ↦ MeasurableSpace.disjoint_countablePartition n hs ht hst + · intro s hs + by_cases h : kernel.fst κ a s = 0 + · simp [h, MeasurableSpace.measurableSet_countablePartition n hs] + · simp [h] + +lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (γ × β)) + (n : ℕ) (a : α) (t : γ) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) : Tendsto (fun m ↦ densityProcess κ (kernel.fst κ) n a t (s m)) atTop (𝓝 (densityProcess κ (kernel.fst κ) n a t univ)) := by simp_rw [densityProcess] refine (ENNReal.tendsto_toReal ?_).comp ?_ · rw [ne_eq, ENNReal.div_eq_top] push_neg - simp_rw [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] + simp_rw [kernel.fst_apply' _ _ (measurableSet_partitionSet _ _)] constructor · refine fun h h0 ↦ h (measure_mono_null (fun x ↦ ?_) h0) simp only [mem_prod, mem_setOf_eq, and_imp] @@ -957,11 +886,11 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) · refine fun h_top ↦ eq_top_mono (measure_mono (fun x ↦ ?_)) h_top simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h - by_cases h0 : kernel.fst κ a (IcoPowTwo n (indexIcoPowTwo n t)) = 0 - · rw [kernel.fst_apply' _ _ (measurableSet_IcoPowTwo _ _)] at h0 ⊢ - suffices ∀ m, κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s m) = 0 by + by_cases h0 : kernel.fst κ a (partitionSet n t) = 0 + · rw [kernel.fst_apply' _ _ (measurableSet_partitionSet _ _)] at h0 ⊢ + suffices ∀ m, κ a (partitionSet n t ×ˢ s m) = 0 by simp only [this, h0, ENNReal.zero_div, tendsto_const_nhds_iff] - suffices κ a (IcoPowTwo n (indexIcoPowTwo n t) ×ˢ univ) = 0 by + suffices κ a (partitionSet n t ×ˢ univ) = 0 by simp only [this, ENNReal.zero_div] convert h0 ext x @@ -971,7 +900,7 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) exact fun h _ ↦ h refine ENNReal.Tendsto.div_const ?_ ?_ · have h := tendsto_measure_iUnion (μ := κ a) - (s := fun m ↦ IcoPowTwo n (indexIcoPowTwo n t) ×ˢ s m) ?_ + (s := fun m ↦ partitionSet n t ×ˢ s m) ?_ swap · intro m m' hmm' simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] @@ -980,7 +909,7 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (ℝ × β)) rw [← prod_iUnion, hs_iUnion] · exact Or.inr h0 -lemma tendsto_densityProcess_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] +lemma tendsto_densityProcess_atTop_ae_of_monotone (κ : kernel α (γ × β)) [IsFiniteKernel κ] (n : ℕ) (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ densityProcess κ (kernel.fst κ) n a t (s m)) atTop (𝓝 1) := by @@ -988,7 +917,7 @@ lemma tendsto_densityProcess_atTop_ae_of_monotone (κ : kernel α (ℝ × β)) [ rw [← ht] exact tendsto_densityProcess_atTop_univ_of_monotone κ n a t s hs hs_iUnion -lemma mLimsup_univ (κ : kernel α (ℝ × β)) [IsFiniteKernel κ] (a : α) : +lemma mLimsup_univ (κ : kernel α (γ × β)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), MLimsup κ (kernel.fst κ) a t univ = 1 := by have h := fun n ↦ densityProcess_univ_ae κ n a rw [← ae_all_iff] at h @@ -1016,7 +945,7 @@ lemma tendsto_mLimsup_atTop_ae_of_monotone [IsFiniteKernel κ] exact absurd (hr.trans (h_le_one r t)) one_lt_two.not_le · exact h_tendsto -- let `F` be the pointwise limit of `fun m ↦ MLimsup κ a (s m) t` for all `t` - let F : ℝ → ℝ := fun t ↦ (h_exists t).choose + let F : γ → ℝ := fun t ↦ (h_exists t).choose have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 (F t)) := fun t ↦ (h_exists t).choose_spec have hF_nonneg : ∀ t, 0 ≤ F t := diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index 432948403de51..95ce159581fe3 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -34,7 +34,7 @@ variable {α β Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β section Lintegral -variable [CountableOrStandardBorel α β] {κ : kernel α (β × Ω)} [IsFiniteKernel κ] +variable [CountableOrCountablyGenerated α β] {κ : kernel α (β × Ω)} [IsFiniteKernel κ] {f : β × Ω → ℝ≥0∞} lemma lintegral_condKernel_mem (a : α) {s : Set (β × Ω)} (hs : MeasurableSet s) : @@ -84,7 +84,7 @@ end Lintegral section Integral -variable [CountableOrStandardBorel α β] {κ : kernel α (β × Ω)} [IsFiniteKernel κ] +variable [CountableOrCountablyGenerated α β] {κ : kernel α (β × Ω)} [IsFiniteKernel κ] {E : Type*} {f : β × Ω → E} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] lemma _root_.MeasureTheory.AEStronglyMeasurable.integral_kernel_condKernel (a : α) @@ -131,7 +131,7 @@ variable {α β Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β section Lintegral -variable [CountableOrStandardBorel α β] {ρ : Measure (β × Ω)} [IsFiniteMeasure ρ] +variable [CountableOrCountablyGenerated α β] {ρ : Measure (β × Ω)} [IsFiniteMeasure ρ] {f : β × Ω → ℝ≥0∞} lemma lintegral_condKernel_mem {s : Set (β × Ω)} (hs : MeasurableSet s) : diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean index 9613594348142..924ce6b6fd58a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean @@ -31,22 +31,23 @@ open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory -variable {α : Type*} {mα : MeasurableSpace α} {κ : kernel α (ℝ × ℝ)} {ν : kernel α ℝ} +variable {α γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} + [MeasurableSpace.CountablyGenerated γ] {κ : kernel α (γ × ℝ)} {ν : kernel α γ} noncomputable -def mLimsupIic (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) (a : α) (t : ℝ) (q : ℚ) : ℝ := +def mLimsupIic (κ : kernel α (γ × ℝ)) (ν : kernel α γ) (a : α) (t : γ) (q : ℚ) : ℝ := MLimsup κ ν a t (Set.Iic q) -lemma measurable_mLimsupIic (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) : - Measurable (fun p : α × ℝ ↦ mLimsupIic κ ν p.1 p.2) := by +lemma measurable_mLimsupIic (κ : kernel α (γ × ℝ)) (ν : kernel α γ) : + Measurable (fun p : α × γ ↦ mLimsupIic κ ν p.1 p.2) := by rw [measurable_pi_iff] exact fun _ ↦ measurable_mLimsup κ ν measurableSet_Iic -lemma measurable_mLimsupIic_right (κ : kernel α (ℝ × ℝ)) (ν : kernel α ℝ) (a : α) (q : ℚ) : +lemma measurable_mLimsupIic_right (κ : kernel α (γ × ℝ)) (ν : kernel α γ) (a : α) (q : ℚ) : Measurable (fun t ↦ mLimsupIic κ ν a t q) := measurable_mLimsup_right _ _ measurableSet_Iic _ -lemma monotone_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) : +lemma monotone_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) : Monotone (mLimsupIic κ ν a t) := by intro q r hqr rw [mLimsupIic, mLimsupIic] @@ -54,15 +55,15 @@ lemma monotone_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) : refine Iic_subset_Iic.mpr ?_ exact_mod_cast hqr -lemma mLimsupIic_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : +lemma mLimsupIic_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (q : ℚ) : 0 ≤ mLimsupIic κ ν a t q := mLimsup_nonneg hκν a t _ -lemma mLimsupIic_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : +lemma mLimsupIic_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (q : ℚ) : mLimsupIic κ ν a t q ≤ 1 := mLimsup_le_one hκν a t _ -lemma tendsto_atTop_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : +lemma tendsto_atTop_mLimsupIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ (kernel.fst κ) a t q) atTop (𝓝 1) := by let ν := kernel.fst κ suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ ν a t n) atTop (𝓝 1) by @@ -124,7 +125,7 @@ lemma tendsto_atBot_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ simp lemma set_integral_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) (q : ℚ) {A : Set ℝ} (hA : MeasurableSet A) : + (a : α) (q : ℚ) {A : Set γ} (hA : MeasurableSet A) : ∫ t in A, mLimsupIic κ ν a t q ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := set_integral_mLimsup hκν a measurableSet_Iic hA @@ -133,7 +134,7 @@ lemma integrable_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] Integrable (fun t ↦ mLimsupIic κ ν a t q) (ν a) := integrable_mLimsup hκν _ measurableSet_Iic -lemma bddBelow_range_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : ℝ) (q : ℚ) : +lemma bddBelow_range_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (q : ℚ) : BddBelow (range fun (r : Ioi q) ↦ mLimsupIic κ ν a t r) := by refine ⟨0, ?_⟩ rw [mem_lowerBounds] @@ -158,7 +159,7 @@ lemma integrable_iInf_rat_gt_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFinite lemma set_integral_iInf_rat_gt_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) (q : ℚ) {A : Set ℝ} (hA : MeasurableSet A) : + (a : α) (q : ℚ) {A : Set γ} (hA : MeasurableSet A) : ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ ν a t r ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by refine le_antisymm ?_ ?_ @@ -193,27 +194,27 @@ lemma iInf_rat_gt_mLimsupIic_eq (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel · intro s hs _ rw [set_integral_mLimsupIic hκν _ _ hs, set_integral_iInf_rat_gt_mLimsupIic hκν _ _ hs] -lemma isRatStieltjesPoint_mLimsupIic_ae (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] (a : α) : +lemma isRatStieltjesPoint_mLimsupIic_ae (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), IsRatStieltjesPoint (fun p q ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2 q) (a, t) := by filter_upwards [tendsto_atTop_mLimsupIic κ a, tendsto_atBot_mLimsupIic le_rfl a, iInf_rat_gt_mLimsupIic_eq le_rfl a] with t ht_top ht_bot ht_iInf exact ⟨monotone_mLimsupIic le_rfl a t, ht_top, ht_bot, ht_iInf⟩ -lemma isRatKernelCDF_mLimsupIic (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : - IsRatKernelCDF (fun p : α × ℝ ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2) κ (kernel.fst κ) where +lemma isRatKernelCDF_mLimsupIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsRatKernelCDF (fun p : α × γ ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2) κ (kernel.fst κ) where measurable := measurable_mLimsupIic κ (kernel.fst κ) isRatStieltjesPoint_ae := isRatStieltjesPoint_mLimsupIic_ae κ integrable := integrable_mLimsupIic le_rfl set_integral := fun _ _ hs _ ↦ set_integral_mLimsupIic le_rfl _ _ hs noncomputable -def mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : - α × ℝ → StieltjesFunction := - stieltjesOfMeasurableRat (fun p : α × ℝ ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2) +def mLimsupCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + α × γ → StieltjesFunction := + stieltjesOfMeasurableRat (fun p : α × γ ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2) (isRatKernelCDF_mLimsupIic κ).measurable -lemma isKernelCDF_mLimsupCDF (κ : kernel α (ℝ × ℝ)) [IsFiniteKernel κ] : +lemma isKernelCDF_mLimsupCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : IsKernelCDF (mLimsupCDF κ) κ (kernel.fst κ) := isKernelCDF_stieltjesOfMeasurableRat (isRatKernelCDF_mLimsupIic κ) diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index 40635d7202994..e3d594610bde8 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -26,7 +26,7 @@ variable {α β Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β section Kernel -variable [CountableOrStandardBorel α β] {ρ : kernel α (β × Ω)} [IsFiniteKernel ρ] +variable [CountableOrCountablyGenerated α β] {ρ : kernel α (β × Ω)} [IsFiniteKernel ρ] /-! ### Uniqueness @@ -222,7 +222,7 @@ end Measure section KernelAndMeasure -lemma kernel.condKernel_apply_eq_condKernel [CountableOrStandardBorel α β] +lemma kernel.condKernel_apply_eq_condKernel [CountableOrCountablyGenerated α β] (κ : kernel α (β × Ω)) [IsFiniteKernel κ] (a : α) : ∀ᵐ b ∂(kernel.fst κ a), kernel.condKernel κ (a, b) = (κ a).condKernel b := by have : κ a = (κ a).fst @@ -236,7 +236,7 @@ lemma kernel.condKernel_apply_eq_condKernel [CountableOrStandardBorel α β] filter_upwards [h] with b hb rw [← hb, kernel.comap_apply] -lemma condKernel_const [CountableOrStandardBorel α β] (ρ : Measure (β × Ω)) [IsFiniteMeasure ρ] +lemma condKernel_const [CountableOrCountablyGenerated α β] (ρ : Measure (β × Ω)) [IsFiniteMeasure ρ] (a : α) : ∀ᵐ b ∂ρ.fst, kernel.condKernel (kernel.const α ρ) (a, b) = ρ.condKernel b := by have h := kernel.condKernel_apply_eq_condKernel (kernel.const α ρ) a diff --git a/Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean similarity index 88% rename from Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean rename to Mathlib/Probability/Kernel/RadonNikodym.lean index cd1efd595cea6..302346b736630 100644 --- a/Mathlib/Probability/Kernel/Disintegration/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -49,7 +49,7 @@ section Real variable {κ ν : kernel α ℝ} -lemma fst_map_le_of_le (hκν : κ ≤ ν) : +lemma kernel.fst_map_prod_le_of_le (hκν : κ ≤ ν) : kernel.fst (kernel.map κ (fun a ↦ (a, ())) (@measurable_prod_mk_right ℝ Unit _ inferInstance _)) ≤ ν := by rwa [kernel.fst_map_prod _ measurable_id' measurable_const, kernel.map_id'] @@ -60,10 +60,10 @@ def g (κ ν : kernel α ℝ) (a : α) (x : ℝ) : ℝ := (@measurable_prod_mk_right ℝ Unit _ inferInstance _)) ν a x univ lemma g_nonneg (hκν : κ ≤ ν) {a : α} {x : ℝ} : 0 ≤ g κ ν a x := - mLimsup_nonneg (fst_map_le_of_le hκν) _ _ _ + mLimsup_nonneg (kernel.fst_map_prod_le_of_le hκν) _ _ _ lemma g_le_one (hκν : κ ≤ ν) {a : α} {x : ℝ} : g κ ν a x ≤ 1 := - mLimsup_le_one (fst_map_le_of_le hκν) _ _ _ + mLimsup_le_one (kernel.fst_map_prod_le_of_le hκν) _ _ _ lemma measurable_g (κ : kernel α ℝ) (ν : kernel α ℝ) : Measurable (fun p : α × ℝ ↦ g κ ν p.1 p.2) := @@ -101,13 +101,13 @@ lemma withDensity_g (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFiniteKernel simp_rw [this] rw [← ofReal_integral_eq_lintegral_ofReal] · unfold g - rw [set_integral_mLimsup (fst_map_le_of_le h_le) a MeasurableSet.univ hs, + rw [set_integral_mLimsup (kernel.fst_map_prod_le_of_le h_le) a MeasurableSet.univ hs, ENNReal.ofReal_toReal, kernel.map_apply' _ _ _ (hs.prod MeasurableSet.univ)] · congr with x simp · exact measure_ne_top _ _ - · exact (integrable_mLimsup (fst_map_le_of_le h_le) a MeasurableSet.univ).restrict - · exact ae_of_all _ (fun x ↦ mLimsup_nonneg (fst_map_le_of_le h_le) _ _ _) + · exact (integrable_mLimsup (kernel.fst_map_prod_le_of_le h_le) a MeasurableSet.univ).restrict + · exact ae_of_all _ (fun x ↦ mLimsup_nonneg (kernel.fst_map_prod_le_of_le h_le) _ _ _) lemma withDensity_one_sub_g (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFiniteKernel ν] : kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (1 - g κ (κ + ν) a x)) = ν := by @@ -124,7 +124,7 @@ lemma withDensity_one_sub_g (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFinit at h rwa [withDensity_g, add_comm, ENNReal.add_right_inj (measure_ne_top _ _)] at h have h_nonneg : ∀ a x, 0 ≤ g κ (κ + ν) a x := - fun _ _ ↦ mLimsup_nonneg (fst_map_le_of_le h_le) _ _ _ + fun _ _ ↦ mLimsup_nonneg (kernel.fst_map_prod_le_of_le h_le) _ _ _ have : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl simp_rw [this, ENNReal.ofReal_sub _ (h_nonneg _ _), ENNReal.ofReal_one] rw [kernel.withDensity_sub_add] @@ -139,13 +139,13 @@ lemma withDensity_one_sub_g (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFinit _ < ⊤ := by rw [MeasureTheory.lintegral_const, one_mul]; exact measure_lt_top _ _ · refine fun a ↦ ae_of_all _ (fun x ↦ ?_) simp only [ENNReal.ofReal_le_one] - exact mLimsup_le_one (fst_map_le_of_le h_le) _ _ _ + exact mLimsup_le_one (kernel.fst_map_prod_le_of_le h_le) _ _ _ noncomputable def kernel.singularPartReal (κ ν : kernel α ℝ) [IsSFiniteKernel κ] [IsSFiniteKernel ν] : kernel α ℝ := - kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (g κ (κ + ν) a x) - - Real.toNNReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x) + kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (g κ (κ + ν) a x) + - Real.toNNReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x) lemma kernel.mutuallySingular_singularPartReal (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : @@ -166,25 +166,22 @@ lemma kernel.mutuallySingular_singularPartReal (κ ν : kernel α ℝ) refine ae_of_all _ (fun x hx ↦ ?_) simp only [mem_setOf_eq] at hx simp [hx] - · rw [kernel.singularPartReal, kernel.withDensity_apply', lintegral_eq_zero_iff, EventuallyEq, + · have h_meas' : Measurable <| Function.uncurry fun a b ↦ ENNReal.ofReal (g κ (κ + ν) a b) + - ENNReal.ofReal (1 - g κ (κ + ν) a b) * rnDerivReal κ ν a b := by + refine (measurable_g _ _).ennreal_ofReal.sub (Measurable.mul ?_ (measurable_rnDerivReal _ _)) + exact (measurable_const.sub (measurable_g _ _)).ennreal_ofReal + have h_meas : Measurable fun b ↦ ENNReal.ofReal (g κ (κ + ν) a b) + - ENNReal.ofReal (1 - g κ (κ + ν) a b) * rnDerivReal κ ν a b := by + change Measurable ((Function.uncurry fun a b ↦ ENNReal.ofReal (g κ (κ + ν) a b) + - ENNReal.ofReal (1 - g κ (κ + ν) a b) * rnDerivReal κ ν a b) ∘ (fun b ↦ (a, b))) + exact h_meas'.comp measurable_prod_mk_left + rw [kernel.singularPartReal, kernel.withDensity_apply', lintegral_eq_zero_iff, EventuallyEq, ae_restrict_iff] all_goals simp_rw [h_coe] rotate_left - · refine measurableSet_preimage (Measurable.sub ?_ ?_) (measurableSet_singleton _) - · exact (measurable_g_right _ _ _).ennreal_ofReal - · refine Measurable.mul ?_ ?_ - · exact (measurable_const.sub (measurable_g_right _ _ _)).ennreal_ofReal - · exact measurable_rnDerivReal_right _ _ _ - · refine Measurable.sub ?_ ?_ - · exact (measurable_g_right _ _ _).ennreal_ofReal - · refine Measurable.mul ?_ ?_ - · exact (measurable_const.sub (measurable_g_right _ _ _)).ennreal_ofReal - · exact measurable_rnDerivReal_right _ _ _ - · refine Measurable.sub ?_ ?_ - · exact (measurable_g _ _).ennreal_ofReal - · refine Measurable.mul ?_ ?_ - · exact (measurable_const.sub (measurable_g _ _)).ennreal_ofReal - · exact measurable_rnDerivReal _ _ + · exact measurableSet_preimage h_meas (measurableSet_singleton _) + · exact h_meas + · exact h_meas' refine ae_of_all _ (fun x hx ↦ ?_) simp only [mem_compl_iff, mem_setOf_eq] at hx simp_rw [rnDerivReal] From 774b8623e3a026b117d50f94dbbf8f97a762b448 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 21 Feb 2024 10:22:12 +0100 Subject: [PATCH 042/129] renames and other stuff --- .../MeasureTheory/MeasurableSpace/Defs.lean | 4 + Mathlib/Probability/Kernel/CondDistrib.lean | 1 - .../Kernel/Disintegration/AuxLemmas.lean | 17 + .../Kernel/Disintegration/Basic.lean | 6 +- .../Kernel/Disintegration/CDFKernel.lean | 1 - .../Kernel/Disintegration/CondCdf.lean | 156 ++++---- .../Kernel/Disintegration/Density.lean | 363 +++++++++--------- .../Kernel/Disintegration/KernelCDFReal.lean | 190 ++++----- Mathlib/Probability/Kernel/RadonNikodym.lean | 313 ++++++--------- 9 files changed, 481 insertions(+), 570 deletions(-) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean b/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean index e1a665fc570a7..83a1901b67561 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean @@ -457,6 +457,10 @@ theorem generateFrom_sup_generateFrom {s t : Set (Set α)} : (@giGenerateFrom α).gc.l_sup.symm #align measurable_space.generate_from_sup_generate_from MeasurableSpace.generateFrom_sup_generateFrom +lemma iSup_generateFrom {ι : Type*} (s : ι → Set (Set α)) : + ⨆ i, MeasurableSpace.generateFrom (s i) = MeasurableSpace.generateFrom (⋃ i, s i) := + (@MeasurableSpace.giGenerateFrom α).gc.l_iSup.symm + theorem generateFrom_singleton_empty : generateFrom {∅} = (⊥ : MeasurableSpace α) := bot_unique <| generateFrom_le <| by simp [@MeasurableSet.empty α ⊥] #align measurable_space.generate_from_singleton_empty MeasurableSpace.generateFrom_singleton_empty diff --git a/Mathlib/Probability/Kernel/CondDistrib.lean b/Mathlib/Probability/Kernel/CondDistrib.lean index a3a537b9d7dff..9f9582fcb2e03 100644 --- a/Mathlib/Probability/Kernel/CondDistrib.lean +++ b/Mathlib/Probability/Kernel/CondDistrib.lean @@ -3,7 +3,6 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.Disintegration.Integral import Mathlib.Probability.Kernel.Disintegration.Unique import Mathlib.Probability.Notation diff --git a/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean b/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean index 6960ec563a1d5..0222614c4779a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean +++ b/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean @@ -40,3 +40,20 @@ lemma measurableSet_tendsto_fun {β γ ι : Type*} [MeasurableSpace β] rw [tendsto_iff_dist_tendsto_zero] rw [this] exact measurableSet_tendsto_nhds (fun n ↦ (hf n).dist hg) 0 + +lemma Measure.iInf_Iic_gt_prod {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] + {s : Set α} (hs : MeasurableSet s) (t : ℚ) : + ⨅ r : { r' : ℚ // t < r' }, ρ (s ×ˢ Iic (r : ℝ)) = ρ (s ×ˢ Iic (t : ℝ)) := by + rw [← measure_iInter_eq_iInf] + · rw [← prod_iInter] + congr with x : 1 + simp only [mem_iInter, mem_Iic, Subtype.forall, Subtype.coe_mk] + refine ⟨fun h ↦ ?_, fun h a hta ↦ h.trans ?_⟩ + · refine le_of_forall_lt_rat_imp_le fun q htq ↦ h q ?_ + exact mod_cast htq + · exact mod_cast hta.le + · exact fun _ => hs.prod measurableSet_Iic + · refine Monotone.directed_ge fun r r' hrr' ↦ prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, ?_⟩) + refine Iic_subset_Iic.mpr ?_ + exact mod_cast hrr' + · exact ⟨⟨t + 1, lt_add_one _⟩, measure_ne_top ρ _⟩ diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 7bde52b4639f6..bcdcdbf0ead72 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -211,21 +211,21 @@ section StandardBorel noncomputable def kernel.condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : kernel (α × γ) ℝ := - cdfKernel _ (isKernelCDF_mLimsupCDF κ) + cdfKernel _ (isKernelCDF_kernelCDF κ) instance (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : IsMarkovKernel (kernel.condKernelReal κ) := by unfold kernel.condKernelReal; infer_instance lemma kernel.eq_compProd_condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : κ = kernel.fst κ ⊗ₖ kernel.condKernelReal κ := - kernel.eq_compProd_cdfKernel (isKernelCDF_mLimsupCDF κ) + kernel.eq_compProd_cdfKernel (isKernelCDF_kernelCDF κ) noncomputable def condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : kernel (α × γ) Ω := let f := measurableEmbedding_real Ω let hf := measurableEmbedding_measurableEmbedding_real Ω let κ' := kernel.map κ (Prod.map (id : γ → γ) f) (measurable_id.prod_map hf.measurable) - condKernelBorelSnd κ (isKernelCDF_mLimsupCDF κ') + condKernelBorelSnd κ (isKernelCDF_kernelCDF κ') instance instIsMarkovKernel_condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : IsMarkovKernel (condKernelBorel κ) := by diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean index af6dd8420fdc4..2f966c4628090 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean @@ -5,7 +5,6 @@ Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes import Mathlib.Probability.Kernel.MeasureCompProd -import Mathlib.Probability.Kernel.Disintegration.AuxLemmas /-! diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index a7c121a79def1..804c6af0d84fc 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -70,7 +70,7 @@ theorem IicSnd_univ (r : ℝ) : ρ.IicSnd r univ = ρ (univ ×ˢ Iic r) := theorem IicSnd_mono {r r' : ℝ} (h_le : r ≤ r') : ρ.IicSnd r ≤ ρ.IicSnd r' := by intro s hs simp_rw [IicSnd_apply ρ _ hs] - refine' measure_mono (prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, Iic_subset_Iic.mpr _⟩)) + refine measure_mono (prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, Iic_subset_Iic.mpr ?_⟩)) exact mod_cast h_le #align measure_theory.measure.Iic_snd_mono MeasureTheory.Measure.IicSnd_mono @@ -91,54 +91,41 @@ theorem IsFiniteMeasure.IicSnd {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] ( theorem iInf_IicSnd_gt (t : ℚ) {s : Set α} (hs : MeasurableSet s) [IsFiniteMeasure ρ] : ⨅ r : { r' : ℚ // t < r' }, ρ.IicSnd r s = ρ.IicSnd t s := by - simp_rw [ρ.IicSnd_apply _ hs] - rw [← measure_iInter_eq_iInf] - · rw [← prod_iInter] - congr with x : 1 - simp only [mem_iInter, mem_Iic, Subtype.forall, Subtype.coe_mk] - refine' ⟨fun h => _, fun h a hta => h.trans _⟩ - · refine' le_of_forall_lt_rat_imp_le fun q htq => h q _ - exact mod_cast htq - · exact mod_cast hta.le - · exact fun _ => hs.prod measurableSet_Iic - · refine' Monotone.directed_ge fun r r' hrr' => prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, _⟩) - refine' Iic_subset_Iic.mpr _ - exact mod_cast hrr' - · exact ⟨⟨t + 1, lt_add_one _⟩, measure_ne_top ρ _⟩ + simp_rw [ρ.IicSnd_apply _ hs, Measure.iInf_Iic_gt_prod hs] #align measure_theory.measure.infi_Iic_snd_gt MeasureTheory.Measure.iInf_IicSnd_gt theorem tendsto_IicSnd_atTop {s : Set α} (hs : MeasurableSet s) : - Tendsto (fun r : ℚ => ρ.IicSnd r s) atTop (𝓝 (ρ.fst s)) := by + Tendsto (fun r : ℚ ↦ ρ.IicSnd r s) atTop (𝓝 (ρ.fst s)) := by simp_rw [ρ.IicSnd_apply _ hs, fst_apply hs, ← prod_univ] rw [← Real.iUnion_Iic_rat, prod_iUnion] - refine' tendsto_measure_iUnion fun r q hr_le_q x => _ + refine tendsto_measure_iUnion fun r q hr_le_q x ↦ ?_ simp only [mem_prod, mem_Iic, and_imp] - refine' fun hxs hxr => ⟨hxs, hxr.trans _⟩ + refine fun hxs hxr ↦ ⟨hxs, hxr.trans ?_⟩ exact mod_cast hr_le_q #align measure_theory.measure.tendsto_Iic_snd_at_top MeasureTheory.Measure.tendsto_IicSnd_atTop theorem tendsto_IicSnd_atBot [IsFiniteMeasure ρ] {s : Set α} (hs : MeasurableSet s) : - Tendsto (fun r : ℚ => ρ.IicSnd r s) atBot (𝓝 0) := by + Tendsto (fun r : ℚ ↦ ρ.IicSnd r s) atBot (𝓝 0) := by simp_rw [ρ.IicSnd_apply _ hs] have h_empty : ρ (s ×ˢ ∅) = 0 := by simp only [prod_empty, measure_empty] rw [← h_empty, ← Real.iInter_Iic_rat, prod_iInter] suffices h_neg : - Tendsto (fun r : ℚ => ρ (s ×ˢ Iic ↑(-r))) atTop (𝓝 (ρ (⋂ r : ℚ, s ×ˢ Iic ↑(-r)))) by + Tendsto (fun r : ℚ ↦ ρ (s ×ˢ Iic ↑(-r))) atTop (𝓝 (ρ (⋂ r : ℚ, s ×ˢ Iic ↑(-r)))) by have h_inter_eq : ⋂ r : ℚ, s ×ˢ Iic ↑(-r) = ⋂ r : ℚ, s ×ˢ Iic (r : ℝ) := by ext1 x simp only [Rat.cast_eq_id, id.def, mem_iInter, mem_prod, mem_Iic] - refine' ⟨fun h i => ⟨(h i).1, _⟩, fun h i => ⟨(h i).1, _⟩⟩ <;> have h' := h (-i) + refine ⟨fun h i ↦ ⟨(h i).1, ?_⟩, fun h i ↦ ⟨(h i).1, ?_⟩⟩ <;> have h' := h (-i) · rw [neg_neg] at h'; exact h'.2 · exact h'.2 rw [h_inter_eq] at h_neg - have h_fun_eq : (fun r : ℚ => ρ (s ×ˢ Iic (r : ℝ))) = fun r : ℚ => ρ (s ×ˢ Iic ↑(- -r)) := by + have h_fun_eq : (fun r : ℚ ↦ ρ (s ×ˢ Iic (r : ℝ))) = fun r : ℚ ↦ ρ (s ×ˢ Iic ↑(- -r)) := by simp_rw [neg_neg] rw [h_fun_eq] exact h_neg.comp tendsto_neg_atBot_atTop - refine' tendsto_measure_iInter (fun q => hs.prod measurableSet_Iic) _ ⟨0, measure_ne_top ρ _⟩ - refine' fun q r hqr => prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, fun x hx => _⟩) + refine tendsto_measure_iInter (fun q ↦ hs.prod measurableSet_Iic) ?_ ⟨0, measure_ne_top ρ _⟩ + refine fun q r hqr ↦ prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, fun x hx ↦ ?_⟩) simp only [Rat.cast_neg, mem_Iic] at hx ⊢ - refine' hx.trans (neg_le_neg _) + refine hx.trans (neg_le_neg ?_) exact mod_cast hqr #align measure_theory.measure.tendsto_Iic_snd_at_bot MeasureTheory.Measure.tendsto_IicSnd_atBot @@ -190,47 +177,45 @@ theorem set_lintegral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) {s : Set rw [this, ← set_lintegral_withDensity_eq_set_lintegral_mul _ measurable_preCDF _ hs] · simp only [withDensity_preCDF ρ r, Pi.one_apply, lintegral_one, Measure.restrict_apply, MeasurableSet.univ, univ_inter] - · rw [(_ : (1 : α → ℝ≥0∞) = fun _ => 1)] + · rw [(_ : (1 : α → ℝ≥0∞) = fun _ ↦ 1)] exacts [measurable_const, rfl] #align probability_theory.set_lintegral_pre_cdf_fst ProbabilityTheory.set_lintegral_preCDF_fst theorem monotone_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - ∀ᵐ a ∂ρ.fst, Monotone fun r => preCDF ρ r a := by + ∀ᵐ a ∂ρ.fst, Monotone fun r ↦ preCDF ρ r a := by simp_rw [Monotone, ae_all_iff] - refine' fun r r' hrr' => - ae_le_of_forall_set_lintegral_le_of_sigmaFinite measurable_preCDF measurable_preCDF - fun s hs _ => _ + refine fun r r' hrr' ↦ ae_le_of_forall_set_lintegral_le_of_sigmaFinite measurable_preCDF + measurable_preCDF fun s hs _ ↦ ?_ rw [set_lintegral_preCDF_fst ρ r hs, set_lintegral_preCDF_fst ρ r' hs] - refine' Measure.IicSnd_mono ρ _ s hs + refine Measure.IicSnd_mono ρ ?_ s hs exact mod_cast hrr' #align probability_theory.monotone_pre_cdf ProbabilityTheory.monotone_preCDF theorem set_lintegral_iInf_gt_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (t : ℚ) {s : Set α} (hs : MeasurableSet s) : ∫⁻ x in s, ⨅ r : Ioi t, preCDF ρ r x ∂ρ.fst = ρ.IicSnd t s := by - refine' le_antisymm _ _ + refine le_antisymm ?_ ?_ · have h : ∀ q : Ioi t, ∫⁻ x in s, ⨅ r : Ioi t, preCDF ρ r x ∂ρ.fst ≤ ρ.IicSnd q s := by intro q rw [← set_lintegral_preCDF_fst ρ _ hs] - refine' set_lintegral_mono_ae _ measurable_preCDF _ - · exact measurable_iInf fun _ => measurable_preCDF + refine set_lintegral_mono_ae ?_ measurable_preCDF ?_ + · exact measurable_iInf fun _ ↦ measurable_preCDF · filter_upwards [monotone_preCDF _] with a _ - exact fun _ => iInf_le _ q + exact fun _ ↦ iInf_le _ q calc ∫⁻ x in s, ⨅ r : Ioi t, preCDF ρ r x ∂ρ.fst ≤ ⨅ q : Ioi t, ρ.IicSnd q s := le_iInf h _ = ρ.IicSnd t s := Measure.iInf_IicSnd_gt ρ t hs · rw [(set_lintegral_preCDF_fst ρ t hs).symm] - refine' set_lintegral_mono_ae measurable_preCDF _ _ - · exact measurable_iInf fun _ => measurable_preCDF + refine set_lintegral_mono_ae measurable_preCDF ?_ ?_ + · exact measurable_iInf fun _ ↦ measurable_preCDF · filter_upwards [monotone_preCDF _] with a ha_mono - exact fun _ => le_iInf fun r => ha_mono (le_of_lt r.prop) + exact fun _ ↦ le_iInf fun r ↦ ha_mono (le_of_lt r.prop) #align probability_theory.set_lintegral_infi_gt_pre_cdf ProbabilityTheory.set_lintegral_iInf_gt_preCDF theorem preCDF_le_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : ∀ᵐ a ∂ρ.fst, ∀ r, preCDF ρ r a ≤ 1 := by rw [ae_all_iff] - refine' fun r => - ae_le_of_forall_set_lintegral_le_of_sigmaFinite measurable_preCDF measurable_const - fun s hs _ => _ + refine fun r ↦ ae_le_of_forall_set_lintegral_le_of_sigmaFinite measurable_preCDF + measurable_const fun s hs _ ↦ ?_ rw [set_lintegral_preCDF_fst ρ r hs] simp only [Pi.one_apply, lintegral_one, Measure.restrict_apply, MeasurableSet.univ, univ_inter] exact Measure.IicSnd_le_fst ρ r s hs @@ -247,20 +232,20 @@ theorem set_integral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) {s : Set α exact (ha r).trans_lt ENNReal.one_lt_top theorem tendsto_lintegral_preCDF_atTop (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - Tendsto (fun r => ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (ρ univ)) := by + Tendsto (fun r ↦ ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (ρ univ)) := by convert ρ.tendsto_IicSnd_atTop MeasurableSet.univ · rw [← set_lintegral_univ, set_lintegral_preCDF_fst ρ _ MeasurableSet.univ] · exact Measure.fst_univ.symm #align probability_theory.tendsto_lintegral_pre_cdf_at_top ProbabilityTheory.tendsto_lintegral_preCDF_atTop theorem tendsto_lintegral_preCDF_atBot (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - Tendsto (fun r => ∫⁻ a, preCDF ρ r a ∂ρ.fst) atBot (𝓝 0) := by + Tendsto (fun r ↦ ∫⁻ a, preCDF ρ r a ∂ρ.fst) atBot (𝓝 0) := by convert ρ.tendsto_IicSnd_atBot MeasurableSet.univ rw [← set_lintegral_univ, set_lintegral_preCDF_fst ρ _ MeasurableSet.univ] #align probability_theory.tendsto_lintegral_pre_cdf_at_bot ProbabilityTheory.tendsto_lintegral_preCDF_atBot theorem tendsto_preCDF_atTop_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - ∀ᵐ a ∂ρ.fst, Tendsto (fun r => preCDF ρ r a) atTop (𝓝 1) := by + ∀ᵐ a ∂ρ.fst, Tendsto (fun r ↦ preCDF ρ r a) atTop (𝓝 1) := by -- We show first that `preCDF` has a limit almost everywhere. That limit has to be at most 1. -- We then show that the integral of `preCDF` tends to the integral of 1, and that it also tends -- to the integral of the limit. Since the limit is at most 1 and has same integral as 1, it is @@ -268,21 +253,21 @@ theorem tendsto_preCDF_atTop_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] have h_mono := monotone_preCDF ρ have h_le_one := preCDF_le_one ρ -- `preCDF` has a limit a.e. - have h_exists : ∀ᵐ a ∂ρ.fst, ∃ l, Tendsto (fun r => preCDF ρ r a) atTop (𝓝 l) := by + have h_exists : ∀ᵐ a ∂ρ.fst, ∃ l, Tendsto (fun r ↦ preCDF ρ r a) atTop (𝓝 l) := by filter_upwards [h_mono] with a ha_mono exact ⟨_, tendsto_atTop_iSup ha_mono⟩ classical -- let `F` be the pointwise limit of `preCDF` where it exists, and 0 elsewhere. - let F : α → ℝ≥0∞ := fun a => - if h : ∃ l, Tendsto (fun r => preCDF ρ r a) atTop (𝓝 l) then h.choose else 0 - have h_tendsto_ℚ : ∀ᵐ a ∂ρ.fst, Tendsto (fun r => preCDF ρ r a) atTop (𝓝 (F a)) := by + let F : α → ℝ≥0∞ := fun a ↦ + if h : ∃ l, Tendsto (fun r ↦ preCDF ρ r a) atTop (𝓝 l) then h.choose else 0 + have h_tendsto_ℚ : ∀ᵐ a ∂ρ.fst, Tendsto (fun r ↦ preCDF ρ r a) atTop (𝓝 (F a)) := by filter_upwards [h_exists] with a ha simp_rw [dif_pos ha] exact ha.choose_spec - have h_tendsto_ℕ : ∀ᵐ a ∂ρ.fst, Tendsto (fun n : ℕ => preCDF ρ n a) atTop (𝓝 (F a)) := by + have h_tendsto_ℕ : ∀ᵐ a ∂ρ.fst, Tendsto (fun n : ℕ ↦ preCDF ρ n a) atTop (𝓝 (F a)) := by filter_upwards [h_tendsto_ℚ] with a ha using ha.comp tendsto_nat_cast_atTop_atTop have hF_ae_meas : AEMeasurable F ρ.fst := by - refine' aemeasurable_of_tendsto_metrizable_ae _ (fun n => _) h_tendsto_ℚ + refine aemeasurable_of_tendsto_metrizable_ae _ (fun n ↦ ?_) h_tendsto_ℚ exact measurable_preCDF.aemeasurable have hF_le_one : ∀ᵐ a ∂ρ.fst, F a ≤ 1 := by filter_upwards [h_tendsto_ℚ, h_le_one] with a ha ha_le using le_of_tendsto' ha ha_le @@ -294,17 +279,14 @@ theorem tendsto_preCDF_atTop_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] -- us that `F` is 1 a.e. have h_lintegral_eq : ∫⁻ a, F a ∂ρ.fst = ∫⁻ _, 1 ∂ρ.fst := by have h_lintegral : - Tendsto (fun r : ℕ => ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (∫⁻ a, F a ∂ρ.fst)) := by - refine' - lintegral_tendsto_of_tendsto_of_monotone - (-- does this exist only for ℕ? - fun _ => measurable_preCDF.aemeasurable) - _ h_tendsto_ℕ + Tendsto (fun r : ℕ ↦ ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (∫⁻ a, F a ∂ρ.fst)) := by + refine lintegral_tendsto_of_tendsto_of_monotone + (fun _ ↦ measurable_preCDF.aemeasurable) ?_ h_tendsto_ℕ filter_upwards [h_mono] with a ha - refine' fun n m hnm => ha _ + refine fun n m hnm ↦ ha ?_ exact mod_cast hnm have h_lintegral' : - Tendsto (fun r : ℕ => ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (∫⁻ _, 1 ∂ρ.fst)) := by + Tendsto (fun r : ℕ ↦ ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (∫⁻ _, 1 ∂ρ.fst)) := by rw [lintegral_one, Measure.fst_univ] exact (tendsto_lintegral_preCDF_atTop ρ).comp tendsto_nat_cast_atTop_atTop exact tendsto_nhds_unique h_lintegral h_lintegral' @@ -322,46 +304,45 @@ theorem tendsto_preCDF_atTop_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] #align probability_theory.tendsto_pre_cdf_at_top_one ProbabilityTheory.tendsto_preCDF_atTop_one theorem tendsto_preCDF_atBot_zero (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - ∀ᵐ a ∂ρ.fst, Tendsto (fun r => preCDF ρ r a) atBot (𝓝 0) := by + ∀ᵐ a ∂ρ.fst, Tendsto (fun r ↦ preCDF ρ r a) atBot (𝓝 0) := by -- We show first that `preCDF` has a limit in ℝ≥0∞ almost everywhere. -- We then show that the integral of `preCDF` tends to 0, and that it also tends -- to the integral of the limit. Since the limit has integral 0, it is equal to 0 a.e. - suffices ∀ᵐ a ∂ρ.fst, Tendsto (fun r => preCDF ρ (-r) a) atTop (𝓝 0) by + suffices ∀ᵐ a ∂ρ.fst, Tendsto (fun r ↦ preCDF ρ (-r) a) atTop (𝓝 0) by filter_upwards [this] with a ha - have h_eq_neg : (fun r : ℚ => preCDF ρ r a) = fun r : ℚ => preCDF ρ (- -r) a := by + have h_eq_neg : (fun r : ℚ ↦ preCDF ρ r a) = fun r : ℚ ↦ preCDF ρ (- -r) a := by simp_rw [neg_neg] rw [h_eq_neg] exact ha.comp tendsto_neg_atBot_atTop - have h_exists : ∀ᵐ a ∂ρ.fst, ∃ l, Tendsto (fun r => preCDF ρ (-r) a) atTop (𝓝 l) := by + have h_exists : ∀ᵐ a ∂ρ.fst, ∃ l, Tendsto (fun r ↦ preCDF ρ (-r) a) atTop (𝓝 l) := by filter_upwards [monotone_preCDF ρ] with a ha - have h_anti : Antitone fun r => preCDF ρ (-r) a := fun p q hpq => ha (neg_le_neg hpq) + have h_anti : Antitone fun r ↦ preCDF ρ (-r) a := fun p q hpq ↦ ha (neg_le_neg hpq) exact ⟨_, tendsto_atTop_iInf h_anti⟩ classical - let F : α → ℝ≥0∞ := fun a => - if h : ∃ l, Tendsto (fun r => preCDF ρ (-r) a) atTop (𝓝 l) then h.choose else 0 - have h_tendsto : ∀ᵐ a ∂ρ.fst, Tendsto (fun r => preCDF ρ (-r) a) atTop (𝓝 (F a)) := by + let F : α → ℝ≥0∞ := fun a ↦ + if h : ∃ l, Tendsto (fun r ↦ preCDF ρ (-r) a) atTop (𝓝 l) then h.choose else 0 + have h_tendsto : ∀ᵐ a ∂ρ.fst, Tendsto (fun r ↦ preCDF ρ (-r) a) atTop (𝓝 (F a)) := by filter_upwards [h_exists] with a ha simp_rw [dif_pos ha] exact ha.choose_spec suffices h_lintegral_eq : ∫⁻ a, F a ∂ρ.fst = 0 by have hF_ae_meas : AEMeasurable F ρ.fst := by - refine' aemeasurable_of_tendsto_metrizable_ae _ (fun n => _) h_tendsto + refine aemeasurable_of_tendsto_metrizable_ae _ (fun n ↦ ?_) h_tendsto exact measurable_preCDF.aemeasurable rw [lintegral_eq_zero_iff' hF_ae_meas] at h_lintegral_eq filter_upwards [h_tendsto, h_lintegral_eq] with a ha_tendsto ha_eq rwa [ha_eq] at ha_tendsto have h_lintegral : - Tendsto (fun r => ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) atTop (𝓝 (∫⁻ a, F a ∂ρ.fst)) := by - refine' - tendsto_lintegral_filter_of_dominated_convergence (fun _ => 1) - (eventually_of_forall fun _ => measurable_preCDF) (eventually_of_forall fun _ => _) _ - h_tendsto + Tendsto (fun r ↦ ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) atTop (𝓝 (∫⁻ a, F a ∂ρ.fst)) := by + refine tendsto_lintegral_filter_of_dominated_convergence (fun _ ↦ 1) + (eventually_of_forall fun _ ↦ measurable_preCDF) (eventually_of_forall fun _ ↦ ?_) ?_ + h_tendsto · filter_upwards [preCDF_le_one ρ] with a ha using ha _ · rw [lintegral_one] exact measure_ne_top _ _ - have h_lintegral' : Tendsto (fun r => ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) atTop (𝓝 0) := by - have h_lintegral_eq : - (fun r => ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) = fun r : ℚ => ρ (univ ×ˢ Iic (-r : ℝ)) := by + have h_lintegral' : Tendsto (fun r ↦ ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) atTop (𝓝 0) := by + have h_lintegral_eq : (fun r ↦ ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) + = fun r : ℚ ↦ ρ (univ ×ˢ Iic (-r : ℝ)) := by ext1 n rw [← set_lintegral_univ, set_lintegral_preCDF_fst ρ _ MeasurableSet.univ, Measure.IicSnd_univ] @@ -374,11 +355,10 @@ theorem tendsto_preCDF_atBot_zero (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ simp_rw [neg_lt] exact exists_rat_gt _ rw [h_zero_eq_measure_iInter] - refine' - tendsto_measure_iInter (fun n => MeasurableSet.univ.prod measurableSet_Iic) - (fun i j hij x => _) ⟨0, measure_ne_top ρ _⟩ + refine tendsto_measure_iInter (fun n ↦ MeasurableSet.univ.prod measurableSet_Iic) + (fun i j hij x ↦ ?_) ⟨0, measure_ne_top ρ _⟩ simp only [mem_prod, mem_univ, mem_Iic, true_and_iff] - refine' fun hxj => hxj.trans (neg_le_neg _) + refine fun hxj ↦ hxj.trans (neg_le_neg ?_) exact mod_cast hij exact tendsto_nhds_unique h_lintegral h_lintegral' #align probability_theory.tendsto_pre_cdf_at_bot_zero ProbabilityTheory.tendsto_preCDF_atBot_zero @@ -386,8 +366,8 @@ theorem tendsto_preCDF_atBot_zero (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ theorem inf_gt_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : ∀ᵐ a ∂ρ.fst, ∀ t : ℚ, ⨅ r : Ioi t, preCDF ρ r a = preCDF ρ t a := by rw [ae_all_iff] - refine' fun t => ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite _ measurable_preCDF _ - · exact measurable_iInf fun i => measurable_preCDF + refine fun t ↦ ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite ?_ measurable_preCDF ?_ + · exact measurable_iInf fun i ↦ measurable_preCDF intro s hs _ rw [set_lintegral_iInf_gt_preCDF ρ t hs, set_lintegral_preCDF_fst ρ t hs] #align probability_theory.inf_gt_pre_cdf ProbabilityTheory.inf_gt_preCDF @@ -445,7 +425,7 @@ lemma isRatStieltjesPoint_ae (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : theorem integrable_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℚ) : Integrable (fun a ↦ (preCDF ρ x a).toReal) ρ.fst := by - refine' integrable_of_forall_fin_meas_le _ (measure_lt_top ρ.fst univ) _ fun t _ _ ↦ _ + refine integrable_of_forall_fin_meas_le _ (measure_lt_top ρ.fst univ) ?_ fun t _ _ ↦ ?_ · exact measurable_preCDF.ennreal_toReal.aestronglyMeasurable · simp_rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_of_nonneg ENNReal.toReal_nonneg] rw [← lintegral_one] @@ -503,14 +483,14 @@ theorem condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) #align probability_theory.cond_cdf_ae_eq ProbabilityTheory.condCDF_ae_eq theorem ofReal_condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) : - (fun a => ENNReal.ofReal (condCDF ρ a r)) =ᵐ[ρ.fst] preCDF ρ r := by + (fun a ↦ ENNReal.ofReal (condCDF ρ a r)) =ᵐ[ρ.fst] preCDF ρ r := by filter_upwards [condCDF_ae_eq ρ r, preCDF_le_one ρ] with a ha ha_le_one rw [ha, ENNReal.ofReal_toReal] exact ((ha_le_one r).trans_lt ENNReal.one_lt_top).ne #align probability_theory.of_real_cond_cdf_ae_eq ProbabilityTheory.ofReal_condCDF_ae_eq /-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ -theorem measurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : Measurable fun a => condCDF ρ a x := +theorem measurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : Measurable fun a ↦ condCDF ρ a x := measurable_stieltjesOfMeasurableRat _ _ #align probability_theory.measurable_cond_cdf ProbabilityTheory.measurable_condCDF @@ -534,12 +514,12 @@ theorem lintegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : /-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ theorem stronglyMeasurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : - StronglyMeasurable fun a => condCDF ρ a x := stronglyMeasurable_stieltjesOfMeasurableRat _ _ + StronglyMeasurable fun a ↦ condCDF ρ a x := stronglyMeasurable_stieltjesOfMeasurableRat _ _ #align probability_theory.strongly_measurable_cond_cdf ProbabilityTheory.stronglyMeasurable_condCDF theorem integrable_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) : Integrable (fun a ↦ condCDF ρ a x) ρ.fst := by - refine' integrable_of_forall_fin_meas_le _ (measure_lt_top ρ.fst univ) _ fun t _ _ => _ + refine integrable_of_forall_fin_meas_le _ (measure_lt_top ρ.fst univ) ?_ fun t _ _ ↦ ?_ · exact (stronglyMeasurable_condCDF ρ _).aestronglyMeasurable · have : ∀ y, (‖condCDF ρ y x‖₊ : ℝ≥0∞) ≤ 1 := by intro y @@ -559,9 +539,9 @@ theorem set_integral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x have h := set_lintegral_condCDF ρ x hs rw [← ofReal_integral_eq_lintegral_ofReal] at h · rw [← h, ENNReal.toReal_ofReal] - exact integral_nonneg fun _ => condCDF_nonneg _ _ _ + exact integral_nonneg fun _ ↦ condCDF_nonneg _ _ _ · exact (integrable_condCDF _ _).integrableOn - · exact eventually_of_forall fun _ => condCDF_nonneg _ _ _ + · exact eventually_of_forall fun _ ↦ condCDF_nonneg _ _ _ #align probability_theory.set_integral_cond_cdf ProbabilityTheory.set_integral_condCDF theorem integral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) : diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index f97954d067c75..350c4e677d49b 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -6,7 +6,6 @@ Authors: Rémy Degenne import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Martingale.Convergence import Mathlib.Analysis.SpecialFunctions.Log.Base -import Mathlib.Probability.Kernel.Disintegration.AuxLemmas /-! # Kernel density @@ -66,36 +65,19 @@ open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory -lemma Measure.iInf_Iic_gt_prod {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] - {s : Set α} (hs : MeasurableSet s) (t : ℚ) : - ⨅ r : { r' : ℚ // t < r' }, ρ (s ×ˢ Iic (r : ℝ)) = ρ (s ×ˢ Iic (t : ℝ)) := by - rw [← measure_iInter_eq_iInf] - · rw [← prod_iInter] - congr with x : 1 - simp only [mem_iInter, mem_Iic, Subtype.forall, Subtype.coe_mk] - refine' ⟨fun h => _, fun h a hta => h.trans _⟩ - · refine' le_of_forall_lt_rat_imp_le fun q htq => h q _ - exact mod_cast htq - · exact mod_cast hta.le - · exact fun _ => hs.prod measurableSet_Iic - · refine' Monotone.directed_ge fun r r' hrr' => prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, _⟩) - refine' Iic_subset_Iic.mpr _ - exact mod_cast hrr' - · exact ⟨⟨t + 1, lt_add_one _⟩, measure_ne_top ρ _⟩ - variable {α β γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} [MeasurableSpace.CountablyGenerated γ] -- todo : `Filtration` should be renamed to `filtration` -def ℱ (γ : Type*) [m : MeasurableSpace γ] [MeasurableSpace.CountablyGenerated γ] : +def partitionFiltration (γ : Type*) [m : MeasurableSpace γ] [MeasurableSpace.CountablyGenerated γ] : Filtration ℕ m where seq := fun n ↦ MeasurableSpace.generateFrom <| MeasurableSpace.countablePartition γ n mono' := monotone_nat_of_le_succ (MeasurableSpace.generateFrom_countablePartition_le_succ _) le' := MeasurableSpace.generateFrom_countablePartition_le γ -lemma measurableSet_ℱ_countablePartition (n : ℕ) {s : Set γ} +lemma measurableSet_partitionFiltration_countablePartition (n : ℕ) {s : Set γ} (hs : s ∈ MeasurableSpace.countablePartition γ n) : - MeasurableSet[ℱ γ n] s := + MeasurableSet[partitionFiltration γ n] s := MeasurableSpace.measurableSet_generateFrom hs lemma existsPartitionSet_mem (n : ℕ) (t : γ) : @@ -133,24 +115,25 @@ lemma partitionSet_of_mem {n : ℕ} {t : γ} {s : Set γ} partitionSet n t = s := by rwa [← mem_countablePartition_iff n t hs] -lemma measurableSet_ℱ_partitionSet (n : ℕ) (t : γ) : - MeasurableSet[ℱ γ n] (partitionSet n t) := - measurableSet_ℱ_countablePartition n (partitionSet_mem n t) +lemma measurableSet_partitionFiltration_partitionSet (n : ℕ) (t : γ) : + MeasurableSet[partitionFiltration γ n] (partitionSet n t) := + measurableSet_partitionFiltration_countablePartition n (partitionSet_mem n t) lemma measurableSet_partitionSet (n : ℕ) (t : γ) : MeasurableSet (partitionSet n t) := - (ℱ γ).le n _ (measurableSet_ℱ_partitionSet n t) + (partitionFiltration γ).le n _ (measurableSet_partitionFiltration_partitionSet n t) instance todo_move (n : ℕ) : Countable (MeasurableSpace.countablePartition γ n) := Set.Finite.countable (MeasurableSpace.finite_countablePartition _ _) -lemma measurable_aux (n : ℕ) (m : MeasurableSpace (MeasurableSpace.countablePartition γ n)) : - @Measurable γ (MeasurableSpace.countablePartition γ n) (ℱ γ n) m +lemma measurable_partitionSet_aux (n : ℕ) + (m : MeasurableSpace (MeasurableSpace.countablePartition γ n)) : + @Measurable γ (MeasurableSpace.countablePartition γ n) (partitionFiltration γ n) m (fun c : γ ↦ ⟨partitionSet n c, partitionSet_mem n c⟩) := by refine @measurable_to_countable' (MeasurableSpace.countablePartition γ n) γ m _ - (ℱ γ n) _ (fun t ↦ ?_) + (partitionFiltration γ n) _ (fun t ↦ ?_) rcases t with ⟨t, ht⟩ - suffices MeasurableSet[ℱ γ n] {x | partitionSet n x = t} by + suffices MeasurableSet[partitionFiltration γ n] {x | partitionSet n x = t} by convert this ext x simp @@ -158,25 +141,16 @@ lemma measurable_aux (n : ℕ) (m : MeasurableSpace (MeasurableSpace.countablePa ext x rw [mem_setOf_eq, mem_countablePartition_iff n x ht] rw [this] - exact measurableSet_ℱ_countablePartition _ ht - -lemma measurable_partitionSet (n : ℕ) : Measurable[ℱ γ n] (partitionSet n) := by - have : partitionSet n = (fun s : MeasurableSpace.countablePartition γ n ↦ (s : Set γ)) - ∘ (fun t ↦ ⟨partitionSet n t, partitionSet_mem n t⟩) := rfl - rw [this] - refine Measurable.comp - (?_ : Measurable (fun s : MeasurableSpace.countablePartition γ n ↦ (s : Set γ))) ?_ - · measurability - exact measurable_aux _ _ + exact measurableSet_partitionFiltration_countablePartition _ ht -lemma MeasurableSpace.generateFrom_iSup {ι : Type*} (s : ι → Set (Set α)) : - MeasurableSpace.generateFrom (⋃ i, s i) = ⨆ i, MeasurableSpace.generateFrom (s i) := - (@MeasurableSpace.giGenerateFrom α).gc.l_iSup +lemma measurable_partitionSet (n : ℕ) : Measurable[partitionFiltration γ n] (partitionSet n) := + measurable_subtype_coe.comp (measurable_partitionSet_aux _ _) -lemma iSup_ℱ (β : Type*) [m : MeasurableSpace β] [MeasurableSpace.CountablyGenerated β] : - ⨆ n, ℱ β n = m := by +lemma iSup_partitionFiltration (β : Type*) [m : MeasurableSpace β] + [MeasurableSpace.CountablyGenerated β] : + ⨆ n, partitionFiltration β n = m := by conv_rhs => rw [← MeasurableSpace.generateFrom_iUnion_countablePartition β] - rw [MeasurableSpace.generateFrom_iSup] + rw [← MeasurableSpace.iSup_generateFrom] rfl variable [MeasurableSpace β] @@ -192,14 +166,13 @@ def densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a lemma densityProcess_def (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (s : Set β) : (fun t ↦ densityProcess κ ν n a t s) - = fun t ↦ (κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t)).toReal := - rfl + = fun t ↦ (κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t)).toReal := rfl -lemma measurable_densityProcess_ℱ_aux (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) - {s : Set β} (hs : MeasurableSet s) : - Measurable[MeasurableSpace.prod mα (ℱ γ n)] (fun (p : α × γ) ↦ +lemma measurable_densityProcess_partitionFiltration_aux (κ : kernel α (γ × β)) (ν : kernel α γ) + (n : ℕ) {s : Set β} (hs : MeasurableSet s) : + Measurable[MeasurableSpace.prod mα (partitionFiltration γ n)] (fun (p : α × γ) ↦ κ p.1 (partitionSet n p.2 ×ˢ s) / ν p.1 (partitionSet n p.2)) := by - change Measurable[MeasurableSpace.prod mα (ℱ γ n)] + change Measurable[MeasurableSpace.prod mα (partitionFiltration γ n)] ((fun (p : α × MeasurableSpace.countablePartition γ n) ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) ∘ (fun (p : α × γ) ↦ (p.1, ⟨partitionSet n p.2, partitionSet_mem n p.2⟩))) have h1 : @Measurable _ _ (MeasurableSpace.prod mα ⊤) _ @@ -217,17 +190,17 @@ lemma measurable_densityProcess_ℱ_aux (κ : kernel α (γ × β)) (ν : kernel refine h1.comp ?_ refine measurable_fst.prod_mk ?_ change @Measurable (α × γ) (MeasurableSpace.countablePartition γ n) - (MeasurableSpace.prod mα (ℱ γ n)) ⊤ + (MeasurableSpace.prod mα (partitionFiltration γ n)) ⊤ ((fun c : γ ↦ ⟨partitionSet n c, partitionSet_mem n c⟩) ∘ (fun p : α × γ ↦ p.2)) refine Measurable.comp ?_ measurable_snd - exact measurable_aux n ⊤ + exact measurable_partitionSet_aux n ⊤ lemma measurable_densityProcess_aux (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) {s : Set β} (hs : MeasurableSet s) : Measurable (fun (p : α × γ) ↦ κ p.1 (partitionSet n p.2 ×ˢ s) / ν p.1 (partitionSet n p.2)) := by - refine Measurable.mono (measurable_densityProcess_ℱ_aux κ ν n hs) ?_ le_rfl - exact sup_le_sup le_rfl (MeasurableSpace.comap_mono ((ℱ γ).le _)) + refine Measurable.mono (measurable_densityProcess_partitionFiltration_aux κ ν n hs) ?_ le_rfl + exact sup_le_sup le_rfl (MeasurableSpace.comap_mono ((partitionFiltration γ).le _)) lemma measurable_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) {s : Set β} (hs : MeasurableSet s) : @@ -244,21 +217,21 @@ lemma measurable_densityProcess_right (κ : kernel α (γ × β)) (ν : kernel Measurable (fun t ↦ densityProcess κ ν n a t s) := (measurable_densityProcess κ ν n hs).comp (measurable_const.prod_mk measurable_id) -lemma measurable_ℱ_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) +lemma measurable_partitionFiltration_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : - Measurable[ℱ γ n] (fun t ↦ densityProcess κ ν n a t s) := by - refine @Measurable.ennreal_toReal _ (ℱ γ n) _ ?_ - exact (measurable_densityProcess_ℱ_aux κ ν n hs).comp measurable_prod_mk_left + Measurable[partitionFiltration γ n] (fun t ↦ densityProcess κ ν n a t s) := by + refine @Measurable.ennreal_toReal _ (partitionFiltration γ n) _ ?_ + exact (measurable_densityProcess_partitionFiltration_aux κ ν n hs).comp measurable_prod_mk_left -lemma stronglyMeasurable_ℱ_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) - (a : α) {s : Set β} (hs : MeasurableSet s) : - StronglyMeasurable[ℱ γ n] (fun t ↦ densityProcess κ ν n a t s) := - (measurable_ℱ_densityProcess κ ν n a hs).stronglyMeasurable +lemma stronglyMeasurable_partitionFiltration_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) + (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : + StronglyMeasurable[partitionFiltration γ n] (fun t ↦ densityProcess κ ν n a t s) := + (measurable_partitionFiltration_densityProcess κ ν n a hs).stronglyMeasurable lemma adapted_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) - {s : Set β} (hs : MeasurableSet s) : - Adapted (ℱ γ) (fun n t ↦ densityProcess κ ν n a t s) := - fun n ↦ stronglyMeasurable_ℱ_densityProcess κ ν n a hs + {s : Set β} (hs : MeasurableSet s) : + Adapted (partitionFiltration γ) (fun n t ↦ densityProcess κ ν n a t s) := + fun n ↦ stronglyMeasurable_partitionFiltration_densityProcess κ ν n a hs lemma densityProcess_nonneg (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (t : γ) (s : Set β) : @@ -364,9 +337,10 @@ lemma integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] rw [set_integral_densityProcess_I hκν _ _ hs k.prop] lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet[ℱ γ n] A) : + (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} + (hA : MeasurableSet[partitionFiltration γ n] A) : ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - refine MeasurableSpace.induction_on_inter (m := ℱ γ n) + refine MeasurableSpace.induction_on_inter (m := partitionFiltration γ n) (s := MeasurableSpace.countablePartition γ n) (C := fun A ↦ ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) rfl ?_ ?_ ?_ ?_ ?_ hA @@ -379,7 +353,7 @@ lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel · rintro u hu rw [set_integral_densityProcess_I hκν _ _ hs hu] · intro A hA hA_eq - have hA' : MeasurableSet A := (ℱ γ).le _ _ hA + have hA' : MeasurableSet A := (partitionFiltration γ).le _ _ hA have h := integral_add_compl hA' (integrable_densityProcess hκν n a hs) rw [hA_eq, integral_densityProcess hκν n a hs] at h have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by @@ -392,10 +366,11 @@ lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] exact measure_mono (by intro x; simp) · intro f hf_disj hf h_eq - rw [integral_iUnion (fun i ↦ (ℱ γ).le n _ (hf i)) hf_disj + rw [integral_iUnion (fun i ↦ (partitionFiltration γ).le n _ (hf i)) hf_disj (integrable_densityProcess hκν _ _ hs).integrableOn] simp_rw [h_eq] - rw [iUnion_prod_const, measure_iUnion _ (fun i ↦ ((ℱ γ).le n _ (hf i)).prod hs)] + rw [iUnion_prod_const, + measure_iUnion _ (fun i ↦ ((partitionFiltration γ).le n _ (hf i)).prod hs)] · rw [ENNReal.tsum_toReal_eq] exact fun _ ↦ measure_ne_top _ _ · intro i j hij @@ -404,13 +379,13 @@ lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel lemma set_integral_densityProcess_of_le (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] {n m : ℕ} (hnm : n ≤ m) (a : α) {s : Set β} (hs : MeasurableSet s) - {A : Set γ} (hA : MeasurableSet[ℱ γ n] A) : + {A : Set γ} (hA : MeasurableSet[partitionFiltration γ n] A) : ∫ t in A, densityProcess κ ν m a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := - set_integral_densityProcess hκν m a hs ((ℱ γ).mono hnm A hA) + set_integral_densityProcess hκν m a hs ((partitionFiltration γ).mono hnm A hA) lemma condexp_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] {i j : ℕ} (hij : i ≤ j) (a : α) {s : Set β} (hs : MeasurableSet s) : - (ν a)[fun t ↦ densityProcess κ ν j a t s | ℱ γ i] + (ν a)[fun t ↦ densityProcess κ ν j a t s | partitionFiltration γ i] =ᵐ[ν a] fun t ↦ densityProcess κ ν i a t s := by symm refine ae_eq_condexp_of_forall_set_integral_eq ?_ ?_ ?_ ?_ ?_ @@ -420,11 +395,12 @@ lemma condexp_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] · intro t ht _ rw [set_integral_densityProcess hκν i a hs ht, set_integral_densityProcess_of_le hκν hij a hs ht] - · exact StronglyMeasurable.aeStronglyMeasurable' (stronglyMeasurable_ℱ_densityProcess κ ν i a hs) + · exact StronglyMeasurable.aeStronglyMeasurable' + (stronglyMeasurable_partitionFiltration_densityProcess κ ν i a hs) lemma martingale_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Martingale (fun n t ↦ densityProcess κ ν n a t s) (ℱ γ) (ν a) := + Martingale (fun n t ↦ densityProcess κ ν n a t s) (partitionFiltration γ) (ν a) := ⟨adapted_densityProcess κ ν a hs, fun _ _ h ↦ condexp_densityProcess hκν h a hs⟩ lemma densityProcess_mono_set (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) @@ -530,7 +506,8 @@ lemma tendsto_densityProcess_atTop_of_antitone (κ : kernel α (γ × β)) (ν : lemma tendsto_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : ∀ᵐ t ∂(ν a), Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop - (𝓝 ((ℱ γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a) t)) := by + (𝓝 ((partitionFiltration γ).limitProcess + (fun n t ↦ densityProcess κ ν n a t s) (ν a) t)) := by refine Submartingale.ae_tendsto_limitProcess (martingale_densityProcess hκν a hs).submartingale (R := (ν a univ).toNNReal) (fun n ↦ ?_) refine (snorm_densityProcess_le hκν n a s).trans_eq ?_ @@ -539,7 +516,8 @@ lemma tendsto_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFini lemma limitProcess_mem_L1 (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Memℒp ((ℱ γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 (ν a) := by + Memℒp ((partitionFiltration γ).limitProcess + (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 (ν a) := by refine Submartingale.memℒp_limitProcess (martingale_densityProcess hκν a hs).submartingale (R := (ν a univ).toNNReal) (fun n ↦ ?_) refine (snorm_densityProcess_le hκν n a s).trans_eq ?_ @@ -549,8 +527,8 @@ lemma limitProcess_mem_L1 (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [Is lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : Tendsto (fun n ↦ snorm ((fun t ↦ densityProcess κ ν n a t s) - - (ℱ γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 (ν a)) - atTop (𝓝 0) := by + - (partitionFiltration γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) + 1 (ν a)) atTop (𝓝 0) := by refine Submartingale.tendsto_snorm_one_limitProcess ?_ ?_ · exact (martingale_densityProcess hκν a hs).submartingale · refine uniformIntegrable_of le_rfl ENNReal.one_ne_top ?_ ?_ @@ -571,98 +549,100 @@ lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : kernel.fst κ ≤ lemma tendsto_snorm_one_restrict_densityProcess_limitProcess [IsFiniteKernel κ] [IsFiniteKernel ν] (hκν : kernel.fst κ ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : Tendsto (fun n ↦ snorm ((fun t ↦ densityProcess κ ν n a t s) - - (ℱ γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 ((ν a).restrict A)) - atTop (𝓝 0) := + - (partitionFiltration γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) + 1 ((ν a).restrict A)) atTop (𝓝 0) := tendsto_snorm_restrict_zero (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) A noncomputable -def MLimsup (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) (t : γ) (s : Set β) : ℝ := +def kernel.density (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) (t : γ) (s : Set β) : ℝ := limsup (fun n ↦ densityProcess κ ν n a t s) atTop -lemma mLimsup_ae_eq_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma density_ae_eq_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - (fun t ↦ MLimsup κ ν a t s) - =ᵐ[ν a] (ℱ γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a) := by + (fun t ↦ kernel.density κ ν a t s) + =ᵐ[ν a] (partitionFiltration γ).limitProcess + (fun n t ↦ densityProcess κ ν n a t s) (ν a) := by filter_upwards [tendsto_densityProcess_limitProcess hκν a hs] with t ht using ht.limsup_eq -lemma tendsto_m_mLimsup (hκν : kernel.fst κ ≤ ν) (a : α) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma tendsto_m_density (hκν : kernel.fst κ ≤ ν) (a : α) [IsFiniteKernel κ] [IsFiniteKernel ν] {s : Set β} (hs : MeasurableSet s) : ∀ᵐ t ∂(ν a), - Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop (𝓝 (MLimsup κ ν a t s)) := by - filter_upwards [tendsto_densityProcess_limitProcess hκν a hs, mLimsup_ae_eq_limitProcess hκν a hs] + Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop (𝓝 (kernel.density κ ν a t s)) := by + filter_upwards [tendsto_densityProcess_limitProcess hκν a hs, density_ae_eq_limitProcess hκν a hs] with t h1 h2 using h2 ▸ h1 -lemma measurable_mLimsup (κ : kernel α (γ × β)) (ν : kernel α γ) +lemma measurable_density (κ : kernel α (γ × β)) (ν : kernel α γ) {s : Set β} (hs : MeasurableSet s) : - Measurable (fun (p : α × γ) ↦ MLimsup κ ν p.1 p.2 s) := + Measurable (fun (p : α × γ) ↦ kernel.density κ ν p.1 p.2 s) := measurable_limsup (fun n ↦ measurable_densityProcess κ ν n hs) -lemma measurable_mLimsup_left (κ : kernel α (γ × β)) (ν : kernel α γ) (t : γ) +lemma measurable_density_left (κ : kernel α (γ × β)) (ν : kernel α γ) (t : γ) {s : Set β} (hs : MeasurableSet s) : - Measurable (fun a ↦ MLimsup κ ν a t s) := by - change Measurable ((fun (p : α × γ) ↦ MLimsup κ ν p.1 p.2 s) ∘ (fun a ↦ (a, t))) - exact (measurable_mLimsup κ ν hs).comp measurable_prod_mk_right + Measurable (fun a ↦ kernel.density κ ν a t s) := by + change Measurable ((fun (p : α × γ) ↦ kernel.density κ ν p.1 p.2 s) ∘ (fun a ↦ (a, t))) + exact (measurable_density κ ν hs).comp measurable_prod_mk_right -lemma measurable_mLimsup_right (κ : kernel α (γ × β)) (ν : kernel α γ) +lemma measurable_density_right (κ : kernel α (γ × β)) (ν : kernel α γ) {s : Set β} (hs : MeasurableSet s) (a : α) : - Measurable (fun t ↦ MLimsup κ ν a t s) := by - change Measurable ((fun (p : α × γ) ↦ MLimsup κ ν p.1 p.2 s) ∘ (fun t ↦ (a, t))) - exact (measurable_mLimsup κ ν hs).comp measurable_prod_mk_left + Measurable (fun t ↦ kernel.density κ ν a t s) := by + change Measurable ((fun (p : α × γ) ↦ kernel.density κ ν p.1 p.2 s) ∘ (fun t ↦ (a, t))) + exact (measurable_density κ ν hs).comp measurable_prod_mk_left -lemma mLimsup_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) {s s' : Set β} (h : s ⊆ s') : - MLimsup κ ν a t s ≤ MLimsup κ ν a t s' := by +lemma density_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) {s s' : Set β} (h : s ⊆ s') : + kernel.density κ ν a t s ≤ kernel.density κ ν a t s' := by refine limsup_le_limsup ?_ ?_ ?_ · exact eventually_of_forall (fun n ↦ densityProcess_mono_set hκν n a t h) · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ -lemma mLimsup_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : - 0 ≤ MLimsup κ ν a t s := by +lemma density_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : + 0 ≤ kernel.density κ ν a t s := by refine le_limsup_of_frequently_le ?_ ?_ · exact frequently_of_forall (fun n ↦ densityProcess_nonneg _ _ _ _ _ _) · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ -lemma mLimsup_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : - MLimsup κ ν a t s ≤ 1 := by +lemma density_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : + kernel.density κ ν a t s ≤ 1 := by refine limsup_le_of_le ?_ ?_ · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) · exact eventually_of_forall (fun n ↦ densityProcess_le_one hκν _ _ _ _) -lemma snorm_mLimsup_le (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) : - snorm (fun t ↦ MLimsup κ ν a t s) 1 (ν a) ≤ ν a univ := by +lemma snorm_density_le (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) : + snorm (fun t ↦ kernel.density κ ν a t s) 1 (ν a) ≤ ν a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ - · simp only [Real.norm_eq_abs, abs_of_nonneg (mLimsup_nonneg hκν a t s), - mLimsup_le_one hκν a t s] + · simp only [Real.norm_eq_abs, abs_of_nonneg (density_nonneg hκν a t s), + density_le_one hκν a t s] · simp -lemma integrable_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] +lemma integrable_density (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Integrable (fun t ↦ MLimsup κ ν a t s) (ν a) := by + Integrable (fun t ↦ kernel.density κ ν a t s) (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ - · exact measurable_mLimsup_right κ ν hs a - · exact (snorm_mLimsup_le hκν a s).trans_lt (measure_lt_top _ _) + · exact measurable_density_right κ ν hs a + · exact (snorm_density_le hκν a s).trans_lt (measure_lt_top _ _) lemma tendsto_set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : Tendsto (fun i ↦ ∫ x in A, densityProcess κ ν i a x s ∂(ν a)) atTop - (𝓝 (∫ x in A, MLimsup κ ν a x s ∂(ν a))) := by - refine tendsto_set_integral_of_L1' (μ := ν a) (fun t ↦ MLimsup κ ν a t s) - (integrable_mLimsup hκν a hs) (F := fun i t ↦ densityProcess κ ν i a t s) (l := atTop) + (𝓝 (∫ x in A, kernel.density κ ν a x s ∂(ν a))) := by + refine tendsto_set_integral_of_L1' (μ := ν a) (fun t ↦ kernel.density κ ν a t s) + (integrable_density hκν a hs) (F := fun i t ↦ densityProcess κ ν i a t s) (l := atTop) (eventually_of_forall (fun n ↦ integrable_densityProcess hκν _ _ hs)) ?_ A refine (tendsto_congr fun n ↦ ?_).mp (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) refine snorm_congr_ae ?_ - exact EventuallyEq.rfl.sub (mLimsup_ae_eq_limitProcess hκν a hs).symm - -/-- Auxiliary lemma for `set_integral_mLimsup`. -/ -lemma set_integral_mLimsup_of_measurableSet (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] - [IsFiniteKernel ν] - (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet[ℱ γ n] A) : - ∫ t in A, MLimsup κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - suffices ∫ t in A, MLimsup κ ν a t s ∂(ν a) = ∫ t in A, densityProcess κ ν n a t s ∂(ν a) by + exact EventuallyEq.rfl.sub (density_ae_eq_limitProcess hκν a hs).symm + +/-- Auxiliary lemma for `set_integral_density`. -/ +lemma set_integral_density_of_measurableSet (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} + (hA : MeasurableSet[partitionFiltration γ n] A) : + ∫ t in A, kernel.density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + suffices ∫ t in A, kernel.density κ ν a t s ∂(ν a) + = ∫ t in A, densityProcess κ ν n a t s ∂(ν a) by rw [this] exact set_integral_densityProcess hκν _ _ hs hA - suffices ∫ t in A, MLimsup κ ν a t s ∂(ν a) + suffices ∫ t in A, kernel.density κ ν a t s ∂(ν a) = limsup (fun i ↦ ∫ t in A, densityProcess κ ν i a t s ∂(ν a)) atTop by rw [this, ← limsup_const (α := ℕ) (f := atTop) (∫ t in A, densityProcess κ ν n a t s ∂(ν a)), limsup_congr] @@ -674,27 +654,28 @@ lemma set_integral_mLimsup_of_measurableSet (hκν : kernel.fst κ ≤ ν) [IsFi have h := tendsto_set_integral_m hκν a hs A rw [h.limsup_eq] -lemma integral_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma integral_density (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - ∫ t, MLimsup κ ν a t s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by - rw [← integral_univ, set_integral_mLimsup_of_measurableSet hκν 0 a hs MeasurableSet.univ] + ∫ t, kernel.density κ ν a t s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by + rw [← integral_univ, set_integral_density_of_measurableSet hκν 0 a hs MeasurableSet.univ] -lemma set_integral_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma set_integral_density (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet A) : - ∫ t in A, MLimsup κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - have hA' : MeasurableSet[⨆ n, ℱ γ n] A := by rwa [iSup_ℱ] - refine MeasurableSpace.induction_on_inter (m := ⨆ n, ℱ γ n) - (C := fun A ↦ ∫ t in A, MLimsup κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) - (MeasurableSpace.measurableSpace_iSup_eq (ℱ γ)) ?_ ?_ ?_ ?_ ?_ hA' + ∫ t in A, kernel.density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + have hA' : MeasurableSet[⨆ n, partitionFiltration γ n] A := by rwa [iSup_partitionFiltration] + refine MeasurableSpace.induction_on_inter (m := ⨆ n, partitionFiltration γ n) + (C := fun A ↦ ∫ t in A, kernel.density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) + (MeasurableSpace.measurableSpace_iSup_eq (partitionFiltration γ)) ?_ ?_ ?_ ?_ ?_ hA' · rintro s ⟨n, hs⟩ t ⟨m, ht⟩ _ - exact ⟨max n m, ((ℱ γ).mono (le_max_left n m) _ hs).inter ((ℱ γ).mono (le_max_right n m) _ ht)⟩ + exact ⟨max n m, ((partitionFiltration γ).mono (le_max_left n m) _ hs).inter + ((partitionFiltration γ).mono (le_max_right n m) _ ht)⟩ · simp · intro A ⟨n, hA⟩ - exact set_integral_mLimsup_of_measurableSet hκν n a hs hA + exact set_integral_density_of_measurableSet hκν n a hs hA · intro A hA hA_eq - rw [iSup_ℱ] at hA - have h := integral_add_compl hA (integrable_mLimsup hκν a hs) - rw [hA_eq, integral_mLimsup hκν a hs] at h + rw [iSup_partitionFiltration] at hA + have h := integral_add_compl hA (integrable_density hκν a hs) + rw [hA_eq, integral_density hκν a hs] at h have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by rw [prod_diff_prod, compl_eq_univ_diff] simp @@ -705,7 +686,7 @@ lemma set_integral_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [I · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] exact measure_mono (by intro x; simp) · intro f hf_disj hf h_eq - rw [integral_iUnion _ hf_disj (integrable_mLimsup hκν _ hs).integrableOn] + rw [integral_iUnion _ hf_disj (integrable_density hκν _ hs).integrableOn] · simp_rw [h_eq] rw [← ENNReal.tsum_toReal_eq (fun _ ↦ measure_ne_top _ _)] congr @@ -713,16 +694,16 @@ lemma set_integral_mLimsup (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [I · intro i j hij rw [Function.onFun, Set.disjoint_prod] exact Or.inl (hf_disj hij) - · rw [iSup_ℱ] at hf + · rw [iSup_partitionFiltration] at hf exact fun i ↦ (hf i).prod hs - · rwa [iSup_ℱ] at hf + · rwa [iSup_partitionFiltration] at hf -lemma tendsto_integral_mLimsup_of_monotone (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] +lemma tendsto_integral_density_of_monotone (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : - Tendsto (fun m ↦ ∫ t, MLimsup κ ν a t (s m) ∂(ν a)) atTop (𝓝 (κ a univ).toReal) := by - simp_rw [integral_mLimsup hκν a (hs_meas _)] + Tendsto (fun m ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop (𝓝 (κ a univ).toReal) := by + simp_rw [integral_density hκν a (hs_meas _)] have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := κ a univ) ?_ swap · rw [mem_nhds_iff] @@ -737,12 +718,12 @@ lemma tendsto_integral_mLimsup_of_monotone (hκν : kernel.fst κ ≤ ν) [IsFin rw [← prod_iUnion, hs_iUnion] simp only [univ_prod_univ, measure_univ] -lemma tendsto_integral_mLimsup_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] +lemma tendsto_integral_density_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : - Tendsto (fun m ↦ ∫ t, MLimsup κ ν a t (s m) ∂(ν a)) atTop (𝓝 0) := by - simp_rw [integral_mLimsup hκν a (hs_meas _)] + Tendsto (fun m ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop (𝓝 0) := by + simp_rw [integral_density hκν a (hs_meas _)] rw [← ENNReal.zero_toReal] have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := 0) ?_ swap @@ -760,37 +741,37 @@ lemma tendsto_integral_mLimsup_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin rw [← prod_iInter, hs_iInter] simp only [ne_eq, prod_empty, OuterMeasure.empty', forall_exists_index] -lemma tendsto_mLimsup_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] +lemma tendsto_density_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : - ∀ᵐ t ∂(ν a), Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 0) := by - have h_anti : ∀ t, Antitone (fun m ↦ MLimsup κ ν a t (s m)) := - fun t n m hnm ↦ mLimsup_mono_set hκν a t (hs hnm) - have h_le_one : ∀ m t, MLimsup κ ν a t (s m) ≤ 1 := fun m t ↦ mLimsup_le_one hκν a t (s m) - -- for all `t`, `fun m ↦ MLimsup κ a (s m) t` has a limit - have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 l) := by + ∀ᵐ t ∂(ν a), Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 0) := by + have h_anti : ∀ t, Antitone (fun m ↦ kernel.density κ ν a t (s m)) := + fun t n m hnm ↦ density_mono_set hκν a t (hs hnm) + have h_le_one : ∀ m t, kernel.density κ ν a t (s m) ≤ 1 := fun m t ↦ density_le_one hκν a t (s m) + -- for all `t`, `fun m ↦ density κ a (s m) t` has a limit + have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 l) := by intro t - have h_tendsto : Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop atBot ∨ - ∃ l, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 l) := + have h_tendsto : Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop atBot ∨ + ∃ l, Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 l) := tendsto_of_antitone (h_anti t) cases' h_tendsto with h_absurd h_tendsto · rw [tendsto_atTop_atBot_iff_of_antitone (h_anti t)] at h_absurd obtain ⟨r, hr⟩ := h_absurd (-1) - have h_nonneg := mLimsup_nonneg hκν a t (s r) + have h_nonneg := density_nonneg hκν a t (s r) linarith · exact h_tendsto - -- let `F` be the pointwise limit of `fun m ↦ MLimsup κ a (s m) t` for all `t` + -- let `F` be the pointwise limit of `fun m ↦ density κ a (s m) t` for all `t` let F : γ → ℝ := fun t ↦ (h_exists t).choose - have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 (F t)) := + have hF_tendsto : ∀ t, Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 (F t)) := fun t ↦ (h_exists t).choose_spec have hF_nonneg : ∀ t, 0 ≤ F t := - fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ mLimsup_nonneg hκν a t (s m)) + fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ density_nonneg hκν a t (s m)) have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) have hF_int : Integrable F (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨?_, ?_⟩ · refine aestronglyMeasurable_of_tendsto_ae atTop (fun n ↦ ?_) (ae_of_all _ hF_tendsto) - exact (measurable_mLimsup_right κ ν (hs_meas _) a).aestronglyMeasurable + exact (measurable_density_right κ ν (hs_meas _) a).aestronglyMeasurable · rw [snorm_one_eq_lintegral_nnnorm] calc ∫⁻ x, ‖F x‖₊ ∂(ν a) ≤ ∫⁻ _, 1 ∂(ν a) := by refine lintegral_mono (fun x ↦ ?_) @@ -809,16 +790,16 @@ lemma tendsto_mLimsup_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin -- `F` is 0 a.e. suffices ∀ᵐ (t : γ) ∂(ν a), 0 = F t by filter_upwards [this] with a ha; simp [ha] refine ae_eq_of_integral_eq_of_ae_le (integrable_const _) hF_int (ae_of_all _ hF_nonneg) ?_ - have h_integral : - Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ t, F t ∂(ν a))) := by + have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop + (𝓝 (∫ t, F t ∂(ν a))) := by refine integral_tendsto_of_tendsto_of_antitone ?_ hF_int ?_ ?_ - · exact fun n ↦ integrable_mLimsup hκν _ (hs_meas n) + · exact fun n ↦ integrable_density hκν _ (hs_meas n) · exact ae_of_all _ h_anti · exact ae_of_all _ hF_tendsto - have h_integral' : - Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ _, 0 ∂(ν a))) := by + have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop + (𝓝 (∫ _, 0 ∂(ν a))) := by simp only [integral_zero] - exact tendsto_integral_mLimsup_of_antitone hκν a s hs hs_iInter hs_meas + exact tendsto_integral_density_of_antitone hκν a s hs hs_iInter hs_meas exact (tendsto_nhds_unique h_integral h_integral').symm section UnivFst @@ -917,45 +898,47 @@ lemma tendsto_densityProcess_atTop_ae_of_monotone (κ : kernel α (γ × β)) [I rw [← ht] exact tendsto_densityProcess_atTop_univ_of_monotone κ n a t s hs hs_iUnion -lemma mLimsup_univ (κ : kernel α (γ × β)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), MLimsup κ (kernel.fst κ) a t univ = 1 := by +lemma density_univ (κ : kernel α (γ × β)) [IsFiniteKernel κ] (a : α) : + ∀ᵐ t ∂(kernel.fst κ a), kernel.density κ (kernel.fst κ) a t univ = 1 := by have h := fun n ↦ densityProcess_univ_ae κ n a rw [← ae_all_iff] at h filter_upwards [h] with t ht - rw [MLimsup] + rw [kernel.density] simp [ht] -lemma tendsto_mLimsup_atTop_ae_of_monotone [IsFiniteKernel κ] +lemma tendsto_density_atTop_ae_of_monotone [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : - ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun m ↦ MLimsup κ (kernel.fst κ) a t (s m)) atTop (𝓝 1) := by + ∀ᵐ t ∂(kernel.fst κ a), + Tendsto (fun m ↦ kernel.density κ (kernel.fst κ) a t (s m)) atTop (𝓝 1) := by let ν := kernel.fst κ - have h_mono : ∀ t, Monotone (fun m ↦ MLimsup κ (kernel.fst κ) a t (s m)) := - fun t n m hnm ↦ mLimsup_mono_set le_rfl a t (hs hnm) - have h_le_one : ∀ m t, MLimsup κ ν a t (s m) ≤ 1 := fun m t ↦ mLimsup_le_one le_rfl a t (s m) - -- for all `t`, `fun m ↦ MLimsup κ a (s m) t` has a limit - have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 l) := by + have h_mono : ∀ t, Monotone (fun m ↦ kernel.density κ (kernel.fst κ) a t (s m)) := + fun t n m hnm ↦ density_mono_set le_rfl a t (hs hnm) + have h_le_one : ∀ m t, kernel.density κ ν a t (s m) ≤ 1 := + fun m t ↦ density_le_one le_rfl a t (s m) + -- for all `t`, `fun m ↦ density κ a (s m) t` has a limit + have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 l) := by intro t - have h_tendsto : Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop atTop ∨ - ∃ l, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 l) := + have h_tendsto : Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop atTop ∨ + ∃ l, Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 l) := tendsto_of_monotone (h_mono t) cases' h_tendsto with h_absurd h_tendsto · rw [tendsto_atTop_atTop_iff_of_monotone (h_mono t)] at h_absurd obtain ⟨r, hr⟩ := h_absurd 2 exact absurd (hr.trans (h_le_one r t)) one_lt_two.not_le · exact h_tendsto - -- let `F` be the pointwise limit of `fun m ↦ MLimsup κ a (s m) t` for all `t` + -- let `F` be the pointwise limit of `fun m ↦ density κ a (s m) t` for all `t` let F : γ → ℝ := fun t ↦ (h_exists t).choose - have hF_tendsto : ∀ t, Tendsto (fun m ↦ MLimsup κ ν a t (s m)) atTop (𝓝 (F t)) := + have hF_tendsto : ∀ t, Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 (F t)) := fun t ↦ (h_exists t).choose_spec have hF_nonneg : ∀ t, 0 ≤ F t := - fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ mLimsup_nonneg le_rfl a t (s m)) + fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ density_nonneg le_rfl a t (s m)) have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) have hF_int : Integrable F (ν a) := by rw [← memℒp_one_iff_integrable] constructor · refine aestronglyMeasurable_of_tendsto_ae atTop (fun n ↦ ?_) (ae_of_all _ hF_tendsto) - exact (measurable_mLimsup_right κ ν (hs_meas _) a).aestronglyMeasurable + exact (measurable_density_right κ ν (hs_meas _) a).aestronglyMeasurable · rw [snorm_one_eq_lintegral_nnnorm] calc ∫⁻ x, ‖F x‖₊ ∂(ν a) ≤ ∫⁻ _, 1 ∂(ν a) := by refine lintegral_mono (fun x ↦ ?_) @@ -971,18 +954,18 @@ lemma tendsto_mLimsup_atTop_ae_of_monotone [IsFiniteKernel κ] -- since `F` is at most 1, proving that its integral is the same as the integral of 1 will tell -- us that `F` is 1 a.e. refine ae_eq_of_integral_eq_of_ae_le hF_int (integrable_const _) (ae_of_all _ hF_le_one) ?_ - have h_integral : - Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ t, F t ∂(ν a))) := by + have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop + (𝓝 (∫ t, F t ∂(ν a))) := by refine integral_tendsto_of_tendsto_of_monotone ?_ hF_int ?_ ?_ - · exact fun n ↦ integrable_mLimsup le_rfl _ (hs_meas n) + · exact fun n ↦ integrable_density le_rfl _ (hs_meas n) · exact ae_of_all _ h_mono · exact ae_of_all _ hF_tendsto - have h_integral' : - Tendsto (fun m : ℕ ↦ ∫ t, MLimsup κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ _, 1 ∂(ν a))) := by + have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop + (𝓝 (∫ _, 1 ∂(ν a))) := by rw [integral_const] simp only [smul_eq_mul, mul_one] rw [kernel.fst_apply' _ _ MeasurableSet.univ] - exact tendsto_integral_mLimsup_of_monotone le_rfl a s hs hs_iUnion hs_meas + exact tendsto_integral_density_of_monotone le_rfl a s hs hs_iUnion hs_meas exact tendsto_nhds_unique h_integral h_integral' end UnivFst diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean index 924ce6b6fd58a..621bd5fcc6954 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean @@ -35,46 +35,47 @@ variable {α γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} [MeasurableSpace.CountablyGenerated γ] {κ : kernel α (γ × ℝ)} {ν : kernel α γ} noncomputable -def mLimsupIic (κ : kernel α (γ × ℝ)) (ν : kernel α γ) (a : α) (t : γ) (q : ℚ) : ℝ := - MLimsup κ ν a t (Set.Iic q) +def kernel.densityIic (κ : kernel α (γ × ℝ)) (ν : kernel α γ) (a : α) (t : γ) (q : ℚ) : ℝ := + kernel.density κ ν a t (Set.Iic q) -lemma measurable_mLimsupIic (κ : kernel α (γ × ℝ)) (ν : kernel α γ) : - Measurable (fun p : α × γ ↦ mLimsupIic κ ν p.1 p.2) := by +lemma measurable_densityIic (κ : kernel α (γ × ℝ)) (ν : kernel α γ) : + Measurable (fun p : α × γ ↦ kernel.densityIic κ ν p.1 p.2) := by rw [measurable_pi_iff] - exact fun _ ↦ measurable_mLimsup κ ν measurableSet_Iic + exact fun _ ↦ measurable_density κ ν measurableSet_Iic -lemma measurable_mLimsupIic_right (κ : kernel α (γ × ℝ)) (ν : kernel α γ) (a : α) (q : ℚ) : - Measurable (fun t ↦ mLimsupIic κ ν a t q) := - measurable_mLimsup_right _ _ measurableSet_Iic _ +lemma measurable_densityIic_right (κ : kernel α (γ × ℝ)) (ν : kernel α γ) (a : α) (q : ℚ) : + Measurable (fun t ↦ kernel.densityIic κ ν a t q) := + measurable_density_right _ _ measurableSet_Iic _ -lemma monotone_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) : - Monotone (mLimsupIic κ ν a t) := by +lemma monotone_densityIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) : + Monotone (kernel.densityIic κ ν a t) := by intro q r hqr - rw [mLimsupIic, mLimsupIic] - refine mLimsup_mono_set hκν a t ?_ + simp_rw [kernel.densityIic] + refine density_mono_set hκν a t ?_ refine Iic_subset_Iic.mpr ?_ exact_mod_cast hqr -lemma mLimsupIic_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (q : ℚ) : - 0 ≤ mLimsupIic κ ν a t q := - mLimsup_nonneg hκν a t _ +lemma densityIic_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (q : ℚ) : + 0 ≤ kernel.densityIic κ ν a t q := + density_nonneg hκν a t _ -lemma mLimsupIic_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (q : ℚ) : - mLimsupIic κ ν a t q ≤ 1 := - mLimsup_le_one hκν a t _ +lemma densityIic_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (q : ℚ) : + kernel.densityIic κ ν a t q ≤ 1 := + density_le_one hκν a t _ -lemma tendsto_atTop_mLimsupIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), Tendsto (fun q ↦ mLimsupIic κ (kernel.fst κ) a t q) atTop (𝓝 1) := by +lemma tendsto_atTop_densityIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] (a : α) : + ∀ᵐ t ∂(kernel.fst κ a), + Tendsto (fun q ↦ kernel.densityIic κ (kernel.fst κ) a t q) atTop (𝓝 1) := by let ν := kernel.fst κ - suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ ν a t n) atTop (𝓝 1) by + suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ kernel.densityIic κ ν a t n) atTop (𝓝 1) by filter_upwards [this] with t ht - let f := fun q : ℚ ↦ mLimsupIic κ ν a t ⌊q⌋₊ - let g := fun q : ℚ ↦ mLimsupIic κ ν a t ⌈q⌉₊ - have hf_le : ∀ᶠ q in atTop, f q ≤ mLimsupIic κ ν a t q := by + let f := fun q : ℚ ↦ kernel.densityIic κ ν a t ⌊q⌋₊ + let g := fun q : ℚ ↦ kernel.densityIic κ ν a t ⌈q⌉₊ + have hf_le : ∀ᶠ q in atTop, f q ≤ kernel.densityIic κ ν a t q := by simp only [eventually_atTop, ge_iff_le] - exact ⟨0, fun q hq ↦ monotone_mLimsupIic le_rfl a t (Nat.floor_le hq)⟩ - have hg_le : ∀ q, mLimsupIic κ ν a t q ≤ g q := - fun q ↦ monotone_mLimsupIic le_rfl a t (Nat.le_ceil _) + exact ⟨0, fun q hq ↦ monotone_densityIic le_rfl a t (Nat.floor_le hq)⟩ + have hg_le : ∀ q, kernel.densityIic κ ν a t q ≤ g q := + fun q ↦ monotone_densityIic le_rfl a t (Nat.le_ceil _) refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ hf_le (eventually_of_forall hg_le) · exact ht.comp tendsto_nat_floor_atTop · exact ht.comp tendsto_nat_ceil_atTop @@ -85,29 +86,30 @@ lemma tendsto_atTop_mLimsupIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] simp only [mem_iUnion, mem_Iic, mem_univ, iff_true] exact ⟨Nat.ceil x, Nat.le_ceil x⟩ have hs_meas : ∀ n, MeasurableSet (s n) := fun _ ↦ measurableSet_Iic - filter_upwards [tendsto_mLimsup_atTop_ae_of_monotone a s hs hs_iUnion hs_meas] + filter_upwards [tendsto_density_atTop_ae_of_monotone a s hs hs_iUnion hs_meas] with x hx using hx -lemma tendsto_atBot_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma tendsto_atBot_densityIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : - ∀ᵐ t ∂(ν a), Tendsto (fun q ↦ mLimsupIic κ ν a t q) atBot (𝓝 0) := by - suffices ∀ᵐ t ∂(ν a), Tendsto (fun q ↦ mLimsupIic κ ν a t (-q)) atTop (𝓝 0) by + ∀ᵐ t ∂(ν a), Tendsto (fun q ↦ kernel.densityIic κ ν a t q) atBot (𝓝 0) := by + suffices ∀ᵐ t ∂(ν a), Tendsto (fun q ↦ kernel.densityIic κ ν a t (-q)) atTop (𝓝 0) by filter_upwards [this] with t ht - have h_eq_neg : (fun q ↦ mLimsupIic κ ν a t q) = fun q ↦ mLimsupIic κ ν a t (- -q) := by + have h_eq_neg : (fun q ↦ kernel.densityIic κ ν a t q) + = fun q ↦ kernel.densityIic κ ν a t (- -q) := by simp_rw [neg_neg] rw [h_eq_neg] exact ht.comp tendsto_neg_atBot_atTop - suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ mLimsupIic κ ν a t (-n)) atTop (𝓝 0) by + suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ kernel.densityIic κ ν a t (-n)) atTop (𝓝 0) by filter_upwards [this] with t ht - let f := fun q : ℚ ↦ mLimsupIic κ ν a t (-⌊q⌋₊) - let g := fun q : ℚ ↦ mLimsupIic κ ν a t (-⌈q⌉₊) - have hf_le : ∀ᶠ q in atTop, mLimsupIic κ ν a t (-q) ≤ f q := by + let f := fun q : ℚ ↦ kernel.densityIic κ ν a t (-⌊q⌋₊) + let g := fun q : ℚ ↦ kernel.densityIic κ ν a t (-⌈q⌉₊) + have hf_le : ∀ᶠ q in atTop, kernel.densityIic κ ν a t (-q) ≤ f q := by simp only [eventually_atTop, ge_iff_le] - refine ⟨0, fun q hq ↦ monotone_mLimsupIic hκν a t ?_⟩ + refine ⟨0, fun q hq ↦ monotone_densityIic hκν a t ?_⟩ rw [neg_le_neg_iff] exact Nat.floor_le hq - have hg_le : ∀ q, g q ≤ mLimsupIic κ ν a t (-q) := by - refine fun q ↦ monotone_mLimsupIic hκν a t ?_ + have hg_le : ∀ q, g q ≤ kernel.densityIic κ ν a t (-q) := by + refine fun q ↦ monotone_densityIic hκν a t ?_ rw [neg_le_neg_iff] exact Nat.le_ceil _ refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ (eventually_of_forall hg_le) hf_le @@ -120,102 +122,102 @@ lemma tendsto_atBot_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, neg_lt] exact exists_nat_gt (-x) have hs_meas : ∀ n, MeasurableSet (s n) := fun _ ↦ measurableSet_Iic - convert tendsto_mLimsup_atTop_ae_of_antitone hκν a s hs hs_iInter hs_meas with x n - rw [mLimsupIic] - simp + convert tendsto_density_atTop_ae_of_antitone hκν a s hs hs_iInter hs_meas with x n + simp [kernel.densityIic] -lemma set_integral_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma set_integral_densityIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (q : ℚ) {A : Set γ} (hA : MeasurableSet A) : - ∫ t in A, mLimsupIic κ ν a t q ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := - set_integral_mLimsup hκν a measurableSet_Iic hA + ∫ t in A, kernel.densityIic κ ν a t q ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := + set_integral_density hκν a measurableSet_Iic hA -lemma integrable_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] +lemma integrable_densityIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] (a : α) (q : ℚ) : - Integrable (fun t ↦ mLimsupIic κ ν a t q) (ν a) := - integrable_mLimsup hκν _ measurableSet_Iic + Integrable (fun t ↦ kernel.densityIic κ ν a t q) (ν a) := + integrable_density hκν _ measurableSet_Iic -lemma bddBelow_range_mLimsupIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (q : ℚ) : - BddBelow (range fun (r : Ioi q) ↦ mLimsupIic κ ν a t r) := by +lemma bddBelow_range_densityIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (q : ℚ) : + BddBelow (range fun (r : Ioi q) ↦ kernel.densityIic κ ν a t r) := by refine ⟨0, ?_⟩ rw [mem_lowerBounds] rintro x ⟨y, rfl⟩ - exact mLimsupIic_nonneg hκν _ _ _ + exact densityIic_nonneg hκν _ _ _ -lemma integrable_iInf_rat_gt_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] +lemma integrable_iInf_rat_gt_densityIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] (a : α) (q : ℚ) : - Integrable (fun t ↦ ⨅ r : Ioi q, mLimsupIic κ ν a t r) (ν a) := by + Integrable (fun t ↦ ⨅ r : Ioi q, kernel.densityIic κ ν a t r) (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ - · exact measurable_iInf fun i ↦ measurable_mLimsupIic_right κ ν a i + · exact measurable_iInf fun i ↦ measurable_densityIic_right κ ν a i refine (?_ : _ ≤ (ν a univ : ℝ≥0∞)).trans_lt (measure_lt_top _ _) refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ · rw [Real.norm_eq_abs, abs_of_nonneg] · refine ciInf_le_of_le ?_ ?_ ?_ - · exact bddBelow_range_mLimsupIic hκν a t _ + · exact bddBelow_range_densityIic hκν a t _ · exact ⟨q + 1, by simp⟩ - · exact mLimsupIic_le_one hκν _ _ _ - · exact le_ciInf fun r ↦ mLimsupIic_nonneg hκν a t r + · exact densityIic_le_one hκν _ _ _ + · exact le_ciInf fun r ↦ densityIic_nonneg hκν a t r · simp -lemma set_integral_iInf_rat_gt_mLimsupIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] - [IsFiniteKernel ν] - (a : α) (q : ℚ) {A : Set γ} (hA : MeasurableSet A) : - ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ ν a t r ∂(ν a) +lemma set_integral_iInf_rat_gt_densityIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) (q : ℚ) {A : Set γ} (hA : MeasurableSet A) : + ∫ t in A, ⨅ r : Ioi q, kernel.densityIic κ ν a t r ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by refine le_antisymm ?_ ?_ - · have h : ∀ r : Ioi q, ∫ t in A, ⨅ r' : Ioi q, mLimsupIic κ ν a t r' ∂(ν a) + · have h : ∀ r : Ioi q, ∫ t in A, ⨅ r' : Ioi q, kernel.densityIic κ ν a t r' ∂(ν a) ≤ (κ a (A ×ˢ Iic (r : ℝ))).toReal := by intro r - rw [← set_integral_mLimsupIic hκν a r hA] + rw [← set_integral_densityIic hκν a r hA] refine set_integral_mono ?_ ?_ ?_ - · exact (integrable_iInf_rat_gt_mLimsupIic hκν _ _).integrableOn - · exact (integrable_mLimsupIic hκν _ _).integrableOn - · exact fun t ↦ ciInf_le (bddBelow_range_mLimsupIic hκν _ _ _) r - calc ∫ t in A, ⨅ r : Ioi q, mLimsupIic κ ν a t r ∂(ν a) + · exact (integrable_iInf_rat_gt_densityIic hκν _ _).integrableOn + · exact (integrable_densityIic hκν _ _).integrableOn + · exact fun t ↦ ciInf_le (bddBelow_range_densityIic hκν _ _ _) r + calc ∫ t in A, ⨅ r : Ioi q, kernel.densityIic κ ν a t r ∂(ν a) ≤ ⨅ r : Ioi q, (κ a (A ×ˢ Iic (r : ℝ))).toReal := le_ciInf h _ = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by rw [← Measure.iInf_Iic_gt_prod hA q] exact (ENNReal.toReal_iInf (fun r ↦ measure_ne_top _ _)).symm - · rw [← set_integral_mLimsupIic hκν a q hA] + · rw [← set_integral_densityIic hκν a q hA] refine set_integral_mono ?_ ?_ ?_ - · exact (integrable_mLimsupIic hκν _ _).integrableOn - · exact (integrable_iInf_rat_gt_mLimsupIic hκν _ _).integrableOn - · exact fun t ↦ le_ciInf (fun r ↦ monotone_mLimsupIic hκν _ _ (le_of_lt r.prop)) + · exact (integrable_densityIic hκν _ _).integrableOn + · exact (integrable_iInf_rat_gt_densityIic hκν _ _).integrableOn + · exact fun t ↦ le_ciInf (fun r ↦ monotone_densityIic hκν _ _ (le_of_lt r.prop)) -lemma iInf_rat_gt_mLimsupIic_eq (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma iInf_rat_gt_densityIic_eq (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : - ∀ᵐ t ∂(ν a), ∀ q : ℚ, ⨅ r : Ioi q, mLimsupIic κ ν a t r = mLimsupIic κ ν a t q := by + ∀ᵐ t ∂(ν a), ∀ q : ℚ, ⨅ r : Ioi q, kernel.densityIic κ ν a t r + = kernel.densityIic κ ν a t q := by rw [ae_all_iff] refine fun q ↦ ae_eq_of_forall_set_integral_eq_of_sigmaFinite (μ := ν a) ?_ ?_ ?_ · intro s _ _ refine Integrable.integrableOn ?_ - exact integrable_iInf_rat_gt_mLimsupIic hκν _ _ - · exact fun s _ _ ↦ (integrable_mLimsupIic hκν a q).integrableOn + exact integrable_iInf_rat_gt_densityIic hκν _ _ + · exact fun s _ _ ↦ (integrable_densityIic hκν a q).integrableOn · intro s hs _ - rw [set_integral_mLimsupIic hκν _ _ hs, set_integral_iInf_rat_gt_mLimsupIic hκν _ _ hs] + rw [set_integral_densityIic hκν _ _ hs, set_integral_iInf_rat_gt_densityIic hκν _ _ hs] -lemma isRatStieltjesPoint_mLimsupIic_ae (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] (a : α) : +lemma isRatStieltjesPoint_densityIic_ae (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] (a : α) : ∀ᵐ t ∂(kernel.fst κ a), - IsRatStieltjesPoint (fun p q ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2 q) (a, t) := by - filter_upwards [tendsto_atTop_mLimsupIic κ a, tendsto_atBot_mLimsupIic le_rfl a, - iInf_rat_gt_mLimsupIic_eq le_rfl a] with t ht_top ht_bot ht_iInf - exact ⟨monotone_mLimsupIic le_rfl a t, ht_top, ht_bot, ht_iInf⟩ - -lemma isRatKernelCDF_mLimsupIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsRatKernelCDF (fun p : α × γ ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2) κ (kernel.fst κ) where - measurable := measurable_mLimsupIic κ (kernel.fst κ) - isRatStieltjesPoint_ae := isRatStieltjesPoint_mLimsupIic_ae κ - integrable := integrable_mLimsupIic le_rfl - set_integral := fun _ _ hs _ ↦ set_integral_mLimsupIic le_rfl _ _ hs + IsRatStieltjesPoint (fun p q ↦ kernel.densityIic κ (kernel.fst κ) p.1 p.2 q) (a, t) := by + filter_upwards [tendsto_atTop_densityIic κ a, tendsto_atBot_densityIic le_rfl a, + iInf_rat_gt_densityIic_eq le_rfl a] with t ht_top ht_bot ht_iInf + exact ⟨monotone_densityIic le_rfl a t, ht_top, ht_bot, ht_iInf⟩ + +lemma isRatKernelCDF_densityIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsRatKernelCDF (fun p : α × γ ↦ kernel.densityIic κ (kernel.fst κ) p.1 p.2) κ (kernel.fst κ) + where + measurable := measurable_densityIic κ (kernel.fst κ) + isRatStieltjesPoint_ae := isRatStieltjesPoint_densityIic_ae κ + integrable := integrable_densityIic le_rfl + set_integral := fun _ _ hs _ ↦ set_integral_densityIic le_rfl _ _ hs noncomputable -def mLimsupCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : +def kernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : α × γ → StieltjesFunction := - stieltjesOfMeasurableRat (fun p : α × γ ↦ mLimsupIic κ (kernel.fst κ) p.1 p.2) - (isRatKernelCDF_mLimsupIic κ).measurable + stieltjesOfMeasurableRat (fun p : α × γ ↦ kernel.densityIic κ (kernel.fst κ) p.1 p.2) + (isRatKernelCDF_densityIic κ).measurable -lemma isKernelCDF_mLimsupCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsKernelCDF (mLimsupCDF κ) κ (kernel.fst κ) := - isKernelCDF_stieltjesOfMeasurableRat (isRatKernelCDF_mLimsupIic κ) +lemma isKernelCDF_kernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsKernelCDF (kernelCDF κ) κ (kernel.fst κ) := + isKernelCDF_stieltjesOfMeasurableRat (isRatKernelCDF_densityIic κ) end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index 302346b736630..c2dceb4b32c5c 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -9,7 +9,7 @@ import Mathlib.Probability.Kernel.WithDensity /-! # Radon-Nikodym derivative and Lebesgue decompusition for kernels -Let `Ω` be a standard Borel space and `κ ν : kernel α Ω`. Then there exists... TODO +Let `γ` be a countably generated measurable space and `κ ν : kernel α γ`. Then there exists... TODO ## Main definitions @@ -43,139 +43,151 @@ open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory -variable {α : Type*} {mα : MeasurableSpace α} - -section Real - -variable {κ ν : kernel α ℝ} +variable {α γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} + [MeasurableSpace.CountablyGenerated γ] {κ ν : kernel α γ} lemma kernel.fst_map_prod_le_of_le (hκν : κ ≤ ν) : kernel.fst (kernel.map κ (fun a ↦ (a, ())) - (@measurable_prod_mk_right ℝ Unit _ inferInstance _)) ≤ ν := by + (@measurable_prod_mk_right γ Unit _ inferInstance _)) ≤ ν := by rwa [kernel.fst_map_prod _ measurable_id' measurable_const, kernel.map_id'] noncomputable -def g (κ ν : kernel α ℝ) (a : α) (x : ℝ) : ℝ := - MLimsup (kernel.map κ (fun a ↦ (a, ())) - (@measurable_prod_mk_right ℝ Unit _ inferInstance _)) ν a x univ +def kernel.rnDerivAux (κ ν : kernel α γ) (a : α) (x : γ) : ℝ := + kernel.density (kernel.map κ (fun a ↦ (a, ())) + (@measurable_prod_mk_right γ Unit _ inferInstance _)) ν a x univ -lemma g_nonneg (hκν : κ ≤ ν) {a : α} {x : ℝ} : 0 ≤ g κ ν a x := - mLimsup_nonneg (kernel.fst_map_prod_le_of_le hκν) _ _ _ +lemma rnDerivAux_nonneg (hκν : κ ≤ ν) {a : α} {x : γ} : 0 ≤ kernel.rnDerivAux κ ν a x := + density_nonneg (kernel.fst_map_prod_le_of_le hκν) _ _ _ -lemma g_le_one (hκν : κ ≤ ν) {a : α} {x : ℝ} : g κ ν a x ≤ 1 := - mLimsup_le_one (kernel.fst_map_prod_le_of_le hκν) _ _ _ +lemma rnDerivAux_le_one (hκν : κ ≤ ν) {a : α} {x : γ} : kernel.rnDerivAux κ ν a x ≤ 1 := + density_le_one (kernel.fst_map_prod_le_of_le hκν) _ _ _ -lemma measurable_g (κ : kernel α ℝ) (ν : kernel α ℝ) : - Measurable (fun p : α × ℝ ↦ g κ ν p.1 p.2) := - measurable_mLimsup _ ν MeasurableSet.univ +lemma measurable_rnDerivAux (κ ν : kernel α γ) : + Measurable (fun p : α × γ ↦ kernel.rnDerivAux κ ν p.1 p.2) := + measurable_density _ ν MeasurableSet.univ -lemma measurable_g_right (κ : kernel α ℝ) (ν : kernel α ℝ) (a : α) : - Measurable (fun x : ℝ ↦ g κ ν a x) := by - change Measurable ((fun p : α × ℝ ↦ g κ ν p.1 p.2) ∘ (fun x ↦ (a, x))) - exact (measurable_g _ _).comp measurable_prod_mk_left +lemma measurable_rnDerivAux_right (κ ν : kernel α γ) (a : α) : + Measurable (fun x : γ ↦ kernel.rnDerivAux κ ν a x) := by + change Measurable ((fun p : α × γ ↦ kernel.rnDerivAux κ ν p.1 p.2) ∘ (fun x ↦ (a, x))) + exact (measurable_rnDerivAux _ _).comp measurable_prod_mk_left noncomputable -def kernel.rnDerivReal (κ ν : kernel α ℝ) (a : α) (x : ℝ) : ℝ≥0∞ := - ENNReal.ofReal (g κ (κ + ν) a x) / ENNReal.ofReal (1 - g κ (κ + ν) a x) - -lemma kernel.rnDerivReal_def (κ ν : kernel α ℝ) : - kernel.rnDerivReal κ ν - = fun a x ↦ ENNReal.ofReal (g κ (κ + ν) a x) / ENNReal.ofReal (1 - g κ (κ + ν) a x) := rfl - -lemma measurable_rnDerivReal (κ ν : kernel α ℝ) : - Measurable (fun p : α × ℝ ↦ kernel.rnDerivReal κ ν p.1 p.2) := - (measurable_g κ _).ennreal_ofReal.div (measurable_const.sub (measurable_g κ _)).ennreal_ofReal - -lemma measurable_rnDerivReal_right (κ ν : kernel α ℝ) (a : α) : - Measurable (fun x : ℝ ↦ kernel.rnDerivReal κ ν a x) := by - change Measurable ((fun p : α × ℝ ↦ kernel.rnDerivReal κ ν p.1 p.2) ∘ (fun x ↦ (a, x))) - exact (measurable_rnDerivReal _ _).comp measurable_prod_mk_left - -lemma withDensity_g (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFiniteKernel ν] : - kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (g κ (κ + ν) a x)) = κ := by +def kernel.rnDeriv (κ ν : kernel α γ) (a : α) (x : γ) : ℝ≥0∞ := + ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) + / ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) + +lemma kernel.rnDeriv_def (κ ν : kernel α γ) : + kernel.rnDeriv κ ν = fun a x ↦ ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) + / ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) := rfl + +lemma measurable_rnDeriv (κ ν : kernel α γ) : + Measurable (fun p : α × γ ↦ kernel.rnDeriv κ ν p.1 p.2) := + (measurable_rnDerivAux κ _).ennreal_ofReal.div + (measurable_const.sub (measurable_rnDerivAux κ _)).ennreal_ofReal + +lemma measurable_rnDeriv_right (κ ν : kernel α γ) (a : α) : + Measurable (fun x : γ ↦ kernel.rnDeriv κ ν a x) := by + change Measurable ((fun p : α × γ ↦ kernel.rnDeriv κ ν p.1 p.2) ∘ (fun x ↦ (a, x))) + exact (measurable_rnDeriv _ _).comp measurable_prod_mk_left + +lemma withDensity_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : + kernel.withDensity (κ + ν) + (fun a x ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x)) = κ := by have h_le : κ ≤ κ + ν := le_add_of_nonneg_right bot_le ext a s hs rw [kernel.withDensity_apply'] - swap; exact (measurable_g _ _).ennreal_ofReal + swap; exact (measurable_rnDerivAux _ _).ennreal_ofReal have : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl simp_rw [this] rw [← ofReal_integral_eq_lintegral_ofReal] - · unfold g - rw [set_integral_mLimsup (kernel.fst_map_prod_le_of_le h_le) a MeasurableSet.univ hs, + · unfold kernel.rnDerivAux + rw [set_integral_density (kernel.fst_map_prod_le_of_le h_le) a MeasurableSet.univ hs, ENNReal.ofReal_toReal, kernel.map_apply' _ _ _ (hs.prod MeasurableSet.univ)] · congr with x simp · exact measure_ne_top _ _ - · exact (integrable_mLimsup (kernel.fst_map_prod_le_of_le h_le) a MeasurableSet.univ).restrict - · exact ae_of_all _ (fun x ↦ mLimsup_nonneg (kernel.fst_map_prod_le_of_le h_le) _ _ _) + · exact (integrable_density (kernel.fst_map_prod_le_of_le h_le) a MeasurableSet.univ).restrict + · exact ae_of_all _ (fun x ↦ density_nonneg (kernel.fst_map_prod_le_of_le h_le) _ _ _) -lemma withDensity_one_sub_g (κ ν : kernel α ℝ) [IsFiniteKernel κ] [IsFiniteKernel ν] : - kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (1 - g κ (κ + ν) a x)) = ν := by +lemma withDensity_one_sub_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : + kernel.withDensity (κ + ν) + (fun a x ↦ Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x)) = ν := by have h_le : κ ≤ κ + ν := le_add_of_nonneg_right bot_le - suffices kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (1 - g κ (κ + ν) a x)) - + kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (g κ (κ + ν) a x)) = κ + ν by + suffices kernel.withDensity (κ + ν) + (fun a x ↦ Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x)) + + kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x)) + = κ + ν by ext a s - have h : (kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (1 - g κ (κ + ν) a x)) - + kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (g κ (κ + ν) a x))) a s + have h : (kernel.withDensity (κ + ν) + (fun a x ↦ Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x)) + + kernel.withDensity (κ + ν) + (fun a x ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x))) a s = κ a s + ν a s := by rw [this] simp simp only [kernel.coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add] at h - rwa [withDensity_g, add_comm, ENNReal.add_right_inj (measure_ne_top _ _)] at h - have h_nonneg : ∀ a x, 0 ≤ g κ (κ + ν) a x := - fun _ _ ↦ mLimsup_nonneg (kernel.fst_map_prod_le_of_le h_le) _ _ _ + rwa [withDensity_rnDerivAux, add_comm, ENNReal.add_right_inj (measure_ne_top _ _)] at h + have h_nonneg : ∀ a x, 0 ≤ kernel.rnDerivAux κ (κ + ν) a x := + fun _ _ ↦ density_nonneg (kernel.fst_map_prod_le_of_le h_le) _ _ _ have : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl simp_rw [this, ENNReal.ofReal_sub _ (h_nonneg _ _), ENNReal.ofReal_one] rw [kernel.withDensity_sub_add] · rw [kernel.withDensity_one'] · exact measurable_const - · exact (measurable_g _ _).ennreal_ofReal + · exact (measurable_rnDerivAux _ _).ennreal_ofReal · refine fun a ↦ ne_of_lt ?_ - calc ∫⁻ x, ENNReal.ofReal (g κ (κ + ν) a x) ∂(κ + ν) a + calc ∫⁻ x, ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) ∂(κ + ν) a ≤ ∫⁻ _, 1 ∂(κ + ν) a := by refine lintegral_mono (fun x ↦ ?_) - simp [g_le_one (le_add_of_nonneg_right bot_le)] + simp [rnDerivAux_le_one (le_add_of_nonneg_right bot_le)] _ < ⊤ := by rw [MeasureTheory.lintegral_const, one_mul]; exact measure_lt_top _ _ · refine fun a ↦ ae_of_all _ (fun x ↦ ?_) simp only [ENNReal.ofReal_le_one] - exact mLimsup_le_one (kernel.fst_map_prod_le_of_le h_le) _ _ _ + exact density_le_one (kernel.fst_map_prod_le_of_le h_le) _ _ _ noncomputable -def kernel.singularPartReal (κ ν : kernel α ℝ) [IsSFiniteKernel κ] [IsSFiniteKernel ν] : - kernel α ℝ := - kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (g κ (κ + ν) a x) - - Real.toNNReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x) +def kernel.singularPart (κ ν : kernel α γ) [IsSFiniteKernel κ] [IsSFiniteKernel ν] : + kernel α γ := + kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x) + - Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x) * kernel.rnDeriv κ ν a x) -lemma kernel.mutuallySingular_singularPartReal (κ ν : kernel α ℝ) +lemma kernel.mutuallySingular_singularPart (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : - kernel.singularPartReal κ ν a ⟂ₘ ν a := by + kernel.singularPart κ ν a ⟂ₘ ν a := by symm have h_coe : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl - refine ⟨{x | g κ (κ + ν) a x = 1}, measurable_g_right _ _ a (measurableSet_singleton _), ?_, ?_⟩ - · suffices kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (1 - g κ (κ + ν) a x)) a - {x | g κ (κ + ν) a x = 1} = 0 by - rwa [withDensity_one_sub_g κ ν] at this + refine ⟨{x | kernel.rnDerivAux κ (κ + ν) a x = 1}, + measurable_rnDerivAux_right _ _ a (measurableSet_singleton _), ?_, ?_⟩ + · suffices kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal + (1 - kernel.rnDerivAux κ (κ + ν) a x)) a {x | kernel.rnDerivAux κ (κ + ν) a x = 1} = 0 by + rwa [withDensity_one_sub_rnDerivAux κ ν] at this simp_rw [h_coe] rw [kernel.withDensity_apply', lintegral_eq_zero_iff, EventuallyEq, ae_restrict_iff] rotate_left - · exact (measurable_const.sub ((measurable_g _ _).comp measurable_prod_mk_left)).ennreal_ofReal + · exact (measurable_const.sub + ((measurable_rnDerivAux _ _).comp measurable_prod_mk_left)).ennreal_ofReal (measurableSet_singleton _) - · exact (measurable_const.sub ((measurable_g _ _).comp measurable_prod_mk_left)).ennreal_ofReal - · exact (measurable_const.sub (measurable_g _ _)).ennreal_ofReal + · exact (measurable_const.sub + ((measurable_rnDerivAux _ _).comp measurable_prod_mk_left)).ennreal_ofReal + · exact (measurable_const.sub (measurable_rnDerivAux _ _)).ennreal_ofReal refine ae_of_all _ (fun x hx ↦ ?_) simp only [mem_setOf_eq] at hx simp [hx] - · have h_meas' : Measurable <| Function.uncurry fun a b ↦ ENNReal.ofReal (g κ (κ + ν) a b) - - ENNReal.ofReal (1 - g κ (κ + ν) a b) * rnDerivReal κ ν a b := by - refine (measurable_g _ _).ennreal_ofReal.sub (Measurable.mul ?_ (measurable_rnDerivReal _ _)) - exact (measurable_const.sub (measurable_g _ _)).ennreal_ofReal - have h_meas : Measurable fun b ↦ ENNReal.ofReal (g κ (κ + ν) a b) - - ENNReal.ofReal (1 - g κ (κ + ν) a b) * rnDerivReal κ ν a b := by - change Measurable ((Function.uncurry fun a b ↦ ENNReal.ofReal (g κ (κ + ν) a b) - - ENNReal.ofReal (1 - g κ (κ + ν) a b) * rnDerivReal κ ν a b) ∘ (fun b ↦ (a, b))) + · have h_meas' : Measurable <| Function.uncurry fun a b ↦ + ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a b) + - ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a b) * rnDeriv κ ν a b := by + refine (measurable_rnDerivAux _ _).ennreal_ofReal.sub + (Measurable.mul ?_ (measurable_rnDeriv _ _)) + exact (measurable_const.sub (measurable_rnDerivAux _ _)).ennreal_ofReal + have h_meas : Measurable fun b ↦ ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a b) + - ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a b) * rnDeriv κ ν a b := by + change Measurable ((Function.uncurry fun a b ↦ + ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a b) + - ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a b) * rnDeriv κ ν a b) + ∘ (fun b ↦ (a, b))) exact h_meas'.comp measurable_prod_mk_left - rw [kernel.singularPartReal, kernel.withDensity_apply', lintegral_eq_zero_iff, EventuallyEq, + rw [kernel.singularPart, kernel.withDensity_apply', lintegral_eq_zero_iff, EventuallyEq, ae_restrict_iff] all_goals simp_rw [h_coe] rotate_left @@ -184,139 +196,54 @@ lemma kernel.mutuallySingular_singularPartReal (κ ν : kernel α ℝ) · exact h_meas' refine ae_of_all _ (fun x hx ↦ ?_) simp only [mem_compl_iff, mem_setOf_eq] at hx - simp_rw [rnDerivReal] + simp_rw [rnDeriv] rw [← ENNReal.ofReal_div_of_pos, div_eq_inv_mul, ← ENNReal.ofReal_mul, ← mul_assoc, mul_inv_cancel, one_mul, tsub_self] · rfl · rw [ne_eq, sub_eq_zero] exact Ne.symm hx - · simp [g_le_one (le_add_of_nonneg_right bot_le)] - · simp [lt_of_le_of_ne (g_le_one (le_add_of_nonneg_right bot_le)) hx] + · simp [rnDerivAux_le_one (le_add_of_nonneg_right bot_le)] + · simp [lt_of_le_of_ne (rnDerivAux_le_one (le_add_of_nonneg_right bot_le)) hx] -lemma kernel.rnDerivReal_add_singularPartReal (κ ν : kernel α ℝ) +lemma kernel.rnDeriv_add_singularPart (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : - kernel.withDensity ν (kernel.rnDerivReal κ ν) + kernel.singularPartReal κ ν = κ := by - have : kernel.withDensity ν (kernel.rnDerivReal κ ν) + kernel.withDensity ν (kernel.rnDeriv κ ν) + kernel.singularPart κ ν = κ := by + have : kernel.withDensity ν (kernel.rnDeriv κ ν) = kernel.withDensity (kernel.withDensity (κ + ν) - (fun a x ↦ Real.toNNReal (1 - g κ (κ + ν) a x))) (kernel.rnDerivReal κ ν) := by - rw [kernel.rnDerivReal_def] + (fun a x ↦ Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x))) (kernel.rnDeriv κ ν) := by + rw [kernel.rnDeriv_def] congr - exact (withDensity_one_sub_g κ ν).symm - rw [this, kernel.singularPartReal, add_comm, ← kernel.withDensity_mul] + exact (withDensity_one_sub_rnDerivAux κ ν).symm + rw [this, kernel.singularPart, add_comm, ← kernel.withDensity_mul] rotate_left - · exact (measurable_const.sub (measurable_g _ _)).real_toNNReal - · exact measurable_rnDerivReal _ _ + · exact (measurable_const.sub (measurable_rnDerivAux _ _)).real_toNNReal + · exact measurable_rnDeriv _ _ have h_coe : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl - have h_le : ∀ a x, ENNReal.ofReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x - ≤ ENNReal.ofReal (g κ (κ + ν) a x) := by + have h_le : ∀ a x, ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) * kernel.rnDeriv κ ν a x + ≤ ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) := by intro a x - unfold kernel.rnDerivReal - by_cases hg_one : g κ (κ + ν) a x = 1 - · simp [hg_one] + unfold kernel.rnDeriv + by_cases h_one : kernel.rnDerivAux κ (κ + ν) a x = 1 + · simp [h_one] rw [← ENNReal.ofReal_div_of_pos, div_eq_inv_mul, ← ENNReal.ofReal_mul, ← mul_assoc, mul_inv_cancel, one_mul] · rw [ne_eq, sub_eq_zero] - exact Ne.symm hg_one - · simp [g_le_one (le_add_of_nonneg_right bot_le)] - · simp [lt_of_le_of_ne (g_le_one (le_add_of_nonneg_right bot_le)) hg_one] - rw [kernel.withDensity_sub_add, withDensity_g] + exact Ne.symm h_one + · simp [rnDerivAux_le_one (le_add_of_nonneg_right bot_le)] + · simp [lt_of_le_of_ne (rnDerivAux_le_one (le_add_of_nonneg_right bot_le)) h_one] + rw [kernel.withDensity_sub_add, withDensity_rnDerivAux] all_goals simp_rw [h_coe] - · exact (measurable_g _ _).ennreal_ofReal - · exact (measurable_const.sub (measurable_g _ _)).ennreal_ofReal.mul (measurable_rnDerivReal _ _) - · intro a - have : ∀ x, ENNReal.ofReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x ≤ 1 := by + · exact (measurable_rnDerivAux _ _).ennreal_ofReal + · exact (measurable_const.sub (measurable_rnDerivAux _ _)).ennreal_ofReal.mul + (measurable_rnDeriv _ _) + · refine fun a ↦ ne_of_lt ?_ + have : ∀ x, + ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) * kernel.rnDeriv κ ν a x ≤ 1 := by refine fun x ↦ (h_le a x).trans ?_ - simp only [ENNReal.ofReal_le_one, g_le_one (le_add_of_nonneg_right bot_le)] - refine ne_of_lt ?_ - calc ∫⁻ x, ENNReal.ofReal (1 - g κ (κ + ν) a x) * kernel.rnDerivReal κ ν a x ∂(κ + ν) a + simp only [ENNReal.ofReal_le_one, rnDerivAux_le_one (le_add_of_nonneg_right bot_le)] + calc ∫⁻ x, ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) * kernel.rnDeriv κ ν a x ∂(κ + ν) a ≤ ∫⁻ _, 1 ∂(κ + ν) a := lintegral_mono this _ < ⊤ := by rw [MeasureTheory.lintegral_const, one_mul]; exact measure_lt_top _ _ · exact fun a ↦ ae_of_all _ (h_le a) -end Real - -section StandardBorel - -variable {Ω : Type*} [MeasurableSpace Ω] [StandardBorelSpace Ω] {κ ν : kernel α Ω} - -protected noncomputable def kernel.rnDeriv (κ ν : kernel α Ω) (a : α) (ω : Ω) : ℝ≥0∞ := - let f := measurableEmbedding_real Ω - let hf := measurableEmbedding_measurableEmbedding_real Ω - kernel.rnDerivReal (kernel.map κ f hf.measurable) (kernel.map ν f hf.measurable) a (f ω) - -protected noncomputable def kernel.singularPart (κ ν : kernel α Ω) - [IsSFiniteKernel κ] [IsSFiniteKernel ν] : kernel α Ω := - let f := measurableEmbedding_real Ω - let hf := measurableEmbedding_measurableEmbedding_real Ω - kernel.comapRight - (kernel.singularPartReal (kernel.map κ f hf.measurable) (kernel.map ν f hf.measurable)) - hf - -lemma kernel.mutuallySingular_singularPart (κ ν : kernel α Ω) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) : - kernel.singularPart κ ν a ⟂ₘ ν a := by - let f := measurableEmbedding_real Ω - let hf := measurableEmbedding_measurableEmbedding_real Ω - change kernel.comapRight - (kernel.singularPartReal (kernel.map κ f hf.measurable) (kernel.map ν f hf.measurable)) - hf a ⟂ₘ ν a - let s := (kernel.mutuallySingular_singularPartReal (kernel.map κ f hf.measurable) - (kernel.map ν f hf.measurable) a).nullSet - have hs : MeasurableSet s := Measure.MutuallySingular.measurableSet_nullSet _ - have hνs := Measure.MutuallySingular.measure_compl_nullSet - (kernel.mutuallySingular_singularPartReal (kernel.map κ f hf.measurable) - (kernel.map ν f hf.measurable) a) - have hsing := Measure.MutuallySingular.measure_nullSet - (kernel.mutuallySingular_singularPartReal (kernel.map κ f hf.measurable) - (kernel.map ν f hf.measurable) a) - refine ⟨f ⁻¹' s, hf.measurable hs, ?_, ?_⟩ - · rw [kernel.comapRight_apply' _ _ _ (hf.measurable hs)] - refine measure_mono_null (fun x hx ↦ ?_) hsing - exact image_preimage_subset _ _ hx - · have : ν = kernel.comapRight (kernel.map ν f hf.measurable) hf := by - ext a - rw [kernel.comapRight_apply _ hf, kernel.map_apply, MeasurableEmbedding.comap_map] - exact hf - rw [this, kernel.comapRight_apply' _ _ _ (hf.measurable hs).compl] - refine measure_mono_null (fun x ↦ ?_) hνs - rw [image_compl_preimage, mem_diff] - exact fun h ↦ h.2 - -lemma kernel.rnDeriv_add_singularPart (κ ν : kernel α Ω) [IsFiniteKernel κ] [IsFiniteKernel ν] : - kernel.withDensity ν (kernel.rnDeriv κ ν) + kernel.singularPart κ ν = κ := by - let f := measurableEmbedding_real Ω - let hf := measurableEmbedding_measurableEmbedding_real Ω - ext a : 1 - simp only [coeFn_add, Pi.add_apply] - change kernel.withDensity ν - (fun a ω ↦ kernel.rnDerivReal (kernel.map κ f hf.measurable) - (kernel.map ν f hf.measurable) a (f ω)) a + kernel.comapRight - (kernel.singularPartReal (kernel.map κ f hf.measurable) (kernel.map ν f hf.measurable)) - hf a = κ a - have h := kernel.rnDerivReal_add_singularPartReal (kernel.map κ f hf.measurable) - (kernel.map ν f hf.measurable) - have : κ = kernel.comapRight (kernel.map κ f hf.measurable) hf := by - ext a - rw [kernel.comapRight_apply _ hf, kernel.map_apply, MeasurableEmbedding.comap_map] - exact hf - conv_rhs => rw [this, ← h] - rw [comapRight_apply, comapRight_apply] - simp only [coeFn_add, Pi.add_apply] - rw [MeasurableEmbedding.comap_add hf] - congr - ext s hs - rw [Measure.comap_apply _ hf.injective (fun s hs ↦ hf.measurableSet_image.mpr hs) _ hs, - kernel.withDensity_apply', kernel.withDensity_apply', kernel.map_apply, - Measure.restrict_map hf.measurable (hf.measurableSet_image.mpr hs), - MeasureTheory.lintegral_map _ hf.measurable, preimage_image_eq _ hf.injective] - · exact measurable_rnDerivReal_right _ _ _ - · exact measurable_rnDerivReal _ _ - · change Measurable ((Function.uncurry fun a x ↦ - rnDerivReal (map κ f hf.measurable) (map ν f hf.measurable) a x) - ∘ (fun p : α × Ω ↦ (p.1, f p.2))) - exact (measurable_rnDerivReal _ _).comp - (measurable_fst.prod_mk (hf.measurable.comp measurable_snd)) - -end StandardBorel - end ProbabilityTheory From 5fb37f907485b1c92a190f5e796ec296e7b80262 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 21 Feb 2024 12:09:29 +0100 Subject: [PATCH 043/129] feat: lemmas about kernel.withDensity --- Mathlib/Probability/Kernel/WithDensity.lean | 58 +++++++++++++++++++++ 1 file changed, 58 insertions(+) diff --git a/Mathlib/Probability/Kernel/WithDensity.lean b/Mathlib/Probability/Kernel/WithDensity.lean index 4c4820fa08c24..389c8ae8eb97d 100644 --- a/Mathlib/Probability/Kernel/WithDensity.lean +++ b/Mathlib/Probability/Kernel/WithDensity.lean @@ -72,6 +72,33 @@ protected theorem withDensity_apply' (κ : kernel α β) [IsSFiniteKernel κ] rw [kernel.withDensity_apply κ hf, withDensity_apply' _ s] #align probability_theory.kernel.with_density_apply' ProbabilityTheory.kernel.withDensity_apply' +nonrec lemma kernel.withDensity_absolutelyContinuous [IsSFiniteKernel κ] + (f : α → β → ℝ≥0∞) (a : α) : + kernel.withDensity κ f a ≪ κ a := by + by_cases hf : Measurable (Function.uncurry f) + · rw [kernel.withDensity_apply _ hf] + exact withDensity_absolutelyContinuous _ _ + · rw [withDensity_of_not_measurable _ hf] + simp [Measure.AbsolutelyContinuous.zero] + +@[simp] +lemma withDensity_one (κ : kernel α β) [IsSFiniteKernel κ] : + kernel.withDensity κ 1 = κ := by + ext; rw [kernel.withDensity_apply _ measurable_const]; simp + +@[simp] +lemma withDensity_one' (κ : kernel α β) [IsSFiniteKernel κ] : + kernel.withDensity κ (fun _ _ ↦ 1) = κ := kernel.withDensity_one _ + +@[simp] +lemma withDensity_zero (κ : kernel α β) [IsSFiniteKernel κ] : + kernel.withDensity κ 0 = 0 := by + ext; rw [kernel.withDensity_apply _ measurable_const]; simp + +@[simp] +lemma withDensity_zero' (κ : kernel α β) [IsSFiniteKernel κ] : + kernel.withDensity κ (fun _ _ ↦ 0) = 0 := kernel.withDensity_zero _ + theorem lintegral_withDensity (κ : kernel α β) [IsSFiniteKernel κ] (hf : Measurable (Function.uncurry f)) (a : α) {g : β → ℝ≥0∞} (hg : Measurable g) : ∫⁻ b, g b ∂withDensity κ f a = ∫⁻ b, f a b * g b ∂κ a := by @@ -134,6 +161,21 @@ theorem withDensity_tsum [Countable ι] (κ : kernel α β) [IsSFiniteKernel κ] rw [kernel.withDensity_apply' _ (hf n) a s] #align probability_theory.kernel.with_density_tsum ProbabilityTheory.kernel.withDensity_tsum +lemma withDensity_sub_add [IsSFiniteKernel κ] {f g : α → β → ℝ≥0∞} + (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) + (hg_int : ∀ a, ∫⁻ x, g a x ∂(κ a) ≠ ∞) (hfg : ∀ a, g a ≤ᵐ[κ a] f a) : + withDensity κ (fun a x ↦ f a x - g a x) + withDensity κ g = withDensity κ f := by + ext a s + simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add] + rw [kernel.withDensity_apply' _ hf, kernel.withDensity_apply' _ hg, kernel.withDensity_apply'] + swap; · exact hf.sub hg + rw [lintegral_sub] + · rw [tsub_add_cancel_iff_le] + exact lintegral_mono_ae (ae_restrict_of_ae (hfg a)) + · exact hg.comp measurable_prod_mk_left + · exact ((set_lintegral_le_lintegral _ _).trans_lt (hg_int a).lt_top).ne + · exact ae_restrict_of_ae (hfg a) + /-- If a kernel `κ` is finite and a function `f : α → β → ℝ≥0∞` is bounded, then `withDensity κ f` is finite. -/ theorem isFiniteKernel_withDensity_of_bounded (κ : kernel α β) [IsFiniteKernel κ] {B : ℝ≥0∞} @@ -230,4 +272,20 @@ instance (κ : kernel α β) [IsSFiniteKernel κ] (f : α → β → ℝ≥0) : IsSFiniteKernel (withDensity κ fun a b => f a b) := IsSFiniteKernel.withDensity κ fun _ _ => ENNReal.coe_ne_top +nonrec lemma withDensity_mul [IsSFiniteKernel κ] {f : α → β → ℝ≥0} {g : α → β → ℝ≥0∞} + (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) : + withDensity κ (fun a x ↦ f a x * g a x) + = withDensity (withDensity κ fun a x ↦ f a x) g := by + ext a : 1 + rw [kernel.withDensity_apply] + swap; · exact (measurable_coe_nnreal_ennreal.comp hf).mul hg + change (Measure.withDensity (κ a) ((fun x ↦ (f a x : ℝ≥0∞)) * (fun x ↦ (g a x : ℝ≥0∞)))) = + (withDensity (withDensity κ fun a x ↦ f a x) g) a + rw [withDensity_mul] + · rw [kernel.withDensity_apply _ hg, kernel.withDensity_apply] + exact measurable_coe_nnreal_ennreal.comp hf + · rw [measurable_coe_nnreal_ennreal_iff] + exact hf.comp measurable_prod_mk_left + · exact hg.comp measurable_prod_mk_left + end ProbabilityTheory.kernel From 45c91cd4803ebc6ac4e289ebc9520468c05e2b91 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 21 Feb 2024 14:00:27 +0100 Subject: [PATCH 044/129] fix --- Mathlib/Probability/Kernel/WithDensity.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Probability/Kernel/WithDensity.lean b/Mathlib/Probability/Kernel/WithDensity.lean index 389c8ae8eb97d..b07c0278c1f51 100644 --- a/Mathlib/Probability/Kernel/WithDensity.lean +++ b/Mathlib/Probability/Kernel/WithDensity.lean @@ -72,7 +72,7 @@ protected theorem withDensity_apply' (κ : kernel α β) [IsSFiniteKernel κ] rw [kernel.withDensity_apply κ hf, withDensity_apply' _ s] #align probability_theory.kernel.with_density_apply' ProbabilityTheory.kernel.withDensity_apply' -nonrec lemma kernel.withDensity_absolutelyContinuous [IsSFiniteKernel κ] +nonrec lemma withDensity_absolutelyContinuous [IsSFiniteKernel κ] (f : α → β → ℝ≥0∞) (a : α) : kernel.withDensity κ f a ≪ κ a := by by_cases hf : Measurable (Function.uncurry f) From 301846de2a0ca47f4b140411df6746d785284d33 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 21 Feb 2024 18:41:56 +0100 Subject: [PATCH 045/129] updated docstring of Density --- .../Kernel/Disintegration/Density.lean | 52 ++++++++++--------- 1 file changed, 28 insertions(+), 24 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 350c4e677d49b..e13c994ddc73a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -5,56 +5,60 @@ Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Martingale.Convergence -import Mathlib.Analysis.SpecialFunctions.Log.Base /-! # Kernel density -Let `κ : kernel α (ℝ × β)` and `ν : kernel α ℝ` be two finite kernels with `kernel.fst κ ≤ ν`. -We build a function `f : α → ℝ → Set β → ℝ` jointly measurable in the first two arguments such that -for all `a : α` and all measurable sets `s : Set β` and `A : Set ℝ`, +Let `κ : kernel α (γ × β)` and `ν : kernel α γ` be two finite kernels with `kernel.fst κ ≤ ν`, +where `γ` has a countably generated σ-algebra (true in particular for standard Borel spaces). +We build a function `f : α → γ → Set β → ℝ` jointly measurable in the first two arguments such that +for all `a : α` and all measurable sets `s : Set β` and `A : Set γ`, `∫ t in A, f a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal`. If we were interested only in a fixed `a : α`, then we could use the Radon-Nikodym derivative to build the function `f`, as follows ``` -def f_a (κ : kernel α (ℝ × β)) (ν : kernel a ℝ) (a : α) (t : ℝ) (s : Set β) : ℝ := +def f (κ : kernel α (γ × β)) (ν : kernel a γ) (a : α) (t : γ) (s : Set β) : ℝ := (((κ a).restrict (univ ×ˢ s)).fst.rnDeriv (ν a) t).toReal ``` However, we can't turn those functions for each `a` into a measurable function of the pair `(a, t)`. -In order to obtain measurability through countability, we discretize the real line. -For each `n : ℕ`, we define the intervals `I n k = [k * 2^-n, (k + 1) * 2^-n)` for `k : ℤ`. -For `t : ℝ`, let `indexI n t = ⌊t * 2^n⌋` be the integer such that `t ∈ I n (indexI n t)`. +In order to obtain measurability through countability, we use the fact that the measurable space `γ` +is countably generated. For each `n : ℕ`, we define a finite partition of `γ`, such that those +partitions are finer as `n` grows, and the σ-algebra generated by the union of all partitions is +the σ-algebra of `γ`. +For `t : γ`, let `partitionSet n t` be the set in the partition such that `t ∈ partitionSet n t`. -For a given `n`, the function `densityProcess κ ν n : α → ℝ → Set β → ℝ` defined by -`fun a t s ↦ (κ a (I n (indexI n t) ×ˢ s) / ν a (I n (indexI n t))).toReal` has the desired +For a given `n`, the function `densityProcess κ ν n : α → γ → Set β → ℝ` defined by +`fun a t s ↦ (κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t)).toReal` has the desired property that `∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal` for all `A` -in the σ-algebra generated by the intervals `I n k` for `k : ℤ` and is measurable in `(a, t)`. +in the σ-algebra generated by the partition at scale `n` and is measurable in `(a, t)`. -Let `ℱ` be the filtration of those σ-algebras for all `n : ℕ`. The functions `densityProcess κ ν n` -described here are a bounded `ν`-martingale for the filtration `ℱ`. By Doob's L1 martingale -convergence theorem, that martingale converges to a limit, which has a product-measurable version -and satisfies the integral equality for all `A` in `⨆ n, ℱ n = borel ℝ`. We have obtained the -desired density function. +Let `partitionFiltration γ` be the filtration of those σ-algebras for all `n : ℕ`. +The functions `densityProcess κ ν n` described here are a bounded `ν`-martingale for the filtration +`partitionFiltration γ`. By Doob's L1 martingale convergence theorem, that martingale converges to +a limit, which has a product-measurable version and satisfies the integral equality for all `A` in +`⨆ n, partitionFiltration γ n`. Finally, the partitions were chosen such that that supremum is equal +to the σ-algebra on `γ`, hence the equality holds for all measurable sets. +We have obtained the desired density function. ## Main definitions -* `FooBar` +* `ProbabilityTheory.kernel.density`: for `k : kernel α (γ × β)` and `ν : kernel α γ` two finite + kernels, `kernel.density κ ν` is a function `α → γ → Set β → ℝ`. ## Main statements -* `fooBar_unique` - -## Implementation details - +* `ProbabilityTheory.set_integral_density`: for all measurable sets `A : Set γ` and `s : Set β`, + `∫ t in A, kernel.density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal`. +* `ProbabilityTheory.measurable_density`: the function `(p : α × γ) ↦ kernel.density κ ν p.1 p.2 s` + is measurable. ## References The construction of the density process in this file follows the proof of Theorem 9.27 in -[O. Kallenberg, Foundations of modern probability][kallenberg2021]. - -TODO: adapted to use countably generated sigma-algebras. +[O. Kallenberg, Foundations of modern probability][kallenberg2021], adapted to use a countably +generated hypothesis instead of specializing to `ℝ`. -/ variable {α β : Type*} {mα : MeasurableSpace α} From d7a760842cbe56d527d32cc03001c5f04cf76283 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 22 Feb 2024 12:58:36 +0100 Subject: [PATCH 046/129] add withDensity_add_right --- Mathlib/Probability/Kernel/WithDensity.lean | 41 +++++++++++++-------- 1 file changed, 26 insertions(+), 15 deletions(-) diff --git a/Mathlib/Probability/Kernel/WithDensity.lean b/Mathlib/Probability/Kernel/WithDensity.lean index b07c0278c1f51..323fe2e04cb0c 100644 --- a/Mathlib/Probability/Kernel/WithDensity.lean +++ b/Mathlib/Probability/Kernel/WithDensity.lean @@ -72,6 +72,13 @@ protected theorem withDensity_apply' (κ : kernel α β) [IsSFiniteKernel κ] rw [kernel.withDensity_apply κ hf, withDensity_apply' _ s] #align probability_theory.kernel.with_density_apply' ProbabilityTheory.kernel.withDensity_apply' +nonrec lemma withDensity_congr_ae (κ : kernel α β) [IsSFiniteKernel κ] {f g : α → β → ℝ≥0∞} + (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) + (hfg : ∀ a, f a =ᵐ[κ a] g a) : + withDensity κ f = withDensity κ g := by + ext a + rw [kernel.withDensity_apply _ hf,kernel.withDensity_apply _ hg, withDensity_congr_ae (hfg a)] + nonrec lemma withDensity_absolutelyContinuous [IsSFiniteKernel κ] (f : α → β → ℝ≥0∞) (a : α) : kernel.withDensity κ f a ≪ κ a := by @@ -138,6 +145,25 @@ theorem withDensity_kernel_sum [Countable ι] (κ : ι → kernel α β) (hκ : exact sum_zero.symm #align probability_theory.kernel.with_density_kernel_sum ProbabilityTheory.kernel.withDensity_kernel_sum +lemma withDensity_add_right [IsSFiniteKernel κ] {f g : α → β → ℝ≥0∞} + (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) : + withDensity κ (f + g) = withDensity κ f + withDensity κ g := by + ext a + rw [coeFn_add, Pi.add_apply, kernel.withDensity_apply _ hf, kernel.withDensity_apply _ hg, + kernel.withDensity_apply,Pi.add_apply, MeasureTheory.withDensity_add_right] + · exact hg.comp measurable_prod_mk_left + · exact hf.add hg + +lemma withDensity_sub_add_cancel [IsSFiniteKernel κ] {f g : α → β → ℝ≥0∞} + (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) + (hfg : ∀ a, g a ≤ᵐ[κ a] f a) : + withDensity κ (fun a x ↦ f a x - g a x) + withDensity κ g = withDensity κ f := by + rw [← withDensity_add_right _ hg] + swap; · exact hf.sub hg + refine withDensity_congr_ae κ ((hf.sub hg).add hg) hf (fun a ↦ ?_) + filter_upwards [hfg a] with x hx + rwa [Pi.add_apply, Pi.add_apply, tsub_add_cancel_iff_le] + theorem withDensity_tsum [Countable ι] (κ : kernel α β) [IsSFiniteKernel κ] {f : ι → α → β → ℝ≥0∞} (hf : ∀ i, Measurable (Function.uncurry (f i))) : withDensity κ (∑' n, f n) = kernel.sum fun n => withDensity κ (f n) := by @@ -161,21 +187,6 @@ theorem withDensity_tsum [Countable ι] (κ : kernel α β) [IsSFiniteKernel κ] rw [kernel.withDensity_apply' _ (hf n) a s] #align probability_theory.kernel.with_density_tsum ProbabilityTheory.kernel.withDensity_tsum -lemma withDensity_sub_add [IsSFiniteKernel κ] {f g : α → β → ℝ≥0∞} - (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) - (hg_int : ∀ a, ∫⁻ x, g a x ∂(κ a) ≠ ∞) (hfg : ∀ a, g a ≤ᵐ[κ a] f a) : - withDensity κ (fun a x ↦ f a x - g a x) + withDensity κ g = withDensity κ f := by - ext a s - simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add] - rw [kernel.withDensity_apply' _ hf, kernel.withDensity_apply' _ hg, kernel.withDensity_apply'] - swap; · exact hf.sub hg - rw [lintegral_sub] - · rw [tsub_add_cancel_iff_le] - exact lintegral_mono_ae (ae_restrict_of_ae (hfg a)) - · exact hg.comp measurable_prod_mk_left - · exact ((set_lintegral_le_lintegral _ _).trans_lt (hg_int a).lt_top).ne - · exact ae_restrict_of_ae (hfg a) - /-- If a kernel `κ` is finite and a function `f : α → β → ℝ≥0∞` is bounded, then `withDensity κ f` is finite. -/ theorem isFiniteKernel_withDensity_of_bounded (κ : kernel α β) [IsFiniteKernel κ] {B : ℝ≥0∞} From 54e85f35bef457a98f109deb325987af622fc9ea Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 22 Feb 2024 14:28:19 +0100 Subject: [PATCH 047/129] split partitioned --- Mathlib/Data/Set/Partitioned.lean | 91 +++++++++++++++++++ .../MeasureTheory/MeasurableSpace/Basic.lean | 64 +------------ 2 files changed, 92 insertions(+), 63 deletions(-) create mode 100644 Mathlib/Data/Set/Partitioned.lean diff --git a/Mathlib/Data/Set/Partitioned.lean b/Mathlib/Data/Set/Partitioned.lean new file mode 100644 index 0000000000000..17c658c31589c --- /dev/null +++ b/Mathlib/Data/Set/Partitioned.lean @@ -0,0 +1,91 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Data.Set.Finite +import Mathlib.Data.Fintype.Card + +/-! +# Partitioned + +## Main definitions + +* `partitioned` + +## Main statements + +* `fooBar_unique` + +-/ + +open Set + +variable {α : Type*} + +def partitioned (f : ℕ → Set α) : ℕ → Set (Set α) + | 0 => {f 0, (f 0)ᶜ} + | n + 1 => ⋃ u ∈ partitioned f n, {u ∩ f (n + 1), u \ f (n + 1)} + +@[simp] +lemma partitioned_zero (f : ℕ → Set α) : partitioned f 0 = {f 0, (f 0)ᶜ} := rfl + +lemma partitioned_succ (f : ℕ → Set α) (n : ℕ) : + partitioned f (n + 1) = ⋃ u ∈ partitioned f n, {u ∩ f (n + 1), u \ f (n + 1)} := rfl + +lemma disjoint_partitioned (f : ℕ → Set α) (n : ℕ) : + ∀ u ∈ partitioned f n, ∀ v ∈ partitioned f n, u ≠ v → Disjoint u v := by + induction n with + | zero => + simp only [Nat.zero_eq, partitioned_zero, union_singleton, mem_insert_iff, + mem_singleton_iff, ne_eq, forall_eq_or_imp, forall_eq, not_true_eq_false, disjoint_self, + bot_eq_empty, compl_empty_iff, IsEmpty.forall_iff, true_and, and_true] + exact ⟨fun _ ↦ disjoint_compl_right, fun _ ↦ disjoint_compl_left⟩ + | succ n ih => + intro u hu v hv huv + rw [partitioned_succ] at hu hv + simp only [union_singleton, mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop] at hu hv + obtain ⟨u', hu', hu'_eq⟩ := hu + obtain ⟨v', hv', hv'_eq⟩ := hv + rcases hu'_eq with rfl | rfl <;> rcases hv'_eq with rfl | rfl + · refine Disjoint.mono (inter_subset_left _ _) (inter_subset_left _ _) (ih u' hu' v' hv' ?_) + exact fun huv' ↦ huv (huv' ▸ rfl) + · by_cases huv' : u' = v' + · rw [huv'] + exact Disjoint.mono_left (inter_subset_right _ _) Set.disjoint_sdiff_right + · exact Disjoint.mono (inter_subset_left _ _) (diff_subset _ _) (ih u' hu' v' hv' huv') + · by_cases huv' : u' = v' + · rw [huv'] + exact Disjoint.mono_right (inter_subset_right _ _) Set.disjoint_sdiff_left + · exact Disjoint.mono (diff_subset _ _) (inter_subset_left _ _) (ih u' hu' v' hv' huv') + · refine Disjoint.mono (diff_subset _ _) (diff_subset _ _) (ih u' hu' v' hv' ?_) + refine fun huv' ↦ huv (huv' ▸ rfl) + +lemma sUnion_partitioned (f : ℕ → Set α) (n : ℕ) : ⋃₀ partitioned f n = univ := by + induction n with + | zero => simp + | succ n ih => + rw [partitioned_succ] + ext x + have : x ∈ ⋃₀ partitioned f n := by simp [ih] + simp only [mem_sUnion, mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop, mem_univ, + iff_true] at this ⊢ + obtain ⟨t, ht, hxt⟩ := this + by_cases hxf : x ∈ f (n + 1) + · exact ⟨t ∩ f (n + 1), ⟨t, ht, Or.inl rfl⟩, hxt, hxf⟩ + · exact ⟨t \ f (n + 1), ⟨t, ht, Or.inr rfl⟩, hxt, hxf⟩ + +lemma partitioned_finite (f : ℕ → Set α) (n : ℕ) : Set.Finite (partitioned f n) := by + induction n with + | zero => simp + | succ n ih => + rw [partitioned_succ] + have : Finite (partitioned f n) := Set.finite_coe_iff.mp ih + rw [← Set.finite_coe_iff] + refine Finite.Set.finite_biUnion (partitioned f n) _ (fun u _ ↦ ?_) + rw [Set.finite_coe_iff] + simp + +noncomputable +instance instFintypePartitioned (f : ℕ → Set α) (n : ℕ) : Fintype (partitioned f n) := + Set.Finite.fintype (partitioned_finite f n) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 0ca1cb8f854ee..a715589823253 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -12,6 +12,7 @@ import Mathlib.Order.Filter.SmallSets import Mathlib.Order.Filter.CountableSeparatingOn import Mathlib.Order.LiminfLimsup import Mathlib.Data.Set.UnionLift +import Mathlib.Data.Set.Partitioned #align_import measure_theory.measurable_space from "leanprover-community/mathlib"@"001ffdc42920050657fd45bd2b8bfbec8eaaeb29" @@ -1994,69 +1995,6 @@ instance [MeasurableSpace α] {s : Set α} [h : CountablyGenerated s] [Measurabl rw [← forall_generateFrom_mem_iff_mem_iff, ← hb] at h simpa using h {y} -def partitioned (f : ℕ → Set α) : ℕ → Set (Set α) - | 0 => {f 0, (f 0)ᶜ} - | n + 1 => ⋃ u ∈ partitioned f n, {u ∩ f (n + 1), u \ f (n + 1)} - -@[simp] -lemma partitioned_zero (f : ℕ → Set α) : partitioned f 0 = {f 0, (f 0)ᶜ} := rfl - -lemma partitioned_succ (f : ℕ → Set α) (n : ℕ) : - partitioned f (n + 1) = ⋃ u ∈ partitioned f n, {u ∩ f (n + 1), u \ f (n + 1)} := rfl - -lemma disjoint_partitioned (f : ℕ → Set α) (n : ℕ) : - ∀ u ∈ partitioned f n, ∀ v ∈ partitioned f n, u ≠ v → Disjoint u v := by - induction n with - | zero => - simp only [Nat.zero_eq, partitioned_zero, union_singleton, mem_insert_iff, - mem_singleton_iff, ne_eq, forall_eq_or_imp, forall_eq, not_true_eq_false, disjoint_self, - bot_eq_empty, compl_empty_iff, IsEmpty.forall_iff, true_and, and_true] - exact ⟨fun _ ↦ disjoint_compl_right, fun _ ↦ disjoint_compl_left⟩ - | succ n ih => - intro u hu v hv huv - rw [partitioned_succ] at hu hv - simp only [union_singleton, mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop] at hu hv - obtain ⟨u', hu', hu'_eq⟩ := hu - obtain ⟨v', hv', hv'_eq⟩ := hv - rcases hu'_eq with rfl | rfl <;> rcases hv'_eq with rfl | rfl - · refine Disjoint.mono (inter_subset_left _ _) (inter_subset_left _ _) (ih u' hu' v' hv' ?_) - exact fun huv' ↦ huv (huv' ▸ rfl) - · by_cases huv' : u' = v' - · rw [huv'] - exact Disjoint.mono_left (inter_subset_right _ _) Set.disjoint_sdiff_right - · exact Disjoint.mono (inter_subset_left _ _) (diff_subset _ _) (ih u' hu' v' hv' huv') - · by_cases huv' : u' = v' - · rw [huv'] - exact Disjoint.mono_right (inter_subset_right _ _) Set.disjoint_sdiff_left - · exact Disjoint.mono (diff_subset _ _) (inter_subset_left _ _) (ih u' hu' v' hv' huv') - · refine Disjoint.mono (diff_subset _ _) (diff_subset _ _) (ih u' hu' v' hv' ?_) - refine fun huv' ↦ huv (huv' ▸ rfl) - -lemma sUnion_partitioned (f : ℕ → Set α) (n : ℕ) : ⋃₀ partitioned f n = univ := by - induction n with - | zero => simp - | succ n ih => - rw [partitioned_succ] - ext x - have : x ∈ ⋃₀ partitioned f n := by simp [ih] - simp only [mem_sUnion, mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop, mem_univ, - iff_true] at this ⊢ - obtain ⟨t, ht, hxt⟩ := this - by_cases hxf : x ∈ f (n + 1) - · exact ⟨t ∩ f (n + 1), ⟨t, ht, Or.inl rfl⟩, hxt, hxf⟩ - · exact ⟨t \ f (n + 1), ⟨t, ht, Or.inr rfl⟩, hxt, hxf⟩ - -lemma partitioned_finite (f : ℕ → Set α) (n : ℕ) : Set.Finite (partitioned f n) := by - induction n with - | zero => simp - | succ n ih => - rw [partitioned_succ] - have : Finite (partitioned f n) := Set.finite_coe_iff.mp ih - rw [← Set.finite_coe_iff] - refine Finite.Set.finite_biUnion (partitioned f n) _ (fun u _ ↦ ?_) - rw [Set.finite_coe_iff] - simp - lemma todo [m : MeasurableSpace α] [h : CountablyGenerated α] : ∃ s : ℕ → Set (Set α), (∀ n, Set.Finite (s n)) ∧ (∀ n, ∀ u ∈ s n, ∀ v ∈ s n, u ≠ v → Disjoint u v) From 0d49ff15cf6ff8e140fdf14cab463eb8cef4e2b5 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 22 Feb 2024 15:28:51 +0100 Subject: [PATCH 048/129] fix --- .../Kernel/Disintegration/CondCdf.lean | 9 ++++---- .../Kernel/Disintegration/Density.lean | 22 ++++++++++--------- 2 files changed, 16 insertions(+), 15 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index 804c6af0d84fc..caee1cd73d9ca 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -68,14 +68,14 @@ theorem IicSnd_univ (r : ℝ) : ρ.IicSnd r univ = ρ (univ ×ˢ Iic r) := #align measure_theory.measure.Iic_snd_univ MeasureTheory.Measure.IicSnd_univ theorem IicSnd_mono {r r' : ℝ} (h_le : r ≤ r') : ρ.IicSnd r ≤ ρ.IicSnd r' := by - intro s hs + refine Measure.le_iff.2 fun s hs ↦ ?_ simp_rw [IicSnd_apply ρ _ hs] refine measure_mono (prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, Iic_subset_Iic.mpr ?_⟩)) exact mod_cast h_le #align measure_theory.measure.Iic_snd_mono MeasureTheory.Measure.IicSnd_mono theorem IicSnd_le_fst (r : ℝ) : ρ.IicSnd r ≤ ρ.fst := by - intro s hs + refine Measure.le_iff.2 fun s hs ↦ ?_ simp_rw [fst_apply hs, IicSnd_apply ρ r hs] exact measure_mono (prod_subset_preimage_fst _ _) #align measure_theory.measure.Iic_snd_le_fst MeasureTheory.Measure.IicSnd_le_fst @@ -187,8 +187,7 @@ theorem monotone_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : refine fun r r' hrr' ↦ ae_le_of_forall_set_lintegral_le_of_sigmaFinite measurable_preCDF measurable_preCDF fun s hs _ ↦ ?_ rw [set_lintegral_preCDF_fst ρ r hs, set_lintegral_preCDF_fst ρ r' hs] - refine Measure.IicSnd_mono ρ ?_ s hs - exact mod_cast hrr' + exact Measure.IicSnd_mono ρ (mod_cast hrr') s #align probability_theory.monotone_pre_cdf ProbabilityTheory.monotone_preCDF theorem set_lintegral_iInf_gt_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (t : ℚ) {s : Set α} @@ -218,7 +217,7 @@ theorem preCDF_le_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : measurable_const fun s hs _ ↦ ?_ rw [set_lintegral_preCDF_fst ρ r hs] simp only [Pi.one_apply, lintegral_one, Measure.restrict_apply, MeasurableSet.univ, univ_inter] - exact Measure.IicSnd_le_fst ρ r s hs + exact Measure.IicSnd_le_fst ρ r s #align probability_theory.pre_cdf_le_one ProbabilityTheory.preCDF_le_one theorem set_integral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) {s : Set α} (hs : MeasurableSet s) diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index e13c994ddc73a..4c67485d1120c 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -250,7 +250,7 @@ lemma apply_partitionSet_le_of_fst_le (hκν : kernel.fst κ ≤ ν) (n : ℕ) ( refine measure_mono (fun x ↦ ?_) simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h - _ ≤ ν a (partitionSet n t) := hκν a _ (measurableSet_partitionSet _ _) + _ ≤ ν a (partitionSet n t) := hκν a _ lemma densityProcess_le_one (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : densityProcess κ ν n a t s ≤ 1 := by @@ -290,7 +290,7 @@ lemma set_integral_densityProcess_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKern by_cases h0 : ν a (partitionSet n t) = 0 · suffices κ a (partitionSet n t ×ˢ s) = 0 by simp [h0, this] have h0' : kernel.fst κ a (partitionSet n t) = 0 := - le_antisymm ((hκν a _ (measurableSet_partitionSet _ _)).trans h0.le) zero_le' + le_antisymm ((hκν a _).trans h0.le) zero_le' rw [kernel.fst_apply' _ _ (measurableSet_partitionSet _ _)] at h0' refine measure_mono_null (fun x ↦ ?_) h0' simp only [mem_prod, mem_setOf_eq, and_imp] @@ -307,7 +307,7 @@ lemma set_integral_densityProcess_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKern by_cases h0 : ν a u = 0 · simp only [h0, mul_zero] have h0' : kernel.fst κ a u = 0 := - le_antisymm ((hκν a _ hu_meas).trans h0.le) zero_le' + le_antisymm ((hκν a _).trans h0.le) zero_le' rw [kernel.fst_apply' _ _ hu_meas] at h0' refine (measure_mono_null ?_ h0').symm intro p @@ -428,7 +428,7 @@ lemma densityProcess_mono_set (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) · exact h_ne_top s' lemma densityProcess_mono_kernel_left {κ' : kernel α (γ × β)} (hκκ' : κ ≤ κ') - (hκ'ν : kernel.fst κ' ≤ ν) (n : ℕ) (a : α) (t : γ) {s : Set β} (hs : MeasurableSet s) : + (hκ'ν : kernel.fst κ' ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : densityProcess κ ν n a t s ≤ densityProcess κ' ν n a t s := by unfold densityProcess by_cases h0 : ν a (partitionSet n t) = 0 @@ -438,11 +438,11 @@ lemma densityProcess_mono_kernel_left {κ' : kernel α (γ × β)} (hκκ' : κ ≤ ν a (partitionSet n t) := apply_partitionSet_le_of_fst_le hκ'ν n a t s rw [ENNReal.toReal_le_toReal] · gcongr - exact hκκ' _ _ ((measurableSet_partitionSet _ _).prod hs) + exact hκκ' _ _ · rw [ne_eq, ENNReal.div_eq_top] simp only [ne_eq, h0, and_false, false_or, not_and, not_not] refine fun h_top ↦ eq_top_mono ?_ h_top - exact (hκκ' _ _ ((measurableSet_partitionSet _ _).prod hs)).trans h_le + exact (hκκ' _ _).trans h_le · rw [ne_eq, ENNReal.div_eq_top] simp only [ne_eq, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top @@ -459,13 +459,13 @@ lemma densityProcess_antitone_kernel_right {ν' : kernel α γ} exact le_antisymm (h_le.trans h0.le) zero_le' have h0' : ν' a (partitionSet n t) ≠ 0 := by refine fun h ↦ h0 (le_antisymm (le_trans ?_ h.le) zero_le') - exact hνν' _ _ (measurableSet_partitionSet _ _) + exact hνν' _ _ rw [ENNReal.toReal_le_toReal] · gcongr - exact hνν' _ _ (measurableSet_partitionSet _ _) + exact hνν' _ _ · simp only [ne_eq, ENNReal.div_eq_top, h0', and_false, false_or, not_and, not_not] refine fun h_top ↦ eq_top_mono ?_ h_top - exact h_le.trans (hνν' _ _ (measurableSet_partitionSet _ _)) + exact h_le.trans (hνν' _ _) · simp only [ne_eq, ENNReal.div_eq_top, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top @@ -555,7 +555,9 @@ lemma tendsto_snorm_one_restrict_densityProcess_limitProcess [IsFiniteKernel κ] Tendsto (fun n ↦ snorm ((fun t ↦ densityProcess κ ν n a t s) - (partitionFiltration γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 ((ν a).restrict A)) atTop (𝓝 0) := - tendsto_snorm_restrict_zero (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) A + tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds + (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) (fun _ ↦ zero_le') + (fun _ ↦ snorm_restrict_le _ _ _ _) noncomputable def kernel.density (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) (t : γ) (s : Set β) : ℝ := From bc9ebaceab2a693cc94996a9d026277bd97b4eb4 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 22 Feb 2024 15:46:31 +0100 Subject: [PATCH 049/129] fix --- Mathlib/MeasureTheory/Integral/Bochner.lean | 14 -------------- .../Kernel/Disintegration/Density.lean | 15 +++++++-------- 2 files changed, 7 insertions(+), 22 deletions(-) diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index b3aff1066d313..1d731b59e1f21 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1302,20 +1302,6 @@ theorem integral_eq_zero_iff_of_nonneg_ae {f : α → ℝ} (hf : 0 ≤ᵐ[μ] f) · exact (ENNReal.measurable_ofReal.comp_aemeasurable hfi.1.aemeasurable) #align measure_theory.integral_eq_zero_iff_of_nonneg_ae MeasureTheory.integral_eq_zero_iff_of_nonneg_ae -lemma ae_eq_of_integral_eq_of_ae_le {f g : α → ℝ} (hf : Integrable f μ) - (hg : Integrable g μ) (h_le : f ≤ᵐ[μ] g) (h_eq : ∫ a, f a ∂μ = ∫ a, g a ∂μ) : - f =ᵐ[μ] g := by - suffices g - f =ᵐ[μ] 0 by - filter_upwards [this] with a ha - symm - simpa only [Pi.sub_apply, Pi.zero_apply, sub_eq_zero] using ha - have h_eq' : ∫ a, (g - f) a ∂μ = 0 := by - simp_rw [Pi.sub_apply] - rwa [integral_sub hg hf, sub_eq_zero, eq_comm] - rwa [integral_eq_zero_iff_of_nonneg_ae _ (hg.sub hf)] at h_eq' - filter_upwards [h_le] with a ha - simpa - theorem integral_eq_zero_iff_of_nonneg {f : α → ℝ} (hf : 0 ≤ f) (hfi : Integrable f μ) : ∫ x, f x ∂μ = 0 ↔ f =ᵐ[μ] 0 := integral_eq_zero_iff_of_nonneg_ae (eventually_of_forall hf) hfi diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 4c67485d1120c..e9a685cafbb61 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -788,14 +788,14 @@ lemma tendsto_density_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin simp only [lintegral_const, one_mul] exact measure_lt_top _ _ -- it suffices to show that the limit `F` is 0 a.e. - suffices ∀ᵐ t ∂(ν a), F t = 0 by + suffices F=ᵐ[ν a] 0 by filter_upwards [this] with t ht_eq + simp only [Pi.zero_apply] at ht_eq rw [← ht_eq] exact hF_tendsto t -- since `F` is nonnegative, proving that its integral is 0 is sufficient to get that -- `F` is 0 a.e. - suffices ∀ᵐ (t : γ) ∂(ν a), 0 = F t by filter_upwards [this] with a ha; simp [ha] - refine ae_eq_of_integral_eq_of_ae_le (integrable_const _) hF_int (ae_of_all _ hF_nonneg) ?_ + rw [← integral_eq_zero_iff_of_nonneg hF_nonneg hF_int] have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ t, F t ∂(ν a))) := by refine integral_tendsto_of_tendsto_of_antitone ?_ hF_int ?_ ?_ @@ -803,10 +803,9 @@ lemma tendsto_density_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin · exact ae_of_all _ h_anti · exact ae_of_all _ hF_tendsto have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop - (𝓝 (∫ _, 0 ∂(ν a))) := by - simp only [integral_zero] + (𝓝 0) := by exact tendsto_integral_density_of_antitone hκν a s hs hs_iInter hs_meas - exact (tendsto_nhds_unique h_integral h_integral').symm + exact tendsto_nhds_unique h_integral h_integral' section UnivFst @@ -953,13 +952,13 @@ lemma tendsto_density_atTop_ae_of_monotone [IsFiniteKernel κ] exact hF_le_one _ _ < ⊤ := by simp only [lintegral_const, measure_univ, one_mul, measure_lt_top] -- it suffices to show that the limit `F` is 1 a.e. - suffices ∀ᵐ t ∂(ν a), F t = 1 by + suffices F =ᵐ[ν a] (fun _ ↦ 1) by filter_upwards [this] with t ht_eq rw [← ht_eq] exact hF_tendsto t -- since `F` is at most 1, proving that its integral is the same as the integral of 1 will tell -- us that `F` is 1 a.e. - refine ae_eq_of_integral_eq_of_ae_le hF_int (integrable_const _) (ae_of_all _ hF_le_one) ?_ + rw [← integral_eq_iff_of_ae_le hF_int (integrable_const _) (ae_of_all _ hF_le_one)] have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ t, F t ∂(ν a))) := by refine integral_tendsto_of_tendsto_of_monotone ?_ hF_int ?_ ?_ From 4ecd880a0d0cfc72e7245f2872e9cf435d4fe068 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 22 Feb 2024 16:35:20 +0100 Subject: [PATCH 050/129] split file --- Mathlib/Data/Set/MemPartition.lean | 91 +++++++++ Mathlib/Data/Set/Partitioned.lean | 91 --------- .../MeasureTheory/MeasurableSpace/Basic.lean | 181 ------------------ .../MeasurableSpace/CountablyGenerated.lean | 173 +++++++++++++++++ 4 files changed, 264 insertions(+), 272 deletions(-) create mode 100644 Mathlib/Data/Set/MemPartition.lean delete mode 100644 Mathlib/Data/Set/Partitioned.lean create mode 100644 Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean diff --git a/Mathlib/Data/Set/MemPartition.lean b/Mathlib/Data/Set/MemPartition.lean new file mode 100644 index 0000000000000..3f0a0e71e259b --- /dev/null +++ b/Mathlib/Data/Set/MemPartition.lean @@ -0,0 +1,91 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Data.Set.Finite +import Mathlib.Data.Fintype.Card + +/-! +# Partitions based on membership of a sequence of sets + +Let `f : ℕ → Set α` be a sequence of sets. For `n : ℕ`, we can form the sets of points that are in +`f 0 ∪ f 1 ∪ ... ∪ f n`; then the sets of points in `(f 0)ᶜ ∪ f 1 ∪ ... ∪ f n` and so on for +all 2^(n+1) choices of a set or its complement. The at most 2^(n+1) sets we obtain form a partition +of `univ : Set α`. We call that partition `memPartition f n` (the membership partition of `f`). + +The partition `memPartition f (n + 1)` is finer than `memPartition f n`. + +## Main definitions + +* `memPartition` + +## Main statements + +* `fooBar_unique` + +-/ + +open Set + +variable {α : Type*} + +def memPartition (f : ℕ → Set α) : ℕ → Set (Set α) + | 0 => {f 0, (f 0)ᶜ} + | n + 1 => {s | ∃ u ∈ memPartition f n, s = u ∩ f (n + 1) ∨ s = u \ f (n + 1)} + +@[simp] +lemma memPartition_zero (f : ℕ → Set α) : memPartition f 0 = {f 0, (f 0)ᶜ} := rfl + +lemma memPartition_succ (f : ℕ → Set α) (n : ℕ) : + memPartition f (n + 1) = {s | ∃ u ∈ memPartition f n, s = u ∩ f (n + 1) ∨ s = u \ f (n + 1)} := + rfl + +lemma disjoint_memPartition (f : ℕ → Set α) (n : ℕ) {u v : Set α} + (hu : u ∈ memPartition f n) (hv : v ∈ memPartition f n) (huv : u ≠ v) : + Disjoint u v := by + revert u v + induction n with + | zero => + intro u v hu hv huv + simp only [Nat.zero_eq, memPartition_zero, mem_insert_iff, mem_singleton_iff] at hu hv + rcases hu with rfl | rfl <;> rcases hv with rfl | rfl + exacts [absurd rfl huv, disjoint_compl_right, disjoint_compl_left, absurd rfl huv] + | succ n ih => + intro u v hu hv huv + rw [memPartition_succ] at hu hv + obtain ⟨u', hu', hu'_eq⟩ := hu + obtain ⟨v', hv', hv'_eq⟩ := hv + rcases hu'_eq with rfl | rfl <;> rcases hv'_eq with rfl | rfl + · refine Disjoint.mono (inter_subset_left _ _) (inter_subset_left _ _) (ih hu' hv' ?_) + exact fun huv' ↦ huv (huv' ▸ rfl) + · exact Disjoint.mono_left (inter_subset_right _ _) Set.disjoint_sdiff_right + · exact Disjoint.mono_right (inter_subset_right _ _) Set.disjoint_sdiff_left + · refine Disjoint.mono (diff_subset _ _) (diff_subset _ _) (ih hu' hv' ?_) + exact fun huv' ↦ huv (huv' ▸ rfl) + +lemma sUnion_memPartition (f : ℕ → Set α) (n : ℕ) : ⋃₀ memPartition f n = univ := by + induction n with + | zero => simp + | succ n ih => + rw [memPartition_succ] + ext x + have : x ∈ ⋃₀ memPartition f n := by simp [ih] + simp only [mem_sUnion, mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop, mem_univ, + iff_true] at this ⊢ + obtain ⟨t, ht, hxt⟩ := this + by_cases hxf : x ∈ f (n + 1) + · exact ⟨t ∩ f (n + 1), ⟨t, ht, Or.inl rfl⟩, hxt, hxf⟩ + · exact ⟨t \ f (n + 1), ⟨t, ht, Or.inr rfl⟩, hxt, hxf⟩ + +lemma memPartition_finite (f : ℕ → Set α) (n : ℕ) : Set.Finite (memPartition f n) := by + induction n with + | zero => simp + | succ n ih => + rw [memPartition_succ] + have : Finite (memPartition f n) := Set.finite_coe_iff.mp ih + rw [← Set.finite_coe_iff] + simp_rw [setOf_exists, ← exists_prop, setOf_exists, setOf_or] + refine Finite.Set.finite_biUnion (memPartition f n) _ (fun u _ ↦ ?_) + rw [Set.finite_coe_iff] + simp diff --git a/Mathlib/Data/Set/Partitioned.lean b/Mathlib/Data/Set/Partitioned.lean deleted file mode 100644 index 17c658c31589c..0000000000000 --- a/Mathlib/Data/Set/Partitioned.lean +++ /dev/null @@ -1,91 +0,0 @@ -/- -Copyright (c) 2024 Rémy Degenne. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Rémy Degenne --/ -import Mathlib.Data.Set.Finite -import Mathlib.Data.Fintype.Card - -/-! -# Partitioned - -## Main definitions - -* `partitioned` - -## Main statements - -* `fooBar_unique` - --/ - -open Set - -variable {α : Type*} - -def partitioned (f : ℕ → Set α) : ℕ → Set (Set α) - | 0 => {f 0, (f 0)ᶜ} - | n + 1 => ⋃ u ∈ partitioned f n, {u ∩ f (n + 1), u \ f (n + 1)} - -@[simp] -lemma partitioned_zero (f : ℕ → Set α) : partitioned f 0 = {f 0, (f 0)ᶜ} := rfl - -lemma partitioned_succ (f : ℕ → Set α) (n : ℕ) : - partitioned f (n + 1) = ⋃ u ∈ partitioned f n, {u ∩ f (n + 1), u \ f (n + 1)} := rfl - -lemma disjoint_partitioned (f : ℕ → Set α) (n : ℕ) : - ∀ u ∈ partitioned f n, ∀ v ∈ partitioned f n, u ≠ v → Disjoint u v := by - induction n with - | zero => - simp only [Nat.zero_eq, partitioned_zero, union_singleton, mem_insert_iff, - mem_singleton_iff, ne_eq, forall_eq_or_imp, forall_eq, not_true_eq_false, disjoint_self, - bot_eq_empty, compl_empty_iff, IsEmpty.forall_iff, true_and, and_true] - exact ⟨fun _ ↦ disjoint_compl_right, fun _ ↦ disjoint_compl_left⟩ - | succ n ih => - intro u hu v hv huv - rw [partitioned_succ] at hu hv - simp only [union_singleton, mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop] at hu hv - obtain ⟨u', hu', hu'_eq⟩ := hu - obtain ⟨v', hv', hv'_eq⟩ := hv - rcases hu'_eq with rfl | rfl <;> rcases hv'_eq with rfl | rfl - · refine Disjoint.mono (inter_subset_left _ _) (inter_subset_left _ _) (ih u' hu' v' hv' ?_) - exact fun huv' ↦ huv (huv' ▸ rfl) - · by_cases huv' : u' = v' - · rw [huv'] - exact Disjoint.mono_left (inter_subset_right _ _) Set.disjoint_sdiff_right - · exact Disjoint.mono (inter_subset_left _ _) (diff_subset _ _) (ih u' hu' v' hv' huv') - · by_cases huv' : u' = v' - · rw [huv'] - exact Disjoint.mono_right (inter_subset_right _ _) Set.disjoint_sdiff_left - · exact Disjoint.mono (diff_subset _ _) (inter_subset_left _ _) (ih u' hu' v' hv' huv') - · refine Disjoint.mono (diff_subset _ _) (diff_subset _ _) (ih u' hu' v' hv' ?_) - refine fun huv' ↦ huv (huv' ▸ rfl) - -lemma sUnion_partitioned (f : ℕ → Set α) (n : ℕ) : ⋃₀ partitioned f n = univ := by - induction n with - | zero => simp - | succ n ih => - rw [partitioned_succ] - ext x - have : x ∈ ⋃₀ partitioned f n := by simp [ih] - simp only [mem_sUnion, mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop, mem_univ, - iff_true] at this ⊢ - obtain ⟨t, ht, hxt⟩ := this - by_cases hxf : x ∈ f (n + 1) - · exact ⟨t ∩ f (n + 1), ⟨t, ht, Or.inl rfl⟩, hxt, hxf⟩ - · exact ⟨t \ f (n + 1), ⟨t, ht, Or.inr rfl⟩, hxt, hxf⟩ - -lemma partitioned_finite (f : ℕ → Set α) (n : ℕ) : Set.Finite (partitioned f n) := by - induction n with - | zero => simp - | succ n ih => - rw [partitioned_succ] - have : Finite (partitioned f n) := Set.finite_coe_iff.mp ih - rw [← Set.finite_coe_iff] - refine Finite.Set.finite_biUnion (partitioned f n) _ (fun u _ ↦ ?_) - rw [Set.finite_coe_iff] - simp - -noncomputable -instance instFintypePartitioned (f : ℕ → Set α) (n : ℕ) : Fintype (partitioned f n) := - Set.Finite.fintype (partitioned_finite f n) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 27daff9c4f7bd..245d2085eaed1 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -12,7 +12,6 @@ import Mathlib.Order.Filter.SmallSets import Mathlib.Order.Filter.CountableSeparatingOn import Mathlib.Order.LiminfLimsup import Mathlib.Data.Set.UnionLift -import Mathlib.Data.Set.Partitioned #align_import measure_theory.measurable_space from "leanprover-community/mathlib"@"001ffdc42920050657fd45bd2b8bfbec8eaaeb29" @@ -1949,186 +1948,6 @@ theorem MeasurableSpace.comap_compl {m' : MeasurableSpace β} [BooleanAlgebra β MeasurableSpace.comap_compl (fun _ _ ↦ measurableSet_top) _ #align measurable_space.comap_not MeasurableSpace.comap_not -section CountablyGenerated - -namespace MeasurableSpace - -variable (α) - -/-- We say a measurable space is countably generated -if it can be generated by a countable set of sets. -/ -class CountablyGenerated [m : MeasurableSpace α] : Prop where - isCountablyGenerated : ∃ b : Set (Set α), b.Countable ∧ m = generateFrom b -#align measurable_space.countably_generated MeasurableSpace.CountablyGenerated - -variable {α} - -theorem CountablyGenerated.comap [m : MeasurableSpace β] [h : CountablyGenerated β] (f : α → β) : - @CountablyGenerated α (.comap f m) := by - rcases h with ⟨⟨b, hbc, rfl⟩⟩ - rw [comap_generateFrom] - letI := generateFrom (preimage f '' b) - exact ⟨_, hbc.image _, rfl⟩ - -theorem CountablyGenerated.sup {m₁ m₂ : MeasurableSpace β} (h₁ : @CountablyGenerated β m₁) - (h₂ : @CountablyGenerated β m₂) : @CountablyGenerated β (m₁ ⊔ m₂) := by - rcases h₁ with ⟨⟨b₁, hb₁c, rfl⟩⟩ - rcases h₂ with ⟨⟨b₂, hb₂c, rfl⟩⟩ - exact @mk _ (_ ⊔ _) ⟨_, hb₁c.union hb₂c, generateFrom_sup_generateFrom⟩ - -instance (priority := 100) [MeasurableSpace α] [Finite α] : CountablyGenerated α where - isCountablyGenerated := - ⟨{s | MeasurableSet s}, Set.to_countable _, generateFrom_measurableSet.symm⟩ - -instance [MeasurableSpace α] [CountablyGenerated α] {p : α → Prop} : - CountablyGenerated { x // p x } := .comap _ - -instance [MeasurableSpace α] [CountablyGenerated α] [MeasurableSpace β] [CountablyGenerated β] : - CountablyGenerated (α × β) := - .sup (.comap Prod.fst) (.comap Prod.snd) - -instance [MeasurableSpace α] {s : Set α} [h : CountablyGenerated s] [MeasurableSingletonClass s] : - HasCountableSeparatingOn α MeasurableSet s := by - suffices HasCountableSeparatingOn s MeasurableSet univ from this.of_subtype fun _ ↦ id - rcases h.1 with ⟨b, hbc, hb⟩ - refine ⟨⟨b, hbc, fun t ht ↦ hb.symm ▸ .basic t ht, fun x _ y _ h ↦ ?_⟩⟩ - rw [← forall_generateFrom_mem_iff_mem_iff, ← hb] at h - simpa using h {y} - -lemma todo [m : MeasurableSpace α] [h : CountablyGenerated α] : - ∃ s : ℕ → Set (Set α), (∀ n, Set.Finite (s n)) - ∧ (∀ n, ∀ u ∈ s n, ∀ v ∈ s n, u ≠ v → Disjoint u v) - ∧ (∀ n, ⋃₀ s n = univ) - ∧ (∀ n, ∀ t ∈ s n, MeasurableSet[generateFrom (s (n + 1))] t) - ∧ m = generateFrom (⋃ n, s n) := by - obtain ⟨b, hb_count, hb_gen⟩ := h.isCountablyGenerated - cases Set.eq_empty_or_nonempty b with - | inl h_empty => - refine ⟨fun _ ↦ {univ}, by simp, by simp, by simp, ?_⟩ - simp only [iUnion_singleton_eq_range, range_const, generateFrom_singleton, mem_univ, - comap_const] - rw [hb_gen] - simp [h_empty] - | inr h_nonempty => - let s₀ := h_nonempty.choose - let t : ℕ → Set α := enumerateCountable hb_count s₀ - have ht_mem : ∀ n, t n ∈ b := enumerateCountable_mem hb_count h_nonempty.choose_spec - refine ⟨partitioned t, partitioned_finite t, disjoint_partitioned t, sUnion_partitioned t, ?_, ?_⟩ - · intro n u hun - rw [← diff_union_inter u (t (n + 1))] - refine MeasurableSet.union ?_ ?_ <;> - · refine measurableSet_generateFrom ?_ - rw [partitioned_succ] - simp_rw [mem_iUnion] - exact ⟨u, hun, by simp⟩ - · rw [hb_gen] - refine le_antisymm (generateFrom_le fun u hu ↦ ?_) (generateFrom_le fun u hu ↦ ?_) - · obtain ⟨n, rfl⟩ : ∃ n, u = t n := by - have h := subset_range_enumerate hb_count s₀ hu - obtain ⟨n, hn⟩ := mem_range.mpr h - exact ⟨n, hn.symm⟩ - suffices MeasurableSet[generateFrom (partitioned t n)] (t n) by - suffices h_le : generateFrom (partitioned t n) ≤ generateFrom (⋃ n, partitioned t n) by - exact h_le _ this - exact generateFrom_mono (subset_iUnion _ _) - cases n with - | zero => - simp only [Nat.zero_eq, partitioned_zero] - refine measurableSet_generateFrom ?_ - simp - | succ n => - have : t (n + 1) = ⋃ u ∈ partitioned t n, u ∩ t (n + 1) := by - simp_rw [← iUnion_inter] - suffices ⋃ u ∈ partitioned t n, u = univ by simp only [this, univ_inter] - rw [← sUnion_eq_biUnion] - exact sUnion_partitioned _ _ - rw [this] - refine MeasurableSet.biUnion (Finite.countable (partitioned_finite _ _)) (fun v hv ↦ ?_) - refine measurableSet_generateFrom ?_ - rw [partitioned_succ] - simp only [mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop] - exact ⟨v, hv, Or.inl rfl⟩ - · simp only [mem_iUnion] at hu - obtain ⟨n, hun⟩ := hu - induction n generalizing u with - | zero => - simp only [Nat.zero_eq, partitioned_zero, mem_insert_iff, mem_singleton_iff] at hun - rcases hun with rfl | rfl - · exact measurableSet_generateFrom (ht_mem _) - · exact (measurableSet_generateFrom (ht_mem _)).compl - | succ n ih => - simp only [partitioned_succ, mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop] - at hun - obtain ⟨v, hv, huv⟩ := hun - rcases huv with rfl | rfl - · exact MeasurableSet.inter (ih v hv) (measurableSet_generateFrom (ht_mem _)) - · exact MeasurableSet.diff (ih v hv) (measurableSet_generateFrom (ht_mem _)) - -def countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] : ℕ → Set (Set α) := - todo.choose - -lemma finite_countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : - Set.Finite (countablePartition α n) := - todo.choose_spec.1 n - -lemma disjoint_countablePartition [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) {s t : Set α} - (hs : s ∈ countablePartition α n) (ht : t ∈ countablePartition α n) (hst : s ≠ t) : - Disjoint s t := - todo.choose_spec.2.1 n s hs t ht hst - -lemma sUnion_countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : - ⋃₀ countablePartition α n = univ := - todo.choose_spec.2.2.1 n - -lemma measurableSet_succ_countablePartition [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) - {s : Set α} (hs : s ∈ countablePartition α n) : - MeasurableSet[generateFrom (countablePartition α (n + 1))] s := - todo.choose_spec.2.2.2.1 n s hs - -lemma generateFrom_iUnion_countablePartition (α : Type*) [m : MeasurableSpace α] - [CountablyGenerated α] : - generateFrom (⋃ n, countablePartition α n) = m := - todo.choose_spec.2.2.2.2.symm - -lemma generateFrom_countablePartition_le_succ (α : Type*) [MeasurableSpace α] [CountablyGenerated α] - (n : ℕ) : - generateFrom (countablePartition α n) ≤ generateFrom (countablePartition α (n + 1)) := - generateFrom_le (fun _ hs ↦ measurableSet_succ_countablePartition n hs) - -lemma generateFrom_countablePartition_le (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] - (n : ℕ) : - generateFrom (countablePartition α n) ≤ m := by - conv_rhs => rw [← generateFrom_iUnion_countablePartition α] - exact generateFrom_mono (subset_iUnion _ _) - -lemma measurableSet_countablePartition [m : MeasurableSpace α] [CountablyGenerated α] (n : ℕ) - {s : Set α} (hs : s ∈ countablePartition α n) : - MeasurableSet s := by - have : MeasurableSet[generateFrom (countablePartition α n)] s := - measurableSet_generateFrom hs - exact generateFrom_countablePartition_le α n _ this - -variable (α) - -open Classical - -/-- If a measurable space is countably generated and separates points, it admits a measurable -injection into the Cantor space `ℕ → Bool` (equipped with the product sigma algebra). -/ -theorem measurable_injection_nat_bool_of_countablyGenerated [MeasurableSpace α] - [HasCountableSeparatingOn α MeasurableSet univ] : - ∃ f : α → ℕ → Bool, Measurable f ∧ Function.Injective f := by - rcases exists_seq_separating α MeasurableSet.empty univ with ⟨e, hem, he⟩ - refine ⟨(· ∈ e ·), ?_, ?_⟩ - · rw [measurable_pi_iff] - refine fun n ↦ measurable_to_bool ?_ - simpa only [preimage, mem_singleton_iff, Bool.decide_iff, setOf_mem_eq] using hem n - · exact fun x y h ↦ he x trivial y trivial fun n ↦ decide_eq_decide.1 <| congr_fun h _ -#align measurable_space.measurable_injection_nat_bool_of_countably_generated MeasurableSpace.measurable_injection_nat_bool_of_countablyGenerated - -end MeasurableSpace - -end CountablyGenerated - namespace Filter variable [MeasurableSpace α] diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean new file mode 100644 index 0000000000000..2d9b5bdb972cb --- /dev/null +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -0,0 +1,173 @@ +/- +Copyright (c) 2023 Felix Weilacher. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Felix Weilacher, Yury Kudryashov, Rémy Degenne +-/ +import Mathlib.MeasureTheory.MeasurableSpace.Basic +import Mathlib.Data.Set.memPartition + +open Set MeasureTheory + +namespace MeasurableSpace + +variable {α β : Type*} + +/-- We say a measurable space is countably generated +if it can be generated by a countable set of sets. -/ +class CountablyGenerated (α : Type*) [m : MeasurableSpace α] : Prop where + isCountablyGenerated : ∃ b : Set (Set α), b.Countable ∧ m = generateFrom b +#align measurable_space.countably_generated MeasurableSpace.CountablyGenerated + +theorem CountablyGenerated.comap [m : MeasurableSpace β] [h : CountablyGenerated β] (f : α → β) : + @CountablyGenerated α (.comap f m) := by + rcases h with ⟨⟨b, hbc, rfl⟩⟩ + rw [comap_generateFrom] + letI := generateFrom (preimage f '' b) + exact ⟨_, hbc.image _, rfl⟩ + +theorem CountablyGenerated.sup {m₁ m₂ : MeasurableSpace β} (h₁ : @CountablyGenerated β m₁) + (h₂ : @CountablyGenerated β m₂) : @CountablyGenerated β (m₁ ⊔ m₂) := by + rcases h₁ with ⟨⟨b₁, hb₁c, rfl⟩⟩ + rcases h₂ with ⟨⟨b₂, hb₂c, rfl⟩⟩ + exact @mk _ (_ ⊔ _) ⟨_, hb₁c.union hb₂c, generateFrom_sup_generateFrom⟩ + +instance (priority := 100) [MeasurableSpace α] [Finite α] : CountablyGenerated α where + isCountablyGenerated := + ⟨{s | MeasurableSet s}, Set.to_countable _, generateFrom_measurableSet.symm⟩ + +instance [MeasurableSpace α] [CountablyGenerated α] {p : α → Prop} : + CountablyGenerated { x // p x } := .comap _ + +instance [MeasurableSpace α] [CountablyGenerated α] [MeasurableSpace β] [CountablyGenerated β] : + CountablyGenerated (α × β) := + .sup (.comap Prod.fst) (.comap Prod.snd) + +instance [MeasurableSpace α] {s : Set α} [h : CountablyGenerated s] [MeasurableSingletonClass s] : + HasCountableSeparatingOn α MeasurableSet s := by + suffices HasCountableSeparatingOn s MeasurableSet univ from this.of_subtype fun _ ↦ id + rcases h.1 with ⟨b, hbc, hb⟩ + refine ⟨⟨b, hbc, fun t ht ↦ hb.symm ▸ .basic t ht, fun x _ y _ h ↦ ?_⟩⟩ + rw [← forall_generateFrom_mem_iff_mem_iff, ← hb] at h + simpa using h {y} + +lemma exists_countablePartition [m : MeasurableSpace α] [h : CountablyGenerated α] : + ∃ s : ℕ → Set (Set α), (∀ n, Set.Finite (s n)) + ∧ (∀ n, ∀ {u v : Set α}, u ∈ s n → v ∈ s n → u ≠ v → Disjoint u v) + ∧ (∀ n, ⋃₀ s n = univ) + ∧ (∀ n, ∀ t ∈ s n, MeasurableSet[generateFrom (s (n + 1))] t) + ∧ m = generateFrom (⋃ n, s n) := by + obtain ⟨b, hb_count, hb_gen⟩ := h.isCountablyGenerated + let b' := insert ∅ b + have hb'_count : Set.Countable b' := Countable.insert _ hb_count + have hb'_gen : m = generateFrom b' := by rwa [generateFrom_insert_empty] + let t : ℕ → Set α := enumerateCountable hb'_count ∅ + -- `ht_mem` is the reason we inserted `∅` into `b` to make `b'` + have ht_mem : ∀ n, t n ∈ b' := enumerateCountable_mem hb'_count (mem_insert _ _) + refine ⟨memPartition t, memPartition_finite t, disjoint_memPartition t, + sUnion_memPartition t, ?_, ?_⟩ + · intro n u hun + rw [← diff_union_inter u (t (n + 1))] + refine MeasurableSet.union ?_ ?_ <;> + · refine measurableSet_generateFrom ?_ + rw [memPartition_succ] + exact ⟨u, hun, by simp⟩ + · rw [hb'_gen] + refine le_antisymm (generateFrom_le fun u hu ↦ ?_) (generateFrom_le fun u hu ↦ ?_) + · obtain ⟨n, rfl⟩ : ∃ n, u = t n := by + have h := subset_range_enumerate hb'_count ∅ hu + obtain ⟨n, hn⟩ := mem_range.mpr h + exact ⟨n, hn.symm⟩ + suffices MeasurableSet[generateFrom (memPartition t n)] (t n) by + suffices h_le : generateFrom (memPartition t n) ≤ generateFrom (⋃ n, memPartition t n) by + exact h_le _ this + exact generateFrom_mono (subset_iUnion _ _) + cases n with + | zero => + simp only [Nat.zero_eq, memPartition_zero] + refine measurableSet_generateFrom ?_ + simp + | succ n => + have : t (n + 1) = ⋃ u ∈ memPartition t n, u ∩ t (n + 1) := by + simp_rw [← iUnion_inter, ← sUnion_eq_biUnion, sUnion_memPartition, univ_inter] + rw [this] + refine MeasurableSet.biUnion (Finite.countable (memPartition_finite _ _)) (fun v hv ↦ ?_) + refine measurableSet_generateFrom ?_ + rw [memPartition_succ] + exact ⟨v, hv, Or.inl rfl⟩ + · simp only [mem_iUnion] at hu + obtain ⟨n, hun⟩ := hu + induction n generalizing u with + | zero => + simp only [Nat.zero_eq, memPartition_zero, mem_insert_iff, mem_singleton_iff] at hun + rcases hun with rfl | rfl + · exact measurableSet_generateFrom (ht_mem _) + · exact (measurableSet_generateFrom (ht_mem _)).compl + | succ n ih => + simp only [memPartition_succ, mem_setOf_eq] at hun + obtain ⟨v, hv, huv⟩ := hun + rcases huv with rfl | rfl + · exact (ih v hv).inter (measurableSet_generateFrom (ht_mem _)) + · exact (ih v hv).diff (measurableSet_generateFrom (ht_mem _)) + +def countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] : ℕ → Set (Set α) := + exists_countablePartition.choose + +lemma finite_countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : + Set.Finite (countablePartition α n) := + exists_countablePartition.choose_spec.1 n + +lemma disjoint_countablePartition [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) {s t : Set α} + (hs : s ∈ countablePartition α n) (ht : t ∈ countablePartition α n) (hst : s ≠ t) : + Disjoint s t := + exists_countablePartition.choose_spec.2.1 n hs ht hst + +lemma sUnion_countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : + ⋃₀ countablePartition α n = univ := + exists_countablePartition.choose_spec.2.2.1 n + +lemma measurableSet_succ_countablePartition [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) + {s : Set α} (hs : s ∈ countablePartition α n) : + MeasurableSet[generateFrom (countablePartition α (n + 1))] s := + exists_countablePartition.choose_spec.2.2.2.1 n s hs + +lemma generateFrom_iUnion_countablePartition (α : Type*) [m : MeasurableSpace α] + [CountablyGenerated α] : + generateFrom (⋃ n, countablePartition α n) = m := + exists_countablePartition.choose_spec.2.2.2.2.symm + +lemma generateFrom_countablePartition_le_succ (α : Type*) [MeasurableSpace α] [CountablyGenerated α] + (n : ℕ) : + generateFrom (countablePartition α n) ≤ generateFrom (countablePartition α (n + 1)) := + generateFrom_le (fun _ hs ↦ measurableSet_succ_countablePartition n hs) + +lemma generateFrom_countablePartition_le (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] + (n : ℕ) : + generateFrom (countablePartition α n) ≤ m := by + conv_rhs => rw [← generateFrom_iUnion_countablePartition α] + exact generateFrom_mono (subset_iUnion _ _) + +lemma measurableSet_countablePartition [m : MeasurableSpace α] [CountablyGenerated α] (n : ℕ) + {s : Set α} (hs : s ∈ countablePartition α n) : + MeasurableSet s := by + have : MeasurableSet[generateFrom (countablePartition α n)] s := + measurableSet_generateFrom hs + exact generateFrom_countablePartition_le α n _ this + +variable (α) + +open Classical + +/-- If a measurable space is countably generated and separates points, it admits a measurable +injection into the Cantor space `ℕ → Bool` (equipped with the product sigma algebra). -/ +theorem measurable_injection_nat_bool_of_countablyGenerated [MeasurableSpace α] + [HasCountableSeparatingOn α MeasurableSet univ] : + ∃ f : α → ℕ → Bool, Measurable f ∧ Function.Injective f := by + rcases exists_seq_separating α MeasurableSet.empty univ with ⟨e, hem, he⟩ + refine ⟨(· ∈ e ·), ?_, ?_⟩ + · rw [measurable_pi_iff] + refine fun n ↦ measurable_to_bool ?_ + simpa only [preimage, mem_singleton_iff, Bool.decide_iff, setOf_mem_eq] using hem n + · exact fun x y h ↦ he x trivial y trivial fun n ↦ decide_eq_decide.1 <| congr_fun h _ +#align measurable_space.measurable_injection_nat_bool_of_countably_generated MeasurableSpace.measurable_injection_nat_bool_of_countablyGenerated + +end MeasurableSpace From 6342f08197f4d77b33d2246a022f015bdc7bd4ed Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 22 Feb 2024 16:48:24 +0100 Subject: [PATCH 051/129] docstring --- Mathlib/Data/Set/MemPartition.lean | 7 +++--- .../MeasurableSpace/CountablyGenerated.lean | 24 +++++++++++++++++++ .../MeasureTheory/Measure/AEMeasurable.lean | 1 + 3 files changed, 29 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Set/MemPartition.lean b/Mathlib/Data/Set/MemPartition.lean index 3f0a0e71e259b..fb771eb72e4f9 100644 --- a/Mathlib/Data/Set/MemPartition.lean +++ b/Mathlib/Data/Set/MemPartition.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.Data.Set.Finite -import Mathlib.Data.Fintype.Card /-! # Partitions based on membership of a sequence of sets @@ -18,11 +17,13 @@ The partition `memPartition f (n + 1)` is finer than `memPartition f n`. ## Main definitions -* `memPartition` +* `memPartition f n`: the membership partition of the first `n+1` sets in `f`. ## Main statements -* `fooBar_unique` +* `disjoint_memPartition`: the sets in `memPartition f n` are disjoint +* `sUnion_memPartition`: the union of the sets in `memPartition f n` is `univ` +* `memPartition_finite`: `memPartition f n` is finite -/ diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index 2d9b5bdb972cb..427d2d0a7beff 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -6,6 +6,30 @@ Authors: Felix Weilacher, Yury Kudryashov, Rémy Degenne import Mathlib.MeasureTheory.MeasurableSpace.Basic import Mathlib.Data.Set.memPartition +/-! +# Countably generated measurable spaces + +We say a measurable space is countably generated if it can be generated by a countable set of sets. + +In such a space, we can also build a sequence of finer and finer finite measurable partitions of +the space such that the measurable space is generated by the union of all partitions. + +## Main definitions + +* `MeasurableSpace.CountablyGenerated`: class stating that a measurable space is countably + generated. +* `MeasurableSpace.countablePartition`: sequences of finer and finer partitions of + a countably generated space. + +## Main statements + +* `MeasurableSpace.measurable_injection_nat_bool_of_countablyGenerated`: if a measurable space is + countably generated and separates points, it admits a measurable injection into the Cantor space + `ℕ → Bool` (equipped with the product sigma algebra). + +-/ + + open Set MeasureTheory namespace MeasurableSpace diff --git a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean index 147b07c63e53d..af2d691e61daf 100644 --- a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.MeasureTheory.Measure.Trim +import Mathlib.MeasureTheory.MeasurableSpace.CountablyGenerated #align_import measure_theory.measure.ae_measurable from "leanprover-community/mathlib"@"3310acfa9787aa171db6d4cba3945f6f275fe9f2" From 6cdfdb436431d2a4bfb393400e8782ad4c60c3ea Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Thu, 22 Feb 2024 18:08:17 +0100 Subject: [PATCH 052/129] fix --- .../Kernel/Disintegration/Unique.lean | 34 +++++++++---------- 1 file changed, 17 insertions(+), 17 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index e3d594610bde8..16d8c718a3f74 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -57,9 +57,9 @@ lemma eq_condKernel_of_measure_eq_compProd_real {ρ : kernel α (β × ℝ)} [Is ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) = kernel.condKernel ρ (a, x) := by have huniv : ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) Set.univ = kernel.condKernel ρ (a, x) Set.univ := eq_condKernel_of_measure_eq_compProd' hκ MeasurableSet.univ a - suffices : ∀ᵐ x ∂(kernel.fst ρ a), - ∀ ⦃t⦄, MeasurableSet t → κ (a, x) t = kernel.condKernel ρ (a, x) t - · filter_upwards [this] with x hx + suffices ∀ᵐ x ∂(kernel.fst ρ a), + ∀ ⦃t⦄, MeasurableSet t → κ (a, x) t = kernel.condKernel ρ (a, x) t by + filter_upwards [this] with x hx ext t ht; exact hx ht apply MeasurableSpace.ae_induction_on_inter Real.borel_eq_generateFrom_Iic_rat Real.isPiSystem_Iic_rat @@ -83,8 +83,8 @@ theorem eq_condKernel_of_kernel_eq_compProd (κ : kernel (α × β) Ω) [IsFinit let hf := measurableEmbedding_measurableEmbedding_real Ω set ρ' : kernel α (β × ℝ) := kernel.map ρ (Prod.map id f) (measurable_id.prod_map hf.measurable) with hρ'def - have hρ' : kernel.fst ρ' = kernel.fst ρ - · ext a s hs + have hρ' : kernel.fst ρ' = kernel.fst ρ := by + ext a s hs rw [hρ'def, kernel.fst_apply' _ _ hs, kernel.fst_apply' _ _ hs, kernel.map_apply'] exacts [rfl, measurable_fst hs] have hρ'' : ∀ᵐ x ∂(kernel.fst ρ a), @@ -109,8 +109,8 @@ theorem eq_condKernel_of_kernel_eq_compProd (κ : kernel (α × β) Ω) [IsFinit rw [← Set.preimage_image_eq s hf.injective, ← Measure.map_apply hf.measurable <| hf.measurableSet_image.2 hs, hx, h _ <| hf.measurableSet_image.2 hs] - suffices : ρ' = (kernel.fst ρ ⊗ₖ (kernel.map (kernel.condKernel ρ) f hf.measurable)) - · rw [← hρ'] at this + suffices ρ' = (kernel.fst ρ ⊗ₖ (kernel.map (kernel.condKernel ρ) f hf.measurable)) by + rw [← hρ'] at this have heq := eq_condKernel_of_measure_eq_compProd_real this a rw [hρ'] at heq filter_upwards [heq] with x hx s hs @@ -156,8 +156,8 @@ lemma Measure.eq_condKernel_of_measure_eq_compProd_real {ρ : Measure (α × ℝ ∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by have huniv : ∀ᵐ x ∂ρ.fst, κ x Set.univ = ρ.condKernel x Set.univ := eq_condKernel_of_measure_eq_compProd' κ hκ MeasurableSet.univ - suffices : ∀ᵐ x ∂ρ.fst, ∀ ⦃t⦄, MeasurableSet t → κ x t = ρ.condKernel x t - · filter_upwards [this] with x hx + suffices ∀ᵐ x ∂ρ.fst, ∀ ⦃t⦄, MeasurableSet t → κ x t = ρ.condKernel x t by + filter_upwards [this] with x hx ext t ht; exact hx ht apply MeasurableSpace.ae_induction_on_inter Real.borel_eq_generateFrom_Iic_rat Real.isPiSystem_Iic_rat @@ -181,13 +181,13 @@ theorem Measure.eq_condKernel_of_measure_eq_compProd (κ : kernel α Ω) [IsFini let f := measurableEmbedding_real Ω let hf := measurableEmbedding_measurableEmbedding_real Ω set ρ' : Measure (α × ℝ) := ρ.map (Prod.map id f) with hρ'def - have hρ' : ρ'.fst = ρ.fst - · ext s hs + have hρ' : ρ'.fst = ρ.fst := by + ext s hs rw [hρ'def, Measure.fst_apply, Measure.fst_apply, Measure.map_apply] exacts [rfl, Measurable.prod measurable_fst <| hf.measurable.comp measurable_snd, measurable_fst hs, hs, hs] - have hρ'' : ∀ᵐ x ∂ρ.fst, kernel.map κ f hf.measurable x = ρ'.condKernel x - · rw [← hρ'] + have hρ'' : ∀ᵐ x ∂ρ.fst, kernel.map κ f hf.measurable x = ρ'.condKernel x := by + rw [← hρ'] refine eq_condKernel_of_measure_eq_compProd_real (kernel.map κ f hf.measurable) ?_ ext s hs conv_lhs => rw [hρ'def, hκ] @@ -196,15 +196,15 @@ theorem Measure.eq_condKernel_of_measure_eq_compProd (κ : kernel α Ω) [IsFini congr with a rw [kernel.map_apply'] exacts [rfl, measurable_prod_mk_left hs] - suffices : ∀ᵐ x ∂ρ.fst, ∀ s, MeasurableSet s → ρ'.condKernel x s = ρ.condKernel x (f ⁻¹' s) - · filter_upwards [hρ'', this] with x hx h + suffices ∀ᵐ x ∂ρ.fst, ∀ s, MeasurableSet s → ρ'.condKernel x s = ρ.condKernel x (f ⁻¹' s) by + filter_upwards [hρ'', this] with x hx h rw [kernel.map_apply] at hx ext s hs rw [← Set.preimage_image_eq s hf.injective, ← Measure.map_apply hf.measurable <| hf.measurableSet_image.2 hs, hx, h _ <| hf.measurableSet_image.2 hs] - suffices : ρ.map (Prod.map id f) = (ρ.fst ⊗ₘ (kernel.map ρ.condKernel f hf.measurable)) - · rw [← hρ'] at this + suffices ρ.map (Prod.map id f) = (ρ.fst ⊗ₘ (kernel.map ρ.condKernel f hf.measurable)) by + rw [← hρ'] at this have heq := eq_condKernel_of_measure_eq_compProd_real _ this rw [hρ'] at heq filter_upwards [heq] with x hx s hs From 570d2435255d92a6c5beb538e3115791dabddc80 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 23 Feb 2024 11:25:21 +0100 Subject: [PATCH 053/129] measurability of absolutelyContinuous and mutuallySingular --- Mathlib/Probability/Kernel/RadonNikodym.lean | 401 ++++++++++++++----- Mathlib/Probability/Kernel/WithDensity.lean | 2 +- 2 files changed, 292 insertions(+), 111 deletions(-) diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index c2dceb4b32c5c..b9a1751c626b5 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -71,34 +71,10 @@ lemma measurable_rnDerivAux_right (κ ν : kernel α γ) (a : α) : change Measurable ((fun p : α × γ ↦ kernel.rnDerivAux κ ν p.1 p.2) ∘ (fun x ↦ (a, x))) exact (measurable_rnDerivAux _ _).comp measurable_prod_mk_left -noncomputable -def kernel.rnDeriv (κ ν : kernel α γ) (a : α) (x : γ) : ℝ≥0∞ := - ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) - / ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) - -lemma kernel.rnDeriv_def (κ ν : kernel α γ) : - kernel.rnDeriv κ ν = fun a x ↦ ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) - / ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) := rfl - -lemma measurable_rnDeriv (κ ν : kernel α γ) : - Measurable (fun p : α × γ ↦ kernel.rnDeriv κ ν p.1 p.2) := - (measurable_rnDerivAux κ _).ennreal_ofReal.div - (measurable_const.sub (measurable_rnDerivAux κ _)).ennreal_ofReal - -lemma measurable_rnDeriv_right (κ ν : kernel α γ) (a : α) : - Measurable (fun x : γ ↦ kernel.rnDeriv κ ν a x) := by - change Measurable ((fun p : α × γ ↦ kernel.rnDeriv κ ν p.1 p.2) ∘ (fun x ↦ (a, x))) - exact (measurable_rnDeriv _ _).comp measurable_prod_mk_left - -lemma withDensity_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : - kernel.withDensity (κ + ν) - (fun a x ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x)) = κ := by +lemma set_lintegral_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set γ} (hs : MeasurableSet s) : + ∫⁻ x in s, ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) ∂(κ + ν) a = κ a s := by have h_le : κ ≤ κ + ν := le_add_of_nonneg_right bot_le - ext a s hs - rw [kernel.withDensity_apply'] - swap; exact (measurable_rnDerivAux _ _).ennreal_ofReal - have : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl - simp_rw [this] rw [← ofReal_integral_eq_lintegral_ofReal] · unfold kernel.rnDerivAux rw [set_integral_density (kernel.fst_map_prod_le_of_le h_le) a MeasurableSet.univ hs, @@ -109,6 +85,16 @@ lemma withDensity_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFinit · exact (integrable_density (kernel.fst_map_prod_le_of_le h_le) a MeasurableSet.univ).restrict · exact ae_of_all _ (fun x ↦ density_nonneg (kernel.fst_map_prod_le_of_le h_le) _ _ _) +lemma withDensity_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : + kernel.withDensity (κ + ν) + (fun a x ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x)) = κ := by + ext a s hs + rw [kernel.withDensity_apply'] + swap; exact (measurable_rnDerivAux _ _).ennreal_ofReal + have : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl + simp_rw [this] + exact set_lintegral_rnDerivAux κ ν a hs + lemma withDensity_one_sub_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x)) = ν := by @@ -146,104 +132,299 @@ lemma withDensity_one_sub_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] simp only [ENNReal.ofReal_le_one] exact density_le_one (kernel.fst_map_prod_le_of_le h_le) _ _ _ +def kernel.mutuallySingularSet (κ ν : kernel α γ) : Set (α × γ) := + {p | kernel.rnDerivAux κ (κ + ν) p.1 p.2 = 1} + +def kernel.mutuallySingularSetSlice (κ ν : kernel α γ) (a : α) : Set γ := + {x | kernel.rnDerivAux κ (κ + ν) a x = 1} + +lemma measurableSet_mutuallySingularSet (κ ν : kernel α γ) : + MeasurableSet (kernel.mutuallySingularSet κ ν) := + measurable_rnDerivAux κ (κ + ν) (measurableSet_singleton 1) + +lemma measurableSet_mutuallySingularSetSlice (κ ν : kernel α γ) (a : α) : + MeasurableSet (kernel.mutuallySingularSetSlice κ ν a) := + measurable_prod_mk_left (measurableSet_mutuallySingularSet κ ν) + +lemma measure_mutuallySingularSetSlice (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) : + ν a (kernel.mutuallySingularSetSlice κ ν a) = 0 := by + have h_coe : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl + suffices kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal + (1 - kernel.rnDerivAux κ (κ + ν) a x)) a {x | kernel.rnDerivAux κ (κ + ν) a x = 1} = 0 by + rwa [withDensity_one_sub_rnDerivAux κ ν] at this + simp_rw [h_coe] + rw [kernel.withDensity_apply', lintegral_eq_zero_iff, EventuallyEq, ae_restrict_iff] + rotate_left + · exact (measurable_const.sub + ((measurable_rnDerivAux _ _).comp measurable_prod_mk_left)).ennreal_ofReal + (measurableSet_singleton _) + · exact (measurable_const.sub + ((measurable_rnDerivAux _ _).comp measurable_prod_mk_left)).ennreal_ofReal + · exact (measurable_const.sub (measurable_rnDerivAux _ _)).ennreal_ofReal + refine ae_of_all _ (fun x hx ↦ ?_) + simp only [mem_setOf_eq] at hx + simp [hx] + +noncomputable +def kernel.rnDeriv (κ ν : kernel α γ) (a : α) (x : γ) : ℝ≥0∞ := + ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) + / ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) + +lemma kernel.rnDeriv_def (κ ν : kernel α γ) : + kernel.rnDeriv κ ν = fun a x ↦ ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) + / ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) := rfl + +lemma measurable_rnDeriv (κ ν : kernel α γ) : + Measurable (fun p : α × γ ↦ kernel.rnDeriv κ ν p.1 p.2) := + (measurable_rnDerivAux κ _).ennreal_ofReal.div + (measurable_const.sub (measurable_rnDerivAux κ _)).ennreal_ofReal + +lemma measurable_rnDeriv_right (κ ν : kernel α γ) (a : α) : + Measurable (fun x : γ ↦ kernel.rnDeriv κ ν a x) := by + change Measurable ((fun p : α × γ ↦ kernel.rnDeriv κ ν p.1 p.2) ∘ (fun x ↦ (a, x))) + exact (measurable_rnDeriv _ _).comp measurable_prod_mk_left + +lemma rnDeriv_eq_top_iff (κ ν : kernel α γ) (a : α) (x : γ) : + kernel.rnDeriv κ ν a x = ∞ ↔ (a, x) ∈ kernel.mutuallySingularSet κ ν := by + simp only [kernel.rnDeriv, ENNReal.div_eq_top, ne_eq, ENNReal.ofReal_eq_zero, not_le, + tsub_le_iff_right, zero_add, ENNReal.ofReal_ne_top, not_false_eq_true, and_true, or_false, + kernel.mutuallySingularSet, mem_setOf_eq] + exact ⟨fun h ↦ le_antisymm (rnDerivAux_le_one (le_add_of_nonneg_right bot_le)) h.2, + fun h ↦ by simp [h]⟩ + noncomputable def kernel.singularPart (κ ν : kernel α γ) [IsSFiniteKernel κ] [IsSFiniteKernel ν] : kernel α γ := kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x) - Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x) * kernel.rnDeriv κ ν a x) -lemma kernel.mutuallySingular_singularPart (κ ν : kernel α γ) - [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : - kernel.singularPart κ ν a ⟂ₘ ν a := by - symm +lemma measurable_singularPart_fun (κ ν : kernel α γ) [IsSFiniteKernel κ] [IsSFiniteKernel ν] : + Measurable (fun p : α × γ ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) p.1 p.2) + - Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) p.1 p.2) * kernel.rnDeriv κ ν p.1 p.2) := by + refine (measurable_rnDerivAux _ _).ennreal_ofReal.sub + (Measurable.mul ?_ (measurable_rnDeriv _ _)) + exact (measurable_const.sub (measurable_rnDerivAux _ _)).ennreal_ofReal + +lemma measurable_singularPart_fun_right (κ ν : kernel α γ) [IsSFiniteKernel κ] [IsSFiniteKernel ν] + (a : α) : + Measurable (fun x : γ ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x) + - Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x) * kernel.rnDeriv κ ν a x) := by + change Measurable ((Function.uncurry fun a b ↦ + ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a b) + - ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a b) * kernel.rnDeriv κ ν a b) + ∘ (fun b ↦ (a, b))) + exact (measurable_singularPart_fun κ ν).comp measurable_prod_mk_left + +lemma singularPart_compl_mutuallySingularSetSlice (κ ν : kernel α γ) [IsSFiniteKernel κ] + [IsSFiniteKernel ν] (a : α) : + kernel.singularPart κ ν a (kernel.mutuallySingularSetSlice κ ν a)ᶜ = 0 := by + have h_coe : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl + rw [kernel.singularPart, kernel.withDensity_apply', lintegral_eq_zero_iff, EventuallyEq, + ae_restrict_iff] + all_goals simp_rw [h_coe] + rotate_left + · exact measurableSet_preimage (measurable_singularPart_fun_right κ ν a) + (measurableSet_singleton _) + · exact measurable_singularPart_fun_right κ ν a + · exact measurable_singularPart_fun κ ν + refine ae_of_all _ (fun x hx ↦ ?_) + simp only [mem_compl_iff, mem_setOf_eq] at hx + simp_rw [kernel.rnDeriv] + rw [← ENNReal.ofReal_div_of_pos, div_eq_inv_mul, ← ENNReal.ofReal_mul, ← mul_assoc, + mul_inv_cancel, one_mul, tsub_self] + · rfl + · rw [ne_eq, sub_eq_zero] + exact Ne.symm hx + · simp [rnDerivAux_le_one (le_add_of_nonneg_right bot_le)] + · simp [lt_of_le_of_ne (rnDerivAux_le_one (le_add_of_nonneg_right bot_le)) hx] + +lemma singularPart_subset_compl_mutuallySingularSetSlice (κ ν : kernel α γ) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) (s : Set γ) + (hs : s ⊆ (kernel.mutuallySingularSetSlice κ ν a)ᶜ) : + kernel.singularPart κ ν a s = 0 := by + exact measure_mono_null hs (singularPart_compl_mutuallySingularSetSlice κ ν a) + +lemma singularPart_subset_mutuallySingularSetSlice (κ ν : kernel α γ) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) {s : Set γ} (hsm : MeasurableSet s) + (hs : s ⊆ kernel.mutuallySingularSetSlice κ ν a) : + kernel.singularPart κ ν a s = κ a s := by + have hs' : ∀ x ∈ s, kernel.rnDerivAux κ (κ + ν) a x = 1 := fun _ hx ↦ hs hx + rw [kernel.singularPart, kernel.withDensity_apply'] + swap; · exact measurable_singularPart_fun κ ν + calc + ∫⁻ x in s, ↑(Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x)) - + ↑(Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x)) * kernel.rnDeriv κ ν a x + ∂(κ + ν) a + = ∫⁻ _ in s, 1 ∂(κ + ν) a := by + refine set_lintegral_congr_fun hsm (ae_of_all _ fun x hx ↦ ?_) + simp [hs' x hx] + _ = (κ + ν) a s := by simp + _ = κ a s := by + suffices ν a s = 0 by simp [this] + exact measure_mono_null hs (measure_mutuallySingularSetSlice κ ν a) + +lemma withDensity_rnDeriv_mutuallySingularSetSlice (κ ν : kernel α γ) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) : + kernel.withDensity ν (kernel.rnDeriv κ ν) a (kernel.mutuallySingularSetSlice κ ν a) = 0 := by + rw [kernel.withDensity_apply'] + swap; · exact measurable_rnDeriv κ ν + refine le_antisymm ?_ zero_le' + calc ∫⁻ x in kernel.mutuallySingularSetSlice κ ν a, kernel.rnDeriv κ ν a x ∂(ν a) + ≤ ∫⁻ _ in kernel.mutuallySingularSetSlice κ ν a, ∞ ∂(ν a) := + set_lintegral_mono (measurable_rnDeriv_right κ ν a) measurable_const (fun _ _ ↦ le_top) + _ = ∞ * ν a (kernel.mutuallySingularSetSlice κ ν a) := by simp + _ = 0 := by + simp only [mul_eq_zero, ENNReal.top_ne_zero, false_or] + exact measure_mutuallySingularSetSlice κ ν a + +lemma withDensity_rnDeriv_subset_mutuallySingularSetSlice (κ ν : kernel α γ) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) (s : Set γ) + (hs : s ⊆ kernel.mutuallySingularSetSlice κ ν a) : + kernel.withDensity ν (kernel.rnDeriv κ ν) a s = 0 := + measure_mono_null hs (withDensity_rnDeriv_mutuallySingularSetSlice κ ν a) + +lemma withDensity_rnDeriv_subset_compl_mutuallySingularSetSlice (κ ν : kernel α γ) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set γ} (hsm : MeasurableSet s) + (hs : s ⊆ (kernel.mutuallySingularSetSlice κ ν a)ᶜ) : + kernel.withDensity ν (kernel.rnDeriv κ ν) a s = κ a s := by have h_coe : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl - refine ⟨{x | kernel.rnDerivAux κ (κ + ν) a x = 1}, - measurable_rnDerivAux_right _ _ a (measurableSet_singleton _), ?_, ?_⟩ - · suffices kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal - (1 - kernel.rnDerivAux κ (κ + ν) a x)) a {x | kernel.rnDerivAux κ (κ + ν) a x = 1} = 0 by - rwa [withDensity_one_sub_rnDerivAux κ ν] at this - simp_rw [h_coe] - rw [kernel.withDensity_apply', lintegral_eq_zero_iff, EventuallyEq, ae_restrict_iff] - rotate_left - · exact (measurable_const.sub - ((measurable_rnDerivAux _ _).comp measurable_prod_mk_left)).ennreal_ofReal - (measurableSet_singleton _) - · exact (measurable_const.sub - ((measurable_rnDerivAux _ _).comp measurable_prod_mk_left)).ennreal_ofReal - · exact (measurable_const.sub (measurable_rnDerivAux _ _)).ennreal_ofReal - refine ae_of_all _ (fun x hx ↦ ?_) - simp only [mem_setOf_eq] at hx - simp [hx] - · have h_meas' : Measurable <| Function.uncurry fun a b ↦ - ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a b) - - ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a b) * rnDeriv κ ν a b := by - refine (measurable_rnDerivAux _ _).ennreal_ofReal.sub - (Measurable.mul ?_ (measurable_rnDeriv _ _)) - exact (measurable_const.sub (measurable_rnDerivAux _ _)).ennreal_ofReal - have h_meas : Measurable fun b ↦ ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a b) - - ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a b) * rnDeriv κ ν a b := by - change Measurable ((Function.uncurry fun a b ↦ - ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a b) - - ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a b) * rnDeriv κ ν a b) - ∘ (fun b ↦ (a, b))) - exact h_meas'.comp measurable_prod_mk_left - rw [kernel.singularPart, kernel.withDensity_apply', lintegral_eq_zero_iff, EventuallyEq, - ae_restrict_iff] - all_goals simp_rw [h_coe] - rotate_left - · exact measurableSet_preimage h_meas (measurableSet_singleton _) - · exact h_meas - · exact h_meas' - refine ae_of_all _ (fun x hx ↦ ?_) - simp only [mem_compl_iff, mem_setOf_eq] at hx - simp_rw [rnDeriv] - rw [← ENNReal.ofReal_div_of_pos, div_eq_inv_mul, ← ENNReal.ofReal_mul, ← mul_assoc, - mul_inv_cancel, one_mul, tsub_self] - · rfl - · rw [ne_eq, sub_eq_zero] - exact Ne.symm hx - · simp [rnDerivAux_le_one (le_add_of_nonneg_right bot_le)] - · simp [lt_of_le_of_ne (rnDerivAux_le_one (le_add_of_nonneg_right bot_le)) hx] - -lemma kernel.rnDeriv_add_singularPart (κ ν : kernel α γ) - [IsFiniteKernel κ] [IsFiniteKernel ν] : - kernel.withDensity ν (kernel.rnDeriv κ ν) + kernel.singularPart κ ν = κ := by have : kernel.withDensity ν (kernel.rnDeriv κ ν) = kernel.withDensity (kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x))) (kernel.rnDeriv κ ν) := by rw [kernel.rnDeriv_def] congr exact (withDensity_one_sub_rnDerivAux κ ν).symm - rw [this, kernel.singularPart, add_comm, ← kernel.withDensity_mul] + rw [this, ← kernel.withDensity_mul, kernel.withDensity_apply'] rotate_left + · exact ((measurable_const.sub (measurable_rnDerivAux _ _)).ennreal_ofReal.mul + (measurable_rnDeriv _ _)) · exact (measurable_const.sub (measurable_rnDerivAux _ _)).real_toNNReal · exact measurable_rnDeriv _ _ - have h_coe : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl - have h_le : ∀ a x, ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) * kernel.rnDeriv κ ν a x - ≤ ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) := by - intro a x - unfold kernel.rnDeriv - by_cases h_one : kernel.rnDerivAux κ (κ + ν) a x = 1 - · simp [h_one] - rw [← ENNReal.ofReal_div_of_pos, div_eq_inv_mul, ← ENNReal.ofReal_mul, ← mul_assoc, - mul_inv_cancel, one_mul] - · rw [ne_eq, sub_eq_zero] - exact Ne.symm h_one - · simp [rnDerivAux_le_one (le_add_of_nonneg_right bot_le)] - · simp [lt_of_le_of_ne (rnDerivAux_le_one (le_add_of_nonneg_right bot_le)) h_one] - rw [kernel.withDensity_sub_add, withDensity_rnDerivAux] - all_goals simp_rw [h_coe] - · exact (measurable_rnDerivAux _ _).ennreal_ofReal - · exact (measurable_const.sub (measurable_rnDerivAux _ _)).ennreal_ofReal.mul - (measurable_rnDeriv _ _) - · refine fun a ↦ ne_of_lt ?_ - have : ∀ x, - ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) * kernel.rnDeriv κ ν a x ≤ 1 := by - refine fun x ↦ (h_le a x).trans ?_ - simp only [ENNReal.ofReal_le_one, rnDerivAux_le_one (le_add_of_nonneg_right bot_le)] - calc ∫⁻ x, ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) * kernel.rnDeriv κ ν a x ∂(κ + ν) a - ≤ ∫⁻ _, 1 ∂(κ + ν) a := lintegral_mono this - _ < ⊤ := by rw [MeasureTheory.lintegral_const, one_mul]; exact measure_lt_top _ _ - · exact fun a ↦ ae_of_all _ (h_le a) + simp_rw [kernel.rnDeriv] + have hs' : ∀ x ∈ s, kernel.rnDerivAux κ (κ + ν) a x < 1 := + fun x hx ↦ lt_of_le_of_ne (rnDerivAux_le_one (le_add_of_nonneg_right bot_le)) (hs hx) + calc + ∫⁻ x in s, ↑(Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x)) * + (ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) / + ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x)) ∂(κ + ν) a + _ = ∫⁻ x in s, ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) ∂(κ + ν) a := by + refine set_lintegral_congr_fun hsm (ae_of_all _ fun x hx ↦ ?_) + rw [h_coe, ← ENNReal.ofReal_div_of_pos, div_eq_inv_mul, ← ENNReal.ofReal_mul, ← mul_assoc, + mul_inv_cancel, one_mul] + · rw [ne_eq, sub_eq_zero] + exact (hs' x hx).ne' + · simp [(hs' x hx).le] + · simp [hs' x hx] + _ = κ a s := set_lintegral_rnDerivAux κ ν a hsm + +lemma kernel.mutuallySingular_singularPart (κ ν : kernel α γ) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + kernel.singularPart κ ν a ⟂ₘ ν a := by + symm + exact ⟨mutuallySingularSetSlice κ ν a, measurableSet_mutuallySingularSetSlice κ ν a, + measure_mutuallySingularSetSlice κ ν a, singularPart_compl_mutuallySingularSetSlice κ ν a⟩ + +lemma kernel.rnDeriv_add_singularPart (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : + kernel.withDensity ν (kernel.rnDeriv κ ν) + kernel.singularPart κ ν = κ := by + ext a s hs + rw [← inter_union_diff s (mutuallySingularSetSlice κ ν a)] + simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, + OuterMeasure.coe_add] + have hm := measurableSet_mutuallySingularSetSlice κ ν a + simp only [measure_union (Disjoint.mono (inter_subset_right _ _) subset_rfl disjoint_sdiff_right) + (hs.diff hm)] + rw [singularPart_subset_mutuallySingularSetSlice _ _ _ (hs.inter hm) (inter_subset_right _ _), + singularPart_subset_compl_mutuallySingularSetSlice _ _ _ _ (diff_subset_iff.mpr (by simp)), + add_zero, withDensity_rnDeriv_subset_mutuallySingularSetSlice _ _ _ _ (inter_subset_right _ _), + zero_add, withDensity_rnDeriv_subset_compl_mutuallySingularSetSlice _ _ _ (hs.diff hm) + (diff_subset_iff.mpr (by simp)), add_comm] + +lemma singularPart_eq_zero_iff_apply_eq_zero (κ ν : kernel α γ) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) : + kernel.singularPart κ ν a = 0 + ↔ kernel.singularPart κ ν a (kernel.mutuallySingularSetSlice κ ν a) = 0 := by + rw [← Measure.measure_univ_eq_zero] + have : univ = (kernel.mutuallySingularSetSlice κ ν a) + ∪ (kernel.mutuallySingularSetSlice κ ν a)ᶜ := by simp + rw [this, measure_union disjoint_compl_right (measurableSet_mutuallySingularSetSlice κ ν a).compl, + singularPart_compl_mutuallySingularSetSlice, add_zero] + +lemma withDensity_rnDeriv_eq_zero_iff_apply_eq_zero (κ ν : kernel α γ) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) : + kernel.withDensity ν (kernel.rnDeriv κ ν) a = 0 + ↔ kernel.withDensity ν (kernel.rnDeriv κ ν) a + (kernel.mutuallySingularSetSlice κ ν a)ᶜ = 0 := by + rw [← Measure.measure_univ_eq_zero] + have : univ = (kernel.mutuallySingularSetSlice κ ν a) + ∪ (kernel.mutuallySingularSetSlice κ ν a)ᶜ := by simp + rw [this, measure_union disjoint_compl_right (measurableSet_mutuallySingularSetSlice κ ν a).compl, + withDensity_rnDeriv_mutuallySingularSetSlice, zero_add] + +lemma singularPart_eq_zero_iff_absolutelyContinuous (κ ν : kernel α γ) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + kernel.singularPart κ ν a = 0 ↔ κ a ≪ ν a := by + conv_rhs => rw [← kernel.rnDeriv_add_singularPart κ ν] + simp only [kernel.coeFn_add, Pi.add_apply] + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · rw [h, add_zero] + exact kernel.withDensity_absolutelyContinuous _ _ + rw [singularPart_eq_zero_iff_apply_eq_zero] + specialize h (measure_mutuallySingularSetSlice κ ν a) + simpa only [Measure.add_toOuterMeasure, OuterMeasure.coe_add, Pi.add_apply, + withDensity_rnDeriv_mutuallySingularSetSlice κ ν, zero_add] using h + +lemma withDensity_rnDeriv_eq_zero_iff_mutuallySingular (κ ν : kernel α γ) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + kernel.withDensity ν (kernel.rnDeriv κ ν) a = 0 ↔ κ a ⟂ₘ ν a := by + conv_rhs => rw [← kernel.rnDeriv_add_singularPart κ ν] + simp only [kernel.coeFn_add, Pi.add_apply] + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · rw [h, zero_add] + exact kernel.mutuallySingular_singularPart _ _ _ + simp only [Measure.MutuallySingular.add_left_iff] at h + have h_ac := kernel.withDensity_absolutelyContinuous (κ := ν) (kernel.rnDeriv κ ν) a + have h_ms := h.1 + rw [← Measure.MutuallySingular.self_iff] + exact h_ms.mono_ac Measure.AbsolutelyContinuous.rfl h_ac + +lemma singularPart_eq_zero_iff_measure_eq_zero (κ ν : kernel α γ) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + kernel.singularPart κ ν a = 0 ↔ κ a (kernel.mutuallySingularSetSlice κ ν a) = 0 := by + have h_eq_add := kernel.rnDeriv_add_singularPart κ ν + simp_rw [kernel.ext_iff, Measure.ext_iff] at h_eq_add + specialize h_eq_add a (kernel.mutuallySingularSetSlice κ ν a) + (measurableSet_mutuallySingularSetSlice κ ν a) + simp only [kernel.coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add, + withDensity_rnDeriv_mutuallySingularSetSlice κ ν, zero_add] at h_eq_add + rw [← h_eq_add] + exact singularPart_eq_zero_iff_apply_eq_zero κ ν a + +lemma withDensity_rnDeriv_eq_zero_iff_measure_eq_zero (κ ν : kernel α γ) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + kernel.withDensity ν (kernel.rnDeriv κ ν) a = 0 + ↔ κ a (kernel.mutuallySingularSetSlice κ ν a)ᶜ = 0 := by + have h_eq_add := kernel.rnDeriv_add_singularPart κ ν + simp_rw [kernel.ext_iff, Measure.ext_iff] at h_eq_add + specialize h_eq_add a (kernel.mutuallySingularSetSlice κ ν a)ᶜ + (measurableSet_mutuallySingularSetSlice κ ν a).compl + simp only [kernel.coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add, + singularPart_compl_mutuallySingularSetSlice κ ν, add_zero] at h_eq_add + rw [← h_eq_add] + exact withDensity_rnDeriv_eq_zero_iff_apply_eq_zero κ ν a + +lemma measurableSet_absolutelyContinuous (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : + MeasurableSet {a | κ a ≪ ν a} := by + simp_rw [← singularPart_eq_zero_iff_absolutelyContinuous, + singularPart_eq_zero_iff_measure_eq_zero] + exact kernel.measurable_kernel_prod_mk_left (measurableSet_mutuallySingularSet κ ν) + (measurableSet_singleton 0) + +lemma measurableSet_mutuallySingular (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : + MeasurableSet {a | κ a ⟂ₘ ν a} := by + simp_rw [← withDensity_rnDeriv_eq_zero_iff_mutuallySingular, + withDensity_rnDeriv_eq_zero_iff_measure_eq_zero] + exact kernel.measurable_kernel_prod_mk_left (measurableSet_mutuallySingularSet κ ν).compl + (measurableSet_singleton 0) end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/WithDensity.lean b/Mathlib/Probability/Kernel/WithDensity.lean index efbe335f6bc44..ca40a5faa9d3f 100644 --- a/Mathlib/Probability/Kernel/WithDensity.lean +++ b/Mathlib/Probability/Kernel/WithDensity.lean @@ -72,7 +72,7 @@ protected theorem withDensity_apply' (κ : kernel α β) [IsSFiniteKernel κ] rw [kernel.withDensity_apply κ hf, withDensity_apply' _ s] #align probability_theory.kernel.with_density_apply' ProbabilityTheory.kernel.withDensity_apply' -nonrec lemma kernel.withDensity_absolutelyContinuous [IsSFiniteKernel κ] +nonrec lemma withDensity_absolutelyContinuous [IsSFiniteKernel κ] (f : α → β → ℝ≥0∞) (a : α) : kernel.withDensity κ f a ≪ κ a := by by_cases hf : Measurable (Function.uncurry f) From 90d8e8b29743847a41c7cfb43fd235af40f1206a Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 23 Feb 2024 12:03:40 +0100 Subject: [PATCH 054/129] rename --- .../MeasurableSpace/CountablyGenerated.lean | 2 +- .../Kernel/Disintegration/Basic.lean | 20 +-- .../Kernel/Disintegration/CDFKernel.lean | 117 +++++++++--------- Mathlib/Probability/Kernel/RadonNikodym.lean | 2 +- 4 files changed, 71 insertions(+), 70 deletions(-) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index 427d2d0a7beff..5f69b8db99da8 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Felix Weilacher. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Felix Weilacher, Yury Kudryashov, Rémy Degenne +Authors: Felix Weilacher, Yury G. Kudryashov, Rémy Degenne -/ import Mathlib.MeasureTheory.MeasurableSpace.Basic import Mathlib.Data.Set.memPartition diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index bcdcdbf0ead72..0be3c7ad2367a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -110,8 +110,8 @@ def condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFu let he := measurableEmbedding_measurableEmbedding_real Ω let x₀ := (range_nonempty e).choose kernel.comapRight - (kernel.piecewise (measurableSet_cdfKernel_eq_one hf he.measurableSet_range) - (cdfKernel f hf) (kernel.deterministic (fun _ ↦ x₀) measurable_const)) + (kernel.piecewise (measurableSet_toKernel_eq_one hf he.measurableSet_range) + (hf.toKernel f) (kernel.deterministic (fun _ ↦ x₀) measurable_const)) he instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω)) @@ -166,19 +166,19 @@ lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω)) [IsFiniteKerne rwa [he.injective he_eq] at h_mem conv_rhs => rw [this] unfold_let κ' - conv_rhs => rw [kernel.eq_compProd_cdfKernel hf] + conv_rhs => rw [kernel.eq_compProd_toKernel hf] change kernel.fst κ ⊗ₖ condKernelBorelSnd κ hf - = kernel.comapRight (kernel.fst κ' ⊗ₖ cdfKernel f hf) h_prod_embed + = kernel.comapRight (kernel.fst κ' ⊗ₖ hf.toKernel f) h_prod_embed ext c t ht : 2 rw [kernel.comapRight_apply' _ _ _ ht, kernel.compProd_apply _ _ _ (h_prod_embed.measurableSet_image.mpr ht)] simp_rw [h_fst, kernel.compProd_apply _ _ _ ht] refine lintegral_congr_ae ?_ - let ρ_set := {p : α × β | cdfKernel f hf p (range e) = 1} + let ρ_set := {p : α × β | hf.toKernel f p (range e) = 1} have h_ae : ∀ a, ∀ᵐ t ∂(kernel.fst κ a), (a, t) ∈ ρ_set := by intro a rw [← h_fst] - refine ae_cdfKernel_eq_one hf a he.measurableSet_range ?_ + refine ae_toKernel_eq_one hf a he.measurableSet_range ?_ simp only [mem_compl_iff, mem_range, not_exists] rw [kernel.map_apply'] · have h_empty : {a : β × Ω | ∀ (x : Ω), ¬e x = e a.2} = ∅ := by @@ -211,14 +211,14 @@ section StandardBorel noncomputable def kernel.condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : kernel (α × γ) ℝ := - cdfKernel _ (isKernelCDF_kernelCDF κ) + (isKernelCDF_kernelCDF κ).toKernel _ instance (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : IsMarkovKernel (kernel.condKernelReal κ) := by unfold kernel.condKernelReal; infer_instance lemma kernel.eq_compProd_condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : κ = kernel.fst κ ⊗ₖ kernel.condKernelReal κ := - kernel.eq_compProd_cdfKernel (isKernelCDF_kernelCDF κ) + kernel.eq_compProd_toKernel (isKernelCDF_kernelCDF κ) noncomputable def condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : kernel (α × γ) Ω := @@ -244,7 +244,7 @@ section Real noncomputable def condKernelUnitReal (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : kernel (Unit × α) ℝ := - cdfKernel (fun (p : Unit × α) ↦ condCDF (ρ ()) p.2) (isKernelCDF_condCDF (ρ ())) + (isKernelCDF_condCDF (ρ ())).toKernel (fun (p : Unit × α) ↦ condCDF (ρ ()) p.2) instance (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : IsMarkovKernel (condKernelUnitReal ρ) := by rw [condKernelUnitReal]; infer_instance @@ -252,7 +252,7 @@ instance (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : IsMarkovKernel (co lemma fst_compProd_condKernelUnitReal (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : kernel.fst ρ ⊗ₖ condKernelUnitReal ρ = ρ := by have : ρ = kernel.const Unit (ρ ()) := by ext; simp - conv_rhs => rw [this, kernel.eq_compProd_cdfKernel (isKernelCDF_condCDF (ρ ()))] + conv_rhs => rw [this, kernel.eq_compProd_toKernel (isKernelCDF_condCDF (ρ ()))] end Real diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean index 2f966c4628090..135b4e42f6f37 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean @@ -242,20 +242,21 @@ lemma StieltjesFunction.measurable_measure {f : α → StieltjesFunction} exact Measurable.ennreal_tsum hf_cd_meas noncomputable -def cdfKernel (f : α × β → StieltjesFunction) (hf : IsKernelCDF f μ ν) : kernel (α × β) ℝ where +def IsKernelCDF.toKernel (f : α × β → StieltjesFunction) (hf : IsKernelCDF f μ ν) : + kernel (α × β) ℝ where val p := (f p).measure property := StieltjesFunction.measurable_measure hf.measurable hf.tendsto_atBot_zero hf.tendsto_atTop_one -lemma cdfKernel_apply (p : α × β) : cdfKernel f hf p = (f p).measure := rfl +lemma IsKernelCDF.toKernel_apply (p : α × β) : hf.toKernel f p = (f p).measure := rfl -instance instIsMarkovKernel_cdfKernel : IsMarkovKernel (cdfKernel f hf) := +instance instIsMarkovKernel_toKernel : IsMarkovKernel (hf.toKernel f) := ⟨fun _ ↦ isProbabilityMeasure_stieltjesFunction (hf.tendsto_atBot_zero _) (hf.tendsto_atTop_one _)⟩ -lemma cdfKernel_Iic (p : α × β) (x : ℝ) : - cdfKernel f hf p (Iic x) = ENNReal.ofReal (f p x) := by - rw [cdfKernel_apply p, (f p).measure_Iic (hf.tendsto_atBot_zero p)] +lemma IsKernelCDF.toKernel_Iic (p : α × β) (x : ℝ) : + hf.toKernel f p (Iic x) = ENNReal.ofReal (f p x) := by + rw [IsKernelCDF.toKernel_apply p, (f p).measure_Iic (hf.tendsto_atBot_zero p)] simp end kernel @@ -266,15 +267,15 @@ variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} {f : α × β → StieltjesFunction} {μ : kernel α (β × ℝ)} {ν : kernel α β} {hf : IsKernelCDF f μ ν} -lemma set_lintegral_cdfKernel_Iic [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) +lemma set_lintegral_toKernel_Iic [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ t in s, cdfKernel f hf (a, t) (Iic x) ∂(ν a) = μ a (s ×ˢ Iic x) := by - simp_rw [cdfKernel_Iic] + ∫⁻ t in s, hf.toKernel f (a, t) (Iic x) ∂(ν a) = μ a (s ×ˢ Iic x) := by + simp_rw [IsKernelCDF.toKernel_Iic] exact hf.set_lintegral _ hs _ -lemma set_lintegral_cdfKernel_univ [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) +lemma set_lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ t in s, cdfKernel f hf (a, t) univ ∂(ν a) = μ a (s ×ˢ univ) := by + ∫⁻ t in s, hf.toKernel f (a, t) univ ∂(ν a) = μ a (s ×ˢ univ) := by rw [← Real.iUnion_Iic_rat, prod_iUnion] have h_dir : Directed (fun x y ↦ x ⊆ y) fun q : ℚ ↦ Iic (q : ℝ) := by refine Monotone.directed_le fun r r' hrr' ↦ Iic_subset_Iic.mpr ?_ @@ -285,39 +286,39 @@ lemma set_lintegral_cdfKernel_univ [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν exact mod_cast hij simp_rw [measure_iUnion_eq_iSup h_dir, measure_iUnion_eq_iSup h_dir_prod] rw [lintegral_iSup_directed] - · simp_rw [set_lintegral_cdfKernel_Iic hf _ _ hs] + · simp_rw [set_lintegral_toKernel_Iic hf _ _ hs] · refine fun q ↦ Measurable.aemeasurable ?_ exact (kernel.measurable_coe _ measurableSet_Iic).comp measurable_prod_mk_left · refine Monotone.directed_le fun i j hij t ↦ measure_mono (Iic_subset_Iic.mpr ?_) exact mod_cast hij -lemma lintegral_cdfKernel_univ [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) : - ∫⁻ t, cdfKernel f hf (a, t) univ ∂(ν a) = μ a univ := by - rw [← set_lintegral_univ, set_lintegral_cdfKernel_univ hf a MeasurableSet.univ, univ_prod_univ] +lemma lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) : + ∫⁻ t, hf.toKernel f (a, t) univ ∂(ν a) = μ a univ := by + rw [← set_lintegral_univ, set_lintegral_toKernel_univ hf a MeasurableSet.univ, univ_prod_univ] -lemma set_lintegral_cdfKernel_prod [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) +lemma set_lintegral_toKernel_prod [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set ℝ} (ht : MeasurableSet t) : - ∫⁻ x in s, cdfKernel f hf (a, x) t ∂(ν a) = μ a (s ×ˢ t) := by - -- `set_lintegral_cdfKernel_Iic` gives the result for `t = Iic x`. These sets form a + ∫⁻ x in s, hf.toKernel f (a, x) t ∂(ν a) = μ a (s ×ˢ t) := by + -- `set_lintegral_toKernel_Iic` gives the result for `t = Iic x`. These sets form a -- π-system that generates the Borel σ-algebra, hence we can get the same equality for any -- measurable set `t`. apply MeasurableSpace.induction_on_inter (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic _ _ _ _ ht · simp only [measure_empty, lintegral_const, zero_mul, prod_empty] · rintro t ⟨q, rfl⟩ - exact set_lintegral_cdfKernel_Iic hf a _ hs + exact set_lintegral_toKernel_Iic hf a _ hs · intro t ht ht_lintegral - calc ∫⁻ x in s, cdfKernel f hf (a, x) tᶜ ∂(ν a) - = ∫⁻ x in s, cdfKernel f hf (a, x) univ - cdfKernel f hf (a, x) t ∂(ν a) := by - congr with x; rw [measure_compl ht (measure_ne_top (cdfKernel f hf (a, x)) _)] - _ = ∫⁻ x in s, cdfKernel f hf (a, x) univ ∂(ν a) - - ∫⁻ x in s, cdfKernel f hf (a, x) t ∂(ν a) := by + calc ∫⁻ x in s, hf.toKernel f (a, x) tᶜ ∂(ν a) + = ∫⁻ x in s, hf.toKernel f (a, x) univ - hf.toKernel f (a, x) t ∂(ν a) := by + congr with x; rw [measure_compl ht (measure_ne_top (hf.toKernel f (a, x)) _)] + _ = ∫⁻ x in s, hf.toKernel f (a, x) univ ∂(ν a) + - ∫⁻ x in s, hf.toKernel f (a, x) t ∂(ν a) := by rw [lintegral_sub] - · exact (kernel.measurable_coe (cdfKernel f hf) ht).comp measurable_prod_mk_left + · exact (kernel.measurable_coe (hf.toKernel f) ht).comp measurable_prod_mk_left · rw [ht_lintegral] exact measure_ne_top _ _ · exact eventually_of_forall fun a ↦ measure_mono (subset_univ _) _ = μ a (s ×ˢ univ) - μ a (s ×ˢ t) := by - rw [set_lintegral_cdfKernel_univ hf a hs, ht_lintegral] + rw [set_lintegral_toKernel_univ hf a hs, ht_lintegral] _ = μ a (s ×ˢ tᶜ) := by rw [← measure_diff _ (hs.prod ht) (measure_ne_top _ _)] · rw [prod_diff_prod, compl_eq_univ_diff] @@ -335,9 +336,9 @@ lemma set_lintegral_cdfKernel_prod [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν · exact fun i ↦ ((kernel.measurable_coe _ (hf_meas i)).comp measurable_prod_mk_left).aemeasurable.restrict -lemma lintegral_cdfKernel_mem [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) +lemma lintegral_toKernel_mem [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) {s : Set (β × ℝ)} (hs : MeasurableSet s) : - ∫⁻ x, cdfKernel f hf (a, x) {y | (x, y) ∈ s} ∂(ν a) = μ a s := by + ∫⁻ x, hf.toKernel f (a, x) {y | (x, y) ∈ s} ∂(ν a) = μ a s := by -- `set_lintegral_cdfKernel_prod` gives the result for sets of the form `t₁ × t₂`. These -- sets form a π-system that generates the product σ-algebra, hence we can get the same equality -- for any measurable set `s`. @@ -350,40 +351,40 @@ lemma lintegral_cdfKernel_mem [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) intro a ha simp only [ha, prod_mk_mem_set_prod_eq, true_and_iff, setOf_mem_eq] rw [← lintegral_add_compl _ ht₁] - have h_eq1 : ∫⁻ x in t₁, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) - = ∫⁻ x in t₁, cdfKernel f hf (a, x) t₂ ∂(ν a) := by + have h_eq1 : ∫⁻ x in t₁, hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) + = ∫⁻ x in t₁, hf.toKernel f (a, x) t₂ ∂(ν a) := by refine set_lintegral_congr_fun ht₁ (eventually_of_forall fun a ha ↦ ?_) rw [h_prod_eq_snd a ha] have h_eq2 : - ∫⁻ x in t₁ᶜ, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) = 0 := by + ∫⁻ x in t₁ᶜ, hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) = 0 := by suffices h_eq_zero : - ∀ x ∈ t₁ᶜ, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} = 0 by + ∀ x ∈ t₁ᶜ, hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} = 0 by rw [set_lintegral_congr_fun ht₁.compl (eventually_of_forall h_eq_zero)] simp only [lintegral_const, zero_mul] intro a hat₁ rw [mem_compl_iff] at hat₁ simp only [hat₁, prod_mk_mem_set_prod_eq, false_and_iff, setOf_false, measure_empty] rw [h_eq1, h_eq2, add_zero] - exact set_lintegral_cdfKernel_prod hf a ht₁ ht₂ + exact set_lintegral_toKernel_prod hf a ht₁ ht₂ · intro t ht ht_eq - calc ∫⁻ x, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ tᶜ} ∂(ν a) - = ∫⁻ x, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t}ᶜ ∂(ν a) := rfl - _ = ∫⁻ x, cdfKernel f hf (a, x) univ - - cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t} ∂(ν a) := by + calc ∫⁻ x, hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ tᶜ} ∂(ν a) + = ∫⁻ x, hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t}ᶜ ∂(ν a) := rfl + _ = ∫⁻ x, hf.toKernel f (a, x) univ + - hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t} ∂(ν a) := by congr with x : 1 exact measure_compl (measurable_prod_mk_left ht) - (measure_ne_top (cdfKernel f hf (a, x)) _) - _ = ∫⁻ x, cdfKernel f hf (a, x) univ ∂(ν a) - - ∫⁻ x, cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t} ∂(ν a) := by - have h_le : (fun x ↦ cdfKernel f hf (a, x) {y : ℝ | (x, y) ∈ t}) - ≤ᵐ[ν a] fun x ↦ cdfKernel f hf (a, x) univ := + (measure_ne_top (hf.toKernel f (a, x)) _) + _ = ∫⁻ x, hf.toKernel f (a, x) univ ∂(ν a) - + ∫⁻ x, hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t} ∂(ν a) := by + have h_le : (fun x ↦ hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t}) + ≤ᵐ[ν a] fun x ↦ hf.toKernel f (a, x) univ := eventually_of_forall fun _ ↦ measure_mono (subset_univ _) rw [lintegral_sub _ _ h_le] · exact kernel.measurable_kernel_prod_mk_left' ht a refine ((lintegral_mono_ae h_le).trans_lt ?_).ne - rw [lintegral_cdfKernel_univ hf] + rw [lintegral_toKernel_univ hf] exact measure_lt_top _ univ - _ = μ a univ - μ a t := by rw [ht_eq, lintegral_cdfKernel_univ hf] + _ = μ a univ - μ a t := by rw [ht_eq, lintegral_toKernel_univ hf] _ = μ a tᶜ := (measure_compl ht (measure_ne_top _ _)).symm · intro f' hf_disj hf_meas hf_eq have h_eq : ∀ a, {x | (a, x) ∈ ⋃ i, f' i} = ⋃ i, {x | (a, x) ∈ f' i} := by @@ -398,41 +399,41 @@ lemma lintegral_cdfKernel_mem [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) intro h_mem_both suffices (a, x) ∈ ∅ by rwa [mem_empty_iff_false] at this rwa [← h_disj, mem_inter_iff] - calc ∫⁻ x, cdfKernel f hf (a, x) (⋃ i, {y | (x, y) ∈ f' i}) ∂(ν a) - = ∫⁻ x, ∑' i, cdfKernel f hf (a, x) {y | (x, y) ∈ f' i} ∂(ν a) := by + calc ∫⁻ x, hf.toKernel f (a, x) (⋃ i, {y | (x, y) ∈ f' i}) ∂(ν a) + = ∫⁻ x, ∑' i, hf.toKernel f (a, x) {y | (x, y) ∈ f' i} ∂(ν a) := by congr with x : 1 rw [measure_iUnion (h_disj x) fun i ↦ measurable_prod_mk_left (hf_meas i)] - _ = ∑' i, ∫⁻ x, cdfKernel f hf (a, x) {y | (x, y) ∈ f' i} ∂(ν a) := + _ = ∑' i, ∫⁻ x, hf.toKernel f (a, x) {y | (x, y) ∈ f' i} ∂(ν a) := lintegral_tsum fun i ↦ (kernel.measurable_kernel_prod_mk_left' (hf_meas i) a).aemeasurable _ = ∑' i, μ a (f' i) := by simp_rw [hf_eq] _ = μ a (iUnion f') := (measure_iUnion hf_disj hf_meas).symm -lemma kernel.eq_compProd_cdfKernel [IsFiniteKernel μ] [IsSFiniteKernel ν] +lemma kernel.eq_compProd_toKernel [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsKernelCDF f μ ν) : - μ = ν ⊗ₖ cdfKernel f hf := by + μ = ν ⊗ₖ hf.toKernel f := by ext a s hs - rw [kernel.compProd_apply _ _ _ hs, lintegral_cdfKernel_mem hf a hs] + rw [kernel.compProd_apply _ _ _ hs, lintegral_toKernel_mem hf a hs] -lemma ae_cdfKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsKernelCDF f μ ν) (a : α) +lemma ae_toKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsKernelCDF f μ ν) (a : α) {s : Set ℝ} (hs : MeasurableSet s) (hμs : μ a {x | x.snd ∈ sᶜ} = 0) : - ∀ᵐ t ∂(ν a), cdfKernel f hf (a, t) s = 1 := by - have h_eq : μ = ν ⊗ₖ cdfKernel f hf := kernel.eq_compProd_cdfKernel hf - have h : μ a {x | x.snd ∈ sᶜ} = (ν ⊗ₖ cdfKernel f hf) a {x | x.snd ∈ sᶜ} := by + ∀ᵐ t ∂(ν a), hf.toKernel f (a, t) s = 1 := by + have h_eq : μ = ν ⊗ₖ hf.toKernel f := kernel.eq_compProd_toKernel hf + have h : μ a {x | x.snd ∈ sᶜ} = (ν ⊗ₖ hf.toKernel f) a {x | x.snd ∈ sᶜ} := by rw [← h_eq] rw [hμs, kernel.compProd_apply] at h swap; · exact measurable_snd hs.compl rw [eq_comm, lintegral_eq_zero_iff] at h swap · simp only [mem_compl_iff, mem_setOf_eq] - change Measurable ((fun p ↦ cdfKernel f _ p {c | c ∉ s}) ∘ (fun b : β ↦ (a, b))) + change Measurable ((fun p ↦ hf.toKernel f p {c | c ∉ s}) ∘ (fun b : β ↦ (a, b))) exact (kernel.measurable_coe _ hs.compl).comp measurable_prod_mk_left simp only [mem_compl_iff, mem_setOf_eq, kernel.prodMkLeft_apply'] at h filter_upwards [h] with t ht - change cdfKernel f hf (a, t) sᶜ = 0 at ht + change hf.toKernel f (a, t) sᶜ = 0 at ht rwa [prob_compl_eq_zero_iff hs] at ht -lemma measurableSet_cdfKernel_eq_one (hf : IsKernelCDF f μ ν) {s : Set ℝ} (hs : MeasurableSet s) : - MeasurableSet {p | cdfKernel f hf p s = 1} := +lemma measurableSet_toKernel_eq_one (hf : IsKernelCDF f μ ν) {s : Set ℝ} (hs : MeasurableSet s) : + MeasurableSet {p | hf.toKernel f p s = 1} := (kernel.measurable_coe _ hs) (measurableSet_singleton 1) end diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index b9a1751c626b5..9afab82e9061f 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.Disintegration.KernelCDFReal +import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.WithDensity /-! From 64ed76adf82689d87355e54f17ec6410d939fab7 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 23 Feb 2024 14:29:25 +0100 Subject: [PATCH 055/129] polish Density.lean --- .../MeasurableSpace/CountablyGenerated.lean | 51 ++- .../Kernel/Disintegration/Density.lean | 395 ++++++++---------- .../Kernel/Disintegration/KernelCDFReal.lean | 2 + Mathlib/Probability/Kernel/RadonNikodym.lean | 2 + 4 files changed, 226 insertions(+), 224 deletions(-) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index 5f69b8db99da8..fcc1527c43a9a 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -74,7 +74,9 @@ instance [MeasurableSpace α] {s : Set α} [h : CountablyGenerated s] [Measurabl rw [← forall_generateFrom_mem_iff_mem_iff, ← hb] at h simpa using h {y} -lemma exists_countablePartition [m : MeasurableSpace α] [h : CountablyGenerated α] : +variable [m : MeasurableSpace α] [h : CountablyGenerated α] + +lemma exists_countablePartition : ∃ s : ℕ → Set (Set α), (∀ n, Set.Finite (s n)) ∧ (∀ n, ∀ {u v : Set α}, u ∈ s n → v ∈ s n → u ≠ v → Disjoint u v) ∧ (∀ n, ⋃₀ s n = univ) @@ -140,7 +142,10 @@ lemma finite_countablePartition (α : Type*) [MeasurableSpace α] [CountablyGene Set.Finite (countablePartition α n) := exists_countablePartition.choose_spec.1 n -lemma disjoint_countablePartition [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) {s t : Set α} +instance instCountable_countablePartition (n : ℕ) : Countable (countablePartition α n) := + Set.Finite.countable (MeasurableSpace.finite_countablePartition _ _) + +lemma disjoint_countablePartition (n : ℕ) {s t : Set α} (hs : s ∈ countablePartition α n) (ht : t ∈ countablePartition α n) (hst : s ≠ t) : Disjoint s t := exists_countablePartition.choose_spec.2.1 n hs ht hst @@ -149,8 +154,7 @@ lemma sUnion_countablePartition (α : Type*) [MeasurableSpace α] [CountablyGene ⋃₀ countablePartition α n = univ := exists_countablePartition.choose_spec.2.2.1 n -lemma measurableSet_succ_countablePartition [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) - {s : Set α} (hs : s ∈ countablePartition α n) : +lemma measurableSet_succ_countablePartition (n : ℕ) {s : Set α} (hs : s ∈ countablePartition α n) : MeasurableSet[generateFrom (countablePartition α (n + 1))] s := exists_countablePartition.choose_spec.2.2.2.1 n s hs @@ -170,13 +174,48 @@ lemma generateFrom_countablePartition_le (α : Type*) [m : MeasurableSpace α] [ conv_rhs => rw [← generateFrom_iUnion_countablePartition α] exact generateFrom_mono (subset_iUnion _ _) -lemma measurableSet_countablePartition [m : MeasurableSpace α] [CountablyGenerated α] (n : ℕ) - {s : Set α} (hs : s ∈ countablePartition α n) : +lemma measurableSet_countablePartition (n : ℕ) {s : Set α} (hs : s ∈ countablePartition α n) : MeasurableSet s := by have : MeasurableSet[generateFrom (countablePartition α n)] s := measurableSet_generateFrom hs exact generateFrom_countablePartition_le α n _ this +lemma existsPartitionSet_mem (n : ℕ) (t : α) : + ∃ s, s ∈ countablePartition α n ∧ t ∈ s := by + have h_univ := sUnion_countablePartition α n + have h_mem_univ := mem_univ t + rw [← h_univ] at h_mem_univ + simpa only [mem_sUnion] using h_mem_univ + +/-- The set in `countablePartition α n` to which `t : α` belongs. -/ +def partitionSet (n : ℕ) (t : α) : Set α := (existsPartitionSet_mem n t).choose + +lemma partitionSet_mem (n : ℕ) (t : α) : partitionSet n t ∈ countablePartition α n := + (existsPartitionSet_mem n t).choose_spec.1 + +lemma mem_partitionSet (n : ℕ) (t : α) : t ∈ partitionSet n t := + (existsPartitionSet_mem n t).choose_spec.2 + +lemma mem_countablePartition_iff (n : ℕ) (t : α) {s : Set α} (hs : s ∈ countablePartition α n) : + t ∈ s ↔ partitionSet n t = s := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · by_contra h_ne + have h_disj : Disjoint s (partitionSet n t) := + disjoint_countablePartition n hs (partitionSet_mem n t) (Ne.symm h_ne) + refine absurd h_disj ?_ + rw [not_disjoint_iff_nonempty_inter] + exact ⟨t, h, mem_partitionSet n t⟩ + · rw [← h] + exact mem_partitionSet n t + +lemma partitionSet_of_mem {n : ℕ} {t : α} {s : Set α} (hs : s ∈ countablePartition α n) + (ht : t ∈ s) : + partitionSet n t = s := by + rwa [← mem_countablePartition_iff n t hs] + +lemma measurableSet_partitionSet (n : ℕ) (t : α) : MeasurableSet (partitionSet n t) := + measurableSet_countablePartition n (partitionSet_mem n t) + variable (α) open Classical diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index e9a685cafbb61..cb3bd230f5416 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -16,7 +16,7 @@ for all `a : α` and all measurable sets `s : Set β` and `A : Set γ`, `∫ t in A, f a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal`. If we were interested only in a fixed `a : α`, then we could use the Radon-Nikodym derivative to -build the function `f`, as follows +build the function `f`, as follows. ``` def f (κ : kernel α (γ × β)) (ν : kernel a γ) (a : α) (t : γ) (s : Set β) : ℝ := (((κ a).restrict (univ ×ˢ s)).fst.rnDeriv (ν a) t).toReal @@ -36,7 +36,7 @@ in the σ-algebra generated by the partition at scale `n` and is measurable in ` Let `partitionFiltration γ` be the filtration of those σ-algebras for all `n : ℕ`. The functions `densityProcess κ ν n` described here are a bounded `ν`-martingale for the filtration -`partitionFiltration γ`. By Doob's L1 martingale convergence theorem, that martingale converges to +`partitionFiltration γ`. By Doob's martingale L1 convergence theorem, that martingale converges to a limit, which has a product-measurable version and satisfies the integral equality for all `A` in `⨆ n, partitionFiltration γ n`. Finally, the partitions were chosen such that that supremum is equal to the σ-algebra on `γ`, hence the equality holds for all measurable sets. @@ -49,10 +49,10 @@ We have obtained the desired density function. ## Main statements -* `ProbabilityTheory.set_integral_density`: for all measurable sets `A : Set γ` and `s : Set β`, - `∫ t in A, kernel.density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal`. -* `ProbabilityTheory.measurable_density`: the function `(p : α × γ) ↦ kernel.density κ ν p.1 p.2 s` - is measurable. +* `ProbabilityTheory.kernel.set_integral_density`: for all measurable sets `A : Set γ` and + `s : Set β`, `∫ t in A, kernel.density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal`. +* `ProbabilityTheory.kernel.measurable_density`: the function + `p : α × γ ↦ kernel.density κ ν p.1 p.2 s` is measurable. ## References @@ -61,80 +61,37 @@ The construction of the density process in this file follows the proof of Theore generated hypothesis instead of specializing to `ℝ`. -/ -variable {α β : Type*} {mα : MeasurableSpace α} - -open MeasureTheory Set Filter +open MeasureTheory Set Filter MeasurableSpace open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory -variable {α β γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} - [MeasurableSpace.CountablyGenerated γ] +variable {α β γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} [CountablyGenerated γ] + +section partitionFiltration --- todo : `Filtration` should be renamed to `filtration` -def partitionFiltration (γ : Type*) [m : MeasurableSpace γ] [MeasurableSpace.CountablyGenerated γ] : +/-- A filtration built from the measurable spaces generated by `countablePartition γ n` for +all `n : ℕ`. -/ +def partitionFiltration (γ : Type*) [m : MeasurableSpace γ] [CountablyGenerated γ] : Filtration ℕ m where - seq := fun n ↦ MeasurableSpace.generateFrom <| MeasurableSpace.countablePartition γ n - mono' := monotone_nat_of_le_succ (MeasurableSpace.generateFrom_countablePartition_le_succ _) - le' := MeasurableSpace.generateFrom_countablePartition_le γ + seq := fun n ↦ generateFrom <| countablePartition γ n + mono' := monotone_nat_of_le_succ (generateFrom_countablePartition_le_succ _) + le' := generateFrom_countablePartition_le γ lemma measurableSet_partitionFiltration_countablePartition (n : ℕ) {s : Set γ} - (hs : s ∈ MeasurableSpace.countablePartition γ n) : + (hs : s ∈ countablePartition γ n) : MeasurableSet[partitionFiltration γ n] s := - MeasurableSpace.measurableSet_generateFrom hs - -lemma existsPartitionSet_mem (n : ℕ) (t : γ) : - ∃ s, s ∈ MeasurableSpace.countablePartition γ n ∧ t ∈ s := by - have h_univ := MeasurableSpace.sUnion_countablePartition γ n - have h_mem_univ := mem_univ t - rw [← h_univ] at h_mem_univ - simpa only [mem_sUnion] using h_mem_univ - -def partitionSet (n : ℕ) (t : γ) : Set γ := - (existsPartitionSet_mem n t).choose - -lemma partitionSet_mem (n : ℕ) (t : γ) : - partitionSet n t ∈ MeasurableSpace.countablePartition γ n := - (existsPartitionSet_mem n t).choose_spec.1 - -lemma mem_partitionSet (n : ℕ) (t : γ) : t ∈ partitionSet n t := - (existsPartitionSet_mem n t).choose_spec.2 - -lemma mem_countablePartition_iff (n : ℕ) (t : γ) {s : Set γ} - (hs : s ∈ MeasurableSpace.countablePartition γ n) : - t ∈ s ↔ partitionSet n t = s := by - refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ - · by_contra h_ne - have h_disj : Disjoint s (partitionSet n t) := - MeasurableSpace.disjoint_countablePartition n hs (partitionSet_mem n t) (Ne.symm h_ne) - refine absurd h_disj ?_ - rw [not_disjoint_iff_nonempty_inter] - exact ⟨t, h, mem_partitionSet n t⟩ - · rw [← h] - exact mem_partitionSet n t - -lemma partitionSet_of_mem {n : ℕ} {t : γ} {s : Set γ} - (hs : s ∈ MeasurableSpace.countablePartition γ n) (ht : t ∈ s) : - partitionSet n t = s := by - rwa [← mem_countablePartition_iff n t hs] + measurableSet_generateFrom hs lemma measurableSet_partitionFiltration_partitionSet (n : ℕ) (t : γ) : MeasurableSet[partitionFiltration γ n] (partitionSet n t) := measurableSet_partitionFiltration_countablePartition n (partitionSet_mem n t) -lemma measurableSet_partitionSet (n : ℕ) (t : γ) : - MeasurableSet (partitionSet n t) := - (partitionFiltration γ).le n _ (measurableSet_partitionFiltration_partitionSet n t) - -instance todo_move (n : ℕ) : Countable (MeasurableSpace.countablePartition γ n) := - Set.Finite.countable (MeasurableSpace.finite_countablePartition _ _) - -lemma measurable_partitionSet_aux (n : ℕ) - (m : MeasurableSpace (MeasurableSpace.countablePartition γ n)) : - @Measurable γ (MeasurableSpace.countablePartition γ n) (partitionFiltration γ n) m +lemma measurable_partitionSet_aux (n : ℕ) (m : MeasurableSpace (countablePartition γ n)) : + @Measurable γ (countablePartition γ n) (partitionFiltration γ n) m (fun c : γ ↦ ⟨partitionSet n c, partitionSet_mem n c⟩) := by - refine @measurable_to_countable' (MeasurableSpace.countablePartition γ n) γ m _ + refine @measurable_to_countable' (countablePartition γ n) γ m _ (partitionFiltration γ n) _ (fun t ↦ ?_) rcases t with ⟨t, ht⟩ suffices MeasurableSet[partitionFiltration γ n] {x | partitionSet n x = t} by @@ -150,19 +107,23 @@ lemma measurable_partitionSet_aux (n : ℕ) lemma measurable_partitionSet (n : ℕ) : Measurable[partitionFiltration γ n] (partitionSet n) := measurable_subtype_coe.comp (measurable_partitionSet_aux _ _) -lemma iSup_partitionFiltration (β : Type*) [m : MeasurableSpace β] - [MeasurableSpace.CountablyGenerated β] : +lemma iSup_partitionFiltration (β : Type*) [m : MeasurableSpace β] [CountablyGenerated β] : ⨆ n, partitionFiltration β n = m := by - conv_rhs => rw [← MeasurableSpace.generateFrom_iUnion_countablePartition β] - rw [← MeasurableSpace.iSup_generateFrom] + conv_rhs => rw [← generateFrom_iUnion_countablePartition β] + rw [← iSup_generateFrom] rfl -variable [MeasurableSpace β] +end partitionFiltration -section DensityProcess +namespace kernel + +variable [MeasurableSpace β] {κ : kernel α (γ × β)} {ν : kernel α γ} -variable {κ : kernel α (γ × β)} {ν : kernel α γ} +section DensityProcess +/-- An `ℕ`-indexed martingale that is a density for `κ` with respect to `ν` on the sets in +`countablePartition γ n`. Used to define its limit `ProbabilityTheory.kernel.density`, which is +a density for those kernels for all measurable sets. -/ noncomputable def densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (t : γ) (s : Set β) : ℝ := @@ -177,34 +138,28 @@ lemma measurable_densityProcess_partitionFiltration_aux (κ : kernel α (γ × Measurable[MeasurableSpace.prod mα (partitionFiltration γ n)] (fun (p : α × γ) ↦ κ p.1 (partitionSet n p.2 ×ˢ s) / ν p.1 (partitionSet n p.2)) := by change Measurable[MeasurableSpace.prod mα (partitionFiltration γ n)] - ((fun (p : α × MeasurableSpace.countablePartition γ n) ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) + ((fun (p : α × countablePartition γ n) ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) ∘ (fun (p : α × γ) ↦ (p.1, ⟨partitionSet n p.2, partitionSet_mem n p.2⟩))) have h1 : @Measurable _ _ (MeasurableSpace.prod mα ⊤) _ - (fun (p : α × MeasurableSpace.countablePartition γ n) ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) := by + (fun (p : α × countablePartition γ n) ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) := by refine Measurable.div ?_ ?_ + · refine measurable_from_prod_countable (fun t ↦ ?_) + exact kernel.measurable_coe _ ((measurableSet_countablePartition _ t.prop).prod hs) · refine measurable_from_prod_countable ?_ rintro ⟨t, ht⟩ - simp only - refine kernel.measurable_coe _ (MeasurableSet.prod ?_ hs) - exact MeasurableSpace.measurableSet_countablePartition _ ht - · refine measurable_from_prod_countable ?_ - rintro ⟨t, ht⟩ - simp only - exact kernel.measurable_coe _ (MeasurableSpace.measurableSet_countablePartition _ ht) - refine h1.comp ?_ - refine measurable_fst.prod_mk ?_ - change @Measurable (α × γ) (MeasurableSpace.countablePartition γ n) + exact kernel.measurable_coe _ (measurableSet_countablePartition _ ht) + refine h1.comp (measurable_fst.prod_mk ?_) + change @Measurable (α × γ) (countablePartition γ n) (MeasurableSpace.prod mα (partitionFiltration γ n)) ⊤ ((fun c : γ ↦ ⟨partitionSet n c, partitionSet_mem n c⟩) ∘ (fun p : α × γ ↦ p.2)) - refine Measurable.comp ?_ measurable_snd - exact measurable_partitionSet_aux n ⊤ + exact (measurable_partitionSet_aux n ⊤).comp measurable_snd lemma measurable_densityProcess_aux (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) {s : Set β} (hs : MeasurableSet s) : Measurable (fun (p : α × γ) ↦ κ p.1 (partitionSet n p.2 ×ˢ s) / ν p.1 (partitionSet n p.2)) := by refine Measurable.mono (measurable_densityProcess_partitionFiltration_aux κ ν n hs) ?_ le_rfl - exact sup_le_sup le_rfl (MeasurableSpace.comap_mono ((partitionFiltration γ).le _)) + exact sup_le_sup le_rfl (comap_mono ((partitionFiltration γ).le _)) lemma measurable_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) {s : Set β} (hs : MeasurableSet s) : @@ -242,30 +197,30 @@ lemma densityProcess_nonneg (κ : kernel α (γ × β)) (ν : kernel α γ) (n : 0 ≤ densityProcess κ ν n a t s := ENNReal.toReal_nonneg -lemma apply_partitionSet_le_of_fst_le (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : +lemma apply_partitionSet_le_of_fst_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : κ a (partitionSet n t ×ˢ s) ≤ ν a (partitionSet n t) := by calc κ a (partitionSet n t ×ˢ s) - ≤ kernel.fst κ a (partitionSet n t) := by - rw [kernel.fst_apply' _ _ (measurableSet_partitionSet _ _)] + ≤ fst κ a (partitionSet n t) := by + rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] refine measure_mono (fun x ↦ ?_) simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h _ ≤ ν a (partitionSet n t) := hκν a _ -lemma densityProcess_le_one (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : +lemma densityProcess_le_one (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : densityProcess κ ν n a t s ≤ 1 := by refine ENNReal.toReal_le_of_le_ofReal zero_le_one (ENNReal.div_le_of_le_mul ?_) rw [ENNReal.ofReal_one, one_mul] exact apply_partitionSet_le_of_fst_le hκν n a t s -lemma snorm_densityProcess_le (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (s : Set β) : +lemma snorm_densityProcess_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (s : Set β) : snorm (fun t ↦ densityProcess κ ν n a t s) 1 (ν a) ≤ ν a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun x ↦ ?_))).trans ?_ · simp only [Real.norm_eq_abs, abs_of_nonneg (densityProcess_nonneg κ ν n a x s), densityProcess_le_one hκν n a x s] · simp -lemma integrable_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] (n : ℕ) +lemma integrable_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : Integrable (fun t ↦ densityProcess κ ν n a t s) (ν a) := by rw [← memℒp_one_iff_integrable] @@ -273,11 +228,11 @@ lemma integrable_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel · exact measurable_densityProcess_right κ ν n a hs · exact (snorm_densityProcess_le hκν n a s).trans_lt (measure_lt_top _ _) -lemma set_integral_densityProcess_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma set_integral_densityProcess_I (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {u : Set γ} - (hu : u ∈ MeasurableSpace.countablePartition γ n) : + (hu : u ∈ countablePartition γ n) : ∫ t in u, densityProcess κ ν n a t s ∂(ν a) = (κ a (u ×ˢ s)).toReal := by - have hu_meas : MeasurableSet u := MeasurableSpace.measurableSet_countablePartition n hu + have hu_meas : MeasurableSet u := measurableSet_countablePartition n hu simp_rw [densityProcess] rw [integral_toReal] rotate_left @@ -289,9 +244,9 @@ lemma set_integral_densityProcess_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKern · refine ae_of_all _ (fun t ↦ ?_) by_cases h0 : ν a (partitionSet n t) = 0 · suffices κ a (partitionSet n t ×ˢ s) = 0 by simp [h0, this] - have h0' : kernel.fst κ a (partitionSet n t) = 0 := + have h0' : fst κ a (partitionSet n t) = 0 := le_antisymm ((hκν a _).trans h0.le) zero_le' - rw [kernel.fst_apply' _ _ (measurableSet_partitionSet _ _)] at h0' + rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] at h0' refine measure_mono_null (fun x ↦ ?_) h0' simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h @@ -303,12 +258,11 @@ lemma set_integral_densityProcess_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKern refine set_lintegral_congr_fun hu_meas (ae_of_all _ (fun t ht ↦ ?_)) rw [partitionSet_of_mem hu ht] rw [this] - simp only [lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter] + simp only [MeasureTheory.lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter] by_cases h0 : ν a u = 0 · simp only [h0, mul_zero] - have h0' : kernel.fst κ a u = 0 := - le_antisymm ((hκν a _).trans h0.le) zero_le' - rw [kernel.fst_apply' _ _ hu_meas] at h0' + have h0' : fst κ a u = 0 := le_antisymm ((hκν a _).trans h0.le) zero_le' + rw [fst_apply' _ _ hu_meas] at h0' refine (measure_mono_null ?_ h0').symm intro p simp only [mem_prod, mem_setOf_eq, and_imp] @@ -316,23 +270,23 @@ lemma set_integral_densityProcess_I (hκν : kernel.fst κ ≤ ν) [IsFiniteKern rw [div_eq_mul_inv, mul_assoc, ENNReal.inv_mul_cancel h0, mul_one] exact measure_ne_top _ _ -lemma integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : ∫ t, densityProcess κ ν n a t s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by - rw [← integral_univ, ← MeasurableSpace.sUnion_countablePartition γ n, sUnion_eq_iUnion, + rw [← integral_univ, ← sUnion_countablePartition γ n, sUnion_eq_iUnion, iUnion_prod_const, measure_iUnion] rotate_left · intro i j hij simp only [Set.disjoint_prod, disjoint_self, bot_eq_empty] - refine Or.inl (MeasurableSpace.disjoint_countablePartition n i.prop j.prop ?_) + refine Or.inl (disjoint_countablePartition n i.prop j.prop ?_) rw [ne_eq, Subtype.coe_inj] exact hij - · exact fun k ↦ (MeasurableSpace.measurableSet_countablePartition n k.prop).prod hs + · exact fun k ↦ (measurableSet_countablePartition n k.prop).prod hs rw [integral_iUnion] rotate_left - · exact fun k ↦ MeasurableSpace.measurableSet_countablePartition n k.prop + · exact fun k ↦ measurableSet_countablePartition n k.prop · intro i j hij - refine MeasurableSpace.disjoint_countablePartition n i.prop j.prop ?_ + refine disjoint_countablePartition n i.prop j.prop ?_ rw [ne_eq, Subtype.coe_inj] exact hij · exact (integrable_densityProcess hκν n a hs).integrableOn @@ -340,19 +294,18 @@ lemma integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] congr with k rw [set_integral_densityProcess_I hκν _ _ hs k.prop] -lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet[partitionFiltration γ n] A) : ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - refine MeasurableSpace.induction_on_inter (m := partitionFiltration γ n) - (s := MeasurableSpace.countablePartition γ n) + refine induction_on_inter (m := partitionFiltration γ n) (s := countablePartition γ n) (C := fun A ↦ ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) rfl ?_ ?_ ?_ ?_ ?_ hA · rintro s hs t ht hst suffices s = t by rwa [this, inter_self] by_contra h_ne rw [← not_disjoint_iff_nonempty_inter] at hst - exact hst <| MeasurableSpace.disjoint_countablePartition n hs ht h_ne + exact hst <| disjoint_countablePartition n hs ht h_ne · simp · rintro u hu rw [set_integral_densityProcess_I hκν _ _ hs hu] @@ -381,13 +334,13 @@ lemma set_integral_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel rw [Function.onFun, Set.disjoint_prod] exact Or.inl (hf_disj hij) -lemma set_integral_densityProcess_of_le (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] +lemma set_integral_densityProcess_of_le (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] {n m : ℕ} (hnm : n ≤ m) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet[partitionFiltration γ n] A) : ∫ t in A, densityProcess κ ν m a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := set_integral_densityProcess hκν m a hs ((partitionFiltration γ).mono hnm A hA) -lemma condexp_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma condexp_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] {i j : ℕ} (hij : i ≤ j) (a : α) {s : Set β} (hs : MeasurableSet s) : (ν a)[fun t ↦ densityProcess κ ν j a t s | partitionFiltration γ i] =ᵐ[ν a] fun t ↦ densityProcess κ ν i a t s := by @@ -402,12 +355,12 @@ lemma condexp_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] · exact StronglyMeasurable.aeStronglyMeasurable' (stronglyMeasurable_partitionFiltration_densityProcess κ ν i a hs) -lemma martingale_densityProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma martingale_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : Martingale (fun n t ↦ densityProcess κ ν n a t s) (partitionFiltration γ) (ν a) := ⟨adapted_densityProcess κ ν a hs, fun _ _ h ↦ condexp_densityProcess hκν h a hs⟩ -lemma densityProcess_mono_set (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) +lemma densityProcess_mono_set (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) {s s' : Set β} (h : s ⊆ s') : densityProcess κ ν n a t s ≤ densityProcess κ ν n a t s' := by unfold densityProcess @@ -428,7 +381,7 @@ lemma densityProcess_mono_set (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) · exact h_ne_top s' lemma densityProcess_mono_kernel_left {κ' : kernel α (γ × β)} (hκκ' : κ ≤ κ') - (hκ'ν : kernel.fst κ' ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : + (hκ'ν : fst κ' ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : densityProcess κ ν n a t s ≤ densityProcess κ' ν n a t s := by unfold densityProcess by_cases h0 : ν a (partitionSet n t) = 0 @@ -448,7 +401,7 @@ lemma densityProcess_mono_kernel_left {κ' : kernel α (γ × β)} (hκκ' : κ exact fun h_top ↦ eq_top_mono h_le h_top lemma densityProcess_antitone_kernel_right {ν' : kernel α γ} - (hνν' : ν ≤ ν') (hκν : kernel.fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : + (hνν' : ν ≤ ν') (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : densityProcess κ ν' n a t s ≤ densityProcess κ ν n a t s := by unfold densityProcess have h_le : κ a (partitionSet n t ×ˢ s) @@ -469,6 +422,7 @@ lemma densityProcess_antitone_kernel_right {ν' : kernel α γ} · simp only [ne_eq, ENNReal.div_eq_top, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top +@[simp] lemma densityProcess_empty (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (t : γ) : densityProcess κ ν n a t ∅ = 0 := by simp [densityProcess] @@ -507,7 +461,7 @@ lemma tendsto_densityProcess_atTop_of_antitone (κ : kernel α (γ × β)) (ν : rw [← densityProcess_empty κ ν n a t] exact tendsto_densityProcess_atTop_empty_of_antitone κ ν n a t s hs hs_iInter hs_meas -lemma tendsto_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] +lemma tendsto_densityProcess_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : ∀ᵐ t ∂(ν a), Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop (𝓝 ((partitionFiltration γ).limitProcess @@ -518,7 +472,7 @@ lemma tendsto_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFini rw [ENNReal.coe_toNNReal] exact measure_ne_top _ _ -lemma limitProcess_mem_L1 (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma limitProcess_densityProcess_mem_L1 (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : Memℒp ((partitionFiltration γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 (ν a) := by @@ -528,11 +482,11 @@ lemma limitProcess_mem_L1 (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [Is rw [ENNReal.coe_toNNReal] exact measure_ne_top _ _ -lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] +lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : Tendsto (fun n ↦ snorm ((fun t ↦ densityProcess κ ν n a t s) - - (partitionFiltration γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) - 1 (ν a)) atTop (𝓝 0) := by + - (partitionFiltration γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) + 1 (ν a)) atTop (𝓝 0) := by refine Submartingale.tendsto_snorm_one_limitProcess ?_ ?_ · exact (martingale_densityProcess hκν a hs).submartingale · refine uniformIntegrable_of le_rfl ENNReal.one_ne_top ?_ ?_ @@ -551,88 +505,95 @@ lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : kernel.fst κ ≤ · simp lemma tendsto_snorm_one_restrict_densityProcess_limitProcess [IsFiniteKernel κ] [IsFiniteKernel ν] - (hκν : kernel.fst κ ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : + (hκν : fst κ ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : Tendsto (fun n ↦ snorm ((fun t ↦ densityProcess κ ν n a t s) - - (partitionFiltration γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) - 1 ((ν a).restrict A)) atTop (𝓝 0) := + - (partitionFiltration γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) + 1 ((ν a).restrict A)) atTop (𝓝 0) := tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) (fun _ ↦ zero_le') (fun _ ↦ snorm_restrict_le _ _ _ _) +end DensityProcess + +section Density + +/-- Density of the kernel `κ` with respect to `ν`. This is a function `α → γ → Set β → ℝ` which +is measurable on `α × γ` for all measurable sets `s : Set β` and satisfies that +`∫ t in A, density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal` for all measurable `A : Set γ`. -/ noncomputable -def kernel.density (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) (t : γ) (s : Set β) : ℝ := +def density (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) (t : γ) (s : Set β) : ℝ := limsup (fun n ↦ densityProcess κ ν n a t s) atTop -lemma density_ae_eq_limitProcess (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma density_ae_eq_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - (fun t ↦ kernel.density κ ν a t s) + (fun t ↦ density κ ν a t s) =ᵐ[ν a] (partitionFiltration γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a) := by filter_upwards [tendsto_densityProcess_limitProcess hκν a hs] with t ht using ht.limsup_eq -lemma tendsto_m_density (hκν : kernel.fst κ ≤ ν) (a : α) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma tendsto_m_density (hκν : fst κ ≤ ν) (a : α) [IsFiniteKernel κ] [IsFiniteKernel ν] {s : Set β} (hs : MeasurableSet s) : ∀ᵐ t ∂(ν a), - Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop (𝓝 (kernel.density κ ν a t s)) := by + Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop (𝓝 (density κ ν a t s)) := by filter_upwards [tendsto_densityProcess_limitProcess hκν a hs, density_ae_eq_limitProcess hκν a hs] with t h1 h2 using h2 ▸ h1 lemma measurable_density (κ : kernel α (γ × β)) (ν : kernel α γ) {s : Set β} (hs : MeasurableSet s) : - Measurable (fun (p : α × γ) ↦ kernel.density κ ν p.1 p.2 s) := + Measurable (fun (p : α × γ) ↦ density κ ν p.1 p.2 s) := measurable_limsup (fun n ↦ measurable_densityProcess κ ν n hs) lemma measurable_density_left (κ : kernel α (γ × β)) (ν : kernel α γ) (t : γ) {s : Set β} (hs : MeasurableSet s) : - Measurable (fun a ↦ kernel.density κ ν a t s) := by - change Measurable ((fun (p : α × γ) ↦ kernel.density κ ν p.1 p.2 s) ∘ (fun a ↦ (a, t))) + Measurable (fun a ↦ density κ ν a t s) := by + change Measurable ((fun (p : α × γ) ↦ density κ ν p.1 p.2 s) ∘ (fun a ↦ (a, t))) exact (measurable_density κ ν hs).comp measurable_prod_mk_right lemma measurable_density_right (κ : kernel α (γ × β)) (ν : kernel α γ) {s : Set β} (hs : MeasurableSet s) (a : α) : - Measurable (fun t ↦ kernel.density κ ν a t s) := by - change Measurable ((fun (p : α × γ) ↦ kernel.density κ ν p.1 p.2 s) ∘ (fun t ↦ (a, t))) + Measurable (fun t ↦ density κ ν a t s) := by + change Measurable ((fun (p : α × γ) ↦ density κ ν p.1 p.2 s) ∘ (fun t ↦ (a, t))) exact (measurable_density κ ν hs).comp measurable_prod_mk_left -lemma density_mono_set (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) {s s' : Set β} (h : s ⊆ s') : - kernel.density κ ν a t s ≤ kernel.density κ ν a t s' := by +lemma density_mono_set (hκν : fst κ ≤ ν) (a : α) (t : γ) {s s' : Set β} (h : s ⊆ s') : + density κ ν a t s ≤ density κ ν a t s' := by refine limsup_le_limsup ?_ ?_ ?_ · exact eventually_of_forall (fun n ↦ densityProcess_mono_set hκν n a t h) · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ -lemma density_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : - 0 ≤ kernel.density κ ν a t s := by +lemma density_nonneg (hκν : fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : + 0 ≤ density κ ν a t s := by refine le_limsup_of_frequently_le ?_ ?_ · exact frequently_of_forall (fun n ↦ densityProcess_nonneg _ _ _ _ _ _) · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ -lemma density_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : - kernel.density κ ν a t s ≤ 1 := by +lemma density_le_one (hκν : fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : + density κ ν a t s ≤ 1 := by refine limsup_le_of_le ?_ ?_ · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) · exact eventually_of_forall (fun n ↦ densityProcess_le_one hκν _ _ _ _) -lemma snorm_density_le (hκν : kernel.fst κ ≤ ν) (a : α) (s : Set β) : - snorm (fun t ↦ kernel.density κ ν a t s) 1 (ν a) ≤ ν a univ := by +lemma snorm_density_le (hκν : fst κ ≤ ν) (a : α) (s : Set β) : + snorm (fun t ↦ density κ ν a t s) 1 (ν a) ≤ ν a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ · simp only [Real.norm_eq_abs, abs_of_nonneg (density_nonneg hκν a t s), density_le_one hκν a t s] · simp -lemma integrable_density (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] +lemma integrable_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Integrable (fun t ↦ kernel.density κ ν a t s) (ν a) := by + Integrable (fun t ↦ density κ ν a t s) (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ · exact measurable_density_right κ ν hs a · exact (snorm_density_le hκν a s).trans_lt (measure_lt_top _ _) -lemma tendsto_set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : +lemma tendsto_set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : Tendsto (fun i ↦ ∫ x in A, densityProcess κ ν i a x s ∂(ν a)) atTop - (𝓝 (∫ x in A, kernel.density κ ν a x s ∂(ν a))) := by - refine tendsto_set_integral_of_L1' (μ := ν a) (fun t ↦ kernel.density κ ν a t s) + (𝓝 (∫ x in A, density κ ν a x s ∂(ν a))) := by + refine tendsto_set_integral_of_L1' (μ := ν a) (fun t ↦ density κ ν a t s) (integrable_density hκν a hs) (F := fun i t ↦ densityProcess κ ν i a t s) (l := atTop) (eventually_of_forall (fun n ↦ integrable_densityProcess hκν _ _ hs)) ?_ A refine (tendsto_congr fun n ↦ ?_).mp (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) @@ -640,15 +601,15 @@ lemma tendsto_set_integral_m (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] exact EventuallyEq.rfl.sub (density_ae_eq_limitProcess hκν a hs).symm /-- Auxiliary lemma for `set_integral_density`. -/ -lemma set_integral_density_of_measurableSet (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] +lemma set_integral_density_of_measurableSet (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet[partitionFiltration γ n] A) : - ∫ t in A, kernel.density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - suffices ∫ t in A, kernel.density κ ν a t s ∂(ν a) + ∫ t in A, density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + suffices ∫ t in A, density κ ν a t s ∂(ν a) = ∫ t in A, densityProcess κ ν n a t s ∂(ν a) by rw [this] exact set_integral_densityProcess hκν _ _ hs hA - suffices ∫ t in A, kernel.density κ ν a t s ∂(ν a) + suffices ∫ t in A, density κ ν a t s ∂(ν a) = limsup (fun i ↦ ∫ t in A, densityProcess κ ν i a t s ∂(ν a)) atTop by rw [this, ← limsup_const (α := ℕ) (f := atTop) (∫ t in A, densityProcess κ ν n a t s ∂(ν a)), limsup_congr] @@ -657,21 +618,21 @@ lemma set_integral_density_of_measurableSet (hκν : kernel.fst κ ≤ ν) [IsFi rw [set_integral_densityProcess_of_le hκν hnm _ hs hA, set_integral_densityProcess hκν _ _ hs hA] -- use L1 convergence - have h := tendsto_set_integral_m hκν a hs A + have h := tendsto_set_integral_densityProcess hκν a hs A rw [h.limsup_eq] -lemma integral_density (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - ∫ t, kernel.density κ ν a t s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by + ∫ t, density κ ν a t s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by rw [← integral_univ, set_integral_density_of_measurableSet hκν 0 a hs MeasurableSet.univ] -lemma set_integral_density (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma set_integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet A) : - ∫ t in A, kernel.density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + ∫ t in A, density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by have hA' : MeasurableSet[⨆ n, partitionFiltration γ n] A := by rwa [iSup_partitionFiltration] - refine MeasurableSpace.induction_on_inter (m := ⨆ n, partitionFiltration γ n) - (C := fun A ↦ ∫ t in A, kernel.density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) - (MeasurableSpace.measurableSpace_iSup_eq (partitionFiltration γ)) ?_ ?_ ?_ ?_ ?_ hA' + refine induction_on_inter (m := ⨆ n, partitionFiltration γ n) + (C := fun A ↦ ∫ t in A, density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) + (measurableSpace_iSup_eq (partitionFiltration γ)) ?_ ?_ ?_ ?_ ?_ hA' · rintro s ⟨n, hs⟩ t ⟨m, ht⟩ _ exact ⟨max n m, ((partitionFiltration γ).mono (le_max_left n m) _ hs).inter ((partitionFiltration γ).mono (le_max_right n m) _ ht)⟩ @@ -704,11 +665,10 @@ lemma set_integral_density (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [I exact fun i ↦ (hf i).prod hs · rwa [iSup_partitionFiltration] at hf -lemma tendsto_integral_density_of_monotone (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] - [IsFiniteKernel ν] +lemma tendsto_integral_density_of_monotone (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : - Tendsto (fun m ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop (𝓝 (κ a univ).toReal) := by + Tendsto (fun m ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop (𝓝 (κ a univ).toReal) := by simp_rw [integral_density hκν a (hs_meas _)] have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := κ a univ) ?_ swap @@ -724,11 +684,10 @@ lemma tendsto_integral_density_of_monotone (hκν : kernel.fst κ ≤ ν) [IsFin rw [← prod_iUnion, hs_iUnion] simp only [univ_prod_univ, measure_univ] -lemma tendsto_integral_density_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] - [IsFiniteKernel ν] - (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) +lemma tendsto_integral_density_of_antitone (hκν : fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : - Tendsto (fun m ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop (𝓝 0) := by + Tendsto (fun m ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop (𝓝 0) := by simp_rw [integral_density hκν a (hs_meas _)] rw [← ENNReal.zero_toReal] have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := 0) ?_ @@ -747,18 +706,18 @@ lemma tendsto_integral_density_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin rw [← prod_iInter, hs_iInter] simp only [ne_eq, prod_empty, OuterMeasure.empty', forall_exists_index] -lemma tendsto_density_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] +lemma tendsto_density_atTop_ae_of_antitone (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : - ∀ᵐ t ∂(ν a), Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 0) := by - have h_anti : ∀ t, Antitone (fun m ↦ kernel.density κ ν a t (s m)) := + ∀ᵐ t ∂(ν a), Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 0) := by + have h_anti : ∀ t, Antitone (fun m ↦ density κ ν a t (s m)) := fun t n m hnm ↦ density_mono_set hκν a t (hs hnm) - have h_le_one : ∀ m t, kernel.density κ ν a t (s m) ≤ 1 := fun m t ↦ density_le_one hκν a t (s m) + have h_le_one : ∀ m t, density κ ν a t (s m) ≤ 1 := fun m t ↦ density_le_one hκν a t (s m) -- for all `t`, `fun m ↦ density κ a (s m) t` has a limit - have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 l) := by + have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := by intro t - have h_tendsto : Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop atBot ∨ - ∃ l, Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 l) := + have h_tendsto : Tendsto (fun m ↦ density κ ν a t (s m)) atTop atBot ∨ + ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := tendsto_of_antitone (h_anti t) cases' h_tendsto with h_absurd h_tendsto · rw [tendsto_atTop_atBot_iff_of_antitone (h_anti t)] at h_absurd @@ -768,7 +727,7 @@ lemma tendsto_density_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin · exact h_tendsto -- let `F` be the pointwise limit of `fun m ↦ density κ a (s m) t` for all `t` let F : γ → ℝ := fun t ↦ (h_exists t).choose - have hF_tendsto : ∀ t, Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 (F t)) := + have hF_tendsto : ∀ t, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 (F t)) := fun t ↦ (h_exists t).choose_spec have hF_nonneg : ∀ t, 0 ≤ F t := fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ density_nonneg hκν a t (s m)) @@ -785,7 +744,7 @@ lemma tendsto_density_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin abs_of_nonneg (hF_nonneg _)] exact hF_le_one _ _ < ⊤ := by - simp only [lintegral_const, one_mul] + simp only [MeasureTheory.lintegral_const, one_mul] exact measure_lt_top _ _ -- it suffices to show that the limit `F` is 0 a.e. suffices F=ᵐ[ν a] 0 by @@ -796,75 +755,76 @@ lemma tendsto_density_atTop_ae_of_antitone (hκν : kernel.fst κ ≤ ν) [IsFin -- since `F` is nonnegative, proving that its integral is 0 is sufficient to get that -- `F` is 0 a.e. rw [← integral_eq_zero_iff_of_nonneg hF_nonneg hF_int] - have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop + have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ t, F t ∂(ν a))) := by refine integral_tendsto_of_tendsto_of_antitone ?_ hF_int ?_ ?_ · exact fun n ↦ integrable_density hκν _ (hs_meas n) · exact ae_of_all _ h_anti · exact ae_of_all _ hF_tendsto - have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop + have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop (𝓝 0) := by exact tendsto_integral_density_of_antitone hκν a s hs hs_iInter hs_meas exact tendsto_nhds_unique h_integral h_integral' section UnivFst +/-! We specialize to `ν = fst κ`, for which `density κ (fst κ) a t univ = 1` almost everywhere. -/ + lemma densityProcess_univ [IsFiniteKernel κ] (n : ℕ) (a : α) (t : γ) : - densityProcess κ (kernel.fst κ) n a t univ - = if kernel.fst κ a (partitionSet n t) = 0 then 0 else 1 := by + densityProcess κ (fst κ) n a t univ = if fst κ a (partitionSet n t) = 0 then 0 else 1 := by rw [densityProcess] - by_cases h : kernel.fst κ a (partitionSet n t) = 0 + by_cases h : fst κ a (partitionSet n t) = 0 · simp [h] by_cases h' : κ a (partitionSet n t ×ˢ univ) = 0 · simp [h'] · rw [ENNReal.div_zero h'] simp · simp only [h, ite_false] - rw [kernel.fst_apply' _ _ (measurableSet_partitionSet _ _)] + rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] have : partitionSet n t ×ˢ univ = {p : γ × β | p.1 ∈ partitionSet n t} := by ext x simp rw [this, ENNReal.div_self] · simp - · rwa [kernel.fst_apply' _ _ (measurableSet_partitionSet _ _)] at h + · rwa [fst_apply' _ _ (measurableSet_partitionSet _ _)] at h · exact measure_ne_top _ _ lemma densityProcess_univ_ae (κ : kernel α (γ × β)) [IsFiniteKernel κ] (n : ℕ) (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), densityProcess κ (kernel.fst κ) n a t univ = 1 := by + ∀ᵐ t ∂(fst κ a), densityProcess κ (fst κ) n a t univ = 1 := by rw [ae_iff] - have : {t | ¬ densityProcess κ (kernel.fst κ) n a t univ = 1} - ⊆ {t | kernel.fst κ a (partitionSet n t) = 0} := by + have : {t | ¬ densityProcess κ (fst κ) n a t univ = 1} + ⊆ {t | fst κ a (partitionSet n t) = 0} := by intro t ht simp only [mem_setOf_eq] at ht ⊢ rw [densityProcess_univ] at ht simpa using ht refine measure_mono_null this ?_ - have : {t | kernel.fst κ a (partitionSet n t) = 0} - ⊆ ⋃ (u) (_ : u ∈ MeasurableSpace.countablePartition γ n) (_ : kernel.fst κ a u = 0), u := by + have : {t | fst κ a (partitionSet n t) = 0} + ⊆ ⋃ (u) (_ : u ∈ countablePartition γ n) (_ : fst κ a u = 0), u := by intro t ht simp only [mem_setOf_eq, mem_iUnion, exists_prop] at ht ⊢ exact ⟨partitionSet n t, partitionSet_mem _ _, ht, mem_partitionSet _ _⟩ refine measure_mono_null this ?_ rw [measure_biUnion] · simp - · exact (MeasurableSpace.finite_countablePartition _ _).countable + · exact (finite_countablePartition _ _).countable · intro s hs t ht hst simp only [disjoint_iUnion_right, disjoint_iUnion_left] - exact fun _ _ ↦ MeasurableSpace.disjoint_countablePartition n hs ht hst + exact fun _ _ ↦ disjoint_countablePartition n hs ht hst · intro s hs - by_cases h : kernel.fst κ a s = 0 - · simp [h, MeasurableSpace.measurableSet_countablePartition n hs] + by_cases h : fst κ a s = 0 + · simp [h, measurableSet_countablePartition n hs] · simp [h] lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (γ × β)) (n : ℕ) (a : α) (t : γ) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) : - Tendsto (fun m ↦ densityProcess κ (kernel.fst κ) n a t (s m)) atTop - (𝓝 (densityProcess κ (kernel.fst κ) n a t univ)) := by + Tendsto (fun m ↦ densityProcess κ (fst κ) n a t (s m)) atTop + (𝓝 (densityProcess κ (fst κ) n a t univ)) := by simp_rw [densityProcess] refine (ENNReal.tendsto_toReal ?_).comp ?_ · rw [ne_eq, ENNReal.div_eq_top] push_neg - simp_rw [kernel.fst_apply' _ _ (measurableSet_partitionSet _ _)] + simp_rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] constructor · refine fun h h0 ↦ h (measure_mono_null (fun x ↦ ?_) h0) simp only [mem_prod, mem_setOf_eq, and_imp] @@ -872,8 +832,8 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (γ × β)) · refine fun h_top ↦ eq_top_mono (measure_mono (fun x ↦ ?_)) h_top simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h - by_cases h0 : kernel.fst κ a (partitionSet n t) = 0 - · rw [kernel.fst_apply' _ _ (measurableSet_partitionSet _ _)] at h0 ⊢ + by_cases h0 : fst κ a (partitionSet n t) = 0 + · rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] at h0 ⊢ suffices ∀ m, κ a (partitionSet n t ×ˢ s m) = 0 by simp only [this, h0, ENNReal.zero_div, tendsto_const_nhds_iff] suffices κ a (partitionSet n t ×ˢ univ) = 0 by @@ -897,35 +857,32 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (γ × β)) lemma tendsto_densityProcess_atTop_ae_of_monotone (κ : kernel α (γ × β)) [IsFiniteKernel κ] (n : ℕ) (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) : - ∀ᵐ t ∂(kernel.fst κ a), - Tendsto (fun m ↦ densityProcess κ (kernel.fst κ) n a t (s m)) atTop (𝓝 1) := by + ∀ᵐ t ∂(fst κ a), Tendsto (fun m ↦ densityProcess κ (fst κ) n a t (s m)) atTop (𝓝 1) := by filter_upwards [densityProcess_univ_ae κ n a] with t ht rw [← ht] exact tendsto_densityProcess_atTop_univ_of_monotone κ n a t s hs hs_iUnion lemma density_univ (κ : kernel α (γ × β)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), kernel.density κ (kernel.fst κ) a t univ = 1 := by + ∀ᵐ t ∂(fst κ a), density κ (fst κ) a t univ = 1 := by have h := fun n ↦ densityProcess_univ_ae κ n a rw [← ae_all_iff] at h filter_upwards [h] with t ht - rw [kernel.density] - simp [ht] + simp [density, ht] lemma tendsto_density_atTop_ae_of_monotone [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : - ∀ᵐ t ∂(kernel.fst κ a), - Tendsto (fun m ↦ kernel.density κ (kernel.fst κ) a t (s m)) atTop (𝓝 1) := by - let ν := kernel.fst κ - have h_mono : ∀ t, Monotone (fun m ↦ kernel.density κ (kernel.fst κ) a t (s m)) := + ∀ᵐ t ∂(fst κ a), Tendsto (fun m ↦ density κ (fst κ) a t (s m)) atTop (𝓝 1) := by + let ν := fst κ + have h_mono : ∀ t, Monotone (fun m ↦ density κ (fst κ) a t (s m)) := fun t n m hnm ↦ density_mono_set le_rfl a t (hs hnm) - have h_le_one : ∀ m t, kernel.density κ ν a t (s m) ≤ 1 := + have h_le_one : ∀ m t, density κ ν a t (s m) ≤ 1 := fun m t ↦ density_le_one le_rfl a t (s m) -- for all `t`, `fun m ↦ density κ a (s m) t` has a limit - have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 l) := by + have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := by intro t - have h_tendsto : Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop atTop ∨ - ∃ l, Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 l) := + have h_tendsto : Tendsto (fun m ↦ density κ ν a t (s m)) atTop atTop ∨ + ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := tendsto_of_monotone (h_mono t) cases' h_tendsto with h_absurd h_tendsto · rw [tendsto_atTop_atTop_iff_of_monotone (h_mono t)] at h_absurd @@ -934,7 +891,7 @@ lemma tendsto_density_atTop_ae_of_monotone [IsFiniteKernel κ] · exact h_tendsto -- let `F` be the pointwise limit of `fun m ↦ density κ a (s m) t` for all `t` let F : γ → ℝ := fun t ↦ (h_exists t).choose - have hF_tendsto : ∀ t, Tendsto (fun m ↦ kernel.density κ ν a t (s m)) atTop (𝓝 (F t)) := + have hF_tendsto : ∀ t, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 (F t)) := fun t ↦ (h_exists t).choose_spec have hF_nonneg : ∀ t, 0 ≤ F t := fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ density_nonneg le_rfl a t (s m)) @@ -950,7 +907,7 @@ lemma tendsto_density_atTop_ae_of_monotone [IsFiniteKernel κ] rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, ENNReal.ofReal_le_one, abs_of_nonneg (hF_nonneg _)] exact hF_le_one _ - _ < ⊤ := by simp only [lintegral_const, measure_univ, one_mul, measure_lt_top] + _ < ⊤ := by simp only [MeasureTheory.lintegral_const, measure_univ, one_mul, measure_lt_top] -- it suffices to show that the limit `F` is 1 a.e. suffices F =ᵐ[ν a] (fun _ ↦ 1) by filter_upwards [this] with t ht_eq @@ -959,22 +916,24 @@ lemma tendsto_density_atTop_ae_of_monotone [IsFiniteKernel κ] -- since `F` is at most 1, proving that its integral is the same as the integral of 1 will tell -- us that `F` is 1 a.e. rw [← integral_eq_iff_of_ae_le hF_int (integrable_const _) (ae_of_all _ hF_le_one)] - have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop + have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ t, F t ∂(ν a))) := by refine integral_tendsto_of_tendsto_of_monotone ?_ hF_int ?_ ?_ · exact fun n ↦ integrable_density le_rfl _ (hs_meas n) · exact ae_of_all _ h_mono · exact ae_of_all _ hF_tendsto - have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, kernel.density κ ν a t (s m) ∂(ν a)) atTop + have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop (𝓝 (∫ _, 1 ∂(ν a))) := by - rw [integral_const] + rw [MeasureTheory.integral_const] simp only [smul_eq_mul, mul_one] - rw [kernel.fst_apply' _ _ MeasurableSet.univ] + rw [fst_apply' _ _ MeasurableSet.univ] exact tendsto_integral_density_of_monotone le_rfl a s hs hs_iUnion hs_meas exact tendsto_nhds_unique h_integral h_integral' end UnivFst -end DensityProcess +end Density + +end kernel end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean index 621bd5fcc6954..ae996ad99b16c 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean @@ -31,6 +31,8 @@ open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory +open ProbabilityTheory.kernel + variable {α γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} [MeasurableSpace.CountablyGenerated γ] {κ : kernel α (γ × ℝ)} {ν : kernel α γ} diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index 9afab82e9061f..660ca62a1d168 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -43,6 +43,8 @@ open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory +open ProbabilityTheory.kernel + variable {α γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} [MeasurableSpace.CountablyGenerated γ] {κ ν : kernel α γ} From 534fa29478810f91f1ee53e30bb5a07f78488d41 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 23 Feb 2024 15:29:22 +0100 Subject: [PATCH 056/129] docstrings --- .../MeasureTheory/MeasurableSpace/CountablyGenerated.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index fcc1527c43a9a..1a6bbc4b0a3d8 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -20,6 +20,8 @@ the space such that the measurable space is generated by the union of all partit generated. * `MeasurableSpace.countablePartition`: sequences of finer and finer partitions of a countably generated space. +* `MeasurableSpace.partitionSet`: `partitionSet n t` is the set in the partition + `countablePartition α n` to which `t` belongs. ## Main statements @@ -135,6 +137,9 @@ lemma exists_countablePartition : · exact (ih v hv).inter (measurableSet_generateFrom (ht_mem _)) · exact (ih v hv).diff (measurableSet_generateFrom (ht_mem _)) +/-- For each `n : ℕ`, `countablePartition α n` is a partition of the space in at most +`2^(n+1)` sets. Each partition is finer than the preceeding one. The measurable space generated by +the union of all those partions is the measurable space `α`. -/ def countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] : ℕ → Set (Set α) := exists_countablePartition.choose From 2aa10513cef5b6f9d875ba453c481a1b2e0ca054 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 23 Feb 2024 15:30:21 +0100 Subject: [PATCH 057/129] typo --- Mathlib/Data/Set/MemPartition.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Set/MemPartition.lean b/Mathlib/Data/Set/MemPartition.lean index fb771eb72e4f9..f74879fbb65ff 100644 --- a/Mathlib/Data/Set/MemPartition.lean +++ b/Mathlib/Data/Set/MemPartition.lean @@ -8,8 +8,8 @@ import Mathlib.Data.Set.Finite /-! # Partitions based on membership of a sequence of sets -Let `f : ℕ → Set α` be a sequence of sets. For `n : ℕ`, we can form the sets of points that are in -`f 0 ∪ f 1 ∪ ... ∪ f n`; then the sets of points in `(f 0)ᶜ ∪ f 1 ∪ ... ∪ f n` and so on for +Let `f : ℕ → Set α` be a sequence of sets. For `n : ℕ`, we can form the set of points that are in +`f 0 ∪ f 1 ∪ ... ∪ f n`; then the set of points in `(f 0)ᶜ ∪ f 1 ∪ ... ∪ f n` and so on for all 2^(n+1) choices of a set or its complement. The at most 2^(n+1) sets we obtain form a partition of `univ : Set α`. We call that partition `memPartition f n` (the membership partition of `f`). From 777f745ecbe2fd6fc7e92e134f6609826f2dfd64 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 23 Feb 2024 16:39:16 +0100 Subject: [PATCH 058/129] more CountablyGenerated --- Mathlib/Data/Set/MemPartition.lean | 3 + .../MeasurableSpace/CountablyGenerated.lean | 155 +++++++++--------- 2 files changed, 81 insertions(+), 77 deletions(-) diff --git a/Mathlib/Data/Set/MemPartition.lean b/Mathlib/Data/Set/MemPartition.lean index f74879fbb65ff..57c4cd2e2e9aa 100644 --- a/Mathlib/Data/Set/MemPartition.lean +++ b/Mathlib/Data/Set/MemPartition.lean @@ -90,3 +90,6 @@ lemma memPartition_finite (f : ℕ → Set α) (n : ℕ) : Set.Finite (memPartit refine Finite.Set.finite_biUnion (memPartition f n) _ (fun u _ ↦ ?_) rw [Set.finite_coe_iff] simp + +instance instFinite_memPartition (f : ℕ → Set α) (n : ℕ) : Finite (memPartition f n) := + Set.finite_coe_iff.mp (memPartition_finite _ _) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index 1a6bbc4b0a3d8..bda069f2d2ad2 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -44,6 +44,23 @@ class CountablyGenerated (α : Type*) [m : MeasurableSpace α] : Prop where isCountablyGenerated : ∃ b : Set (Set α), b.Countable ∧ m = generateFrom b #align measurable_space.countably_generated MeasurableSpace.CountablyGenerated +/-- A countable set of sets that generate the measurable space. +We insert `∅` to ensure it is nonempty. -/ +def generatingSet (α : Type*) [MeasurableSpace α] [h : CountablyGenerated α] : Set (Set α) := + insert ∅ h.isCountablyGenerated.choose + +lemma countable_generatingSet [MeasurableSpace α] [h : CountablyGenerated α] : + Countable (generatingSet α) := + Countable.insert _ h.isCountablyGenerated.choose_spec.1 + +lemma generateFrom_generatingSet [m : MeasurableSpace α] [h : CountablyGenerated α] : + generateFrom (generatingSet α) = m := + (generateFrom_insert_empty _).trans <| h.isCountablyGenerated.choose_spec.2.symm + +lemma nonempty_generatingSet [MeasurableSpace α] [CountablyGenerated α] : + Set.Nonempty (generatingSet α) := + ⟨∅, mem_insert _ _⟩ + theorem CountablyGenerated.comap [m : MeasurableSpace β] [h : CountablyGenerated β] (f : α → β) : @CountablyGenerated α (.comap f m) := by rcases h with ⟨⟨b, hbc, rfl⟩⟩ @@ -78,101 +95,86 @@ instance [MeasurableSpace α] {s : Set α} [h : CountablyGenerated s] [Measurabl variable [m : MeasurableSpace α] [h : CountablyGenerated α] -lemma exists_countablePartition : - ∃ s : ℕ → Set (Set α), (∀ n, Set.Finite (s n)) - ∧ (∀ n, ∀ {u v : Set α}, u ∈ s n → v ∈ s n → u ≠ v → Disjoint u v) - ∧ (∀ n, ⋃₀ s n = univ) - ∧ (∀ n, ∀ t ∈ s n, MeasurableSet[generateFrom (s (n + 1))] t) - ∧ m = generateFrom (⋃ n, s n) := by - obtain ⟨b, hb_count, hb_gen⟩ := h.isCountablyGenerated - let b' := insert ∅ b - have hb'_count : Set.Countable b' := Countable.insert _ hb_count - have hb'_gen : m = generateFrom b' := by rwa [generateFrom_insert_empty] - let t : ℕ → Set α := enumerateCountable hb'_count ∅ - -- `ht_mem` is the reason we inserted `∅` into `b` to make `b'` - have ht_mem : ∀ n, t n ∈ b' := enumerateCountable_mem hb'_count (mem_insert _ _) - refine ⟨memPartition t, memPartition_finite t, disjoint_memPartition t, - sUnion_memPartition t, ?_, ?_⟩ - · intro n u hun - rw [← diff_union_inter u (t (n + 1))] - refine MeasurableSet.union ?_ ?_ <;> - · refine measurableSet_generateFrom ?_ - rw [memPartition_succ] - exact ⟨u, hun, by simp⟩ - · rw [hb'_gen] - refine le_antisymm (generateFrom_le fun u hu ↦ ?_) (generateFrom_le fun u hu ↦ ?_) - · obtain ⟨n, rfl⟩ : ∃ n, u = t n := by - have h := subset_range_enumerate hb'_count ∅ hu - obtain ⟨n, hn⟩ := mem_range.mpr h - exact ⟨n, hn.symm⟩ - suffices MeasurableSet[generateFrom (memPartition t n)] (t n) by - suffices h_le : generateFrom (memPartition t n) ≤ generateFrom (⋃ n, memPartition t n) by - exact h_le _ this - exact generateFrom_mono (subset_iUnion _ _) - cases n with - | zero => - simp only [Nat.zero_eq, memPartition_zero] - refine measurableSet_generateFrom ?_ - simp - | succ n => - have : t (n + 1) = ⋃ u ∈ memPartition t n, u ∩ t (n + 1) := by - simp_rw [← iUnion_inter, ← sUnion_eq_biUnion, sUnion_memPartition, univ_inter] - rw [this] - refine MeasurableSet.biUnion (Finite.countable (memPartition_finite _ _)) (fun v hv ↦ ?_) - refine measurableSet_generateFrom ?_ - rw [memPartition_succ] - exact ⟨v, hv, Or.inl rfl⟩ - · simp only [mem_iUnion] at hu - obtain ⟨n, hun⟩ := hu - induction n generalizing u with - | zero => - simp only [Nat.zero_eq, memPartition_zero, mem_insert_iff, mem_singleton_iff] at hun - rcases hun with rfl | rfl - · exact measurableSet_generateFrom (ht_mem _) - · exact (measurableSet_generateFrom (ht_mem _)).compl - | succ n ih => - simp only [memPartition_succ, mem_setOf_eq] at hun - obtain ⟨v, hv, huv⟩ := hun - rcases huv with rfl | rfl - · exact (ih v hv).inter (measurableSet_generateFrom (ht_mem _)) - · exact (ih v hv).diff (measurableSet_generateFrom (ht_mem _)) - /-- For each `n : ℕ`, `countablePartition α n` is a partition of the space in at most `2^(n+1)` sets. Each partition is finer than the preceeding one. The measurable space generated by the union of all those partions is the measurable space `α`. -/ def countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] : ℕ → Set (Set α) := - exists_countablePartition.choose + memPartition (enumerateCountable countable_generatingSet ∅) lemma finite_countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : Set.Finite (countablePartition α n) := - exists_countablePartition.choose_spec.1 n + memPartition_finite _ n -instance instCountable_countablePartition (n : ℕ) : Countable (countablePartition α n) := - Set.Finite.countable (MeasurableSpace.finite_countablePartition _ _) +instance instFinite_countablePartition (n : ℕ) : Finite (countablePartition α n) := + Set.finite_coe_iff.mp (MeasurableSpace.finite_countablePartition _ _) lemma disjoint_countablePartition (n : ℕ) {s t : Set α} (hs : s ∈ countablePartition α n) (ht : t ∈ countablePartition α n) (hst : s ≠ t) : Disjoint s t := - exists_countablePartition.choose_spec.2.1 n hs ht hst + disjoint_memPartition _ n hs ht hst lemma sUnion_countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : ⋃₀ countablePartition α n = univ := - exists_countablePartition.choose_spec.2.2.1 n + sUnion_memPartition _ n lemma measurableSet_succ_countablePartition (n : ℕ) {s : Set α} (hs : s ∈ countablePartition α n) : - MeasurableSet[generateFrom (countablePartition α (n + 1))] s := - exists_countablePartition.choose_spec.2.2.2.1 n s hs - -lemma generateFrom_iUnion_countablePartition (α : Type*) [m : MeasurableSpace α] - [CountablyGenerated α] : - generateFrom (⋃ n, countablePartition α n) = m := - exists_countablePartition.choose_spec.2.2.2.2.symm + MeasurableSet[generateFrom (countablePartition α (n + 1))] s := by + let t : ℕ → Set α := enumerateCountable countable_generatingSet ∅ + rw [← diff_union_inter s (t (n + 1))] + refine MeasurableSet.union ?_ ?_ <;> + · refine measurableSet_generateFrom ?_ + rw [countablePartition, memPartition_succ] + exact ⟨s, hs, by simp⟩ lemma generateFrom_countablePartition_le_succ (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : generateFrom (countablePartition α n) ≤ generateFrom (countablePartition α (n + 1)) := generateFrom_le (fun _ hs ↦ measurableSet_succ_countablePartition n hs) +lemma generateFrom_iUnion_countablePartition (α : Type*) [m : MeasurableSpace α] + [CountablyGenerated α] : + generateFrom (⋃ n, countablePartition α n) = m := by + let t : ℕ → Set α := enumerateCountable countable_generatingSet ∅ + have ht_mem : ∀ n, t n ∈ generatingSet α := + enumerateCountable_mem countable_generatingSet (mem_insert _ _) + conv_rhs => rw [← generateFrom_generatingSet (α := α)] + refine le_antisymm (generateFrom_le fun u hu ↦ ?_) (generateFrom_le fun u hu ↦ ?_) + · simp only [mem_iUnion] at hu + obtain ⟨n, hun⟩ := hu + induction n generalizing u with + | zero => + simp only [Nat.zero_eq, memPartition_zero, mem_insert_iff, mem_singleton_iff] at hun + rcases hun with rfl | rfl + · exact measurableSet_generateFrom (ht_mem _) + · exact (measurableSet_generateFrom (ht_mem _)).compl + | succ n ih => + simp only [countablePartition, memPartition_succ, mem_setOf_eq] at hun + obtain ⟨v, hv, huv⟩ := hun + rcases huv with rfl | rfl + · exact (ih v hv).inter (measurableSet_generateFrom (ht_mem _)) + · exact (ih v hv).diff (measurableSet_generateFrom (ht_mem _)) + · obtain ⟨n, rfl⟩ : ∃ n, u = t n := by + have h := subset_range_enumerate countable_generatingSet ∅ hu + obtain ⟨n, hn⟩ := mem_range.mpr h + exact ⟨n, hn.symm⟩ + suffices MeasurableSet[generateFrom (memPartition t n)] (t n) by + suffices h_le : generateFrom (memPartition t n) ≤ generateFrom (⋃ n, memPartition t n) by + exact h_le _ this + exact generateFrom_mono (subset_iUnion _ _) + cases n with + | zero => + simp only [Nat.zero_eq, memPartition_zero] + refine measurableSet_generateFrom ?_ + simp + | succ n => + have : t (n + 1) = ⋃ u ∈ memPartition t n, u ∩ t (n + 1) := by + simp_rw [← iUnion_inter, ← sUnion_eq_biUnion, sUnion_memPartition, univ_inter] + rw [this] + refine MeasurableSet.biUnion (Finite.countable (memPartition_finite _ _)) (fun v hv ↦ ?_) + refine measurableSet_generateFrom ?_ + rw [memPartition_succ] + exact ⟨v, hv, Or.inl rfl⟩ + lemma generateFrom_countablePartition_le (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : generateFrom (countablePartition α n) ≤ m := by @@ -185,21 +187,20 @@ lemma measurableSet_countablePartition (n : ℕ) {s : Set α} (hs : s ∈ counta measurableSet_generateFrom hs exact generateFrom_countablePartition_le α n _ this -lemma existsPartitionSet_mem (n : ℕ) (t : α) : - ∃ s, s ∈ countablePartition α n ∧ t ∈ s := by +lemma exists_countablePartition_mem (n : ℕ) (t : α) : ∃ s, s ∈ countablePartition α n ∧ t ∈ s := by have h_univ := sUnion_countablePartition α n have h_mem_univ := mem_univ t rw [← h_univ] at h_mem_univ simpa only [mem_sUnion] using h_mem_univ /-- The set in `countablePartition α n` to which `t : α` belongs. -/ -def partitionSet (n : ℕ) (t : α) : Set α := (existsPartitionSet_mem n t).choose +def partitionSet (n : ℕ) (t : α) : Set α := (exists_countablePartition_mem n t).choose lemma partitionSet_mem (n : ℕ) (t : α) : partitionSet n t ∈ countablePartition α n := - (existsPartitionSet_mem n t).choose_spec.1 + (exists_countablePartition_mem n t).choose_spec.1 lemma mem_partitionSet (n : ℕ) (t : α) : t ∈ partitionSet n t := - (existsPartitionSet_mem n t).choose_spec.2 + (exists_countablePartition_mem n t).choose_spec.2 lemma mem_countablePartition_iff (n : ℕ) (t : α) {s : Set α} (hs : s ∈ countablePartition α n) : t ∈ s ↔ partitionSet n t = s := by From 02fa789a2abc51c367341a3d92a5c6c6216b424e Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sat, 24 Feb 2024 12:15:42 +0100 Subject: [PATCH 059/129] remove aux lemma --- .../Kernel/Disintegration/AuxLemmas.lean | 21 ------------------- .../Disintegration/MeasurableStieltjes.lean | 4 ++-- 2 files changed, 2 insertions(+), 23 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean b/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean index 0222614c4779a..935190b5106ee 100644 --- a/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean +++ b/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean @@ -20,27 +20,6 @@ theorem Real.iInter_Iic_rat : ⋂ r : ℚ, Iic (r : ℝ) = ∅ := by exact exists_rat_lt x #align real.Inter_Iic_rat Real.iInter_Iic_rat -lemma measurableSet_tendsto_nhds {β γ ι : Type*} [MeasurableSpace β] - [TopologicalSpace γ] [PolishSpace γ] [MeasurableSpace γ] - [hγ : OpensMeasurableSpace γ] [Countable ι] {l : Filter ι} - [l.IsCountablyGenerated] {f : ι → β → γ} (hf : ∀ i, Measurable (f i)) (c : γ) : - MeasurableSet { x | Tendsto (fun n ↦ f n x) l (𝓝 c) } := sorry - -lemma measurableSet_tendsto_fun {β γ ι : Type*} [MeasurableSpace β] - [TopologicalSpace γ] [PolishSpace γ] [MeasurableSpace γ] - [hγ : OpensMeasurableSpace γ] [Countable ι] {l : Filter ι} - [l.IsCountablyGenerated] {f : ι → β → γ} (hf : ∀ i, Measurable (f i)) {g : β → γ} - (hg : Measurable g) : - MeasurableSet { x | Tendsto (fun n ↦ f n x) l (𝓝 (g x)) } := by - letI := upgradePolishSpace γ - have : { x | Tendsto (fun n ↦ f n x) l (𝓝 (g x)) } - = { x | Tendsto (fun n ↦ dist (f n x) (g x)) l (𝓝 0) } := by - ext x - simp only [Set.mem_setOf_eq] - rw [tendsto_iff_dist_tendsto_zero] - rw [this] - exact measurableSet_tendsto_nhds (fun n ↦ (hf n).dist hg) 0 - lemma Measure.iInf_Iic_gt_prod {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] {s : Set α} (hs : MeasurableSet s) (t : ℚ) : ⨅ r : { r' : ℚ // t < r' }, ρ (s ×ˢ Iic (r : ℝ)) = ρ (s ×ˢ Iic (t : ℝ)) := by diff --git a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean index 22470d6ff7d0a..3365ad471af36 100644 --- a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean +++ b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean @@ -78,9 +78,9 @@ lemma measurableSet_isRatStieltjesPoint (hf : Measurable f) : refine MeasurableSet.iInter (fun _ ↦ ?_) exact measurableSet_le hf.eval hf.eval have h2 : MeasurableSet {a | Tendsto (f a) atTop (𝓝 1)} := - measurableSet_tendsto_nhds (fun q ↦ hf.eval) 1 + measurableSet_tendsto _ (fun q ↦ hf.eval) have h3 : MeasurableSet {a | Tendsto (f a) atBot (𝓝 0)} := - measurableSet_tendsto_nhds (fun q ↦ hf.eval) 0 + measurableSet_tendsto _ (fun q ↦ hf.eval) have h4 : MeasurableSet {a | ∀ t : ℚ, ⨅ r : Ioi t, f a r = f a t} := by rw [Set.setOf_forall] refine MeasurableSet.iInter (fun q ↦ ?_) From 93049202525f689eda549caf9f23d493c36efba4 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sat, 24 Feb 2024 15:47:00 +0100 Subject: [PATCH 060/129] split CountablyGenerated file about filtration --- Mathlib/Data/Set/MemPartition.lean | 7 +- .../MeasurableSpace/CountablyGenerated.lean | 69 ++++--- .../Kernel/Disintegration/CondCdf.lean | 7 - .../Kernel/Disintegration/Density.lean | 173 +++++++----------- Mathlib/Probability/Kernel/RadonNikodym.lean | 14 +- .../Process/CountablyGenerated.lean | 89 +++++++++ 6 files changed, 191 insertions(+), 168 deletions(-) create mode 100644 Mathlib/Probability/Process/CountablyGenerated.lean diff --git a/Mathlib/Data/Set/MemPartition.lean b/Mathlib/Data/Set/MemPartition.lean index 57c4cd2e2e9aa..3993e7ed37d50 100644 --- a/Mathlib/Data/Set/MemPartition.lean +++ b/Mathlib/Data/Set/MemPartition.lean @@ -23,7 +23,7 @@ The partition `memPartition f (n + 1)` is finer than `memPartition f n`. * `disjoint_memPartition`: the sets in `memPartition f n` are disjoint * `sUnion_memPartition`: the union of the sets in `memPartition f n` is `univ` -* `memPartition_finite`: `memPartition f n` is finite +* `finite_memPartition`: `memPartition f n` is finite -/ @@ -65,6 +65,7 @@ lemma disjoint_memPartition (f : ℕ → Set α) (n : ℕ) {u v : Set α} · refine Disjoint.mono (diff_subset _ _) (diff_subset _ _) (ih hu' hv' ?_) exact fun huv' ↦ huv (huv' ▸ rfl) +@[simp] lemma sUnion_memPartition (f : ℕ → Set α) (n : ℕ) : ⋃₀ memPartition f n = univ := by induction n with | zero => simp @@ -79,7 +80,7 @@ lemma sUnion_memPartition (f : ℕ → Set α) (n : ℕ) : ⋃₀ memPartition f · exact ⟨t ∩ f (n + 1), ⟨t, ht, Or.inl rfl⟩, hxt, hxf⟩ · exact ⟨t \ f (n + 1), ⟨t, ht, Or.inr rfl⟩, hxt, hxf⟩ -lemma memPartition_finite (f : ℕ → Set α) (n : ℕ) : Set.Finite (memPartition f n) := by +lemma finite_memPartition (f : ℕ → Set α) (n : ℕ) : Set.Finite (memPartition f n) := by induction n with | zero => simp | succ n ih => @@ -92,4 +93,4 @@ lemma memPartition_finite (f : ℕ → Set α) (n : ℕ) : Set.Finite (memPartit simp instance instFinite_memPartition (f : ℕ → Set α) (n : ℕ) : Finite (memPartition f n) := - Set.finite_coe_iff.mp (memPartition_finite _ _) + Set.finite_coe_iff.mp (finite_memPartition _ _) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index bda069f2d2ad2..2b819f3880a52 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -97,18 +97,18 @@ variable [m : MeasurableSpace α] [h : CountablyGenerated α] /-- For each `n : ℕ`, `countablePartition α n` is a partition of the space in at most `2^(n+1)` sets. Each partition is finer than the preceeding one. The measurable space generated by -the union of all those partions is the measurable space `α`. -/ +the union of all those partions is the measurable space on `α`. -/ def countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] : ℕ → Set (Set α) := memPartition (enumerateCountable countable_generatingSet ∅) lemma finite_countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : Set.Finite (countablePartition α n) := - memPartition_finite _ n + finite_memPartition _ n instance instFinite_countablePartition (n : ℕ) : Finite (countablePartition α n) := - Set.finite_coe_iff.mp (MeasurableSpace.finite_countablePartition _ _) + Set.finite_coe_iff.mp (finite_countablePartition _ _) -lemma disjoint_countablePartition (n : ℕ) {s t : Set α} +lemma disjoint_countablePartition {n : ℕ} {s t : Set α} (hs : s ∈ countablePartition α n) (ht : t ∈ countablePartition α n) (hst : s ≠ t) : Disjoint s t := disjoint_memPartition _ n hs ht hst @@ -131,6 +131,19 @@ lemma generateFrom_countablePartition_le_succ (α : Type*) [MeasurableSpace α] generateFrom (countablePartition α n) ≤ generateFrom (countablePartition α (n + 1)) := generateFrom_le (fun _ hs ↦ measurableSet_succ_countablePartition n hs) +lemma measurableSet_generateFrom_memPartition (t : ℕ → Set α) (n : ℕ) : + MeasurableSet[generateFrom (memPartition t n)] (t n) := by + cases n with + | zero => exact measurableSet_generateFrom (by simp) + | succ n => + have : t (n + 1) = ⋃ u ∈ memPartition t n, u ∩ t (n + 1) := by + simp_rw [← iUnion_inter, ← sUnion_eq_biUnion, sUnion_memPartition, univ_inter] + rw [this] + refine MeasurableSet.biUnion (finite_memPartition _ _).countable (fun v hv ↦ ?_) + refine measurableSet_generateFrom ?_ + rw [memPartition_succ] + exact ⟨v, hv, Or.inl rfl⟩ + lemma generateFrom_iUnion_countablePartition (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] : generateFrom (⋃ n, countablePartition α n) = m := by @@ -153,27 +166,9 @@ lemma generateFrom_iUnion_countablePartition (α : Type*) [m : MeasurableSpace rcases huv with rfl | rfl · exact (ih v hv).inter (measurableSet_generateFrom (ht_mem _)) · exact (ih v hv).diff (measurableSet_generateFrom (ht_mem _)) - · obtain ⟨n, rfl⟩ : ∃ n, u = t n := by - have h := subset_range_enumerate countable_generatingSet ∅ hu - obtain ⟨n, hn⟩ := mem_range.mpr h - exact ⟨n, hn.symm⟩ - suffices MeasurableSet[generateFrom (memPartition t n)] (t n) by - suffices h_le : generateFrom (memPartition t n) ≤ generateFrom (⋃ n, memPartition t n) by - exact h_le _ this - exact generateFrom_mono (subset_iUnion _ _) - cases n with - | zero => - simp only [Nat.zero_eq, memPartition_zero] - refine measurableSet_generateFrom ?_ - simp - | succ n => - have : t (n + 1) = ⋃ u ∈ memPartition t n, u ∩ t (n + 1) := by - simp_rw [← iUnion_inter, ← sUnion_eq_biUnion, sUnion_memPartition, univ_inter] - rw [this] - refine MeasurableSet.biUnion (Finite.countable (memPartition_finite _ _)) (fun v hv ↦ ?_) - refine measurableSet_generateFrom ?_ - rw [memPartition_succ] - exact ⟨v, hv, Or.inl rfl⟩ + · obtain ⟨n, hn⟩ := mem_range.mpr (subset_range_enumerate countable_generatingSet ∅ hu) + rw [← hn] + exact generateFrom_mono (subset_iUnion _ _) _ (measurableSet_generateFrom_memPartition t n) lemma generateFrom_countablePartition_le (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : @@ -202,22 +197,20 @@ lemma partitionSet_mem (n : ℕ) (t : α) : partitionSet n t ∈ countablePartit lemma mem_partitionSet (n : ℕ) (t : α) : t ∈ partitionSet n t := (exists_countablePartition_mem n t).choose_spec.2 -lemma mem_countablePartition_iff (n : ℕ) (t : α) {s : Set α} (hs : s ∈ countablePartition α n) : - t ∈ s ↔ partitionSet n t = s := by - refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ - · by_contra h_ne - have h_disj : Disjoint s (partitionSet n t) := - disjoint_countablePartition n hs (partitionSet_mem n t) (Ne.symm h_ne) - refine absurd h_disj ?_ - rw [not_disjoint_iff_nonempty_inter] - exact ⟨t, h, mem_partitionSet n t⟩ - · rw [← h] - exact mem_partitionSet n t +lemma partitionSet_eq_iff {n : ℕ} (t : α) {s : Set α} (hs : s ∈ countablePartition α n) : + partitionSet n t = s ↔ t ∈ s := by + refine ⟨fun h ↦ h ▸ mem_partitionSet n t, fun h ↦ ?_⟩ + by_contra h_ne + have h_disj : Disjoint s (partitionSet n t) := + disjoint_countablePartition hs (partitionSet_mem n t) (Ne.symm h_ne) + refine absurd h_disj ?_ + rw [not_disjoint_iff_nonempty_inter] + exact ⟨t, h, mem_partitionSet n t⟩ lemma partitionSet_of_mem {n : ℕ} {t : α} {s : Set α} (hs : s ∈ countablePartition α n) (ht : t ∈ s) : - partitionSet n t = s := by - rwa [← mem_countablePartition_iff n t hs] + partitionSet n t = s := + (partitionSet_eq_iff t hs).mpr ht lemma measurableSet_partitionSet (n : ℕ) (t : α) : MeasurableSet (partitionSet n t) := measurableSet_countablePartition n (partitionSet_mem n t) diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index caee1cd73d9ca..2a1a23ba3ede8 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -36,17 +36,10 @@ The construction of the conditional cdf in this file follows the proof of Theore -/ - open MeasureTheory Set Filter TopologicalSpace open scoped NNReal ENNReal MeasureTheory Topology -section AuxLemmasToBeMoved - -variable {α β ι : Type*} - -end AuxLemmasToBeMoved - namespace MeasureTheory.Measure variable {α β : Type*} {mα : MeasurableSpace α} (ρ : Measure (α × ℝ)) diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index cb3bd230f5416..5b0accef399fa 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -5,6 +5,7 @@ Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Martingale.Convergence +import Mathlib.Probability.Process.CountablyGenerated /-! # Kernel density @@ -65,59 +66,10 @@ open MeasureTheory Set Filter MeasurableSpace open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory -namespace ProbabilityTheory - -variable {α β γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} [CountablyGenerated γ] - -section partitionFiltration - -/-- A filtration built from the measurable spaces generated by `countablePartition γ n` for -all `n : ℕ`. -/ -def partitionFiltration (γ : Type*) [m : MeasurableSpace γ] [CountablyGenerated γ] : - Filtration ℕ m where - seq := fun n ↦ generateFrom <| countablePartition γ n - mono' := monotone_nat_of_le_succ (generateFrom_countablePartition_le_succ _) - le' := generateFrom_countablePartition_le γ - -lemma measurableSet_partitionFiltration_countablePartition (n : ℕ) {s : Set γ} - (hs : s ∈ countablePartition γ n) : - MeasurableSet[partitionFiltration γ n] s := - measurableSet_generateFrom hs - -lemma measurableSet_partitionFiltration_partitionSet (n : ℕ) (t : γ) : - MeasurableSet[partitionFiltration γ n] (partitionSet n t) := - measurableSet_partitionFiltration_countablePartition n (partitionSet_mem n t) - -lemma measurable_partitionSet_aux (n : ℕ) (m : MeasurableSpace (countablePartition γ n)) : - @Measurable γ (countablePartition γ n) (partitionFiltration γ n) m - (fun c : γ ↦ ⟨partitionSet n c, partitionSet_mem n c⟩) := by - refine @measurable_to_countable' (countablePartition γ n) γ m _ - (partitionFiltration γ n) _ (fun t ↦ ?_) - rcases t with ⟨t, ht⟩ - suffices MeasurableSet[partitionFiltration γ n] {x | partitionSet n x = t} by - convert this - ext x - simp - have : {x | partitionSet n x = t} = t := by - ext x - rw [mem_setOf_eq, mem_countablePartition_iff n x ht] - rw [this] - exact measurableSet_partitionFiltration_countablePartition _ ht - -lemma measurable_partitionSet (n : ℕ) : Measurable[partitionFiltration γ n] (partitionSet n) := - measurable_subtype_coe.comp (measurable_partitionSet_aux _ _) - -lemma iSup_partitionFiltration (β : Type*) [m : MeasurableSpace β] [CountablyGenerated β] : - ⨆ n, partitionFiltration β n = m := by - conv_rhs => rw [← generateFrom_iUnion_countablePartition β] - rw [← iSup_generateFrom] - rfl - -end partitionFiltration - -namespace kernel +namespace ProbabilityTheory.kernel -variable [MeasurableSpace β] {κ : kernel α (γ × β)} {ν : kernel α γ} +variable {α β γ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} + [CountablyGenerated γ] {κ : kernel α (γ × β)} {ν : kernel α γ} section DensityProcess @@ -135,13 +87,13 @@ lemma densityProcess_def (κ : kernel α (γ × β)) (ν : kernel α γ) (n : lemma measurable_densityProcess_partitionFiltration_aux (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) {s : Set β} (hs : MeasurableSet s) : - Measurable[MeasurableSpace.prod mα (partitionFiltration γ n)] (fun (p : α × γ) ↦ + Measurable[mα.prod (partitionFiltration γ n)] (fun (p : α × γ) ↦ κ p.1 (partitionSet n p.2 ×ˢ s) / ν p.1 (partitionSet n p.2)) := by - change Measurable[MeasurableSpace.prod mα (partitionFiltration γ n)] + change Measurable[mα.prod (partitionFiltration γ n)] ((fun (p : α × countablePartition γ n) ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) ∘ (fun (p : α × γ) ↦ (p.1, ⟨partitionSet n p.2, partitionSet_mem n p.2⟩))) - have h1 : @Measurable _ _ (MeasurableSpace.prod mα ⊤) _ - (fun (p : α × countablePartition γ n) ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) := by + have h1 : @Measurable _ _ (mα.prod ⊤) _ + (fun p : α × countablePartition γ n ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) := by refine Measurable.div ?_ ?_ · refine measurable_from_prod_countable (fun t ↦ ?_) exact kernel.measurable_coe _ ((measurableSet_countablePartition _ t.prop).prod hs) @@ -149,8 +101,7 @@ lemma measurable_densityProcess_partitionFiltration_aux (κ : kernel α (γ × rintro ⟨t, ht⟩ exact kernel.measurable_coe _ (measurableSet_countablePartition _ ht) refine h1.comp (measurable_fst.prod_mk ?_) - change @Measurable (α × γ) (countablePartition γ n) - (MeasurableSpace.prod mα (partitionFiltration γ n)) ⊤ + change @Measurable (α × γ) (countablePartition γ n) (mα.prod (partitionFiltration γ n)) ⊤ ((fun c : γ ↦ ⟨partitionSet n c, partitionSet_mem n c⟩) ∘ (fun p : α × γ ↦ p.2)) exact (measurable_partitionSet_aux n ⊤).comp measurable_snd @@ -197,7 +148,7 @@ lemma densityProcess_nonneg (κ : kernel α (γ × β)) (ν : kernel α γ) (n : 0 ≤ densityProcess κ ν n a t s := ENNReal.toReal_nonneg -lemma apply_partitionSet_le_of_fst_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : +lemma meas_partitionSet_le_of_fst_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : κ a (partitionSet n t ×ˢ s) ≤ ν a (partitionSet n t) := by calc κ a (partitionSet n t ×ˢ s) ≤ fst κ a (partitionSet n t) := by @@ -211,7 +162,7 @@ lemma densityProcess_le_one (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) densityProcess κ ν n a t s ≤ 1 := by refine ENNReal.toReal_le_of_le_ofReal zero_le_one (ENNReal.div_le_of_le_mul ?_) rw [ENNReal.ofReal_one, one_mul] - exact apply_partitionSet_le_of_fst_le hκν n a t s + exact meas_partitionSet_le_of_fst_le hκν n a t s lemma snorm_densityProcess_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (s : Set β) : snorm (fun t ↦ densityProcess κ ν n a t s) 1 (ν a) ≤ ν a univ := by @@ -237,10 +188,9 @@ lemma set_integral_densityProcess_I (hκν : fst κ ≤ ν) [IsFiniteKernel κ] rw [integral_toReal] rotate_left · refine Measurable.aemeasurable ?_ - have h := measurable_densityProcess_aux κ ν n hs change Measurable ((fun (p : α × _) ↦ κ p.1 (partitionSet n p.2 ×ˢ s) / ν p.1 (partitionSet n p.2)) ∘ (fun t ↦ (a, t))) - exact h.comp measurable_prod_mk_left + exact (measurable_densityProcess_aux κ ν n hs).comp measurable_prod_mk_left · refine ae_of_all _ (fun t ↦ ?_) by_cases h0 : ν a (partitionSet n t) = 0 · suffices κ a (partitionSet n t ×ˢ s) = 0 by simp [h0, this] @@ -250,8 +200,7 @@ lemma set_integral_densityProcess_I (hκν : fst κ ≤ ν) [IsFiniteKernel κ] refine measure_mono_null (fun x ↦ ?_) h0' simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h - · refine ENNReal.div_lt_top ?_ h0 - exact measure_ne_top _ _ + · exact ENNReal.div_lt_top (measure_ne_top _ _) h0 congr have : ∫⁻ t in u, κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t) ∂(ν a) = ∫⁻ _ in u, κ a (u ×ˢ s) / ν a u ∂(ν a) := by @@ -278,15 +227,14 @@ lemma integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFin rotate_left · intro i j hij simp only [Set.disjoint_prod, disjoint_self, bot_eq_empty] - refine Or.inl (disjoint_countablePartition n i.prop j.prop ?_) + refine Or.inl (disjoint_countablePartition i.prop j.prop ?_) rw [ne_eq, Subtype.coe_inj] exact hij · exact fun k ↦ (measurableSet_countablePartition n k.prop).prod hs rw [integral_iUnion] rotate_left · exact fun k ↦ measurableSet_countablePartition n k.prop - · intro i j hij - refine disjoint_countablePartition n i.prop j.prop ?_ + · refine fun i j hij ↦ disjoint_countablePartition i.prop j.prop ?_ rw [ne_eq, Subtype.coe_inj] exact hij · exact (integrable_densityProcess hκν n a hs).integrableOn @@ -305,7 +253,7 @@ lemma set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [I suffices s = t by rwa [this, inter_self] by_contra h_ne rw [← not_disjoint_iff_nonempty_inter] at hst - exact hst <| disjoint_countablePartition n hs ht h_ne + exact hst <| disjoint_countablePartition hs ht h_ne · simp · rintro u hu rw [set_integral_densityProcess_I hκν _ _ hs hu] @@ -313,12 +261,10 @@ lemma set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [I have hA' : MeasurableSet A := (partitionFiltration γ).le _ _ hA have h := integral_add_compl hA' (integrable_densityProcess hκν n a hs) rw [hA_eq, integral_densityProcess hκν n a hs] at h - have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by - rw [prod_diff_prod, compl_eq_univ_diff] - simp + have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by rw [prod_diff_prod, compl_eq_univ_diff]; simp rw [this, measure_diff (by intro x; simp) (hA'.prod hs) (measure_ne_top (κ a) _), - ENNReal.toReal_sub_of_le (measure_mono (by intro x; simp)) (measure_ne_top _ _)] - rw [eq_tsub_iff_add_eq_of_le, add_comm] + ENNReal.toReal_sub_of_le (measure_mono (by intro x; simp)) (measure_ne_top _ _), + eq_tsub_iff_add_eq_of_le, add_comm] · exact h · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] exact measure_mono (by intro x; simp) @@ -344,11 +290,9 @@ lemma condexp_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFini {i j : ℕ} (hij : i ≤ j) (a : α) {s : Set β} (hs : MeasurableSet s) : (ν a)[fun t ↦ densityProcess κ ν j a t s | partitionFiltration γ i] =ᵐ[ν a] fun t ↦ densityProcess κ ν i a t s := by - symm - refine ae_eq_condexp_of_forall_set_integral_eq ?_ ?_ ?_ ?_ ?_ + refine (ae_eq_condexp_of_forall_set_integral_eq ?_ ?_ ?_ ?_ ?_).symm · exact integrable_densityProcess hκν j a hs - · refine fun t _ _ ↦ Integrable.integrableOn ?_ - exact integrable_densityProcess hκν _ _ hs + · exact fun t _ _ ↦ (integrable_densityProcess hκν _ _ hs).integrableOn · intro t ht _ rw [set_integral_densityProcess hκν i a hs ht, set_integral_densityProcess_of_le hκν hij a hs ht] @@ -372,13 +316,10 @@ lemma densityProcess_mono_set (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ rw [ne_eq, ENNReal.div_eq_top] simp only [ne_eq, h0, and_false, false_or, not_and, not_not] refine fun h_top ↦ eq_top_mono ?_ h_top - exact apply_partitionSet_le_of_fst_le hκν n a t s - rw [ENNReal.toReal_le_toReal] - · gcongr - rw [prod_subset_prod_iff] - simp [subset_rfl, h] - · exact h_ne_top s - · exact h_ne_top s' + exact meas_partitionSet_le_of_fst_le hκν n a t s + rw [ENNReal.toReal_le_toReal (h_ne_top s) (h_ne_top s')] + gcongr + simp [prod_subset_prod_iff, subset_rfl, h] lemma densityProcess_mono_kernel_left {κ' : kernel α (γ × β)} (hκκ' : κ ≤ κ') (hκ'ν : fst κ' ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : @@ -387,8 +328,8 @@ lemma densityProcess_mono_kernel_left {κ' : kernel α (γ × β)} (hκκ' : κ by_cases h0 : ν a (partitionSet n t) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] simp - have h_le : κ' a (partitionSet n t ×ˢ s) - ≤ ν a (partitionSet n t) := apply_partitionSet_le_of_fst_le hκ'ν n a t s + have h_le : κ' a (partitionSet n t ×ˢ s) ≤ ν a (partitionSet n t) := + meas_partitionSet_le_of_fst_le hκ'ν n a t s rw [ENNReal.toReal_le_toReal] · gcongr exact hκκ' _ _ @@ -404,15 +345,12 @@ lemma densityProcess_antitone_kernel_right {ν' : kernel α γ} (hνν' : ν ≤ ν') (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : densityProcess κ ν' n a t s ≤ densityProcess κ ν n a t s := by unfold densityProcess - have h_le : κ a (partitionSet n t ×ˢ s) - ≤ ν a (partitionSet n t) := apply_partitionSet_le_of_fst_le hκν n a t s + have h_le : κ a (partitionSet n t ×ˢ s) ≤ ν a (partitionSet n t) := + meas_partitionSet_le_of_fst_le hκν n a t s by_cases h0 : ν a (partitionSet n t) = 0 - · suffices κ a (partitionSet n t ×ˢ s) = 0 by - simp only [this, ENNReal.zero_div, ENNReal.zero_toReal, h0, le_refl] - exact le_antisymm (h_le.trans h0.le) zero_le' - have h0' : ν' a (partitionSet n t) ≠ 0 := by - refine fun h ↦ h0 (le_antisymm (le_trans ?_ h.le) zero_le') - exact hνν' _ _ + · simp [le_antisymm (h_le.trans h0.le) zero_le', h0] + have h0' : ν' a (partitionSet n t) ≠ 0 := + fun h ↦ h0 (le_antisymm ((hνν' _ _).trans h.le) zero_le') rw [ENNReal.toReal_le_toReal] · gcongr exact hνν' _ _ @@ -491,17 +429,13 @@ lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : fst κ ≤ ν) [IsF · exact (martingale_densityProcess hκν a hs).submartingale · refine uniformIntegrable_of le_rfl ENNReal.one_ne_top ?_ ?_ · exact fun n ↦ (measurable_densityProcess_right κ ν n a hs).aestronglyMeasurable - · intro ε _ - refine ⟨2, fun n ↦ ?_⟩ - refine le_of_eq_of_le ?_ (?_ : 0 ≤ ENNReal.ofReal ε) - · have : {x | 2 ≤ ‖densityProcess κ ν n a x s‖₊} = ∅ := by - ext x - simp only [mem_setOf_eq, mem_empty_iff_false, iff_false, not_le] - refine (?_ : _ ≤ (1 : ℝ≥0)).trans_lt one_lt_two - rw [Real.nnnorm_of_nonneg (densityProcess_nonneg _ _ _ _ _ _)] - exact mod_cast (densityProcess_le_one hκν _ _ _ _) - rw [this] - simp + · refine fun ε _ ↦ ⟨2, fun n ↦ le_of_eq_of_le ?_ (?_ : 0 ≤ ENNReal.ofReal ε)⟩ + · suffices {x | 2 ≤ ‖densityProcess κ ν n a x s‖₊} = ∅ by simp [this] + ext x + simp only [mem_setOf_eq, mem_empty_iff_false, iff_false, not_le] + refine (?_ : _ ≤ (1 : ℝ≥0)).trans_lt one_lt_two + rw [Real.nnnorm_of_nonneg (densityProcess_nonneg _ _ _ _ _ _)] + exact mod_cast (densityProcess_le_one hκν _ _ _ _) · simp lemma tendsto_snorm_one_restrict_densityProcess_limitProcess [IsFiniteKernel κ] [IsFiniteKernel ν] @@ -574,6 +508,8 @@ lemma density_le_one (hκν : fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) · exact eventually_of_forall (fun n ↦ densityProcess_le_one hκν _ _ _ _) +section Integral + lemma snorm_density_le (hκν : fst κ ≤ ν) (a : α) (s : Set β) : snorm (fun t ↦ density κ ν a t s) 1 (ν a) ≤ ν a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ @@ -605,15 +541,13 @@ lemma set_integral_density_of_measurableSet (hκν : fst κ ≤ ν) [IsFiniteKer [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet[partitionFiltration γ n] A) : ∫ t in A, density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - suffices ∫ t in A, density κ ν a t s ∂(ν a) - = ∫ t in A, densityProcess κ ν n a t s ∂(ν a) by - rw [this] - exact set_integral_densityProcess hκν _ _ hs hA + suffices ∫ t in A, density κ ν a t s ∂(ν a) = ∫ t in A, densityProcess κ ν n a t s ∂(ν a) by + exact this ▸ set_integral_densityProcess hκν _ _ hs hA suffices ∫ t in A, density κ ν a t s ∂(ν a) = limsup (fun i ↦ ∫ t in A, densityProcess κ ν i a t s ∂(ν a)) atTop by rw [this, ← limsup_const (α := ℕ) (f := atTop) (∫ t in A, densityProcess κ ν n a t s ∂(ν a)), limsup_congr] - simp only [eventually_atTop, ge_iff_le] + simp only [eventually_atTop] refine ⟨n, fun m hnm ↦ ?_⟩ rw [set_integral_densityProcess_of_le hκν hnm _ hs hA, set_integral_densityProcess hκν _ _ hs hA] @@ -665,6 +599,23 @@ lemma set_integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFinite exact fun i ↦ (hf i).prod hs · rwa [iSup_partitionFiltration] at hf +lemma set_lintegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet A) : + ∫⁻ t in A, ENNReal.ofReal (density κ ν a t s) ∂(ν a) = κ a (A ×ˢ s) := by + rw [← ofReal_integral_eq_lintegral_ofReal] + · rw [set_integral_density hκν a hs hA, + ENNReal.ofReal_toReal (measure_ne_top _ _)] + · exact (integrable_density hκν a hs).restrict + · exact ae_of_all _ (fun _ ↦ density_nonneg hκν _ _ _) + +lemma lintegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) : + ∫⁻ t, ENNReal.ofReal (density κ ν a t s) ∂(ν a) = κ a (univ ×ˢ s) := by + rw [← set_lintegral_univ] + exact set_lintegral_density hκν a hs MeasurableSet.univ + +end Integral + lemma tendsto_integral_density_of_monotone (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : @@ -810,7 +761,7 @@ lemma densityProcess_univ_ae (κ : kernel α (γ × β)) [IsFiniteKernel κ] (n · exact (finite_countablePartition _ _).countable · intro s hs t ht hst simp only [disjoint_iUnion_right, disjoint_iUnion_left] - exact fun _ _ ↦ disjoint_countablePartition n hs ht hst + exact fun _ _ ↦ disjoint_countablePartition hs ht hst · intro s hs by_cases h : fst κ a s = 0 · simp [h, measurableSet_countablePartition n hs] diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index 660ca62a1d168..fd6a35c9d2b9e 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -77,15 +77,11 @@ lemma set_lintegral_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFin (a : α) {s : Set γ} (hs : MeasurableSet s) : ∫⁻ x in s, ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) ∂(κ + ν) a = κ a s := by have h_le : κ ≤ κ + ν := le_add_of_nonneg_right bot_le - rw [← ofReal_integral_eq_lintegral_ofReal] - · unfold kernel.rnDerivAux - rw [set_integral_density (kernel.fst_map_prod_le_of_le h_le) a MeasurableSet.univ hs, - ENNReal.ofReal_toReal, kernel.map_apply' _ _ _ (hs.prod MeasurableSet.univ)] - · congr with x - simp - · exact measure_ne_top _ _ - · exact (integrable_density (kernel.fst_map_prod_le_of_le h_le) a MeasurableSet.univ).restrict - · exact ae_of_all _ (fun x ↦ density_nonneg (kernel.fst_map_prod_le_of_le h_le) _ _ _) + simp_rw [kernel.rnDerivAux] + rw [set_lintegral_density (kernel.fst_map_prod_le_of_le h_le) _ MeasurableSet.univ hs, + kernel.map_apply' _ _ _ (hs.prod MeasurableSet.univ)] + congr with x + simp lemma withDensity_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : kernel.withDensity (κ + ν) diff --git a/Mathlib/Probability/Process/CountablyGenerated.lean b/Mathlib/Probability/Process/CountablyGenerated.lean new file mode 100644 index 0000000000000..82df1b4c158bb --- /dev/null +++ b/Mathlib/Probability/Process/CountablyGenerated.lean @@ -0,0 +1,89 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.MeasureTheory.MeasurableSpace.CountablyGenerated +import Mathlib.Probability.Process.Filtration + +/-! +# Filtration built from the finite partitions of a countably generated measurable space + +In a countably generated measurable space `α`, we can build a sequence of finer and finer finite +measurable partitions of the space such that the measurable space is generated by the union of all +partitions. +This sequence of partitions is defined in `MeasureTheory.MeasurableSpace.CountablyGenerated`. + +Here, we build the filtration of the measurable spaces generated by `countablePartition α n` for all +`n : ℕ`, which we call `partitionFiltration α`. +Since each measurable space in the filtration is finite, we can easily build measurable functions on +those spaces. By building a martingale with respect to `partitionFiltration α` and using the +martingale convergence theorems, we can define a measurable function on `α`. + +## Main definitions + +* `ProbabilityTheory.partitionFiltration`: A filtration built from the measurable spaces generated + by `countablePartition α n` for all `n : ℕ`. + +## Main statements + +* `ProbabilityTheory.iSup_partitionFiltration`: `⨆ n, partitionFiltration α n` is the measurable + space on `α`. + +-/ + +open MeasureTheory MeasurableSpace + +namespace ProbabilityTheory + +variable {α : Type*} [MeasurableSpace α] [CountablyGenerated α] + +/-- A filtration built from the measurable spaces generated by `countablePartition α n` for +all `n : ℕ`. -/ +def partitionFiltration (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] : + Filtration ℕ m where + seq := fun n ↦ generateFrom <| countablePartition α n + mono' := monotone_nat_of_le_succ (generateFrom_countablePartition_le_succ _) + le' := generateFrom_countablePartition_le α + +lemma measurableSet_partitionFiltration_of_mem_countablePartition (n : ℕ) {s : Set α} + (hs : s ∈ countablePartition α n) : + MeasurableSet[partitionFiltration α n] s := + measurableSet_generateFrom hs + +lemma measurableSet_partitionFiltration_partitionSet (n : ℕ) (t : α) : + MeasurableSet[partitionFiltration α n] (partitionSet n t) := + measurableSet_partitionFiltration_of_mem_countablePartition n (partitionSet_mem n t) + +lemma measurable_partitionSet_aux (n : ℕ) (m : MeasurableSpace (countablePartition α n)) : + @Measurable α (countablePartition α n) (partitionFiltration α n) m + (fun c : α ↦ ⟨partitionSet n c, partitionSet_mem n c⟩) := by + refine @measurable_to_countable' (countablePartition α n) α m _ + (partitionFiltration α n) _ (fun t ↦ ?_) + rcases t with ⟨t, ht⟩ + suffices MeasurableSet[partitionFiltration α n] {x | partitionSet n x = t} by + convert this + ext x + simp + have : {x | partitionSet n x = t} = t := by + ext x + rw [Set.mem_setOf_eq, ← partitionSet_eq_iff x ht] + rw [this] + exact measurableSet_partitionFiltration_of_mem_countablePartition _ ht + +lemma measurable_partitionFiltration_partitionSet (α : Type*) + [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : + Measurable[partitionFiltration α n] (partitionSet n) := + measurable_subtype_coe.comp (measurable_partitionSet_aux _ _) + +lemma measurable_partitionSet (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : + Measurable (partitionSet (α := α) n) := + (measurable_partitionFiltration_partitionSet α n).mono ((partitionFiltration α).le n) le_rfl + +lemma iSup_partitionFiltration (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] : + ⨆ n, partitionFiltration α n = m := by + conv_rhs => rw [← generateFrom_iUnion_countablePartition α] + rw [← iSup_generateFrom] + rfl + +end ProbabilityTheory From 659f8eb00b70b3360c8486a037986d1d1e200529 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 07:37:22 +0100 Subject: [PATCH 061/129] minor --- Mathlib/Probability/Process/CountablyGenerated.lean | 8 +++----- 1 file changed, 3 insertions(+), 5 deletions(-) diff --git a/Mathlib/Probability/Process/CountablyGenerated.lean b/Mathlib/Probability/Process/CountablyGenerated.lean index 82df1b4c158bb..74f3553a1ff2e 100644 --- a/Mathlib/Probability/Process/CountablyGenerated.lean +++ b/Mathlib/Probability/Process/CountablyGenerated.lean @@ -42,7 +42,7 @@ variable {α : Type*} [MeasurableSpace α] [CountablyGenerated α] all `n : ℕ`. -/ def partitionFiltration (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] : Filtration ℕ m where - seq := fun n ↦ generateFrom <| countablePartition α n + seq n := generateFrom (countablePartition α n) mono' := monotone_nat_of_le_succ (generateFrom_countablePartition_le_succ _) le' := generateFrom_countablePartition_le α @@ -67,7 +67,7 @@ lemma measurable_partitionSet_aux (n : ℕ) (m : MeasurableSpace (countableParti simp have : {x | partitionSet n x = t} = t := by ext x - rw [Set.mem_setOf_eq, ← partitionSet_eq_iff x ht] + exact partitionSet_eq_iff x ht rw [this] exact measurableSet_partitionFiltration_of_mem_countablePartition _ ht @@ -82,8 +82,6 @@ lemma measurable_partitionSet (α : Type*) [MeasurableSpace α] [CountablyGenera lemma iSup_partitionFiltration (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] : ⨆ n, partitionFiltration α n = m := by - conv_rhs => rw [← generateFrom_iUnion_countablePartition α] - rw [← iSup_generateFrom] - rfl + conv_rhs => rw [← generateFrom_iUnion_countablePartition α, ← iSup_generateFrom] end ProbabilityTheory From 4f47e8dd84285bbba2f646c788796fc08e731800 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 07:39:50 +0100 Subject: [PATCH 062/129] docstring --- Mathlib/Data/Set/MemPartition.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/Data/Set/MemPartition.lean b/Mathlib/Data/Set/MemPartition.lean index 3993e7ed37d50..e1b3130ba851e 100644 --- a/Mathlib/Data/Set/MemPartition.lean +++ b/Mathlib/Data/Set/MemPartition.lean @@ -31,6 +31,8 @@ open Set variable {α : Type*} +/-- `memPartition f n` is the partition containing at most `2^(n+1)` sets, where each set contains +the points that for all `i` belong to one of `f i` or its complement. -/ def memPartition (f : ℕ → Set α) : ℕ → Set (Set α) | 0 => {f 0, (f 0)ᶜ} | n + 1 => {s | ∃ u ∈ memPartition f n, s = u ∩ f (n + 1) ∨ s = u \ f (n + 1)} From 0f501bc5fa5ac9ffb902e485c06a7319aa4d90b1 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 07:55:48 +0100 Subject: [PATCH 063/129] feat: isCoboundedUnder_le lemmas --- Mathlib/Order/LiminfLimsup.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 473baac089541..0e5c8eb0a21ac 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -246,6 +246,26 @@ theorem IsBoundedUnder.isCoboundedUnder_ge {u : γ → α} {l : Filter γ} [Preo (h : l.IsBoundedUnder (· ≤ ·) u) : l.IsCoboundedUnder (· ≥ ·) u := h.isCoboundedUnder_flip +lemma isCoboundedUnder_le_of_eventually_le [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} {x : α} + (hf : ∀ᶠ i in l, x ≤ f i) : + IsCoboundedUnder (· ≤ ·) l f := + IsBoundedUnder.isCoboundedUnder_le ⟨x, hf⟩ + +lemma isCoboundedUnder_ge_of_eventually_le [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} {x : α} + (hf : ∀ᶠ i in l, f i ≤ x) : + IsCoboundedUnder (· ≥ ·) l f := + IsBoundedUnder.isCoboundedUnder_ge ⟨x, hf⟩ + +lemma isCoboundedUnder_le_of_le [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} {x : α} + (hf : ∀ i, x ≤ f i) : + IsCoboundedUnder (· ≤ ·) l f := + isCoboundedUnder_le_of_eventually_le l (eventually_of_forall hf) + +lemma isCoboundedUnder_ge_of_le [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} {x : α} + (hf : ∀ i, f i ≤ x) : + IsCoboundedUnder (· ≥ ·) l f := + isCoboundedUnder_ge_of_eventually_le l (eventually_of_forall hf) + theorem isCobounded_bot : IsCobounded r ⊥ ↔ ∃ b, ∀ x, r b x := by simp [IsCobounded] #align filter.is_cobounded_bot Filter.isCobounded_bot From f3e7ed673b8aae59a2bb0f65fb9713419c1367d1 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 07:57:26 +0100 Subject: [PATCH 064/129] feat: countable filtration in a countably generated measurable space --- Mathlib/Data/Set/MemPartition.lean | 98 ++++++++ .../MeasureTheory/MeasurableSpace/Basic.lean | 67 ----- .../MeasurableSpace/CountablyGenerated.lean | 235 ++++++++++++++++++ .../MeasureTheory/Measure/AEMeasurable.lean | 1 + .../Process/CountablyGenerated.lean | 84 +++++++ 5 files changed, 418 insertions(+), 67 deletions(-) create mode 100644 Mathlib/Data/Set/MemPartition.lean create mode 100644 Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean create mode 100644 Mathlib/Probability/Process/CountablyGenerated.lean diff --git a/Mathlib/Data/Set/MemPartition.lean b/Mathlib/Data/Set/MemPartition.lean new file mode 100644 index 0000000000000..e1b3130ba851e --- /dev/null +++ b/Mathlib/Data/Set/MemPartition.lean @@ -0,0 +1,98 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Data.Set.Finite + +/-! +# Partitions based on membership of a sequence of sets + +Let `f : ℕ → Set α` be a sequence of sets. For `n : ℕ`, we can form the set of points that are in +`f 0 ∪ f 1 ∪ ... ∪ f n`; then the set of points in `(f 0)ᶜ ∪ f 1 ∪ ... ∪ f n` and so on for +all 2^(n+1) choices of a set or its complement. The at most 2^(n+1) sets we obtain form a partition +of `univ : Set α`. We call that partition `memPartition f n` (the membership partition of `f`). + +The partition `memPartition f (n + 1)` is finer than `memPartition f n`. + +## Main definitions + +* `memPartition f n`: the membership partition of the first `n+1` sets in `f`. + +## Main statements + +* `disjoint_memPartition`: the sets in `memPartition f n` are disjoint +* `sUnion_memPartition`: the union of the sets in `memPartition f n` is `univ` +* `finite_memPartition`: `memPartition f n` is finite + +-/ + +open Set + +variable {α : Type*} + +/-- `memPartition f n` is the partition containing at most `2^(n+1)` sets, where each set contains +the points that for all `i` belong to one of `f i` or its complement. -/ +def memPartition (f : ℕ → Set α) : ℕ → Set (Set α) + | 0 => {f 0, (f 0)ᶜ} + | n + 1 => {s | ∃ u ∈ memPartition f n, s = u ∩ f (n + 1) ∨ s = u \ f (n + 1)} + +@[simp] +lemma memPartition_zero (f : ℕ → Set α) : memPartition f 0 = {f 0, (f 0)ᶜ} := rfl + +lemma memPartition_succ (f : ℕ → Set α) (n : ℕ) : + memPartition f (n + 1) = {s | ∃ u ∈ memPartition f n, s = u ∩ f (n + 1) ∨ s = u \ f (n + 1)} := + rfl + +lemma disjoint_memPartition (f : ℕ → Set α) (n : ℕ) {u v : Set α} + (hu : u ∈ memPartition f n) (hv : v ∈ memPartition f n) (huv : u ≠ v) : + Disjoint u v := by + revert u v + induction n with + | zero => + intro u v hu hv huv + simp only [Nat.zero_eq, memPartition_zero, mem_insert_iff, mem_singleton_iff] at hu hv + rcases hu with rfl | rfl <;> rcases hv with rfl | rfl + exacts [absurd rfl huv, disjoint_compl_right, disjoint_compl_left, absurd rfl huv] + | succ n ih => + intro u v hu hv huv + rw [memPartition_succ] at hu hv + obtain ⟨u', hu', hu'_eq⟩ := hu + obtain ⟨v', hv', hv'_eq⟩ := hv + rcases hu'_eq with rfl | rfl <;> rcases hv'_eq with rfl | rfl + · refine Disjoint.mono (inter_subset_left _ _) (inter_subset_left _ _) (ih hu' hv' ?_) + exact fun huv' ↦ huv (huv' ▸ rfl) + · exact Disjoint.mono_left (inter_subset_right _ _) Set.disjoint_sdiff_right + · exact Disjoint.mono_right (inter_subset_right _ _) Set.disjoint_sdiff_left + · refine Disjoint.mono (diff_subset _ _) (diff_subset _ _) (ih hu' hv' ?_) + exact fun huv' ↦ huv (huv' ▸ rfl) + +@[simp] +lemma sUnion_memPartition (f : ℕ → Set α) (n : ℕ) : ⋃₀ memPartition f n = univ := by + induction n with + | zero => simp + | succ n ih => + rw [memPartition_succ] + ext x + have : x ∈ ⋃₀ memPartition f n := by simp [ih] + simp only [mem_sUnion, mem_iUnion, mem_insert_iff, mem_singleton_iff, exists_prop, mem_univ, + iff_true] at this ⊢ + obtain ⟨t, ht, hxt⟩ := this + by_cases hxf : x ∈ f (n + 1) + · exact ⟨t ∩ f (n + 1), ⟨t, ht, Or.inl rfl⟩, hxt, hxf⟩ + · exact ⟨t \ f (n + 1), ⟨t, ht, Or.inr rfl⟩, hxt, hxf⟩ + +lemma finite_memPartition (f : ℕ → Set α) (n : ℕ) : Set.Finite (memPartition f n) := by + induction n with + | zero => simp + | succ n ih => + rw [memPartition_succ] + have : Finite (memPartition f n) := Set.finite_coe_iff.mp ih + rw [← Set.finite_coe_iff] + simp_rw [setOf_exists, ← exists_prop, setOf_exists, setOf_or] + refine Finite.Set.finite_biUnion (memPartition f n) _ (fun u _ ↦ ?_) + rw [Set.finite_coe_iff] + simp + +instance instFinite_memPartition (f : ℕ → Set α) (n : ℕ) : Finite (memPartition f n) := + Set.finite_coe_iff.mp (finite_memPartition _ _) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 0701daf449a4c..9c116f1ec26ca 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -1948,73 +1948,6 @@ theorem MeasurableSpace.comap_compl {m' : MeasurableSpace β} [BooleanAlgebra β MeasurableSpace.comap_compl (fun _ _ ↦ measurableSet_top) _ #align measurable_space.comap_not MeasurableSpace.comap_not -section CountablyGenerated - -namespace MeasurableSpace - -variable (α) - -/-- We say a measurable space is countably generated -if it can be generated by a countable set of sets. -/ -class CountablyGenerated [m : MeasurableSpace α] : Prop where - isCountablyGenerated : ∃ b : Set (Set α), b.Countable ∧ m = generateFrom b -#align measurable_space.countably_generated MeasurableSpace.CountablyGenerated - -variable {α} - -theorem CountablyGenerated.comap [m : MeasurableSpace β] [h : CountablyGenerated β] (f : α → β) : - @CountablyGenerated α (.comap f m) := by - rcases h with ⟨⟨b, hbc, rfl⟩⟩ - rw [comap_generateFrom] - letI := generateFrom (preimage f '' b) - exact ⟨_, hbc.image _, rfl⟩ - -theorem CountablyGenerated.sup {m₁ m₂ : MeasurableSpace β} (h₁ : @CountablyGenerated β m₁) - (h₂ : @CountablyGenerated β m₂) : @CountablyGenerated β (m₁ ⊔ m₂) := by - rcases h₁ with ⟨⟨b₁, hb₁c, rfl⟩⟩ - rcases h₂ with ⟨⟨b₂, hb₂c, rfl⟩⟩ - exact @mk _ (_ ⊔ _) ⟨_, hb₁c.union hb₂c, generateFrom_sup_generateFrom⟩ - -instance (priority := 100) [MeasurableSpace α] [Finite α] : CountablyGenerated α where - isCountablyGenerated := - ⟨{s | MeasurableSet s}, Set.to_countable _, generateFrom_measurableSet.symm⟩ - -instance [MeasurableSpace α] [CountablyGenerated α] {p : α → Prop} : - CountablyGenerated { x // p x } := .comap _ - -instance [MeasurableSpace α] [CountablyGenerated α] [MeasurableSpace β] [CountablyGenerated β] : - CountablyGenerated (α × β) := - .sup (.comap Prod.fst) (.comap Prod.snd) - -instance [MeasurableSpace α] {s : Set α} [h : CountablyGenerated s] [MeasurableSingletonClass s] : - HasCountableSeparatingOn α MeasurableSet s := by - suffices HasCountableSeparatingOn s MeasurableSet univ from this.of_subtype fun _ ↦ id - rcases h.1 with ⟨b, hbc, hb⟩ - refine ⟨⟨b, hbc, fun t ht ↦ hb.symm ▸ .basic t ht, fun x _ y _ h ↦ ?_⟩⟩ - rw [← forall_generateFrom_mem_iff_mem_iff, ← hb] at h - simpa using h {y} - -variable (α) - -open Classical - -/-- If a measurable space is countably generated and separates points, it admits a measurable -injection into the Cantor space `ℕ → Bool` (equipped with the product sigma algebra). -/ -theorem measurable_injection_nat_bool_of_countablyGenerated [MeasurableSpace α] - [HasCountableSeparatingOn α MeasurableSet univ] : - ∃ f : α → ℕ → Bool, Measurable f ∧ Function.Injective f := by - rcases exists_seq_separating α MeasurableSet.empty univ with ⟨e, hem, he⟩ - refine ⟨(· ∈ e ·), ?_, ?_⟩ - · rw [measurable_pi_iff] - refine fun n ↦ measurable_to_bool ?_ - simpa only [preimage, mem_singleton_iff, Bool.decide_iff, setOf_mem_eq] using hem n - · exact fun x y h ↦ he x trivial y trivial fun n ↦ decide_eq_decide.1 <| congr_fun h _ -#align measurable_space.measurable_injection_nat_bool_of_countably_generated MeasurableSpace.measurable_injection_nat_bool_of_countablyGenerated - -end MeasurableSpace - -end CountablyGenerated - namespace Filter variable [MeasurableSpace α] diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean new file mode 100644 index 0000000000000..2b819f3880a52 --- /dev/null +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -0,0 +1,235 @@ +/- +Copyright (c) 2023 Felix Weilacher. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Felix Weilacher, Yury G. Kudryashov, Rémy Degenne +-/ +import Mathlib.MeasureTheory.MeasurableSpace.Basic +import Mathlib.Data.Set.memPartition + +/-! +# Countably generated measurable spaces + +We say a measurable space is countably generated if it can be generated by a countable set of sets. + +In such a space, we can also build a sequence of finer and finer finite measurable partitions of +the space such that the measurable space is generated by the union of all partitions. + +## Main definitions + +* `MeasurableSpace.CountablyGenerated`: class stating that a measurable space is countably + generated. +* `MeasurableSpace.countablePartition`: sequences of finer and finer partitions of + a countably generated space. +* `MeasurableSpace.partitionSet`: `partitionSet n t` is the set in the partition + `countablePartition α n` to which `t` belongs. + +## Main statements + +* `MeasurableSpace.measurable_injection_nat_bool_of_countablyGenerated`: if a measurable space is + countably generated and separates points, it admits a measurable injection into the Cantor space + `ℕ → Bool` (equipped with the product sigma algebra). + +-/ + + +open Set MeasureTheory + +namespace MeasurableSpace + +variable {α β : Type*} + +/-- We say a measurable space is countably generated +if it can be generated by a countable set of sets. -/ +class CountablyGenerated (α : Type*) [m : MeasurableSpace α] : Prop where + isCountablyGenerated : ∃ b : Set (Set α), b.Countable ∧ m = generateFrom b +#align measurable_space.countably_generated MeasurableSpace.CountablyGenerated + +/-- A countable set of sets that generate the measurable space. +We insert `∅` to ensure it is nonempty. -/ +def generatingSet (α : Type*) [MeasurableSpace α] [h : CountablyGenerated α] : Set (Set α) := + insert ∅ h.isCountablyGenerated.choose + +lemma countable_generatingSet [MeasurableSpace α] [h : CountablyGenerated α] : + Countable (generatingSet α) := + Countable.insert _ h.isCountablyGenerated.choose_spec.1 + +lemma generateFrom_generatingSet [m : MeasurableSpace α] [h : CountablyGenerated α] : + generateFrom (generatingSet α) = m := + (generateFrom_insert_empty _).trans <| h.isCountablyGenerated.choose_spec.2.symm + +lemma nonempty_generatingSet [MeasurableSpace α] [CountablyGenerated α] : + Set.Nonempty (generatingSet α) := + ⟨∅, mem_insert _ _⟩ + +theorem CountablyGenerated.comap [m : MeasurableSpace β] [h : CountablyGenerated β] (f : α → β) : + @CountablyGenerated α (.comap f m) := by + rcases h with ⟨⟨b, hbc, rfl⟩⟩ + rw [comap_generateFrom] + letI := generateFrom (preimage f '' b) + exact ⟨_, hbc.image _, rfl⟩ + +theorem CountablyGenerated.sup {m₁ m₂ : MeasurableSpace β} (h₁ : @CountablyGenerated β m₁) + (h₂ : @CountablyGenerated β m₂) : @CountablyGenerated β (m₁ ⊔ m₂) := by + rcases h₁ with ⟨⟨b₁, hb₁c, rfl⟩⟩ + rcases h₂ with ⟨⟨b₂, hb₂c, rfl⟩⟩ + exact @mk _ (_ ⊔ _) ⟨_, hb₁c.union hb₂c, generateFrom_sup_generateFrom⟩ + +instance (priority := 100) [MeasurableSpace α] [Finite α] : CountablyGenerated α where + isCountablyGenerated := + ⟨{s | MeasurableSet s}, Set.to_countable _, generateFrom_measurableSet.symm⟩ + +instance [MeasurableSpace α] [CountablyGenerated α] {p : α → Prop} : + CountablyGenerated { x // p x } := .comap _ + +instance [MeasurableSpace α] [CountablyGenerated α] [MeasurableSpace β] [CountablyGenerated β] : + CountablyGenerated (α × β) := + .sup (.comap Prod.fst) (.comap Prod.snd) + +instance [MeasurableSpace α] {s : Set α} [h : CountablyGenerated s] [MeasurableSingletonClass s] : + HasCountableSeparatingOn α MeasurableSet s := by + suffices HasCountableSeparatingOn s MeasurableSet univ from this.of_subtype fun _ ↦ id + rcases h.1 with ⟨b, hbc, hb⟩ + refine ⟨⟨b, hbc, fun t ht ↦ hb.symm ▸ .basic t ht, fun x _ y _ h ↦ ?_⟩⟩ + rw [← forall_generateFrom_mem_iff_mem_iff, ← hb] at h + simpa using h {y} + +variable [m : MeasurableSpace α] [h : CountablyGenerated α] + +/-- For each `n : ℕ`, `countablePartition α n` is a partition of the space in at most +`2^(n+1)` sets. Each partition is finer than the preceeding one. The measurable space generated by +the union of all those partions is the measurable space on `α`. -/ +def countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] : ℕ → Set (Set α) := + memPartition (enumerateCountable countable_generatingSet ∅) + +lemma finite_countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : + Set.Finite (countablePartition α n) := + finite_memPartition _ n + +instance instFinite_countablePartition (n : ℕ) : Finite (countablePartition α n) := + Set.finite_coe_iff.mp (finite_countablePartition _ _) + +lemma disjoint_countablePartition {n : ℕ} {s t : Set α} + (hs : s ∈ countablePartition α n) (ht : t ∈ countablePartition α n) (hst : s ≠ t) : + Disjoint s t := + disjoint_memPartition _ n hs ht hst + +lemma sUnion_countablePartition (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : + ⋃₀ countablePartition α n = univ := + sUnion_memPartition _ n + +lemma measurableSet_succ_countablePartition (n : ℕ) {s : Set α} (hs : s ∈ countablePartition α n) : + MeasurableSet[generateFrom (countablePartition α (n + 1))] s := by + let t : ℕ → Set α := enumerateCountable countable_generatingSet ∅ + rw [← diff_union_inter s (t (n + 1))] + refine MeasurableSet.union ?_ ?_ <;> + · refine measurableSet_generateFrom ?_ + rw [countablePartition, memPartition_succ] + exact ⟨s, hs, by simp⟩ + +lemma generateFrom_countablePartition_le_succ (α : Type*) [MeasurableSpace α] [CountablyGenerated α] + (n : ℕ) : + generateFrom (countablePartition α n) ≤ generateFrom (countablePartition α (n + 1)) := + generateFrom_le (fun _ hs ↦ measurableSet_succ_countablePartition n hs) + +lemma measurableSet_generateFrom_memPartition (t : ℕ → Set α) (n : ℕ) : + MeasurableSet[generateFrom (memPartition t n)] (t n) := by + cases n with + | zero => exact measurableSet_generateFrom (by simp) + | succ n => + have : t (n + 1) = ⋃ u ∈ memPartition t n, u ∩ t (n + 1) := by + simp_rw [← iUnion_inter, ← sUnion_eq_biUnion, sUnion_memPartition, univ_inter] + rw [this] + refine MeasurableSet.biUnion (finite_memPartition _ _).countable (fun v hv ↦ ?_) + refine measurableSet_generateFrom ?_ + rw [memPartition_succ] + exact ⟨v, hv, Or.inl rfl⟩ + +lemma generateFrom_iUnion_countablePartition (α : Type*) [m : MeasurableSpace α] + [CountablyGenerated α] : + generateFrom (⋃ n, countablePartition α n) = m := by + let t : ℕ → Set α := enumerateCountable countable_generatingSet ∅ + have ht_mem : ∀ n, t n ∈ generatingSet α := + enumerateCountable_mem countable_generatingSet (mem_insert _ _) + conv_rhs => rw [← generateFrom_generatingSet (α := α)] + refine le_antisymm (generateFrom_le fun u hu ↦ ?_) (generateFrom_le fun u hu ↦ ?_) + · simp only [mem_iUnion] at hu + obtain ⟨n, hun⟩ := hu + induction n generalizing u with + | zero => + simp only [Nat.zero_eq, memPartition_zero, mem_insert_iff, mem_singleton_iff] at hun + rcases hun with rfl | rfl + · exact measurableSet_generateFrom (ht_mem _) + · exact (measurableSet_generateFrom (ht_mem _)).compl + | succ n ih => + simp only [countablePartition, memPartition_succ, mem_setOf_eq] at hun + obtain ⟨v, hv, huv⟩ := hun + rcases huv with rfl | rfl + · exact (ih v hv).inter (measurableSet_generateFrom (ht_mem _)) + · exact (ih v hv).diff (measurableSet_generateFrom (ht_mem _)) + · obtain ⟨n, hn⟩ := mem_range.mpr (subset_range_enumerate countable_generatingSet ∅ hu) + rw [← hn] + exact generateFrom_mono (subset_iUnion _ _) _ (measurableSet_generateFrom_memPartition t n) + +lemma generateFrom_countablePartition_le (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] + (n : ℕ) : + generateFrom (countablePartition α n) ≤ m := by + conv_rhs => rw [← generateFrom_iUnion_countablePartition α] + exact generateFrom_mono (subset_iUnion _ _) + +lemma measurableSet_countablePartition (n : ℕ) {s : Set α} (hs : s ∈ countablePartition α n) : + MeasurableSet s := by + have : MeasurableSet[generateFrom (countablePartition α n)] s := + measurableSet_generateFrom hs + exact generateFrom_countablePartition_le α n _ this + +lemma exists_countablePartition_mem (n : ℕ) (t : α) : ∃ s, s ∈ countablePartition α n ∧ t ∈ s := by + have h_univ := sUnion_countablePartition α n + have h_mem_univ := mem_univ t + rw [← h_univ] at h_mem_univ + simpa only [mem_sUnion] using h_mem_univ + +/-- The set in `countablePartition α n` to which `t : α` belongs. -/ +def partitionSet (n : ℕ) (t : α) : Set α := (exists_countablePartition_mem n t).choose + +lemma partitionSet_mem (n : ℕ) (t : α) : partitionSet n t ∈ countablePartition α n := + (exists_countablePartition_mem n t).choose_spec.1 + +lemma mem_partitionSet (n : ℕ) (t : α) : t ∈ partitionSet n t := + (exists_countablePartition_mem n t).choose_spec.2 + +lemma partitionSet_eq_iff {n : ℕ} (t : α) {s : Set α} (hs : s ∈ countablePartition α n) : + partitionSet n t = s ↔ t ∈ s := by + refine ⟨fun h ↦ h ▸ mem_partitionSet n t, fun h ↦ ?_⟩ + by_contra h_ne + have h_disj : Disjoint s (partitionSet n t) := + disjoint_countablePartition hs (partitionSet_mem n t) (Ne.symm h_ne) + refine absurd h_disj ?_ + rw [not_disjoint_iff_nonempty_inter] + exact ⟨t, h, mem_partitionSet n t⟩ + +lemma partitionSet_of_mem {n : ℕ} {t : α} {s : Set α} (hs : s ∈ countablePartition α n) + (ht : t ∈ s) : + partitionSet n t = s := + (partitionSet_eq_iff t hs).mpr ht + +lemma measurableSet_partitionSet (n : ℕ) (t : α) : MeasurableSet (partitionSet n t) := + measurableSet_countablePartition n (partitionSet_mem n t) + +variable (α) + +open Classical + +/-- If a measurable space is countably generated and separates points, it admits a measurable +injection into the Cantor space `ℕ → Bool` (equipped with the product sigma algebra). -/ +theorem measurable_injection_nat_bool_of_countablyGenerated [MeasurableSpace α] + [HasCountableSeparatingOn α MeasurableSet univ] : + ∃ f : α → ℕ → Bool, Measurable f ∧ Function.Injective f := by + rcases exists_seq_separating α MeasurableSet.empty univ with ⟨e, hem, he⟩ + refine ⟨(· ∈ e ·), ?_, ?_⟩ + · rw [measurable_pi_iff] + refine fun n ↦ measurable_to_bool ?_ + simpa only [preimage, mem_singleton_iff, Bool.decide_iff, setOf_mem_eq] using hem n + · exact fun x y h ↦ he x trivial y trivial fun n ↦ decide_eq_decide.1 <| congr_fun h _ +#align measurable_space.measurable_injection_nat_bool_of_countably_generated MeasurableSpace.measurable_injection_nat_bool_of_countablyGenerated + +end MeasurableSpace diff --git a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean index 147b07c63e53d..af2d691e61daf 100644 --- a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.MeasureTheory.Measure.Trim +import Mathlib.MeasureTheory.MeasurableSpace.CountablyGenerated #align_import measure_theory.measure.ae_measurable from "leanprover-community/mathlib"@"3310acfa9787aa171db6d4cba3945f6f275fe9f2" diff --git a/Mathlib/Probability/Process/CountablyGenerated.lean b/Mathlib/Probability/Process/CountablyGenerated.lean new file mode 100644 index 0000000000000..0fba01e38d309 --- /dev/null +++ b/Mathlib/Probability/Process/CountablyGenerated.lean @@ -0,0 +1,84 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.MeasureTheory.MeasurableSpace.CountablyGenerated +import Mathlib.Probability.Process.Filtration + +/-! +# Filtration built from the finite partitions of a countably generated measurable space + +In a countably generated measurable space `α`, we can build a sequence of finer and finer finite +measurable partitions of the space such that the measurable space is generated by the union of all +partitions. +This sequence of partitions is defined in `MeasureTheory.MeasurableSpace.CountablyGenerated`. + +Here, we build the filtration of the measurable spaces generated by `countablePartition α n` for all +`n : ℕ`, which we call `partitionFiltration α`. +Since each measurable space in the filtration is finite, we can easily build measurable functions on +those spaces. By building a martingale with respect to `partitionFiltration α` and using the +martingale convergence theorems, we can define a measurable function on `α`. + +## Main definitions + +* `ProbabilityTheory.partitionFiltration`: A filtration built from the measurable spaces generated + by `countablePartition α n` for all `n : ℕ`. + +## Main statements + +* `ProbabilityTheory.iSup_partitionFiltration`: `⨆ n, partitionFiltration α n` is the measurable + space on `α`. + +-/ + +open MeasureTheory MeasurableSpace + +namespace ProbabilityTheory + +variable {α : Type*} [MeasurableSpace α] [CountablyGenerated α] + +/-- A filtration built from the measurable spaces generated by `countablePartition α n` for +all `n : ℕ`. -/ +def partitionFiltration (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] : + Filtration ℕ m where + seq n := generateFrom (countablePartition α n) + mono' := monotone_nat_of_le_succ (generateFrom_countablePartition_le_succ _) + le' := generateFrom_countablePartition_le α + +lemma measurableSet_partitionFiltration_of_mem_countablePartition (n : ℕ) {s : Set α} + (hs : s ∈ countablePartition α n) : + MeasurableSet[partitionFiltration α n] s := + measurableSet_generateFrom hs + +lemma measurableSet_partitionFiltration_partitionSet (n : ℕ) (t : α) : + MeasurableSet[partitionFiltration α n] (partitionSet n t) := + measurableSet_partitionFiltration_of_mem_countablePartition n (partitionSet_mem n t) + +lemma measurable_partitionSet_aux (n : ℕ) (m : MeasurableSpace (countablePartition α n)) : + @Measurable α (countablePartition α n) (partitionFiltration α n) m + (fun c : α ↦ ⟨partitionSet n c, partitionSet_mem n c⟩) := by + refine @measurable_to_countable' (countablePartition α n) α m _ + (partitionFiltration α n) _ (fun t ↦ ?_) + rcases t with ⟨t, ht⟩ + suffices MeasurableSet[partitionFiltration α n] {x | partitionSet n x = t} by + convert this + ext x + simp + simp_rw [partitionSet_eq_iff _ ht] + exact measurableSet_partitionFiltration_of_mem_countablePartition _ ht + +lemma measurable_partitionFiltration_partitionSet (α : Type*) + [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : + Measurable[partitionFiltration α n] (partitionSet n) := + measurable_subtype_coe.comp (measurable_partitionSet_aux _ _) + +lemma measurable_partitionSet (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : + Measurable (partitionSet (α := α) n) := + (measurable_partitionFiltration_partitionSet α n).mono ((partitionFiltration α).le n) le_rfl + +lemma iSup_partitionFiltration (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] : + ⨆ n, partitionFiltration α n = m := by + conv_rhs => rw [← generateFrom_iUnion_countablePartition α, ← iSup_generateFrom] + +end ProbabilityTheory From c88f084147f999055959283591b9e90105f5ea19 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 08:02:50 +0100 Subject: [PATCH 065/129] fix --- Mathlib.lean | 3 +++ Mathlib/Probability/Process/CountablyGenerated.lean | 2 +- 2 files changed, 4 insertions(+), 1 deletion(-) diff --git a/Mathlib.lean b/Mathlib.lean index a1b52bd0d4ed4..44df93f727edd 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2054,6 +2054,7 @@ import Mathlib.Data.Set.Intervals.UnorderedInterval import Mathlib.Data.Set.Intervals.WithBotTop import Mathlib.Data.Set.Lattice import Mathlib.Data.Set.List +import Mathlib.Data.Set.MemPartition import Mathlib.Data.Set.MulAntidiagonal import Mathlib.Data.Set.NAry import Mathlib.Data.Set.Opposite @@ -2775,6 +2776,7 @@ import Mathlib.MeasureTheory.Integral.TorusIntegral import Mathlib.MeasureTheory.Integral.VitaliCaratheodory import Mathlib.MeasureTheory.MeasurableSpace.Basic import Mathlib.MeasureTheory.MeasurableSpace.Card +import Mathlib.MeasureTheory.MeasurableSpace.CountablyGenerated import Mathlib.MeasureTheory.MeasurableSpace.Defs import Mathlib.MeasureTheory.MeasurableSpace.Invariants import Mathlib.MeasureTheory.Measure.AEDisjoint @@ -3145,6 +3147,7 @@ import Mathlib.Probability.ProbabilityMassFunction.Constructions import Mathlib.Probability.ProbabilityMassFunction.Integrals import Mathlib.Probability.ProbabilityMassFunction.Monad import Mathlib.Probability.Process.Adapted +import Mathlib.Probability.Process.CountablyGenerated import Mathlib.Probability.Process.Filtration import Mathlib.Probability.Process.HittingTime import Mathlib.Probability.Process.Stopping diff --git a/Mathlib/Probability/Process/CountablyGenerated.lean b/Mathlib/Probability/Process/CountablyGenerated.lean index 0fba01e38d309..6b77075a26429 100644 --- a/Mathlib/Probability/Process/CountablyGenerated.lean +++ b/Mathlib/Probability/Process/CountablyGenerated.lean @@ -9,7 +9,7 @@ import Mathlib.Probability.Process.Filtration /-! # Filtration built from the finite partitions of a countably generated measurable space -In a countably generated measurable space `α`, we can build a sequence of finer and finer finite +In a countably generated measurable space `α`, we can build a sequence of finer and finer finite measurable partitions of the space such that the measurable space is generated by the union of all partitions. This sequence of partitions is defined in `MeasureTheory.MeasurableSpace.CountablyGenerated`. From 2b81b615a47c3c5fa078db666d2b36efed3c6d37 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 08:04:21 +0100 Subject: [PATCH 066/129] fix --- Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index 2b819f3880a52..ce28cc4850e83 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Felix Weilacher, Yury G. Kudryashov, Rémy Degenne -/ import Mathlib.MeasureTheory.MeasurableSpace.Basic -import Mathlib.Data.Set.memPartition +import Mathlib.Data.Set.MemPartition /-! # Countably generated measurable spaces From 85b299b93c465a1f3bf68d04b0b92d7530990600 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 08:28:48 +0100 Subject: [PATCH 067/129] fix imports --- Mathlib/MeasureTheory/MeasurableSpace/Basic.lean | 1 - Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean | 1 + 2 files changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 9c116f1ec26ca..38ad5c1fc320f 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -9,7 +9,6 @@ import Mathlib.GroupTheory.Coset import Mathlib.Logic.Equiv.Fin import Mathlib.MeasureTheory.MeasurableSpace.Defs import Mathlib.Order.Filter.SmallSets -import Mathlib.Order.Filter.CountableSeparatingOn import Mathlib.Order.LiminfLimsup import Mathlib.Data.Set.UnionLift diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index ce28cc4850e83..32f73843b9636 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -5,6 +5,7 @@ Authors: Felix Weilacher, Yury G. Kudryashov, Rémy Degenne -/ import Mathlib.MeasureTheory.MeasurableSpace.Basic import Mathlib.Data.Set.MemPartition +import Mathlib.Order.Filter.CountableSeparatingOn /-! # Countably generated measurable spaces From 000bd1078d12113812448e673e58aa03479a7275 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 08:30:23 +0100 Subject: [PATCH 068/129] fix imports --- Mathlib/MeasureTheory/MeasurableSpace/Basic.lean | 1 - Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean | 1 + 2 files changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 9c116f1ec26ca..38ad5c1fc320f 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -9,7 +9,6 @@ import Mathlib.GroupTheory.Coset import Mathlib.Logic.Equiv.Fin import Mathlib.MeasureTheory.MeasurableSpace.Defs import Mathlib.Order.Filter.SmallSets -import Mathlib.Order.Filter.CountableSeparatingOn import Mathlib.Order.LiminfLimsup import Mathlib.Data.Set.UnionLift diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index ce28cc4850e83..32f73843b9636 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -5,6 +5,7 @@ Authors: Felix Weilacher, Yury G. Kudryashov, Rémy Degenne -/ import Mathlib.MeasureTheory.MeasurableSpace.Basic import Mathlib.Data.Set.MemPartition +import Mathlib.Order.Filter.CountableSeparatingOn /-! # Countably generated measurable spaces From 02bfb16d53f0054dac1805542b14df26150c98e4 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 08:42:03 +0100 Subject: [PATCH 069/129] feat: density of a finite kernel wrt another kernel --- Mathlib.lean | 1 + .../Kernel/Disintegration/Density.lean | 890 ++++++++++++++++++ 2 files changed, 891 insertions(+) create mode 100644 Mathlib/Probability/Kernel/Disintegration/Density.lean diff --git a/Mathlib.lean b/Mathlib.lean index 44df93f727edd..b6a49df57bc62 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3127,6 +3127,7 @@ import Mathlib.Probability.Kernel.CondCdf import Mathlib.Probability.Kernel.CondDistrib import Mathlib.Probability.Kernel.Condexp import Mathlib.Probability.Kernel.Disintegration +import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.IntegralCompProd import Mathlib.Probability.Kernel.Invariance import Mathlib.Probability.Kernel.MeasurableIntegral diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean new file mode 100644 index 0000000000000..5b0accef399fa --- /dev/null +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -0,0 +1,890 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Probability.Kernel.Composition +import Mathlib.Probability.Martingale.Convergence +import Mathlib.Probability.Process.CountablyGenerated + +/-! +# Kernel density + +Let `κ : kernel α (γ × β)` and `ν : kernel α γ` be two finite kernels with `kernel.fst κ ≤ ν`, +where `γ` has a countably generated σ-algebra (true in particular for standard Borel spaces). +We build a function `f : α → γ → Set β → ℝ` jointly measurable in the first two arguments such that +for all `a : α` and all measurable sets `s : Set β` and `A : Set γ`, +`∫ t in A, f a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal`. + +If we were interested only in a fixed `a : α`, then we could use the Radon-Nikodym derivative to +build the function `f`, as follows. +``` +def f (κ : kernel α (γ × β)) (ν : kernel a γ) (a : α) (t : γ) (s : Set β) : ℝ := + (((κ a).restrict (univ ×ˢ s)).fst.rnDeriv (ν a) t).toReal +``` +However, we can't turn those functions for each `a` into a measurable function of the pair `(a, t)`. + +In order to obtain measurability through countability, we use the fact that the measurable space `γ` +is countably generated. For each `n : ℕ`, we define a finite partition of `γ`, such that those +partitions are finer as `n` grows, and the σ-algebra generated by the union of all partitions is +the σ-algebra of `γ`. +For `t : γ`, let `partitionSet n t` be the set in the partition such that `t ∈ partitionSet n t`. + +For a given `n`, the function `densityProcess κ ν n : α → γ → Set β → ℝ` defined by +`fun a t s ↦ (κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t)).toReal` has the desired +property that `∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal` for all `A` +in the σ-algebra generated by the partition at scale `n` and is measurable in `(a, t)`. + +Let `partitionFiltration γ` be the filtration of those σ-algebras for all `n : ℕ`. +The functions `densityProcess κ ν n` described here are a bounded `ν`-martingale for the filtration +`partitionFiltration γ`. By Doob's martingale L1 convergence theorem, that martingale converges to +a limit, which has a product-measurable version and satisfies the integral equality for all `A` in +`⨆ n, partitionFiltration γ n`. Finally, the partitions were chosen such that that supremum is equal +to the σ-algebra on `γ`, hence the equality holds for all measurable sets. +We have obtained the desired density function. + +## Main definitions + +* `ProbabilityTheory.kernel.density`: for `k : kernel α (γ × β)` and `ν : kernel α γ` two finite + kernels, `kernel.density κ ν` is a function `α → γ → Set β → ℝ`. + +## Main statements + +* `ProbabilityTheory.kernel.set_integral_density`: for all measurable sets `A : Set γ` and + `s : Set β`, `∫ t in A, kernel.density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal`. +* `ProbabilityTheory.kernel.measurable_density`: the function + `p : α × γ ↦ kernel.density κ ν p.1 p.2 s` is measurable. + +## References + +The construction of the density process in this file follows the proof of Theorem 9.27 in +[O. Kallenberg, Foundations of modern probability][kallenberg2021], adapted to use a countably +generated hypothesis instead of specializing to `ℝ`. +-/ + +open MeasureTheory Set Filter MeasurableSpace + +open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory + +namespace ProbabilityTheory.kernel + +variable {α β γ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} + [CountablyGenerated γ] {κ : kernel α (γ × β)} {ν : kernel α γ} + +section DensityProcess + +/-- An `ℕ`-indexed martingale that is a density for `κ` with respect to `ν` on the sets in +`countablePartition γ n`. Used to define its limit `ProbabilityTheory.kernel.density`, which is +a density for those kernels for all measurable sets. -/ +noncomputable +def densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (t : γ) (s : Set β) : + ℝ := + (κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t)).toReal + +lemma densityProcess_def (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (s : Set β) : + (fun t ↦ densityProcess κ ν n a t s) + = fun t ↦ (κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t)).toReal := rfl + +lemma measurable_densityProcess_partitionFiltration_aux (κ : kernel α (γ × β)) (ν : kernel α γ) + (n : ℕ) {s : Set β} (hs : MeasurableSet s) : + Measurable[mα.prod (partitionFiltration γ n)] (fun (p : α × γ) ↦ + κ p.1 (partitionSet n p.2 ×ˢ s) / ν p.1 (partitionSet n p.2)) := by + change Measurable[mα.prod (partitionFiltration γ n)] + ((fun (p : α × countablePartition γ n) ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) + ∘ (fun (p : α × γ) ↦ (p.1, ⟨partitionSet n p.2, partitionSet_mem n p.2⟩))) + have h1 : @Measurable _ _ (mα.prod ⊤) _ + (fun p : α × countablePartition γ n ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) := by + refine Measurable.div ?_ ?_ + · refine measurable_from_prod_countable (fun t ↦ ?_) + exact kernel.measurable_coe _ ((measurableSet_countablePartition _ t.prop).prod hs) + · refine measurable_from_prod_countable ?_ + rintro ⟨t, ht⟩ + exact kernel.measurable_coe _ (measurableSet_countablePartition _ ht) + refine h1.comp (measurable_fst.prod_mk ?_) + change @Measurable (α × γ) (countablePartition γ n) (mα.prod (partitionFiltration γ n)) ⊤ + ((fun c : γ ↦ ⟨partitionSet n c, partitionSet_mem n c⟩) ∘ (fun p : α × γ ↦ p.2)) + exact (measurable_partitionSet_aux n ⊤).comp measurable_snd + +lemma measurable_densityProcess_aux (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) + {s : Set β} (hs : MeasurableSet s) : + Measurable (fun (p : α × γ) ↦ + κ p.1 (partitionSet n p.2 ×ˢ s) / ν p.1 (partitionSet n p.2)) := by + refine Measurable.mono (measurable_densityProcess_partitionFiltration_aux κ ν n hs) ?_ le_rfl + exact sup_le_sup le_rfl (comap_mono ((partitionFiltration γ).le _)) + +lemma measurable_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) + {s : Set β} (hs : MeasurableSet s) : + Measurable (fun (p : α × γ) ↦ densityProcess κ ν n p.1 p.2 s) := + (measurable_densityProcess_aux κ ν n hs).ennreal_toReal + +lemma measurable_densityProcess_left (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) + (t : γ) {s : Set β} (hs : MeasurableSet s) : + Measurable (fun a ↦ densityProcess κ ν n a t s) := + (measurable_densityProcess κ ν n hs).comp (measurable_id.prod_mk measurable_const) + +lemma measurable_densityProcess_right (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) + {s : Set β} (a : α) (hs : MeasurableSet s) : + Measurable (fun t ↦ densityProcess κ ν n a t s) := + (measurable_densityProcess κ ν n hs).comp (measurable_const.prod_mk measurable_id) + +lemma measurable_partitionFiltration_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) + (a : α) {s : Set β} (hs : MeasurableSet s) : + Measurable[partitionFiltration γ n] (fun t ↦ densityProcess κ ν n a t s) := by + refine @Measurable.ennreal_toReal _ (partitionFiltration γ n) _ ?_ + exact (measurable_densityProcess_partitionFiltration_aux κ ν n hs).comp measurable_prod_mk_left + +lemma stronglyMeasurable_partitionFiltration_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) + (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : + StronglyMeasurable[partitionFiltration γ n] (fun t ↦ densityProcess κ ν n a t s) := + (measurable_partitionFiltration_densityProcess κ ν n a hs).stronglyMeasurable + +lemma adapted_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) + {s : Set β} (hs : MeasurableSet s) : + Adapted (partitionFiltration γ) (fun n t ↦ densityProcess κ ν n a t s) := + fun n ↦ stronglyMeasurable_partitionFiltration_densityProcess κ ν n a hs + +lemma densityProcess_nonneg (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) + (a : α) (t : γ) (s : Set β) : + 0 ≤ densityProcess κ ν n a t s := + ENNReal.toReal_nonneg + +lemma meas_partitionSet_le_of_fst_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : + κ a (partitionSet n t ×ˢ s) ≤ ν a (partitionSet n t) := by + calc κ a (partitionSet n t ×ˢ s) + ≤ fst κ a (partitionSet n t) := by + rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] + refine measure_mono (fun x ↦ ?_) + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + _ ≤ ν a (partitionSet n t) := hκν a _ + +lemma densityProcess_le_one (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : + densityProcess κ ν n a t s ≤ 1 := by + refine ENNReal.toReal_le_of_le_ofReal zero_le_one (ENNReal.div_le_of_le_mul ?_) + rw [ENNReal.ofReal_one, one_mul] + exact meas_partitionSet_le_of_fst_le hκν n a t s + +lemma snorm_densityProcess_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (s : Set β) : + snorm (fun t ↦ densityProcess κ ν n a t s) 1 (ν a) ≤ ν a univ := by + refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun x ↦ ?_))).trans ?_ + · simp only [Real.norm_eq_abs, abs_of_nonneg (densityProcess_nonneg κ ν n a x s), + densityProcess_le_one hκν n a x s] + · simp + +lemma integrable_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (n : ℕ) + (a : α) {s : Set β} (hs : MeasurableSet s) : + Integrable (fun t ↦ densityProcess κ ν n a t s) (ν a) := by + rw [← memℒp_one_iff_integrable] + refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ + · exact measurable_densityProcess_right κ ν n a hs + · exact (snorm_densityProcess_le hκν n a s).trans_lt (measure_lt_top _ _) + +lemma set_integral_densityProcess_I (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {u : Set γ} + (hu : u ∈ countablePartition γ n) : + ∫ t in u, densityProcess κ ν n a t s ∂(ν a) = (κ a (u ×ˢ s)).toReal := by + have hu_meas : MeasurableSet u := measurableSet_countablePartition n hu + simp_rw [densityProcess] + rw [integral_toReal] + rotate_left + · refine Measurable.aemeasurable ?_ + change Measurable ((fun (p : α × _) ↦ κ p.1 (partitionSet n p.2 ×ˢ s) + / ν p.1 (partitionSet n p.2)) ∘ (fun t ↦ (a, t))) + exact (measurable_densityProcess_aux κ ν n hs).comp measurable_prod_mk_left + · refine ae_of_all _ (fun t ↦ ?_) + by_cases h0 : ν a (partitionSet n t) = 0 + · suffices κ a (partitionSet n t ×ˢ s) = 0 by simp [h0, this] + have h0' : fst κ a (partitionSet n t) = 0 := + le_antisymm ((hκν a _).trans h0.le) zero_le' + rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] at h0' + refine measure_mono_null (fun x ↦ ?_) h0' + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + · exact ENNReal.div_lt_top (measure_ne_top _ _) h0 + congr + have : ∫⁻ t in u, κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t) ∂(ν a) + = ∫⁻ _ in u, κ a (u ×ˢ s) / ν a u ∂(ν a) := by + refine set_lintegral_congr_fun hu_meas (ae_of_all _ (fun t ht ↦ ?_)) + rw [partitionSet_of_mem hu ht] + rw [this] + simp only [MeasureTheory.lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter] + by_cases h0 : ν a u = 0 + · simp only [h0, mul_zero] + have h0' : fst κ a u = 0 := le_antisymm ((hκν a _).trans h0.le) zero_le' + rw [fst_apply' _ _ hu_meas] at h0' + refine (measure_mono_null ?_ h0').symm + intro p + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + rw [div_eq_mul_inv, mul_assoc, ENNReal.inv_mul_cancel h0, mul_one] + exact measure_ne_top _ _ + +lemma integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : + ∫ t, densityProcess κ ν n a t s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by + rw [← integral_univ, ← sUnion_countablePartition γ n, sUnion_eq_iUnion, + iUnion_prod_const, measure_iUnion] + rotate_left + · intro i j hij + simp only [Set.disjoint_prod, disjoint_self, bot_eq_empty] + refine Or.inl (disjoint_countablePartition i.prop j.prop ?_) + rw [ne_eq, Subtype.coe_inj] + exact hij + · exact fun k ↦ (measurableSet_countablePartition n k.prop).prod hs + rw [integral_iUnion] + rotate_left + · exact fun k ↦ measurableSet_countablePartition n k.prop + · refine fun i j hij ↦ disjoint_countablePartition i.prop j.prop ?_ + rw [ne_eq, Subtype.coe_inj] + exact hij + · exact (integrable_densityProcess hκν n a hs).integrableOn + rw [ENNReal.tsum_toReal_eq (fun _ ↦ measure_ne_top _ _)] + congr with k + rw [set_integral_densityProcess_I hκν _ _ hs k.prop] + +lemma set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} + (hA : MeasurableSet[partitionFiltration γ n] A) : + ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + refine induction_on_inter (m := partitionFiltration γ n) (s := countablePartition γ n) + (C := fun A ↦ ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) rfl + ?_ ?_ ?_ ?_ ?_ hA + · rintro s hs t ht hst + suffices s = t by rwa [this, inter_self] + by_contra h_ne + rw [← not_disjoint_iff_nonempty_inter] at hst + exact hst <| disjoint_countablePartition hs ht h_ne + · simp + · rintro u hu + rw [set_integral_densityProcess_I hκν _ _ hs hu] + · intro A hA hA_eq + have hA' : MeasurableSet A := (partitionFiltration γ).le _ _ hA + have h := integral_add_compl hA' (integrable_densityProcess hκν n a hs) + rw [hA_eq, integral_densityProcess hκν n a hs] at h + have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by rw [prod_diff_prod, compl_eq_univ_diff]; simp + rw [this, measure_diff (by intro x; simp) (hA'.prod hs) (measure_ne_top (κ a) _), + ENNReal.toReal_sub_of_le (measure_mono (by intro x; simp)) (measure_ne_top _ _), + eq_tsub_iff_add_eq_of_le, add_comm] + · exact h + · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] + exact measure_mono (by intro x; simp) + · intro f hf_disj hf h_eq + rw [integral_iUnion (fun i ↦ (partitionFiltration γ).le n _ (hf i)) hf_disj + (integrable_densityProcess hκν _ _ hs).integrableOn] + simp_rw [h_eq] + rw [iUnion_prod_const, + measure_iUnion _ (fun i ↦ ((partitionFiltration γ).le n _ (hf i)).prod hs)] + · rw [ENNReal.tsum_toReal_eq] + exact fun _ ↦ measure_ne_top _ _ + · intro i j hij + rw [Function.onFun, Set.disjoint_prod] + exact Or.inl (hf_disj hij) + +lemma set_integral_densityProcess_of_le (hκν : fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] {n m : ℕ} (hnm : n ≤ m) (a : α) {s : Set β} (hs : MeasurableSet s) + {A : Set γ} (hA : MeasurableSet[partitionFiltration γ n] A) : + ∫ t in A, densityProcess κ ν m a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := + set_integral_densityProcess hκν m a hs ((partitionFiltration γ).mono hnm A hA) + +lemma condexp_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + {i j : ℕ} (hij : i ≤ j) (a : α) {s : Set β} (hs : MeasurableSet s) : + (ν a)[fun t ↦ densityProcess κ ν j a t s | partitionFiltration γ i] + =ᵐ[ν a] fun t ↦ densityProcess κ ν i a t s := by + refine (ae_eq_condexp_of_forall_set_integral_eq ?_ ?_ ?_ ?_ ?_).symm + · exact integrable_densityProcess hκν j a hs + · exact fun t _ _ ↦ (integrable_densityProcess hκν _ _ hs).integrableOn + · intro t ht _ + rw [set_integral_densityProcess hκν i a hs ht, + set_integral_densityProcess_of_le hκν hij a hs ht] + · exact StronglyMeasurable.aeStronglyMeasurable' + (stronglyMeasurable_partitionFiltration_densityProcess κ ν i a hs) + +lemma martingale_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) : + Martingale (fun n t ↦ densityProcess κ ν n a t s) (partitionFiltration γ) (ν a) := + ⟨adapted_densityProcess κ ν a hs, fun _ _ h ↦ condexp_densityProcess hκν h a hs⟩ + +lemma densityProcess_mono_set (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) + {s s' : Set β} (h : s ⊆ s') : + densityProcess κ ν n a t s ≤ densityProcess κ ν n a t s' := by + unfold densityProcess + by_cases h0 : ν a (partitionSet n t) = 0 + · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] + simp + have h_ne_top : ∀ s, κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t) ≠ ⊤ := by + intro s + rw [ne_eq, ENNReal.div_eq_top] + simp only [ne_eq, h0, and_false, false_or, not_and, not_not] + refine fun h_top ↦ eq_top_mono ?_ h_top + exact meas_partitionSet_le_of_fst_le hκν n a t s + rw [ENNReal.toReal_le_toReal (h_ne_top s) (h_ne_top s')] + gcongr + simp [prod_subset_prod_iff, subset_rfl, h] + +lemma densityProcess_mono_kernel_left {κ' : kernel α (γ × β)} (hκκ' : κ ≤ κ') + (hκ'ν : fst κ' ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : + densityProcess κ ν n a t s ≤ densityProcess κ' ν n a t s := by + unfold densityProcess + by_cases h0 : ν a (partitionSet n t) = 0 + · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] + simp + have h_le : κ' a (partitionSet n t ×ˢ s) ≤ ν a (partitionSet n t) := + meas_partitionSet_le_of_fst_le hκ'ν n a t s + rw [ENNReal.toReal_le_toReal] + · gcongr + exact hκκ' _ _ + · rw [ne_eq, ENNReal.div_eq_top] + simp only [ne_eq, h0, and_false, false_or, not_and, not_not] + refine fun h_top ↦ eq_top_mono ?_ h_top + exact (hκκ' _ _).trans h_le + · rw [ne_eq, ENNReal.div_eq_top] + simp only [ne_eq, h0, and_false, false_or, not_and, not_not] + exact fun h_top ↦ eq_top_mono h_le h_top + +lemma densityProcess_antitone_kernel_right {ν' : kernel α γ} + (hνν' : ν ≤ ν') (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : + densityProcess κ ν' n a t s ≤ densityProcess κ ν n a t s := by + unfold densityProcess + have h_le : κ a (partitionSet n t ×ˢ s) ≤ ν a (partitionSet n t) := + meas_partitionSet_le_of_fst_le hκν n a t s + by_cases h0 : ν a (partitionSet n t) = 0 + · simp [le_antisymm (h_le.trans h0.le) zero_le', h0] + have h0' : ν' a (partitionSet n t) ≠ 0 := + fun h ↦ h0 (le_antisymm ((hνν' _ _).trans h.le) zero_le') + rw [ENNReal.toReal_le_toReal] + · gcongr + exact hνν' _ _ + · simp only [ne_eq, ENNReal.div_eq_top, h0', and_false, false_or, not_and, not_not] + refine fun h_top ↦ eq_top_mono ?_ h_top + exact h_le.trans (hνν' _ _) + · simp only [ne_eq, ENNReal.div_eq_top, h0, and_false, false_or, not_and, not_not] + exact fun h_top ↦ eq_top_mono h_le h_top + +@[simp] +lemma densityProcess_empty (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (t : γ) : + densityProcess κ ν n a t ∅ = 0 := by + simp [densityProcess] + +lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (γ × β)) (ν : kernel α γ) + [IsFiniteKernel κ] (n : ℕ) (a : α) (t : γ) + (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) + (hs_meas : ∀ n, MeasurableSet (s n)) : + Tendsto (fun m ↦ densityProcess κ ν n a t (s m)) atTop (𝓝 (densityProcess κ ν n a t ∅)) := by + simp_rw [densityProcess] + by_cases h0 : ν a (partitionSet n t) = 0 + · simp_rw [h0, ENNReal.toReal_div] + simp + refine (ENNReal.tendsto_toReal ?_).comp ?_ + · rw [ne_eq, ENNReal.div_eq_top] + push_neg + simp + refine ENNReal.Tendsto.div_const ?_ ?_ + · have h := tendsto_measure_iInter (μ := κ a) + (s := fun m ↦ partitionSet n t ×ˢ s m) ?_ ?_ ?_ + · convert h + rw [← prod_iInter, hs_iInter] + · exact fun n ↦ MeasurableSet.prod (measurableSet_partitionSet _ _) (hs_meas n) + · intro m m' hmm' + simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] + exact Or.inl <| hs hmm' + · exact ⟨0, measure_ne_top _ _⟩ + · simp only [prod_empty, OuterMeasure.empty', ne_eq, not_true_eq_false, false_or, h0, + not_false_iff] + +lemma tendsto_densityProcess_atTop_of_antitone (κ : kernel α (γ × β)) (ν : kernel α γ) + [IsFiniteKernel κ] (n : ℕ) (a : α) (t : γ) + (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) + (hs_meas : ∀ n, MeasurableSet (s n)) : + Tendsto (fun m ↦ densityProcess κ ν n a t (s m)) atTop (𝓝 0) := by + rw [← densityProcess_empty κ ν n a t] + exact tendsto_densityProcess_atTop_empty_of_antitone κ ν n a t s hs hs_iInter hs_meas + +lemma tendsto_densityProcess_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : + ∀ᵐ t ∂(ν a), Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop + (𝓝 ((partitionFiltration γ).limitProcess + (fun n t ↦ densityProcess κ ν n a t s) (ν a) t)) := by + refine Submartingale.ae_tendsto_limitProcess (martingale_densityProcess hκν a hs).submartingale + (R := (ν a univ).toNNReal) (fun n ↦ ?_) + refine (snorm_densityProcess_le hκν n a s).trans_eq ?_ + rw [ENNReal.coe_toNNReal] + exact measure_ne_top _ _ + +lemma limitProcess_densityProcess_mem_L1 (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) : + Memℒp ((partitionFiltration γ).limitProcess + (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 (ν a) := by + refine Submartingale.memℒp_limitProcess (martingale_densityProcess hκν a hs).submartingale + (R := (ν a univ).toNNReal) (fun n ↦ ?_) + refine (snorm_densityProcess_le hκν n a s).trans_eq ?_ + rw [ENNReal.coe_toNNReal] + exact measure_ne_top _ _ + +lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : + Tendsto (fun n ↦ snorm ((fun t ↦ densityProcess κ ν n a t s) + - (partitionFiltration γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) + 1 (ν a)) atTop (𝓝 0) := by + refine Submartingale.tendsto_snorm_one_limitProcess ?_ ?_ + · exact (martingale_densityProcess hκν a hs).submartingale + · refine uniformIntegrable_of le_rfl ENNReal.one_ne_top ?_ ?_ + · exact fun n ↦ (measurable_densityProcess_right κ ν n a hs).aestronglyMeasurable + · refine fun ε _ ↦ ⟨2, fun n ↦ le_of_eq_of_le ?_ (?_ : 0 ≤ ENNReal.ofReal ε)⟩ + · suffices {x | 2 ≤ ‖densityProcess κ ν n a x s‖₊} = ∅ by simp [this] + ext x + simp only [mem_setOf_eq, mem_empty_iff_false, iff_false, not_le] + refine (?_ : _ ≤ (1 : ℝ≥0)).trans_lt one_lt_two + rw [Real.nnnorm_of_nonneg (densityProcess_nonneg _ _ _ _ _ _)] + exact mod_cast (densityProcess_le_one hκν _ _ _ _) + · simp + +lemma tendsto_snorm_one_restrict_densityProcess_limitProcess [IsFiniteKernel κ] [IsFiniteKernel ν] + (hκν : fst κ ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : + Tendsto (fun n ↦ snorm ((fun t ↦ densityProcess κ ν n a t s) + - (partitionFiltration γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) + 1 ((ν a).restrict A)) atTop (𝓝 0) := + tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds + (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) (fun _ ↦ zero_le') + (fun _ ↦ snorm_restrict_le _ _ _ _) + +end DensityProcess + +section Density + +/-- Density of the kernel `κ` with respect to `ν`. This is a function `α → γ → Set β → ℝ` which +is measurable on `α × γ` for all measurable sets `s : Set β` and satisfies that +`∫ t in A, density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal` for all measurable `A : Set γ`. -/ +noncomputable +def density (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) (t : γ) (s : Set β) : ℝ := + limsup (fun n ↦ densityProcess κ ν n a t s) atTop + +lemma density_ae_eq_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) : + (fun t ↦ density κ ν a t s) + =ᵐ[ν a] (partitionFiltration γ).limitProcess + (fun n t ↦ densityProcess κ ν n a t s) (ν a) := by + filter_upwards [tendsto_densityProcess_limitProcess hκν a hs] with t ht using ht.limsup_eq + +lemma tendsto_m_density (hκν : fst κ ≤ ν) (a : α) [IsFiniteKernel κ] [IsFiniteKernel ν] + {s : Set β} (hs : MeasurableSet s) : + ∀ᵐ t ∂(ν a), + Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop (𝓝 (density κ ν a t s)) := by + filter_upwards [tendsto_densityProcess_limitProcess hκν a hs, density_ae_eq_limitProcess hκν a hs] + with t h1 h2 using h2 ▸ h1 + +lemma measurable_density (κ : kernel α (γ × β)) (ν : kernel α γ) + {s : Set β} (hs : MeasurableSet s) : + Measurable (fun (p : α × γ) ↦ density κ ν p.1 p.2 s) := + measurable_limsup (fun n ↦ measurable_densityProcess κ ν n hs) + +lemma measurable_density_left (κ : kernel α (γ × β)) (ν : kernel α γ) (t : γ) + {s : Set β} (hs : MeasurableSet s) : + Measurable (fun a ↦ density κ ν a t s) := by + change Measurable ((fun (p : α × γ) ↦ density κ ν p.1 p.2 s) ∘ (fun a ↦ (a, t))) + exact (measurable_density κ ν hs).comp measurable_prod_mk_right + +lemma measurable_density_right (κ : kernel α (γ × β)) (ν : kernel α γ) + {s : Set β} (hs : MeasurableSet s) (a : α) : + Measurable (fun t ↦ density κ ν a t s) := by + change Measurable ((fun (p : α × γ) ↦ density κ ν p.1 p.2 s) ∘ (fun t ↦ (a, t))) + exact (measurable_density κ ν hs).comp measurable_prod_mk_left + +lemma density_mono_set (hκν : fst κ ≤ ν) (a : α) (t : γ) {s s' : Set β} (h : s ⊆ s') : + density κ ν a t s ≤ density κ ν a t s' := by + refine limsup_le_limsup ?_ ?_ ?_ + · exact eventually_of_forall (fun n ↦ densityProcess_mono_set hκν n a t h) + · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) + · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ + +lemma density_nonneg (hκν : fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : + 0 ≤ density κ ν a t s := by + refine le_limsup_of_frequently_le ?_ ?_ + · exact frequently_of_forall (fun n ↦ densityProcess_nonneg _ _ _ _ _ _) + · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ + +lemma density_le_one (hκν : fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : + density κ ν a t s ≤ 1 := by + refine limsup_le_of_le ?_ ?_ + · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) + · exact eventually_of_forall (fun n ↦ densityProcess_le_one hκν _ _ _ _) + +section Integral + +lemma snorm_density_le (hκν : fst κ ≤ ν) (a : α) (s : Set β) : + snorm (fun t ↦ density κ ν a t s) 1 (ν a) ≤ ν a univ := by + refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ + · simp only [Real.norm_eq_abs, abs_of_nonneg (density_nonneg hκν a t s), + density_le_one hκν a t s] + · simp + +lemma integrable_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) : + Integrable (fun t ↦ density κ ν a t s) (ν a) := by + rw [← memℒp_one_iff_integrable] + refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ + · exact measurable_density_right κ ν hs a + · exact (snorm_density_le hκν a s).trans_lt (measure_lt_top _ _) + +lemma tendsto_set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : + Tendsto (fun i ↦ ∫ x in A, densityProcess κ ν i a x s ∂(ν a)) atTop + (𝓝 (∫ x in A, density κ ν a x s ∂(ν a))) := by + refine tendsto_set_integral_of_L1' (μ := ν a) (fun t ↦ density κ ν a t s) + (integrable_density hκν a hs) (F := fun i t ↦ densityProcess κ ν i a t s) (l := atTop) + (eventually_of_forall (fun n ↦ integrable_densityProcess hκν _ _ hs)) ?_ A + refine (tendsto_congr fun n ↦ ?_).mp (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) + refine snorm_congr_ae ?_ + exact EventuallyEq.rfl.sub (density_ae_eq_limitProcess hκν a hs).symm + +/-- Auxiliary lemma for `set_integral_density`. -/ +lemma set_integral_density_of_measurableSet (hκν : fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} + (hA : MeasurableSet[partitionFiltration γ n] A) : + ∫ t in A, density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + suffices ∫ t in A, density κ ν a t s ∂(ν a) = ∫ t in A, densityProcess κ ν n a t s ∂(ν a) by + exact this ▸ set_integral_densityProcess hκν _ _ hs hA + suffices ∫ t in A, density κ ν a t s ∂(ν a) + = limsup (fun i ↦ ∫ t in A, densityProcess κ ν i a t s ∂(ν a)) atTop by + rw [this, ← limsup_const (α := ℕ) (f := atTop) (∫ t in A, densityProcess κ ν n a t s ∂(ν a)), + limsup_congr] + simp only [eventually_atTop] + refine ⟨n, fun m hnm ↦ ?_⟩ + rw [set_integral_densityProcess_of_le hκν hnm _ hs hA, + set_integral_densityProcess hκν _ _ hs hA] + -- use L1 convergence + have h := tendsto_set_integral_densityProcess hκν a hs A + rw [h.limsup_eq] + +lemma integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) : + ∫ t, density κ ν a t s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by + rw [← integral_univ, set_integral_density_of_measurableSet hκν 0 a hs MeasurableSet.univ] + +lemma set_integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet A) : + ∫ t in A, density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + have hA' : MeasurableSet[⨆ n, partitionFiltration γ n] A := by rwa [iSup_partitionFiltration] + refine induction_on_inter (m := ⨆ n, partitionFiltration γ n) + (C := fun A ↦ ∫ t in A, density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) + (measurableSpace_iSup_eq (partitionFiltration γ)) ?_ ?_ ?_ ?_ ?_ hA' + · rintro s ⟨n, hs⟩ t ⟨m, ht⟩ _ + exact ⟨max n m, ((partitionFiltration γ).mono (le_max_left n m) _ hs).inter + ((partitionFiltration γ).mono (le_max_right n m) _ ht)⟩ + · simp + · intro A ⟨n, hA⟩ + exact set_integral_density_of_measurableSet hκν n a hs hA + · intro A hA hA_eq + rw [iSup_partitionFiltration] at hA + have h := integral_add_compl hA (integrable_density hκν a hs) + rw [hA_eq, integral_density hκν a hs] at h + have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by + rw [prod_diff_prod, compl_eq_univ_diff] + simp + rw [this, measure_diff (by intro x; simp) (hA.prod hs) (measure_ne_top (κ a) _), + ENNReal.toReal_sub_of_le (measure_mono (by intro x; simp)) (measure_ne_top _ _)] + rw [eq_tsub_iff_add_eq_of_le, add_comm] + · exact h + · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] + exact measure_mono (by intro x; simp) + · intro f hf_disj hf h_eq + rw [integral_iUnion _ hf_disj (integrable_density hκν _ hs).integrableOn] + · simp_rw [h_eq] + rw [← ENNReal.tsum_toReal_eq (fun _ ↦ measure_ne_top _ _)] + congr + rw [iUnion_prod_const, measure_iUnion] + · intro i j hij + rw [Function.onFun, Set.disjoint_prod] + exact Or.inl (hf_disj hij) + · rw [iSup_partitionFiltration] at hf + exact fun i ↦ (hf i).prod hs + · rwa [iSup_partitionFiltration] at hf + +lemma set_lintegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet A) : + ∫⁻ t in A, ENNReal.ofReal (density κ ν a t s) ∂(ν a) = κ a (A ×ˢ s) := by + rw [← ofReal_integral_eq_lintegral_ofReal] + · rw [set_integral_density hκν a hs hA, + ENNReal.ofReal_toReal (measure_ne_top _ _)] + · exact (integrable_density hκν a hs).restrict + · exact ae_of_all _ (fun _ ↦ density_nonneg hκν _ _ _) + +lemma lintegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set β} (hs : MeasurableSet s) : + ∫⁻ t, ENNReal.ofReal (density κ ν a t s) ∂(ν a) = κ a (univ ×ˢ s) := by + rw [← set_lintegral_univ] + exact set_lintegral_density hκν a hs MeasurableSet.univ + +end Integral + +lemma tendsto_integral_density_of_monotone (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) + (hs_meas : ∀ n, MeasurableSet (s n)) : + Tendsto (fun m ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop (𝓝 (κ a univ).toReal) := by + simp_rw [integral_density hκν a (hs_meas _)] + have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := κ a univ) ?_ + swap + · rw [mem_nhds_iff] + refine ⟨Iio (κ a univ + 1), fun x hx ↦ ne_top_of_lt (?_ : x < κ a univ + 1), isOpen_Iio, ?_⟩ + · simpa using hx + · simp only [mem_Iio] + exact ENNReal.lt_add_right (measure_ne_top _ _) one_ne_zero + refine h_cont.tendsto.comp ?_ + have h := tendsto_measure_iUnion (s := fun n ↦ univ ×ˢ s n) (μ := κ a) ?_ + swap; · intro n m hnm x; simp only [mem_prod, mem_univ, true_and]; exact fun h ↦ hs hnm h + convert h + rw [← prod_iUnion, hs_iUnion] + simp only [univ_prod_univ, measure_univ] + +lemma tendsto_integral_density_of_antitone (hκν : fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) + (hs_meas : ∀ n, MeasurableSet (s n)) : + Tendsto (fun m ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop (𝓝 0) := by + simp_rw [integral_density hκν a (hs_meas _)] + rw [← ENNReal.zero_toReal] + have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := 0) ?_ + swap + · rw [mem_nhds_iff] + refine ⟨Iio 1, fun x hx ↦ ne_top_of_lt (?_ : x < 1), isOpen_Iio, ?_⟩ + · simpa using hx + · simp + refine h_cont.tendsto.comp ?_ + have h := tendsto_measure_iInter (s := fun n ↦ univ ×ˢ s n) (μ := κ a) + (fun n ↦ MeasurableSet.univ.prod (hs_meas n)) ?_ ?_ + rotate_left + · intro n m hnm x; simp only [mem_prod, mem_univ, true_and]; exact fun h ↦ hs hnm h + · refine ⟨0, measure_ne_top _ _⟩ + convert h + rw [← prod_iInter, hs_iInter] + simp only [ne_eq, prod_empty, OuterMeasure.empty', forall_exists_index] + +lemma tendsto_density_atTop_ae_of_antitone (hκν : fst κ ≤ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) + (hs_meas : ∀ n, MeasurableSet (s n)) : + ∀ᵐ t ∂(ν a), Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 0) := by + have h_anti : ∀ t, Antitone (fun m ↦ density κ ν a t (s m)) := + fun t n m hnm ↦ density_mono_set hκν a t (hs hnm) + have h_le_one : ∀ m t, density κ ν a t (s m) ≤ 1 := fun m t ↦ density_le_one hκν a t (s m) + -- for all `t`, `fun m ↦ density κ a (s m) t` has a limit + have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := by + intro t + have h_tendsto : Tendsto (fun m ↦ density κ ν a t (s m)) atTop atBot ∨ + ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := + tendsto_of_antitone (h_anti t) + cases' h_tendsto with h_absurd h_tendsto + · rw [tendsto_atTop_atBot_iff_of_antitone (h_anti t)] at h_absurd + obtain ⟨r, hr⟩ := h_absurd (-1) + have h_nonneg := density_nonneg hκν a t (s r) + linarith + · exact h_tendsto + -- let `F` be the pointwise limit of `fun m ↦ density κ a (s m) t` for all `t` + let F : γ → ℝ := fun t ↦ (h_exists t).choose + have hF_tendsto : ∀ t, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 (F t)) := + fun t ↦ (h_exists t).choose_spec + have hF_nonneg : ∀ t, 0 ≤ F t := + fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ density_nonneg hκν a t (s m)) + have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) + have hF_int : Integrable F (ν a) := by + rw [← memℒp_one_iff_integrable] + refine ⟨?_, ?_⟩ + · refine aestronglyMeasurable_of_tendsto_ae atTop (fun n ↦ ?_) (ae_of_all _ hF_tendsto) + exact (measurable_density_right κ ν (hs_meas _) a).aestronglyMeasurable + · rw [snorm_one_eq_lintegral_nnnorm] + calc ∫⁻ x, ‖F x‖₊ ∂(ν a) ≤ ∫⁻ _, 1 ∂(ν a) := by + refine lintegral_mono (fun x ↦ ?_) + rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, ENNReal.ofReal_le_one, + abs_of_nonneg (hF_nonneg _)] + exact hF_le_one _ + _ < ⊤ := by + simp only [MeasureTheory.lintegral_const, one_mul] + exact measure_lt_top _ _ + -- it suffices to show that the limit `F` is 0 a.e. + suffices F=ᵐ[ν a] 0 by + filter_upwards [this] with t ht_eq + simp only [Pi.zero_apply] at ht_eq + rw [← ht_eq] + exact hF_tendsto t + -- since `F` is nonnegative, proving that its integral is 0 is sufficient to get that + -- `F` is 0 a.e. + rw [← integral_eq_zero_iff_of_nonneg hF_nonneg hF_int] + have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop + (𝓝 (∫ t, F t ∂(ν a))) := by + refine integral_tendsto_of_tendsto_of_antitone ?_ hF_int ?_ ?_ + · exact fun n ↦ integrable_density hκν _ (hs_meas n) + · exact ae_of_all _ h_anti + · exact ae_of_all _ hF_tendsto + have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop + (𝓝 0) := by + exact tendsto_integral_density_of_antitone hκν a s hs hs_iInter hs_meas + exact tendsto_nhds_unique h_integral h_integral' + +section UnivFst + +/-! We specialize to `ν = fst κ`, for which `density κ (fst κ) a t univ = 1` almost everywhere. -/ + +lemma densityProcess_univ [IsFiniteKernel κ] (n : ℕ) (a : α) (t : γ) : + densityProcess κ (fst κ) n a t univ = if fst κ a (partitionSet n t) = 0 then 0 else 1 := by + rw [densityProcess] + by_cases h : fst κ a (partitionSet n t) = 0 + · simp [h] + by_cases h' : κ a (partitionSet n t ×ˢ univ) = 0 + · simp [h'] + · rw [ENNReal.div_zero h'] + simp + · simp only [h, ite_false] + rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] + have : partitionSet n t ×ˢ univ = {p : γ × β | p.1 ∈ partitionSet n t} := by + ext x + simp + rw [this, ENNReal.div_self] + · simp + · rwa [fst_apply' _ _ (measurableSet_partitionSet _ _)] at h + · exact measure_ne_top _ _ + +lemma densityProcess_univ_ae (κ : kernel α (γ × β)) [IsFiniteKernel κ] (n : ℕ) (a : α) : + ∀ᵐ t ∂(fst κ a), densityProcess κ (fst κ) n a t univ = 1 := by + rw [ae_iff] + have : {t | ¬ densityProcess κ (fst κ) n a t univ = 1} + ⊆ {t | fst κ a (partitionSet n t) = 0} := by + intro t ht + simp only [mem_setOf_eq] at ht ⊢ + rw [densityProcess_univ] at ht + simpa using ht + refine measure_mono_null this ?_ + have : {t | fst κ a (partitionSet n t) = 0} + ⊆ ⋃ (u) (_ : u ∈ countablePartition γ n) (_ : fst κ a u = 0), u := by + intro t ht + simp only [mem_setOf_eq, mem_iUnion, exists_prop] at ht ⊢ + exact ⟨partitionSet n t, partitionSet_mem _ _, ht, mem_partitionSet _ _⟩ + refine measure_mono_null this ?_ + rw [measure_biUnion] + · simp + · exact (finite_countablePartition _ _).countable + · intro s hs t ht hst + simp only [disjoint_iUnion_right, disjoint_iUnion_left] + exact fun _ _ ↦ disjoint_countablePartition hs ht hst + · intro s hs + by_cases h : fst κ a s = 0 + · simp [h, measurableSet_countablePartition n hs] + · simp [h] + +lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (γ × β)) + (n : ℕ) (a : α) (t : γ) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) : + Tendsto (fun m ↦ densityProcess κ (fst κ) n a t (s m)) atTop + (𝓝 (densityProcess κ (fst κ) n a t univ)) := by + simp_rw [densityProcess] + refine (ENNReal.tendsto_toReal ?_).comp ?_ + · rw [ne_eq, ENNReal.div_eq_top] + push_neg + simp_rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] + constructor + · refine fun h h0 ↦ h (measure_mono_null (fun x ↦ ?_) h0) + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + · refine fun h_top ↦ eq_top_mono (measure_mono (fun x ↦ ?_)) h_top + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + by_cases h0 : fst κ a (partitionSet n t) = 0 + · rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] at h0 ⊢ + suffices ∀ m, κ a (partitionSet n t ×ˢ s m) = 0 by + simp only [this, h0, ENNReal.zero_div, tendsto_const_nhds_iff] + suffices κ a (partitionSet n t ×ˢ univ) = 0 by + simp only [this, ENNReal.zero_div] + convert h0 + ext x + simp only [mem_prod, mem_univ, and_true, mem_setOf_eq] + refine fun m ↦ measure_mono_null (fun x ↦ ?_) h0 + simp only [mem_prod, mem_setOf_eq, and_imp] + exact fun h _ ↦ h + refine ENNReal.Tendsto.div_const ?_ ?_ + · have h := tendsto_measure_iUnion (μ := κ a) + (s := fun m ↦ partitionSet n t ×ˢ s m) ?_ + swap + · intro m m' hmm' + simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] + exact Or.inl <| hs hmm' + convert h + rw [← prod_iUnion, hs_iUnion] + · exact Or.inr h0 + +lemma tendsto_densityProcess_atTop_ae_of_monotone (κ : kernel α (γ × β)) [IsFiniteKernel κ] + (n : ℕ) (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) : + ∀ᵐ t ∂(fst κ a), Tendsto (fun m ↦ densityProcess κ (fst κ) n a t (s m)) atTop (𝓝 1) := by + filter_upwards [densityProcess_univ_ae κ n a] with t ht + rw [← ht] + exact tendsto_densityProcess_atTop_univ_of_monotone κ n a t s hs hs_iUnion + +lemma density_univ (κ : kernel α (γ × β)) [IsFiniteKernel κ] (a : α) : + ∀ᵐ t ∂(fst κ a), density κ (fst κ) a t univ = 1 := by + have h := fun n ↦ densityProcess_univ_ae κ n a + rw [← ae_all_iff] at h + filter_upwards [h] with t ht + simp [density, ht] + +lemma tendsto_density_atTop_ae_of_monotone [IsFiniteKernel κ] + (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) + (hs_meas : ∀ n, MeasurableSet (s n)) : + ∀ᵐ t ∂(fst κ a), Tendsto (fun m ↦ density κ (fst κ) a t (s m)) atTop (𝓝 1) := by + let ν := fst κ + have h_mono : ∀ t, Monotone (fun m ↦ density κ (fst κ) a t (s m)) := + fun t n m hnm ↦ density_mono_set le_rfl a t (hs hnm) + have h_le_one : ∀ m t, density κ ν a t (s m) ≤ 1 := + fun m t ↦ density_le_one le_rfl a t (s m) + -- for all `t`, `fun m ↦ density κ a (s m) t` has a limit + have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := by + intro t + have h_tendsto : Tendsto (fun m ↦ density κ ν a t (s m)) atTop atTop ∨ + ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := + tendsto_of_monotone (h_mono t) + cases' h_tendsto with h_absurd h_tendsto + · rw [tendsto_atTop_atTop_iff_of_monotone (h_mono t)] at h_absurd + obtain ⟨r, hr⟩ := h_absurd 2 + exact absurd (hr.trans (h_le_one r t)) one_lt_two.not_le + · exact h_tendsto + -- let `F` be the pointwise limit of `fun m ↦ density κ a (s m) t` for all `t` + let F : γ → ℝ := fun t ↦ (h_exists t).choose + have hF_tendsto : ∀ t, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 (F t)) := + fun t ↦ (h_exists t).choose_spec + have hF_nonneg : ∀ t, 0 ≤ F t := + fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ density_nonneg le_rfl a t (s m)) + have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) + have hF_int : Integrable F (ν a) := by + rw [← memℒp_one_iff_integrable] + constructor + · refine aestronglyMeasurable_of_tendsto_ae atTop (fun n ↦ ?_) (ae_of_all _ hF_tendsto) + exact (measurable_density_right κ ν (hs_meas _) a).aestronglyMeasurable + · rw [snorm_one_eq_lintegral_nnnorm] + calc ∫⁻ x, ‖F x‖₊ ∂(ν a) ≤ ∫⁻ _, 1 ∂(ν a) := by + refine lintegral_mono (fun x ↦ ?_) + rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, ENNReal.ofReal_le_one, + abs_of_nonneg (hF_nonneg _)] + exact hF_le_one _ + _ < ⊤ := by simp only [MeasureTheory.lintegral_const, measure_univ, one_mul, measure_lt_top] + -- it suffices to show that the limit `F` is 1 a.e. + suffices F =ᵐ[ν a] (fun _ ↦ 1) by + filter_upwards [this] with t ht_eq + rw [← ht_eq] + exact hF_tendsto t + -- since `F` is at most 1, proving that its integral is the same as the integral of 1 will tell + -- us that `F` is 1 a.e. + rw [← integral_eq_iff_of_ae_le hF_int (integrable_const _) (ae_of_all _ hF_le_one)] + have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop + (𝓝 (∫ t, F t ∂(ν a))) := by + refine integral_tendsto_of_tendsto_of_monotone ?_ hF_int ?_ ?_ + · exact fun n ↦ integrable_density le_rfl _ (hs_meas n) + · exact ae_of_all _ h_mono + · exact ae_of_all _ hF_tendsto + have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop + (𝓝 (∫ _, 1 ∂(ν a))) := by + rw [MeasureTheory.integral_const] + simp only [smul_eq_mul, mul_one] + rw [fst_apply' _ _ MeasurableSet.univ] + exact tendsto_integral_density_of_monotone le_rfl a s hs hs_iUnion hs_meas + exact tendsto_nhds_unique h_integral h_integral' + +end UnivFst + +end Density + +end kernel + +end ProbabilityTheory From 0d8afaf6e999d659027609391b79493d74e72a5a Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 10:00:55 +0100 Subject: [PATCH 070/129] docstrings in RN --- Mathlib/Probability/Kernel/RadonNikodym.lean | 76 +++++++++++--------- 1 file changed, 44 insertions(+), 32 deletions(-) diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index fd6a35c9d2b9e..33bd2e45add85 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -7,54 +7,58 @@ import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.WithDensity /-! -# Radon-Nikodym derivative and Lebesgue decompusition for kernels +# Radon-Nikodym derivative and Lebesgue decomposition for kernels -Let `γ` be a countably generated measurable space and `κ ν : kernel α γ`. Then there exists... TODO +Let `γ` be a countably generated measurable space and `κ ν : kernel α γ` be finite kernels. +Then there exists a function `kernel.rnDeriv κ ν : α → γ → ℝ≥0∞` jointly measurable on `α × γ` +and a kernel `kernel.singularPart κ ν : kernel α γ` such that +* `κ = kernel.withDensity ν (kernel.rnDeriv κ ν) + kernel.singularPart κ ν`, +* for all `a : α`, `kernel.singularPart κ ν a ⟂ₘ ν a`, +* for all `a : α`, `kernel.singularPart κ ν a = 0 ↔ κ a ≪ ν a`, +* for all `a : α`, `kernel.withDensity ν (kernel.rnDeriv κ ν) a = 0 ↔ κ a ⟂ₘ ν a`. + +Furthermore, the sets `{a | κ a ≪ ν a}` and `{a | κ a ⟂ₘ ν a}` are measurable. ## Main definitions -* `ProbabilityTheory.kernel.rnDeriv` -* `ProbabilityTheory.kernel.singularPart` +* `ProbabilityTheory.kernel.rnDeriv`: a function `α → γ → ℝ≥0∞` jointly measurable on `α × γ` +* `ProbabilityTheory.kernel.singularPart`: a `kernel α γ` ## Main statements -* `fooBar_unique` - -## Notation - - - -## Implementation details - - +* `ProbabilityTheory.kernel.mutuallySingular_singularPart`: for all `a : α`, + `kernel.singularPart κ ν a ⟂ₘ ν a` +* `ProbabilityTheory.kernel.rnDeriv_add_singularPart`: + `kernel.withDensity ν (kernel.rnDeriv κ ν) + kernel.singularPart κ ν = κ` +* `ProbabilityTheory.kernel.measurableSet_absolutelyContinuous` : the set `{a | κ a ≪ ν a}` + is Measurable +* `ProbabilityTheory.kernel.measurableSet_mutuallySingular` : the set `{a | κ a ⟂ₘ ν a}` + is Measurable ## References Theorem 1.28 in [O. Kallenberg, Random Measures, Theory and Applications][kallenberg2017]. -## Tags - -Foobars, barfoos -/ open MeasureTheory Set Filter open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory -namespace ProbabilityTheory - -open ProbabilityTheory.kernel +namespace ProbabilityTheory.kernel variable {α γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} [MeasurableSpace.CountablyGenerated γ] {κ ν : kernel α γ} -lemma kernel.fst_map_prod_le_of_le (hκν : κ ≤ ν) : +lemma fst_map_prod_le_of_le (hκν : κ ≤ ν) : kernel.fst (kernel.map κ (fun a ↦ (a, ())) (@measurable_prod_mk_right γ Unit _ inferInstance _)) ≤ ν := by rwa [kernel.fst_map_prod _ measurable_id' measurable_const, kernel.map_id'] +/-- Auxiliary function used to define `ProbabilityTheory.kernel.rnDeriv` and +`ProbabilityTheory.kernel.singularPart`. -/ noncomputable -def kernel.rnDerivAux (κ ν : kernel α γ) (a : α) (x : γ) : ℝ := +def rnDerivAux (κ ν : kernel α γ) (a : α) (x : γ) : ℝ := kernel.density (kernel.map κ (fun a ↦ (a, ())) (@measurable_prod_mk_right γ Unit _ inferInstance _)) ν a x univ @@ -130,10 +134,17 @@ lemma withDensity_one_sub_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] simp only [ENNReal.ofReal_le_one] exact density_le_one (kernel.fst_map_prod_le_of_le h_le) _ _ _ -def kernel.mutuallySingularSet (κ ν : kernel α γ) : Set (α × γ) := +/-- A set of points in `α × γ` related to the absolute continuity / mutual singularity of +`κ` and `ν`. -/ +def mutuallySingularSet (κ ν : kernel α γ) : Set (α × γ) := {p | kernel.rnDerivAux κ (κ + ν) p.1 p.2 = 1} -def kernel.mutuallySingularSetSlice (κ ν : kernel α γ) (a : α) : Set γ := +/-- A set of points in `α × γ` related to the absolute continuity / mutual singularity of +`κ` and `ν`. That is, +* `kernel.withDensity ν (kernel.rnDeriv κ ν) a (kernel.mutuallySingularSetSlice κ ν a) = 0`, +* `kernel.singularPart κ ν a (kernel.mutuallySingularSetSlice κ ν a)ᶜ = 0`. + -/ +def mutuallySingularSetSlice (κ ν : kernel α γ) (a : α) : Set γ := {x | kernel.rnDerivAux κ (κ + ν) a x = 1} lemma measurableSet_mutuallySingularSet (κ ν : kernel α γ) : @@ -164,12 +175,13 @@ lemma measure_mutuallySingularSetSlice (κ ν : kernel α γ) [IsFiniteKernel κ simp only [mem_setOf_eq] at hx simp [hx] +/-- Radon-Nikodym derivative of the kernel `κ` with respect to the kernel `ν`. -/ noncomputable -def kernel.rnDeriv (κ ν : kernel α γ) (a : α) (x : γ) : ℝ≥0∞ := +def rnDeriv (κ ν : kernel α γ) (a : α) (x : γ) : ℝ≥0∞ := ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) / ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) -lemma kernel.rnDeriv_def (κ ν : kernel α γ) : +lemma rnDeriv_def (κ ν : kernel α γ) : kernel.rnDeriv κ ν = fun a x ↦ ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) / ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) := rfl @@ -191,21 +203,21 @@ lemma rnDeriv_eq_top_iff (κ ν : kernel α γ) (a : α) (x : γ) : exact ⟨fun h ↦ le_antisymm (rnDerivAux_le_one (le_add_of_nonneg_right bot_le)) h.2, fun h ↦ by simp [h]⟩ +/-- Singular part of the kernel `κ` with respect to the kernel `ν`. -/ noncomputable -def kernel.singularPart (κ ν : kernel α γ) [IsSFiniteKernel κ] [IsSFiniteKernel ν] : +def singularPart (κ ν : kernel α γ) [IsSFiniteKernel κ] [IsSFiniteKernel ν] : kernel α γ := kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x) - Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x) * kernel.rnDeriv κ ν a x) -lemma measurable_singularPart_fun (κ ν : kernel α γ) [IsSFiniteKernel κ] [IsSFiniteKernel ν] : +lemma measurable_singularPart_fun (κ ν : kernel α γ) : Measurable (fun p : α × γ ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) p.1 p.2) - Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) p.1 p.2) * kernel.rnDeriv κ ν p.1 p.2) := by refine (measurable_rnDerivAux _ _).ennreal_ofReal.sub (Measurable.mul ?_ (measurable_rnDeriv _ _)) exact (measurable_const.sub (measurable_rnDerivAux _ _)).ennreal_ofReal -lemma measurable_singularPart_fun_right (κ ν : kernel α γ) [IsSFiniteKernel κ] [IsSFiniteKernel ν] - (a : α) : +lemma measurable_singularPart_fun_right (κ ν : kernel α γ) (a : α) : Measurable (fun x : γ ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x) - Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x) * kernel.rnDeriv κ ν a x) := by change Measurable ((Function.uncurry fun a b ↦ @@ -316,14 +328,14 @@ lemma withDensity_rnDeriv_subset_compl_mutuallySingularSetSlice (κ ν : kernel · simp [hs' x hx] _ = κ a s := set_lintegral_rnDerivAux κ ν a hsm -lemma kernel.mutuallySingular_singularPart (κ ν : kernel α γ) +lemma mutuallySingular_singularPart (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : kernel.singularPart κ ν a ⟂ₘ ν a := by symm exact ⟨mutuallySingularSetSlice κ ν a, measurableSet_mutuallySingularSetSlice κ ν a, measure_mutuallySingularSetSlice κ ν a, singularPart_compl_mutuallySingularSetSlice κ ν a⟩ -lemma kernel.rnDeriv_add_singularPart (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : +lemma rnDeriv_add_singularPart (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : kernel.withDensity ν (kernel.rnDeriv κ ν) + kernel.singularPart κ ν = κ := by ext a s hs rw [← inter_union_diff s (mutuallySingularSetSlice κ ν a)] @@ -425,4 +437,4 @@ lemma measurableSet_mutuallySingular (κ ν : kernel α γ) [IsFiniteKernel κ] exact kernel.measurable_kernel_prod_mk_left (measurableSet_mutuallySingularSet κ ν).compl (measurableSet_singleton 0) -end ProbabilityTheory +end ProbabilityTheory.kernel From e290aed8bf1e7aaf0bbbc45f0783ee8e5f86e5a3 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 10:05:44 +0100 Subject: [PATCH 071/129] add RN for kernels --- Mathlib.lean | 1 + Mathlib/Probability/Kernel/RadonNikodym.lean | 440 +++++++++++++++++++ docs/references.bib | 13 +- 3 files changed, 453 insertions(+), 1 deletion(-) create mode 100644 Mathlib/Probability/Kernel/RadonNikodym.lean diff --git a/Mathlib.lean b/Mathlib.lean index b6a49df57bc62..e9d727939d69b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3132,6 +3132,7 @@ import Mathlib.Probability.Kernel.IntegralCompProd import Mathlib.Probability.Kernel.Invariance import Mathlib.Probability.Kernel.MeasurableIntegral import Mathlib.Probability.Kernel.MeasureCompProd +import Mathlib.Probability.Kernel.RadonNikodym import Mathlib.Probability.Kernel.WithDensity import Mathlib.Probability.Martingale.Basic import Mathlib.Probability.Martingale.BorelCantelli diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean new file mode 100644 index 0000000000000..33bd2e45add85 --- /dev/null +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -0,0 +1,440 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Probability.Kernel.Disintegration.Density +import Mathlib.Probability.Kernel.WithDensity + +/-! +# Radon-Nikodym derivative and Lebesgue decomposition for kernels + +Let `γ` be a countably generated measurable space and `κ ν : kernel α γ` be finite kernels. +Then there exists a function `kernel.rnDeriv κ ν : α → γ → ℝ≥0∞` jointly measurable on `α × γ` +and a kernel `kernel.singularPart κ ν : kernel α γ` such that +* `κ = kernel.withDensity ν (kernel.rnDeriv κ ν) + kernel.singularPart κ ν`, +* for all `a : α`, `kernel.singularPart κ ν a ⟂ₘ ν a`, +* for all `a : α`, `kernel.singularPart κ ν a = 0 ↔ κ a ≪ ν a`, +* for all `a : α`, `kernel.withDensity ν (kernel.rnDeriv κ ν) a = 0 ↔ κ a ⟂ₘ ν a`. + +Furthermore, the sets `{a | κ a ≪ ν a}` and `{a | κ a ⟂ₘ ν a}` are measurable. + +## Main definitions + +* `ProbabilityTheory.kernel.rnDeriv`: a function `α → γ → ℝ≥0∞` jointly measurable on `α × γ` +* `ProbabilityTheory.kernel.singularPart`: a `kernel α γ` + +## Main statements + +* `ProbabilityTheory.kernel.mutuallySingular_singularPart`: for all `a : α`, + `kernel.singularPart κ ν a ⟂ₘ ν a` +* `ProbabilityTheory.kernel.rnDeriv_add_singularPart`: + `kernel.withDensity ν (kernel.rnDeriv κ ν) + kernel.singularPart κ ν = κ` +* `ProbabilityTheory.kernel.measurableSet_absolutelyContinuous` : the set `{a | κ a ≪ ν a}` + is Measurable +* `ProbabilityTheory.kernel.measurableSet_mutuallySingular` : the set `{a | κ a ⟂ₘ ν a}` + is Measurable + +## References + +Theorem 1.28 in [O. Kallenberg, Random Measures, Theory and Applications][kallenberg2017]. + +-/ + +open MeasureTheory Set Filter + +open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory + +namespace ProbabilityTheory.kernel + +variable {α γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} + [MeasurableSpace.CountablyGenerated γ] {κ ν : kernel α γ} + +lemma fst_map_prod_le_of_le (hκν : κ ≤ ν) : + kernel.fst (kernel.map κ (fun a ↦ (a, ())) + (@measurable_prod_mk_right γ Unit _ inferInstance _)) ≤ ν := by + rwa [kernel.fst_map_prod _ measurable_id' measurable_const, kernel.map_id'] + +/-- Auxiliary function used to define `ProbabilityTheory.kernel.rnDeriv` and +`ProbabilityTheory.kernel.singularPart`. -/ +noncomputable +def rnDerivAux (κ ν : kernel α γ) (a : α) (x : γ) : ℝ := + kernel.density (kernel.map κ (fun a ↦ (a, ())) + (@measurable_prod_mk_right γ Unit _ inferInstance _)) ν a x univ + +lemma rnDerivAux_nonneg (hκν : κ ≤ ν) {a : α} {x : γ} : 0 ≤ kernel.rnDerivAux κ ν a x := + density_nonneg (kernel.fst_map_prod_le_of_le hκν) _ _ _ + +lemma rnDerivAux_le_one (hκν : κ ≤ ν) {a : α} {x : γ} : kernel.rnDerivAux κ ν a x ≤ 1 := + density_le_one (kernel.fst_map_prod_le_of_le hκν) _ _ _ + +lemma measurable_rnDerivAux (κ ν : kernel α γ) : + Measurable (fun p : α × γ ↦ kernel.rnDerivAux κ ν p.1 p.2) := + measurable_density _ ν MeasurableSet.univ + +lemma measurable_rnDerivAux_right (κ ν : kernel α γ) (a : α) : + Measurable (fun x : γ ↦ kernel.rnDerivAux κ ν a x) := by + change Measurable ((fun p : α × γ ↦ kernel.rnDerivAux κ ν p.1 p.2) ∘ (fun x ↦ (a, x))) + exact (measurable_rnDerivAux _ _).comp measurable_prod_mk_left + +lemma set_lintegral_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) {s : Set γ} (hs : MeasurableSet s) : + ∫⁻ x in s, ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) ∂(κ + ν) a = κ a s := by + have h_le : κ ≤ κ + ν := le_add_of_nonneg_right bot_le + simp_rw [kernel.rnDerivAux] + rw [set_lintegral_density (kernel.fst_map_prod_le_of_le h_le) _ MeasurableSet.univ hs, + kernel.map_apply' _ _ _ (hs.prod MeasurableSet.univ)] + congr with x + simp + +lemma withDensity_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : + kernel.withDensity (κ + ν) + (fun a x ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x)) = κ := by + ext a s hs + rw [kernel.withDensity_apply'] + swap; exact (measurable_rnDerivAux _ _).ennreal_ofReal + have : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl + simp_rw [this] + exact set_lintegral_rnDerivAux κ ν a hs + +lemma withDensity_one_sub_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : + kernel.withDensity (κ + ν) + (fun a x ↦ Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x)) = ν := by + have h_le : κ ≤ κ + ν := le_add_of_nonneg_right bot_le + suffices kernel.withDensity (κ + ν) + (fun a x ↦ Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x)) + + kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x)) + = κ + ν by + ext a s + have h : (kernel.withDensity (κ + ν) + (fun a x ↦ Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x)) + + kernel.withDensity (κ + ν) + (fun a x ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x))) a s + = κ a s + ν a s := by + rw [this] + simp + simp only [kernel.coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add] + at h + rwa [withDensity_rnDerivAux, add_comm, ENNReal.add_right_inj (measure_ne_top _ _)] at h + have h_nonneg : ∀ a x, 0 ≤ kernel.rnDerivAux κ (κ + ν) a x := + fun _ _ ↦ density_nonneg (kernel.fst_map_prod_le_of_le h_le) _ _ _ + have : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl + simp_rw [this, ENNReal.ofReal_sub _ (h_nonneg _ _), ENNReal.ofReal_one] + rw [kernel.withDensity_sub_add] + · rw [kernel.withDensity_one'] + · exact measurable_const + · exact (measurable_rnDerivAux _ _).ennreal_ofReal + · refine fun a ↦ ne_of_lt ?_ + calc ∫⁻ x, ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) ∂(κ + ν) a + ≤ ∫⁻ _, 1 ∂(κ + ν) a := by + refine lintegral_mono (fun x ↦ ?_) + simp [rnDerivAux_le_one (le_add_of_nonneg_right bot_le)] + _ < ⊤ := by rw [MeasureTheory.lintegral_const, one_mul]; exact measure_lt_top _ _ + · refine fun a ↦ ae_of_all _ (fun x ↦ ?_) + simp only [ENNReal.ofReal_le_one] + exact density_le_one (kernel.fst_map_prod_le_of_le h_le) _ _ _ + +/-- A set of points in `α × γ` related to the absolute continuity / mutual singularity of +`κ` and `ν`. -/ +def mutuallySingularSet (κ ν : kernel α γ) : Set (α × γ) := + {p | kernel.rnDerivAux κ (κ + ν) p.1 p.2 = 1} + +/-- A set of points in `α × γ` related to the absolute continuity / mutual singularity of +`κ` and `ν`. That is, +* `kernel.withDensity ν (kernel.rnDeriv κ ν) a (kernel.mutuallySingularSetSlice κ ν a) = 0`, +* `kernel.singularPart κ ν a (kernel.mutuallySingularSetSlice κ ν a)ᶜ = 0`. + -/ +def mutuallySingularSetSlice (κ ν : kernel α γ) (a : α) : Set γ := + {x | kernel.rnDerivAux κ (κ + ν) a x = 1} + +lemma measurableSet_mutuallySingularSet (κ ν : kernel α γ) : + MeasurableSet (kernel.mutuallySingularSet κ ν) := + measurable_rnDerivAux κ (κ + ν) (measurableSet_singleton 1) + +lemma measurableSet_mutuallySingularSetSlice (κ ν : kernel α γ) (a : α) : + MeasurableSet (kernel.mutuallySingularSetSlice κ ν a) := + measurable_prod_mk_left (measurableSet_mutuallySingularSet κ ν) + +lemma measure_mutuallySingularSetSlice (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) : + ν a (kernel.mutuallySingularSetSlice κ ν a) = 0 := by + have h_coe : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl + suffices kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal + (1 - kernel.rnDerivAux κ (κ + ν) a x)) a {x | kernel.rnDerivAux κ (κ + ν) a x = 1} = 0 by + rwa [withDensity_one_sub_rnDerivAux κ ν] at this + simp_rw [h_coe] + rw [kernel.withDensity_apply', lintegral_eq_zero_iff, EventuallyEq, ae_restrict_iff] + rotate_left + · exact (measurable_const.sub + ((measurable_rnDerivAux _ _).comp measurable_prod_mk_left)).ennreal_ofReal + (measurableSet_singleton _) + · exact (measurable_const.sub + ((measurable_rnDerivAux _ _).comp measurable_prod_mk_left)).ennreal_ofReal + · exact (measurable_const.sub (measurable_rnDerivAux _ _)).ennreal_ofReal + refine ae_of_all _ (fun x hx ↦ ?_) + simp only [mem_setOf_eq] at hx + simp [hx] + +/-- Radon-Nikodym derivative of the kernel `κ` with respect to the kernel `ν`. -/ +noncomputable +def rnDeriv (κ ν : kernel α γ) (a : α) (x : γ) : ℝ≥0∞ := + ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) + / ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) + +lemma rnDeriv_def (κ ν : kernel α γ) : + kernel.rnDeriv κ ν = fun a x ↦ ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) + / ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x) := rfl + +lemma measurable_rnDeriv (κ ν : kernel α γ) : + Measurable (fun p : α × γ ↦ kernel.rnDeriv κ ν p.1 p.2) := + (measurable_rnDerivAux κ _).ennreal_ofReal.div + (measurable_const.sub (measurable_rnDerivAux κ _)).ennreal_ofReal + +lemma measurable_rnDeriv_right (κ ν : kernel α γ) (a : α) : + Measurable (fun x : γ ↦ kernel.rnDeriv κ ν a x) := by + change Measurable ((fun p : α × γ ↦ kernel.rnDeriv κ ν p.1 p.2) ∘ (fun x ↦ (a, x))) + exact (measurable_rnDeriv _ _).comp measurable_prod_mk_left + +lemma rnDeriv_eq_top_iff (κ ν : kernel α γ) (a : α) (x : γ) : + kernel.rnDeriv κ ν a x = ∞ ↔ (a, x) ∈ kernel.mutuallySingularSet κ ν := by + simp only [kernel.rnDeriv, ENNReal.div_eq_top, ne_eq, ENNReal.ofReal_eq_zero, not_le, + tsub_le_iff_right, zero_add, ENNReal.ofReal_ne_top, not_false_eq_true, and_true, or_false, + kernel.mutuallySingularSet, mem_setOf_eq] + exact ⟨fun h ↦ le_antisymm (rnDerivAux_le_one (le_add_of_nonneg_right bot_le)) h.2, + fun h ↦ by simp [h]⟩ + +/-- Singular part of the kernel `κ` with respect to the kernel `ν`. -/ +noncomputable +def singularPart (κ ν : kernel α γ) [IsSFiniteKernel κ] [IsSFiniteKernel ν] : + kernel α γ := + kernel.withDensity (κ + ν) (fun a x ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x) + - Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x) * kernel.rnDeriv κ ν a x) + +lemma measurable_singularPart_fun (κ ν : kernel α γ) : + Measurable (fun p : α × γ ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) p.1 p.2) + - Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) p.1 p.2) * kernel.rnDeriv κ ν p.1 p.2) := by + refine (measurable_rnDerivAux _ _).ennreal_ofReal.sub + (Measurable.mul ?_ (measurable_rnDeriv _ _)) + exact (measurable_const.sub (measurable_rnDerivAux _ _)).ennreal_ofReal + +lemma measurable_singularPart_fun_right (κ ν : kernel α γ) (a : α) : + Measurable (fun x : γ ↦ Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x) + - Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x) * kernel.rnDeriv κ ν a x) := by + change Measurable ((Function.uncurry fun a b ↦ + ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a b) + - ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a b) * kernel.rnDeriv κ ν a b) + ∘ (fun b ↦ (a, b))) + exact (measurable_singularPart_fun κ ν).comp measurable_prod_mk_left + +lemma singularPart_compl_mutuallySingularSetSlice (κ ν : kernel α γ) [IsSFiniteKernel κ] + [IsSFiniteKernel ν] (a : α) : + kernel.singularPart κ ν a (kernel.mutuallySingularSetSlice κ ν a)ᶜ = 0 := by + have h_coe : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl + rw [kernel.singularPart, kernel.withDensity_apply', lintegral_eq_zero_iff, EventuallyEq, + ae_restrict_iff] + all_goals simp_rw [h_coe] + rotate_left + · exact measurableSet_preimage (measurable_singularPart_fun_right κ ν a) + (measurableSet_singleton _) + · exact measurable_singularPart_fun_right κ ν a + · exact measurable_singularPart_fun κ ν + refine ae_of_all _ (fun x hx ↦ ?_) + simp only [mem_compl_iff, mem_setOf_eq] at hx + simp_rw [kernel.rnDeriv] + rw [← ENNReal.ofReal_div_of_pos, div_eq_inv_mul, ← ENNReal.ofReal_mul, ← mul_assoc, + mul_inv_cancel, one_mul, tsub_self] + · rfl + · rw [ne_eq, sub_eq_zero] + exact Ne.symm hx + · simp [rnDerivAux_le_one (le_add_of_nonneg_right bot_le)] + · simp [lt_of_le_of_ne (rnDerivAux_le_one (le_add_of_nonneg_right bot_le)) hx] + +lemma singularPart_subset_compl_mutuallySingularSetSlice (κ ν : kernel α γ) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) (s : Set γ) + (hs : s ⊆ (kernel.mutuallySingularSetSlice κ ν a)ᶜ) : + kernel.singularPart κ ν a s = 0 := by + exact measure_mono_null hs (singularPart_compl_mutuallySingularSetSlice κ ν a) + +lemma singularPart_subset_mutuallySingularSetSlice (κ ν : kernel α γ) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) {s : Set γ} (hsm : MeasurableSet s) + (hs : s ⊆ kernel.mutuallySingularSetSlice κ ν a) : + kernel.singularPart κ ν a s = κ a s := by + have hs' : ∀ x ∈ s, kernel.rnDerivAux κ (κ + ν) a x = 1 := fun _ hx ↦ hs hx + rw [kernel.singularPart, kernel.withDensity_apply'] + swap; · exact measurable_singularPart_fun κ ν + calc + ∫⁻ x in s, ↑(Real.toNNReal (kernel.rnDerivAux κ (κ + ν) a x)) - + ↑(Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x)) * kernel.rnDeriv κ ν a x + ∂(κ + ν) a + = ∫⁻ _ in s, 1 ∂(κ + ν) a := by + refine set_lintegral_congr_fun hsm (ae_of_all _ fun x hx ↦ ?_) + simp [hs' x hx] + _ = (κ + ν) a s := by simp + _ = κ a s := by + suffices ν a s = 0 by simp [this] + exact measure_mono_null hs (measure_mutuallySingularSetSlice κ ν a) + +lemma withDensity_rnDeriv_mutuallySingularSetSlice (κ ν : kernel α γ) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) : + kernel.withDensity ν (kernel.rnDeriv κ ν) a (kernel.mutuallySingularSetSlice κ ν a) = 0 := by + rw [kernel.withDensity_apply'] + swap; · exact measurable_rnDeriv κ ν + refine le_antisymm ?_ zero_le' + calc ∫⁻ x in kernel.mutuallySingularSetSlice κ ν a, kernel.rnDeriv κ ν a x ∂(ν a) + ≤ ∫⁻ _ in kernel.mutuallySingularSetSlice κ ν a, ∞ ∂(ν a) := + set_lintegral_mono (measurable_rnDeriv_right κ ν a) measurable_const (fun _ _ ↦ le_top) + _ = ∞ * ν a (kernel.mutuallySingularSetSlice κ ν a) := by simp + _ = 0 := by + simp only [mul_eq_zero, ENNReal.top_ne_zero, false_or] + exact measure_mutuallySingularSetSlice κ ν a + +lemma withDensity_rnDeriv_subset_mutuallySingularSetSlice (κ ν : kernel α γ) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) (s : Set γ) + (hs : s ⊆ kernel.mutuallySingularSetSlice κ ν a) : + kernel.withDensity ν (kernel.rnDeriv κ ν) a s = 0 := + measure_mono_null hs (withDensity_rnDeriv_mutuallySingularSetSlice κ ν a) + +lemma withDensity_rnDeriv_subset_compl_mutuallySingularSetSlice (κ ν : kernel α γ) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set γ} (hsm : MeasurableSet s) + (hs : s ⊆ (kernel.mutuallySingularSetSlice κ ν a)ᶜ) : + kernel.withDensity ν (kernel.rnDeriv κ ν) a s = κ a s := by + have h_coe : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl + have : kernel.withDensity ν (kernel.rnDeriv κ ν) + = kernel.withDensity (kernel.withDensity (κ + ν) + (fun a x ↦ Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x))) (kernel.rnDeriv κ ν) := by + rw [kernel.rnDeriv_def] + congr + exact (withDensity_one_sub_rnDerivAux κ ν).symm + rw [this, ← kernel.withDensity_mul, kernel.withDensity_apply'] + rotate_left + · exact ((measurable_const.sub (measurable_rnDerivAux _ _)).ennreal_ofReal.mul + (measurable_rnDeriv _ _)) + · exact (measurable_const.sub (measurable_rnDerivAux _ _)).real_toNNReal + · exact measurable_rnDeriv _ _ + simp_rw [kernel.rnDeriv] + have hs' : ∀ x ∈ s, kernel.rnDerivAux κ (κ + ν) a x < 1 := + fun x hx ↦ lt_of_le_of_ne (rnDerivAux_le_one (le_add_of_nonneg_right bot_le)) (hs hx) + calc + ∫⁻ x in s, ↑(Real.toNNReal (1 - kernel.rnDerivAux κ (κ + ν) a x)) * + (ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) / + ENNReal.ofReal (1 - kernel.rnDerivAux κ (κ + ν) a x)) ∂(κ + ν) a + _ = ∫⁻ x in s, ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) ∂(κ + ν) a := by + refine set_lintegral_congr_fun hsm (ae_of_all _ fun x hx ↦ ?_) + rw [h_coe, ← ENNReal.ofReal_div_of_pos, div_eq_inv_mul, ← ENNReal.ofReal_mul, ← mul_assoc, + mul_inv_cancel, one_mul] + · rw [ne_eq, sub_eq_zero] + exact (hs' x hx).ne' + · simp [(hs' x hx).le] + · simp [hs' x hx] + _ = κ a s := set_lintegral_rnDerivAux κ ν a hsm + +lemma mutuallySingular_singularPart (κ ν : kernel α γ) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + kernel.singularPart κ ν a ⟂ₘ ν a := by + symm + exact ⟨mutuallySingularSetSlice κ ν a, measurableSet_mutuallySingularSetSlice κ ν a, + measure_mutuallySingularSetSlice κ ν a, singularPart_compl_mutuallySingularSetSlice κ ν a⟩ + +lemma rnDeriv_add_singularPart (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : + kernel.withDensity ν (kernel.rnDeriv κ ν) + kernel.singularPart κ ν = κ := by + ext a s hs + rw [← inter_union_diff s (mutuallySingularSetSlice κ ν a)] + simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, + OuterMeasure.coe_add] + have hm := measurableSet_mutuallySingularSetSlice κ ν a + simp only [measure_union (Disjoint.mono (inter_subset_right _ _) subset_rfl disjoint_sdiff_right) + (hs.diff hm)] + rw [singularPart_subset_mutuallySingularSetSlice _ _ _ (hs.inter hm) (inter_subset_right _ _), + singularPart_subset_compl_mutuallySingularSetSlice _ _ _ _ (diff_subset_iff.mpr (by simp)), + add_zero, withDensity_rnDeriv_subset_mutuallySingularSetSlice _ _ _ _ (inter_subset_right _ _), + zero_add, withDensity_rnDeriv_subset_compl_mutuallySingularSetSlice _ _ _ (hs.diff hm) + (diff_subset_iff.mpr (by simp)), add_comm] + +lemma singularPart_eq_zero_iff_apply_eq_zero (κ ν : kernel α γ) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) : + kernel.singularPart κ ν a = 0 + ↔ kernel.singularPart κ ν a (kernel.mutuallySingularSetSlice κ ν a) = 0 := by + rw [← Measure.measure_univ_eq_zero] + have : univ = (kernel.mutuallySingularSetSlice κ ν a) + ∪ (kernel.mutuallySingularSetSlice κ ν a)ᶜ := by simp + rw [this, measure_union disjoint_compl_right (measurableSet_mutuallySingularSetSlice κ ν a).compl, + singularPart_compl_mutuallySingularSetSlice, add_zero] + +lemma withDensity_rnDeriv_eq_zero_iff_apply_eq_zero (κ ν : kernel α γ) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) : + kernel.withDensity ν (kernel.rnDeriv κ ν) a = 0 + ↔ kernel.withDensity ν (kernel.rnDeriv κ ν) a + (kernel.mutuallySingularSetSlice κ ν a)ᶜ = 0 := by + rw [← Measure.measure_univ_eq_zero] + have : univ = (kernel.mutuallySingularSetSlice κ ν a) + ∪ (kernel.mutuallySingularSetSlice κ ν a)ᶜ := by simp + rw [this, measure_union disjoint_compl_right (measurableSet_mutuallySingularSetSlice κ ν a).compl, + withDensity_rnDeriv_mutuallySingularSetSlice, zero_add] + +lemma singularPart_eq_zero_iff_absolutelyContinuous (κ ν : kernel α γ) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + kernel.singularPart κ ν a = 0 ↔ κ a ≪ ν a := by + conv_rhs => rw [← kernel.rnDeriv_add_singularPart κ ν] + simp only [kernel.coeFn_add, Pi.add_apply] + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · rw [h, add_zero] + exact kernel.withDensity_absolutelyContinuous _ _ + rw [singularPart_eq_zero_iff_apply_eq_zero] + specialize h (measure_mutuallySingularSetSlice κ ν a) + simpa only [Measure.add_toOuterMeasure, OuterMeasure.coe_add, Pi.add_apply, + withDensity_rnDeriv_mutuallySingularSetSlice κ ν, zero_add] using h + +lemma withDensity_rnDeriv_eq_zero_iff_mutuallySingular (κ ν : kernel α γ) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + kernel.withDensity ν (kernel.rnDeriv κ ν) a = 0 ↔ κ a ⟂ₘ ν a := by + conv_rhs => rw [← kernel.rnDeriv_add_singularPart κ ν] + simp only [kernel.coeFn_add, Pi.add_apply] + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · rw [h, zero_add] + exact kernel.mutuallySingular_singularPart _ _ _ + simp only [Measure.MutuallySingular.add_left_iff] at h + have h_ac := kernel.withDensity_absolutelyContinuous (κ := ν) (kernel.rnDeriv κ ν) a + have h_ms := h.1 + rw [← Measure.MutuallySingular.self_iff] + exact h_ms.mono_ac Measure.AbsolutelyContinuous.rfl h_ac + +lemma singularPart_eq_zero_iff_measure_eq_zero (κ ν : kernel α γ) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + kernel.singularPart κ ν a = 0 ↔ κ a (kernel.mutuallySingularSetSlice κ ν a) = 0 := by + have h_eq_add := kernel.rnDeriv_add_singularPart κ ν + simp_rw [kernel.ext_iff, Measure.ext_iff] at h_eq_add + specialize h_eq_add a (kernel.mutuallySingularSetSlice κ ν a) + (measurableSet_mutuallySingularSetSlice κ ν a) + simp only [kernel.coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add, + withDensity_rnDeriv_mutuallySingularSetSlice κ ν, zero_add] at h_eq_add + rw [← h_eq_add] + exact singularPart_eq_zero_iff_apply_eq_zero κ ν a + +lemma withDensity_rnDeriv_eq_zero_iff_measure_eq_zero (κ ν : kernel α γ) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + kernel.withDensity ν (kernel.rnDeriv κ ν) a = 0 + ↔ κ a (kernel.mutuallySingularSetSlice κ ν a)ᶜ = 0 := by + have h_eq_add := kernel.rnDeriv_add_singularPart κ ν + simp_rw [kernel.ext_iff, Measure.ext_iff] at h_eq_add + specialize h_eq_add a (kernel.mutuallySingularSetSlice κ ν a)ᶜ + (measurableSet_mutuallySingularSetSlice κ ν a).compl + simp only [kernel.coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add, + singularPart_compl_mutuallySingularSetSlice κ ν, add_zero] at h_eq_add + rw [← h_eq_add] + exact withDensity_rnDeriv_eq_zero_iff_apply_eq_zero κ ν a + +lemma measurableSet_absolutelyContinuous (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : + MeasurableSet {a | κ a ≪ ν a} := by + simp_rw [← singularPart_eq_zero_iff_absolutelyContinuous, + singularPart_eq_zero_iff_measure_eq_zero] + exact kernel.measurable_kernel_prod_mk_left (measurableSet_mutuallySingularSet κ ν) + (measurableSet_singleton 0) + +lemma measurableSet_mutuallySingular (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] : + MeasurableSet {a | κ a ⟂ₘ ν a} := by + simp_rw [← withDensity_rnDeriv_eq_zero_iff_mutuallySingular, + withDensity_rnDeriv_eq_zero_iff_measure_eq_zero] + exact kernel.measurable_kernel_prod_mk_left (measurableSet_mutuallySingularSet κ ν).compl + (measurableSet_singleton 0) + +end ProbabilityTheory.kernel diff --git a/docs/references.bib b/docs/references.bib index 2a887bb746cd6..4a1f97e06f8dc 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1785,8 +1785,19 @@ @Article{ KahnMaltsiniotis2008 url = {https://doi.org/10.1016/j.aim.2008.03.010} } +@Book{ kallenberg2017, + title = {Random Measures, Theory and Applications}, + author = {Kallenberg, Olav}, + series = {Probability Theory and Stochastic Modelling}, + volume = {77}, + year = {2017}, + publisher = {Springer}, + doi = {10.1007/978-3-319-41598-7}, + url = {https://doi.org/10.1007/978-3-319-41598-7} +} + @Book{ kallenberg2021, - author = {Olav Kallenberg}, + author = {Kallenberg, Olav}, title = {Foundations of modern probability}, series = {Probability Theory and Stochastic Modelling}, volume = {99}, From 0b0b32bef52fe9de7f13f35210ef704a80c8d9cd Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 10:10:43 +0100 Subject: [PATCH 072/129] fix --- Mathlib/Probability/Kernel/RadonNikodym.lean | 8 +------- 1 file changed, 1 insertion(+), 7 deletions(-) diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index 33bd2e45add85..621e5a393cb2c 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -120,16 +120,10 @@ lemma withDensity_one_sub_rnDerivAux (κ ν : kernel α γ) [IsFiniteKernel κ] fun _ _ ↦ density_nonneg (kernel.fst_map_prod_le_of_le h_le) _ _ _ have : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl simp_rw [this, ENNReal.ofReal_sub _ (h_nonneg _ _), ENNReal.ofReal_one] - rw [kernel.withDensity_sub_add] + rw [kernel.withDensity_sub_add_cancel] · rw [kernel.withDensity_one'] · exact measurable_const · exact (measurable_rnDerivAux _ _).ennreal_ofReal - · refine fun a ↦ ne_of_lt ?_ - calc ∫⁻ x, ENNReal.ofReal (kernel.rnDerivAux κ (κ + ν) a x) ∂(κ + ν) a - ≤ ∫⁻ _, 1 ∂(κ + ν) a := by - refine lintegral_mono (fun x ↦ ?_) - simp [rnDerivAux_le_one (le_add_of_nonneg_right bot_le)] - _ < ⊤ := by rw [MeasureTheory.lintegral_const, one_mul]; exact measure_lt_top _ _ · refine fun a ↦ ae_of_all _ (fun x ↦ ?_) simp only [ENNReal.ofReal_le_one] exact density_le_one (kernel.fst_map_prod_le_of_le h_le) _ _ _ From 715bd03716660479d349a497fcd4af9b8df5b77d Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 10:15:06 +0100 Subject: [PATCH 073/129] delete duplicate --- Mathlib/Order/LiminfLimsup.lean | 10 ---------- 1 file changed, 10 deletions(-) diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 7ff15e1763630..0e5c8eb0a21ac 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -242,16 +242,6 @@ theorem IsBoundedUnder.isCoboundedUnder_le {u : γ → α} {l : Filter γ} [Preo (h : l.IsBoundedUnder (· ≥ ·) u) : l.IsCoboundedUnder (· ≤ ·) u := h.isCoboundedUnder_flip -lemma isCoboundedUnder_le_of_eventually_le [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} {x : α} - (hf : ∀ᶠ i in l, x ≤ f i) : - IsCoboundedUnder (· ≤ ·) l f := - IsBoundedUnder.isCoboundedUnder_le ⟨x, hf⟩ - -lemma isCoboundedUnder_le_of_le [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} {x : α} - (hf : ∀ i, x ≤ f i) : - IsCoboundedUnder (· ≤ ·) l f := - isCoboundedUnder_le_of_eventually_le l (eventually_of_forall hf) - theorem IsBoundedUnder.isCoboundedUnder_ge {u : γ → α} {l : Filter γ} [Preorder α] [NeBot l] (h : l.IsBoundedUnder (· ≤ ·) u) : l.IsCoboundedUnder (· ≥ ·) u := h.isCoboundedUnder_flip From cb90c8d9616c5d5c0a41ed51e0f79a4a889830d3 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 25 Feb 2024 10:44:00 +0100 Subject: [PATCH 074/129] clean MeasurableStieltjes --- .../Kernel/Disintegration/CDFKernel.lean | 1 + .../Disintegration/MeasurableStieltjes.lean | 23 +++++++++---------- 2 files changed, 12 insertions(+), 12 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean index 135b4e42f6f37..13ab6dea1d14b 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean @@ -5,6 +5,7 @@ Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes import Mathlib.Probability.Kernel.MeasureCompProd +import Mathlib.Probability.Kernel.Disintegration.AuxLemmas /-! diff --git a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean index 3365ad471af36..b0f38ce39beeb 100644 --- a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean +++ b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean @@ -9,15 +9,15 @@ import Mathlib.MeasureTheory.Measure.GiryMonad import Mathlib.MeasureTheory.Measure.Stieltjes import Mathlib.Analysis.Normed.Order.Lattice import Mathlib.MeasureTheory.Function.StronglyMeasurable.Basic -import Mathlib.Probability.Kernel.Disintegration.AuxLemmas /-! -# Cumulative distributions functions of Markov kernels +# Measurable parametric Stieltjes functions We provide tools to build a measurable function `α → StieltjesFunction` with limits 0 at -∞ and 1 at -+∞ for all `a : α` from a measurable function `f : α → ℚ → ℝ`. The reason for going through `ℚ` -instead of defining directly a Stieltjes function is that since `ℚ` is countable, building a -measurable function is much easier. ++∞ for all `a : α` from a measurable function `f : α → ℚ → ℝ`. These measurable parametric Stieltjes +functions are cumulative distribution functions (CDF) of transition kernels. +The reason for going through `ℚ` instead of defining directly a Stieltjes function is that since +`ℚ` is countable, building a measurable function is much easier. This construction will be possible if `f a : ℚ → ℝ` satisfies a package of properties for all `a`: monotonicity, limits at +-∞ at a continuity property. We define `IsRatStieltjesPoint f a` to state @@ -26,8 +26,8 @@ that this is the case at `a` and define the property `IsRatCDF f` that `f` is me The function `α → StieltjesFunction` obtained by extending `f` by continuity from the right is then called `IsRatCDF.stieltjesFunction`. -In applications, we will only have `IsRatStieltjesPoint f a` almost surely with respect to some -measure. We thus define +In applications, we will often only have `IsRatStieltjesPoint f a` almost surely with respect to +some measure. In order to turn that almost everywhere property into an everywhere property we define `toRatCDF (f : α → ℚ → ℝ) := fun a q ↦ if IsRatStieltjesPoint f a then f a q else defaultRatCDF q`, which satisfies the property `IsRatCDF (toRatCDF f)`. @@ -41,7 +41,6 @@ Finally, we define `stieltjesOfMeasurableRat`, composition of `toRatCDF` and -/ - open MeasureTheory Set Filter TopologicalSpace open scoped NNReal ENNReal MeasureTheory Topology @@ -140,8 +139,8 @@ end IsRatCDF section DefaultRatCDF /-- A function with the property `IsRatCDF`. -Used in a `piecewise` construction to convert a function which only satisfies the properties -defining `IsRatCDF` almost everywhere into a true `IsRatCDF`. -/ +Used in a piecewise construction to convert a function which only satisfies the properties +defining `IsRatCDF` on some set into a true `IsRatCDF`. -/ def defaultRatCDF (q : ℚ) := if q < 0 then (0 : ℝ) else 1 lemma monotone_defaultRatCDF : Monotone defaultRatCDF := by @@ -171,7 +170,7 @@ lemma tendsto_defaultRatCDF_atBot : Tendsto defaultRatCDF atBot (𝓝 0) := by refine ⟨-1, fun q hq => (if_pos (hq.trans_lt ?_)).symm⟩ linarith -lemma inf_gt_rat_defaultRatCDF (t : ℚ) : +lemma iInf_rat_gt_defaultRatCDF (t : ℚ) : ⨅ r : Ioi t, defaultRatCDF r = defaultRatCDF t := by simp only [defaultRatCDF] have h_bdd : BddBelow (range fun r : ↥(Ioi t) ↦ ite ((r : ℚ) < 0) (0 : ℝ) 1) := by @@ -202,7 +201,7 @@ lemma isRatStieltjesPoint_defaultRatCDF (a : α) : mono := monotone_defaultRatCDF tendsto_atTop_one := tendsto_defaultRatCDF_atTop tendsto_atBot_zero := tendsto_defaultRatCDF_atBot - iInf_rat_gt_eq := inf_gt_rat_defaultRatCDF + iInf_rat_gt_eq := iInf_rat_gt_defaultRatCDF lemma IsRatCDF_defaultRatCDF (α : Type*) [MeasurableSpace α] : IsRatCDF (fun (_ : α) (q : ℚ) ↦ defaultRatCDF q) where From 768c1b91f5e731b98715790cff7310afc3c07902 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 26 Feb 2024 07:32:06 +0100 Subject: [PATCH 075/129] minor --- Mathlib/MeasureTheory/Measure/Stieltjes.lean | 6 ++++ .../Kernel/Disintegration/Basic.lean | 8 +++-- .../Kernel/Disintegration/CDFKernel.lean | 31 ++++++++----------- .../Kernel/Disintegration/CondCdf.lean | 1 + .../Kernel/Disintegration/KernelCDFReal.lean | 1 + 5 files changed, 26 insertions(+), 21 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Stieltjes.lean b/Mathlib/MeasureTheory/Measure/Stieltjes.lean index cf5e3eaff5cd0..815941839b111 100644 --- a/Mathlib/MeasureTheory/Measure/Stieltjes.lean +++ b/Mathlib/MeasureTheory/Measure/Stieltjes.lean @@ -446,6 +446,12 @@ theorem measure_univ {l u : ℝ} (hfl : Tendsto f atBot (𝓝 l)) (hfu : Tendsto exact ENNReal.tendsto_ofReal (Tendsto.sub_const hfu _) #align stieltjes_function.measure_univ StieltjesFunction.measure_univ +lemma isFiniteMeasure {l u : ℝ} (hfl : Tendsto f atBot (𝓝 l)) (hfu : Tendsto f atTop (𝓝 u)) : + IsFiniteMeasure f.measure := ⟨by simp [f.measure_univ hfl hfu]⟩ + +lemma isProbabilityMeasure (hf_bot : Tendsto f atBot (𝓝 0)) (hf_top : Tendsto f atTop (𝓝 1)) : + IsProbabilityMeasure f.measure := ⟨by simp [f.measure_univ hf_bot hf_top]⟩ + instance instIsLocallyFiniteMeasure : IsLocallyFiniteMeasure f.measure := ⟨fun x => ⟨Ioo (x - 1) (x + 1), Ioo_mem_nhds (by linarith) (by linarith), by simp⟩⟩ #align stieltjes_function.measure.measure_theory.is_locally_finite_measure StieltjesFunction.instIsLocallyFiniteMeasure diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 0be3c7ad2367a..557e71d3ee84f 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ +import Mathlib.Probability.Kernel.MeasureCompProd import Mathlib.Probability.Kernel.Disintegration.CondCdf import Mathlib.Probability.Kernel.Disintegration.KernelCDFReal @@ -10,8 +11,9 @@ import Mathlib.Probability.Kernel.Disintegration.KernelCDFReal # Disintegration of kernels and measures Let `κ : kernel α (β × Ω)` be a finite kernel, where `Ω` is a standard Borel space. Then if `α` is -countable or `β` is a standard Borel space, there exists a `kernel (α × β) Ω`, called conditional -kernel and denoted by `kernel.condKernel κ` such that `κ = kernel.fst κ ⊗ₖ kernel.condKernel κ`. +countable or `β` has a countably generated σ-algebra (for example if it is standard Borel), +there exists a `kernel (α × β) Ω`, called conditional kernel and denoted by `kernel.condKernel κ` +such that `κ = kernel.fst κ ⊗ₖ kernel.condKernel κ`. We also define a conditional kernel for a measure `ρ : Measure (β × Ω)`, where `Ω` is a standard Borel space. This is a `kernel β Ω` denoted by `ρ.condKernel` such that `ρ = ρ.fst ⊗ₘ ρ.condKernel`. @@ -36,7 +38,7 @@ The first step (building the measurable function on `ℚ`) is done differently d * If `α` is not countable, we can't proceed separately for each `a : α` and have to build a function `f : α × β → ℚ → ℝ` which is measurable on the product. We are able to do so if `β` has a countably generated σ-algebra (this is the case in particular for standard Borel spaces). - See the file `KernelCDFBorel.lean`. + See the files `Density.lean` and `KernelCDFReal.lean`. We define a class `CountableOrCountablyGenerated α β` which encodes the property `(Countable α ∧ MeasurableSingletonClass α) ∨ CountablyGenerated β`. The conditional kernel is diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean index 13ab6dea1d14b..f88b5a37ae717 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean @@ -3,16 +3,15 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ +import Mathlib.MeasureTheory.Integral.SetIntegral +import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes -import Mathlib.Probability.Kernel.MeasureCompProd -import Mathlib.Probability.Kernel.Disintegration.AuxLemmas /-! -/ - open MeasureTheory Set Filter TopologicalSpace open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory @@ -89,8 +88,7 @@ lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKern · refine le_of_forall_lt_rat_imp_le fun q hq ↦ h q ?_ exact mod_cast hq · exact fun _ ↦ measurableSet_Iic - · refine Monotone.directed_ge fun r r' hrr' ↦ ?_ - refine Iic_subset_Iic.mpr ?_ + · refine Monotone.directed_ge fun r r' hrr' ↦ Iic_subset_Iic.mpr ?_ exact mod_cast hrr' · obtain ⟨q, hq⟩ := exists_rat_gt x exact ⟨⟨q, hq⟩, measure_ne_top _ _⟩ @@ -206,11 +204,6 @@ section kernel variable {_ : MeasurableSpace β} {f : α × β → StieltjesFunction} {μ : kernel α (β × ℝ)} {ν : kernel α β} {hf : IsKernelCDF f μ ν} -lemma isProbabilityMeasure_stieltjesFunction {f : StieltjesFunction} - (hf_bot : Tendsto f atBot (𝓝 0)) (hf_top : Tendsto f atTop (𝓝 1)) : - IsProbabilityMeasure f.measure := - ⟨by simp [StieltjesFunction.measure_univ _ hf_bot hf_top]⟩ - lemma StieltjesFunction.measurable_measure {f : α → StieltjesFunction} (hf : ∀ q, Measurable fun a ↦ f a q) (hf_bot : ∀ a, Tendsto (f a) atBot (𝓝 0)) @@ -218,10 +211,8 @@ lemma StieltjesFunction.measurable_measure {f : α → StieltjesFunction} Measurable fun a ↦ (f a).measure := by rw [Measure.measurable_measure] have : ∀ a, IsProbabilityMeasure (f a).measure := - fun a ↦ isProbabilityMeasure_stieltjesFunction (hf_bot _) (hf_top _) - refine fun s hs ↦ ?_ - -- Porting note: supplied `C` - refine MeasurableSpace.induction_on_inter + fun a ↦ (f a).isProbabilityMeasure (hf_bot a) (hf_top a) + refine fun s hs ↦ MeasurableSpace.induction_on_inter (C := fun s ↦ Measurable fun b ↦ StieltjesFunction.measure (f b) s) (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic ?_ ?_ ?_ ?_ hs · simp only [measure_empty, measurable_const] @@ -252,8 +243,7 @@ def IsKernelCDF.toKernel (f : α × β → StieltjesFunction) (hf : IsKernelCDF lemma IsKernelCDF.toKernel_apply (p : α × β) : hf.toKernel f p = (f p).measure := rfl instance instIsMarkovKernel_toKernel : IsMarkovKernel (hf.toKernel f) := - ⟨fun _ ↦ isProbabilityMeasure_stieltjesFunction - (hf.tendsto_atBot_zero _) (hf.tendsto_atTop_one _)⟩ + ⟨fun _ ↦ (f _).isProbabilityMeasure (hf.tendsto_atBot_zero _) (hf.tendsto_atTop_one _)⟩ lemma IsKernelCDF.toKernel_Iic (p : α × β) (x : ℝ) : hf.toKernel f p (Iic x) = ENNReal.ofReal (f p x) := by @@ -277,7 +267,12 @@ lemma set_lintegral_toKernel_Iic [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) lemma set_lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) {s : Set β} (hs : MeasurableSet s) : ∫⁻ t in s, hf.toKernel f (a, t) univ ∂(ν a) = μ a (s ×ˢ univ) := by - rw [← Real.iUnion_Iic_rat, prod_iUnion] + have : ⋃ r : ℚ, Iic (r : ℝ) = univ := by + ext1 x + simp only [mem_iUnion, mem_Iic, mem_univ, iff_true_iff] + obtain ⟨r, hr⟩ := exists_rat_gt x + exact ⟨r, hr.le⟩ + rw [← this, prod_iUnion] have h_dir : Directed (fun x y ↦ x ⊆ y) fun q : ℚ ↦ Iic (q : ℝ) := by refine Monotone.directed_le fun r r' hrr' ↦ Iic_subset_Iic.mpr ?_ exact mod_cast hrr' @@ -340,7 +335,7 @@ lemma set_lintegral_toKernel_prod [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) lemma lintegral_toKernel_mem [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) {s : Set (β × ℝ)} (hs : MeasurableSet s) : ∫⁻ x, hf.toKernel f (a, x) {y | (x, y) ∈ s} ∂(ν a) = μ a s := by - -- `set_lintegral_cdfKernel_prod` gives the result for sets of the form `t₁ × t₂`. These + -- `set_lintegral_toKernel_prod` gives the result for sets of the form `t₁ × t₂`. These -- sets form a π-system that generates the product σ-algebra, hence we can get the same equality -- for any measurable set `s`. apply MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod _ _ _ _ hs diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index 2a1a23ba3ede8..15b270f12e11d 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -5,6 +5,7 @@ Authors: Rémy Degenne -/ import Mathlib.MeasureTheory.Decomposition.RadonNikodym import Mathlib.Probability.Kernel.Disintegration.CDFKernel +import Mathlib.Probability.Kernel.Disintegration.AuxLemmas #align_import probability.kernel.cond_cdf from "leanprover-community/mathlib"@"3b88f4005dc2e28d42f974cc1ce838f0dafb39b8" diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean index ae996ad99b16c..8da6e4431704b 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean @@ -5,6 +5,7 @@ Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.Disintegration.CDFKernel import Mathlib.Probability.Kernel.Disintegration.Density +import Mathlib.Probability.Kernel.Disintegration.AuxLemmas /-! # Kernel CDF From ded4b18ee0a4d89d1d04c2216d9ab67c028c6d7a Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 26 Feb 2024 11:46:40 +0100 Subject: [PATCH 076/129] start docstring for CDFKernel --- .../Probability/Kernel/Disintegration/CDFKernel.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean index f88b5a37ae717..5df6295278a92 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean @@ -8,7 +8,18 @@ import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes /-! +# TODO +## Main definitions + +* `ProbabilityTheory.IsRatKernelCDF`: +* `ProbabilityTheory.IsKernelCDF`: +* `ProbabilityTheory.IsKernelCDF.toKernel`: + +## Main statements + +* `ProbabilityTheory.isKernelCDF_stieltjesOfMeasurableRat`: +* `ProbabilityTheory.kernel.eq_compProd_toKernel`: -/ From 18213af0cd4d087eb740c964a8d0494695887c34 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 26 Feb 2024 11:50:05 +0100 Subject: [PATCH 077/129] fix merge --- .../Probability/Kernel/Disintegration/MeasurableStieltjes.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean index 5195f4b78ae89..ce2bcefe52d15 100644 --- a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean +++ b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean @@ -399,7 +399,10 @@ lemma IsMeasurableRatCDF.measure_stieltjesFunction_Iic (a : α) (x : ℝ) : exact (hf.stieltjesFunction a).measure_Iic (tendsto_stieltjesFunction_atBot hf a) _ lemma IsMeasurableRatCDF.measure_stieltjesFunction_univ (a : α) : + (hf.stieltjesFunction a).measure univ = 1 := by + rw [← ENNReal.ofReal_one, ← sub_zero (1 : ℝ)] exact StieltjesFunction.measure_univ _ (tendsto_stieltjesFunction_atBot hf a) + (tendsto_stieltjesFunction_atTop hf a) instance IsMeasurableRatCDF.instIsProbabilityMeasure_stieltjesFunction (a : α) : IsProbabilityMeasure (hf.stieltjesFunction a).measure := From b05613ecbd4716154c157e06aed70ef54edc1a09 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 26 Feb 2024 15:59:31 +0100 Subject: [PATCH 078/129] rename files --- Mathlib.lean | 4 ++-- .../Probability/Kernel/Disintegration/AuxLemmas.lean | 2 +- Mathlib/Probability/Kernel/Disintegration/Basic.lean | 8 ++++---- Mathlib/Probability/Kernel/Disintegration/CondCdf.lean | 4 ++-- .../{KernelCDFReal.lean => CondKernelCdf.lean} | 4 ++-- .../Disintegration/{CDFKernel.lean => KernelCdf.lean} | 10 +++++----- 6 files changed, 16 insertions(+), 16 deletions(-) rename Mathlib/Probability/Kernel/Disintegration/{KernelCDFReal.lean => CondKernelCdf.lean} (98%) rename Mathlib/Probability/Kernel/Disintegration/{CDFKernel.lean => KernelCdf.lean} (99%) diff --git a/Mathlib.lean b/Mathlib.lean index 6401b6a75b643..b537073f27a3b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3138,11 +3138,11 @@ import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Kernel.CondDistrib import Mathlib.Probability.Kernel.Condexp import Mathlib.Probability.Kernel.Disintegration.Basic -import Mathlib.Probability.Kernel.Disintegration.CDFKernel import Mathlib.Probability.Kernel.Disintegration.CondCdf +import Mathlib.Probability.Kernel.Disintegration.CondKernelCdf import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.Disintegration.Integral -import Mathlib.Probability.Kernel.Disintegration.KernelCDFReal +import Mathlib.Probability.Kernel.Disintegration.KernelCdf import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes import Mathlib.Probability.Kernel.Disintegration.Unique import Mathlib.Probability.Kernel.IntegralCompProd diff --git a/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean b/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean index 935190b5106ee..065de63bc988f 100644 --- a/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean +++ b/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean @@ -20,7 +20,7 @@ theorem Real.iInter_Iic_rat : ⋂ r : ℚ, Iic (r : ℝ) = ∅ := by exact exists_rat_lt x #align real.Inter_Iic_rat Real.iInter_Iic_rat -lemma Measure.iInf_Iic_gt_prod {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] +lemma Measure.iInf_rat_gt_prod_Iic {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] {s : Set α} (hs : MeasurableSet s) (t : ℚ) : ⨅ r : { r' : ℚ // t < r' }, ρ (s ×ˢ Iic (r : ℝ)) = ρ (s ×ˢ Iic (t : ℝ)) := by rw [← measure_iInter_eq_iInf] diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 557e71d3ee84f..0a585325dd6a4 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -5,7 +5,7 @@ Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.MeasureCompProd import Mathlib.Probability.Kernel.Disintegration.CondCdf -import Mathlib.Probability.Kernel.Disintegration.KernelCDFReal +import Mathlib.Probability.Kernel.Disintegration.CondKernelCdf /-! # Disintegration of kernels and measures @@ -27,18 +27,18 @@ For `κ : kernel α (β × ℝ)`, the construction of the conditional kernel pro * Extend that function to `(α × β) → StieltjesFunction`. See the file `MeasurableStieltjes.lean`. * Finally obtain from the measurable Stieltjes function a measure on `ℝ` for each element of `α × β` in a measurable way: we have obtained a `kernel (α × β) ℝ`. - See the file `CDFKernel.lean` for that step. + See the file `KernelCdf.lean` for that step. The first step (building the measurable function on `ℚ`) is done differently depending on whether `α` is countable or not. * If `α` is countable, we can provide for each `a : α` a function `f : β → ℚ → ℝ` and proceed as above to obtain a `kernel β ℝ`. Since `α` is countable, measurability is not an issue and we can put those together into a `kernel (α × β) ℝ`. The construction of that `f` is done in - the `CondCDF.lean` file. + the `CondCdf.lean` file. * If `α` is not countable, we can't proceed separately for each `a : α` and have to build a function `f : α × β → ℚ → ℝ` which is measurable on the product. We are able to do so if `β` has a countably generated σ-algebra (this is the case in particular for standard Borel spaces). - See the files `Density.lean` and `KernelCDFReal.lean`. + See the files `Density.lean` and `CondKernelCdf.lean`. We define a class `CountableOrCountablyGenerated α β` which encodes the property `(Countable α ∧ MeasurableSingletonClass α) ∨ CountablyGenerated β`. The conditional kernel is diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index 15b270f12e11d..28dcc6b63e43d 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.MeasureTheory.Decomposition.RadonNikodym -import Mathlib.Probability.Kernel.Disintegration.CDFKernel +import Mathlib.Probability.Kernel.Disintegration.KernelCdf import Mathlib.Probability.Kernel.Disintegration.AuxLemmas #align_import probability.kernel.cond_cdf from "leanprover-community/mathlib"@"3b88f4005dc2e28d42f974cc1ce838f0dafb39b8" @@ -85,7 +85,7 @@ theorem IsFiniteMeasure.IicSnd {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] ( theorem iInf_IicSnd_gt (t : ℚ) {s : Set α} (hs : MeasurableSet s) [IsFiniteMeasure ρ] : ⨅ r : { r' : ℚ // t < r' }, ρ.IicSnd r s = ρ.IicSnd t s := by - simp_rw [ρ.IicSnd_apply _ hs, Measure.iInf_Iic_gt_prod hs] + simp_rw [ρ.IicSnd_apply _ hs, Measure.iInf_rat_gt_prod_Iic hs] #align measure_theory.measure.infi_Iic_snd_gt MeasureTheory.Measure.iInf_IicSnd_gt theorem tendsto_IicSnd_atTop {s : Set α} (hs : MeasurableSet s) : diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean b/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean similarity index 98% rename from Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean rename to Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean index 8da6e4431704b..b0dfc3ff6e524 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCDFReal.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.Disintegration.CDFKernel +import Mathlib.Probability.Kernel.Disintegration.KernelCdf import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.Disintegration.AuxLemmas @@ -177,7 +177,7 @@ lemma set_integral_iInf_rat_gt_densityIic (hκν : kernel.fst κ ≤ ν) [IsFini calc ∫ t in A, ⨅ r : Ioi q, kernel.densityIic κ ν a t r ∂(ν a) ≤ ⨅ r : Ioi q, (κ a (A ×ˢ Iic (r : ℝ))).toReal := le_ciInf h _ = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by - rw [← Measure.iInf_Iic_gt_prod hA q] + rw [← Measure.iInf_rat_gt_prod_Iic hA q] exact (ENNReal.toReal_iInf (fun r ↦ measure_ne_top _ _)).symm · rw [← set_integral_densityIic hκν a q hA] refine set_integral_mono ?_ ?_ ?_ diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCdf.lean similarity index 99% rename from Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean rename to Mathlib/Probability/Kernel/Disintegration/KernelCdf.lean index 5df6295278a92..841c5b5a5c4f7 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCdf.lean @@ -8,7 +8,7 @@ import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes /-! -# TODO +# Cumulative distribution function of a Markov kernel ## Main definitions @@ -210,7 +210,7 @@ lemma isKernelCDF_stieltjesOfMeasurableRat {f : α × β → ℚ → ℝ} (hf : end IsKernelCDF -section kernel +section ToKernel variable {_ : MeasurableSpace β} {f : α × β → StieltjesFunction} {μ : kernel α (β × ℝ)} {ν : kernel α β} {hf : IsKernelCDF f μ ν} @@ -220,10 +220,10 @@ lemma StieltjesFunction.measurable_measure {f : α → StieltjesFunction} (hf_bot : ∀ a, Tendsto (f a) atBot (𝓝 0)) (hf_top : ∀ a, Tendsto (f a) atTop (𝓝 1)) : Measurable fun a ↦ (f a).measure := by - rw [Measure.measurable_measure] + refine Measure.measurable_measure.mpr fun s hs ↦ ?_ have : ∀ a, IsProbabilityMeasure (f a).measure := fun a ↦ (f a).isProbabilityMeasure (hf_bot a) (hf_top a) - refine fun s hs ↦ MeasurableSpace.induction_on_inter + refine MeasurableSpace.induction_on_inter (C := fun s ↦ Measurable fun b ↦ StieltjesFunction.measure (f b) s) (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic ?_ ?_ ?_ ?_ hs · simp only [measure_empty, measurable_const] @@ -261,7 +261,7 @@ lemma IsKernelCDF.toKernel_Iic (p : α × β) (x : ℝ) : rw [IsKernelCDF.toKernel_apply p, (f p).measure_Iic (hf.tendsto_atBot_zero p)] simp -end kernel +end ToKernel section From 7870eb8fddda9d168dbbe2a4e8aa631fc14f0cb9 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 26 Feb 2024 16:36:52 +0100 Subject: [PATCH 079/129] change variable name --- .../Kernel/Disintegration/Density.lean | 384 +++++++++--------- 1 file changed, 191 insertions(+), 193 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 5b0accef399fa..424bd6de20a93 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -14,26 +14,26 @@ Let `κ : kernel α (γ × β)` and `ν : kernel α γ` be two finite kernels wi where `γ` has a countably generated σ-algebra (true in particular for standard Borel spaces). We build a function `f : α → γ → Set β → ℝ` jointly measurable in the first two arguments such that for all `a : α` and all measurable sets `s : Set β` and `A : Set γ`, -`∫ t in A, f a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal`. +`∫ x in A, f a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal`. If we were interested only in a fixed `a : α`, then we could use the Radon-Nikodym derivative to build the function `f`, as follows. ``` -def f (κ : kernel α (γ × β)) (ν : kernel a γ) (a : α) (t : γ) (s : Set β) : ℝ := - (((κ a).restrict (univ ×ˢ s)).fst.rnDeriv (ν a) t).toReal +def f (κ : kernel α (γ × β)) (ν : kernel a γ) (a : α) (x : γ) (s : Set β) : ℝ := + (((κ a).restrict (univ ×ˢ s)).fst.rnDeriv (ν a) x).toReal ``` -However, we can't turn those functions for each `a` into a measurable function of the pair `(a, t)`. +However, we can't turn those functions for each `a` into a measurable function of the pair `(a, x)`. In order to obtain measurability through countability, we use the fact that the measurable space `γ` is countably generated. For each `n : ℕ`, we define a finite partition of `γ`, such that those partitions are finer as `n` grows, and the σ-algebra generated by the union of all partitions is the σ-algebra of `γ`. -For `t : γ`, let `partitionSet n t` be the set in the partition such that `t ∈ partitionSet n t`. +For `x : γ`, let `partitionSet n x` be the set in the partition such that `x ∈ partitionSet n x`. For a given `n`, the function `densityProcess κ ν n : α → γ → Set β → ℝ` defined by -`fun a t s ↦ (κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t)).toReal` has the desired -property that `∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal` for all `A` -in the σ-algebra generated by the partition at scale `n` and is measurable in `(a, t)`. +`fun a x s ↦ (κ a (partitionSet n x ×ˢ s) / ν a (partitionSet n x)).toReal` has the desired +property that `∫ x in A, densityProcess κ ν n a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal` for all `A` +in the σ-algebra generated by the partition at scale `n` and is measurable in `(a, x)`. Let `partitionFiltration γ` be the filtration of those σ-algebras for all `n : ℕ`. The functions `densityProcess κ ν n` described here are a bounded `ν`-martingale for the filtration @@ -51,7 +51,7 @@ We have obtained the desired density function. ## Main statements * `ProbabilityTheory.kernel.set_integral_density`: for all measurable sets `A : Set γ` and - `s : Set β`, `∫ t in A, kernel.density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal`. + `s : Set β`, `∫ x in A, kernel.density κ ν a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal`. * `ProbabilityTheory.kernel.measurable_density`: the function `p : α × γ ↦ kernel.density κ ν p.1 p.2 s` is measurable. @@ -77,9 +77,9 @@ section DensityProcess `countablePartition γ n`. Used to define its limit `ProbabilityTheory.kernel.density`, which is a density for those kernels for all measurable sets. -/ noncomputable -def densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (t : γ) (s : Set β) : +def densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (x : γ) (s : Set β) : ℝ := - (κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t)).toReal + (κ a (partitionSet n x ×ˢ s) / ν a (partitionSet n x)).toReal lemma densityProcess_def (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (s : Set β) : (fun t ↦ densityProcess κ ν n a t s) @@ -118,54 +118,54 @@ lemma measurable_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (measurable_densityProcess_aux κ ν n hs).ennreal_toReal lemma measurable_densityProcess_left (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) - (t : γ) {s : Set β} (hs : MeasurableSet s) : - Measurable (fun a ↦ densityProcess κ ν n a t s) := + (x : γ) {s : Set β} (hs : MeasurableSet s) : + Measurable (fun a ↦ densityProcess κ ν n a x s) := (measurable_densityProcess κ ν n hs).comp (measurable_id.prod_mk measurable_const) lemma measurable_densityProcess_right (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) {s : Set β} (a : α) (hs : MeasurableSet s) : - Measurable (fun t ↦ densityProcess κ ν n a t s) := + Measurable (fun x ↦ densityProcess κ ν n a x s) := (measurable_densityProcess κ ν n hs).comp (measurable_const.prod_mk measurable_id) lemma measurable_partitionFiltration_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : - Measurable[partitionFiltration γ n] (fun t ↦ densityProcess κ ν n a t s) := by + Measurable[partitionFiltration γ n] (fun x ↦ densityProcess κ ν n a x s) := by refine @Measurable.ennreal_toReal _ (partitionFiltration γ n) _ ?_ exact (measurable_densityProcess_partitionFiltration_aux κ ν n hs).comp measurable_prod_mk_left lemma stronglyMeasurable_partitionFiltration_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : - StronglyMeasurable[partitionFiltration γ n] (fun t ↦ densityProcess κ ν n a t s) := + StronglyMeasurable[partitionFiltration γ n] (fun x ↦ densityProcess κ ν n a x s) := (measurable_partitionFiltration_densityProcess κ ν n a hs).stronglyMeasurable lemma adapted_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) {s : Set β} (hs : MeasurableSet s) : - Adapted (partitionFiltration γ) (fun n t ↦ densityProcess κ ν n a t s) := + Adapted (partitionFiltration γ) (fun n x ↦ densityProcess κ ν n a x s) := fun n ↦ stronglyMeasurable_partitionFiltration_densityProcess κ ν n a hs lemma densityProcess_nonneg (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) - (a : α) (t : γ) (s : Set β) : - 0 ≤ densityProcess κ ν n a t s := + (a : α) (x : γ) (s : Set β) : + 0 ≤ densityProcess κ ν n a x s := ENNReal.toReal_nonneg -lemma meas_partitionSet_le_of_fst_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : - κ a (partitionSet n t ×ˢ s) ≤ ν a (partitionSet n t) := by - calc κ a (partitionSet n t ×ˢ s) - ≤ fst κ a (partitionSet n t) := by +lemma meas_partitionSet_le_of_fst_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : + κ a (partitionSet n x ×ˢ s) ≤ ν a (partitionSet n x) := by + calc κ a (partitionSet n x ×ˢ s) + ≤ fst κ a (partitionSet n x) := by rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] refine measure_mono (fun x ↦ ?_) simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h - _ ≤ ν a (partitionSet n t) := hκν a _ + _ ≤ ν a (partitionSet n x) := hκν a _ -lemma densityProcess_le_one (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : - densityProcess κ ν n a t s ≤ 1 := by +lemma densityProcess_le_one (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : + densityProcess κ ν n a x s ≤ 1 := by refine ENNReal.toReal_le_of_le_ofReal zero_le_one (ENNReal.div_le_of_le_mul ?_) rw [ENNReal.ofReal_one, one_mul] - exact meas_partitionSet_le_of_fst_le hκν n a t s + exact meas_partitionSet_le_of_fst_le hκν n a x s lemma snorm_densityProcess_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (s : Set β) : - snorm (fun t ↦ densityProcess κ ν n a t s) 1 (ν a) ≤ ν a univ := by + snorm (fun x ↦ densityProcess κ ν n a x s) 1 (ν a) ≤ ν a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun x ↦ ?_))).trans ?_ · simp only [Real.norm_eq_abs, abs_of_nonneg (densityProcess_nonneg κ ν n a x s), densityProcess_le_one hκν n a x s] @@ -173,7 +173,7 @@ lemma snorm_densityProcess_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (s : Se lemma integrable_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : - Integrable (fun t ↦ densityProcess κ ν n a t s) (ν a) := by + Integrable (fun x ↦ densityProcess κ ν n a x s) (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ · exact measurable_densityProcess_right κ ν n a hs @@ -182,19 +182,19 @@ lemma integrable_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (n : lemma set_integral_densityProcess_I (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {u : Set γ} (hu : u ∈ countablePartition γ n) : - ∫ t in u, densityProcess κ ν n a t s ∂(ν a) = (κ a (u ×ˢ s)).toReal := by + ∫ x in u, densityProcess κ ν n a x s ∂(ν a) = (κ a (u ×ˢ s)).toReal := by have hu_meas : MeasurableSet u := measurableSet_countablePartition n hu simp_rw [densityProcess] rw [integral_toReal] rotate_left · refine Measurable.aemeasurable ?_ change Measurable ((fun (p : α × _) ↦ κ p.1 (partitionSet n p.2 ×ˢ s) - / ν p.1 (partitionSet n p.2)) ∘ (fun t ↦ (a, t))) + / ν p.1 (partitionSet n p.2)) ∘ (fun x ↦ (a, x))) exact (measurable_densityProcess_aux κ ν n hs).comp measurable_prod_mk_left - · refine ae_of_all _ (fun t ↦ ?_) - by_cases h0 : ν a (partitionSet n t) = 0 - · suffices κ a (partitionSet n t ×ˢ s) = 0 by simp [h0, this] - have h0' : fst κ a (partitionSet n t) = 0 := + · refine ae_of_all _ (fun x ↦ ?_) + by_cases h0 : ν a (partitionSet n x) = 0 + · suffices κ a (partitionSet n x ×ˢ s) = 0 by simp [h0, this] + have h0' : fst κ a (partitionSet n x) = 0 := le_antisymm ((hκν a _).trans h0.le) zero_le' rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] at h0' refine measure_mono_null (fun x ↦ ?_) h0' @@ -202,7 +202,7 @@ lemma set_integral_densityProcess_I (hκν : fst κ ≤ ν) [IsFiniteKernel κ] exact fun h _ ↦ h · exact ENNReal.div_lt_top (measure_ne_top _ _) h0 congr - have : ∫⁻ t in u, κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t) ∂(ν a) + have : ∫⁻ x in u, κ a (partitionSet n x ×ˢ s) / ν a (partitionSet n x) ∂(ν a) = ∫⁻ _ in u, κ a (u ×ˢ s) / ν a u ∂(ν a) := by refine set_lintegral_congr_fun hu_meas (ae_of_all _ (fun t ht ↦ ?_)) rw [partitionSet_of_mem hu ht] @@ -221,7 +221,7 @@ lemma set_integral_densityProcess_I (hκν : fst κ ≤ ν) [IsFiniteKernel κ] lemma integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : - ∫ t, densityProcess κ ν n a t s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by + ∫ x, densityProcess κ ν n a x s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by rw [← integral_univ, ← sUnion_countablePartition γ n, sUnion_eq_iUnion, iUnion_prod_const, measure_iUnion] rotate_left @@ -245,7 +245,7 @@ lemma integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFin lemma set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet[partitionFiltration γ n] A) : - ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + ∫ x in A, densityProcess κ ν n a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by refine induction_on_inter (m := partitionFiltration γ n) (s := countablePartition γ n) (C := fun A ↦ ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) rfl ?_ ?_ ?_ ?_ ?_ hA @@ -283,53 +283,53 @@ lemma set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [I lemma set_integral_densityProcess_of_le (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] {n m : ℕ} (hnm : n ≤ m) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet[partitionFiltration γ n] A) : - ∫ t in A, densityProcess κ ν m a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := + ∫ x in A, densityProcess κ ν m a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal := set_integral_densityProcess hκν m a hs ((partitionFiltration γ).mono hnm A hA) lemma condexp_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] {i j : ℕ} (hij : i ≤ j) (a : α) {s : Set β} (hs : MeasurableSet s) : - (ν a)[fun t ↦ densityProcess κ ν j a t s | partitionFiltration γ i] - =ᵐ[ν a] fun t ↦ densityProcess κ ν i a t s := by + (ν a)[fun x ↦ densityProcess κ ν j a x s | partitionFiltration γ i] + =ᵐ[ν a] fun x ↦ densityProcess κ ν i a x s := by refine (ae_eq_condexp_of_forall_set_integral_eq ?_ ?_ ?_ ?_ ?_).symm · exact integrable_densityProcess hκν j a hs - · exact fun t _ _ ↦ (integrable_densityProcess hκν _ _ hs).integrableOn - · intro t ht _ - rw [set_integral_densityProcess hκν i a hs ht, - set_integral_densityProcess_of_le hκν hij a hs ht] + · exact fun _ _ _ ↦ (integrable_densityProcess hκν _ _ hs).integrableOn + · intro x hx _ + rw [set_integral_densityProcess hκν i a hs hx, + set_integral_densityProcess_of_le hκν hij a hs hx] · exact StronglyMeasurable.aeStronglyMeasurable' (stronglyMeasurable_partitionFiltration_densityProcess κ ν i a hs) lemma martingale_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Martingale (fun n t ↦ densityProcess κ ν n a t s) (partitionFiltration γ) (ν a) := + Martingale (fun n x ↦ densityProcess κ ν n a x s) (partitionFiltration γ) (ν a) := ⟨adapted_densityProcess κ ν a hs, fun _ _ h ↦ condexp_densityProcess hκν h a hs⟩ -lemma densityProcess_mono_set (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) +lemma densityProcess_mono_set (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) {s s' : Set β} (h : s ⊆ s') : - densityProcess κ ν n a t s ≤ densityProcess κ ν n a t s' := by + densityProcess κ ν n a x s ≤ densityProcess κ ν n a x s' := by unfold densityProcess - by_cases h0 : ν a (partitionSet n t) = 0 + by_cases h0 : ν a (partitionSet n x) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] simp - have h_ne_top : ∀ s, κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t) ≠ ⊤ := by + have h_ne_top : ∀ s, κ a (partitionSet n x ×ˢ s) / ν a (partitionSet n x) ≠ ⊤ := by intro s rw [ne_eq, ENNReal.div_eq_top] simp only [ne_eq, h0, and_false, false_or, not_and, not_not] refine fun h_top ↦ eq_top_mono ?_ h_top - exact meas_partitionSet_le_of_fst_le hκν n a t s + exact meas_partitionSet_le_of_fst_le hκν n a x s rw [ENNReal.toReal_le_toReal (h_ne_top s) (h_ne_top s')] gcongr simp [prod_subset_prod_iff, subset_rfl, h] lemma densityProcess_mono_kernel_left {κ' : kernel α (γ × β)} (hκκ' : κ ≤ κ') - (hκ'ν : fst κ' ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : - densityProcess κ ν n a t s ≤ densityProcess κ' ν n a t s := by + (hκ'ν : fst κ' ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : + densityProcess κ ν n a x s ≤ densityProcess κ' ν n a x s := by unfold densityProcess - by_cases h0 : ν a (partitionSet n t) = 0 + by_cases h0 : ν a (partitionSet n x) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] simp - have h_le : κ' a (partitionSet n t ×ˢ s) ≤ ν a (partitionSet n t) := - meas_partitionSet_le_of_fst_le hκ'ν n a t s + have h_le : κ' a (partitionSet n x ×ˢ s) ≤ ν a (partitionSet n x) := + meas_partitionSet_le_of_fst_le hκ'ν n a x s rw [ENNReal.toReal_le_toReal] · gcongr exact hκκ' _ _ @@ -342,14 +342,14 @@ lemma densityProcess_mono_kernel_left {κ' : kernel α (γ × β)} (hκκ' : κ exact fun h_top ↦ eq_top_mono h_le h_top lemma densityProcess_antitone_kernel_right {ν' : kernel α γ} - (hνν' : ν ≤ ν') (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (t : γ) (s : Set β) : - densityProcess κ ν' n a t s ≤ densityProcess κ ν n a t s := by + (hνν' : ν ≤ ν') (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : + densityProcess κ ν' n a x s ≤ densityProcess κ ν n a x s := by unfold densityProcess - have h_le : κ a (partitionSet n t ×ˢ s) ≤ ν a (partitionSet n t) := - meas_partitionSet_le_of_fst_le hκν n a t s - by_cases h0 : ν a (partitionSet n t) = 0 + have h_le : κ a (partitionSet n x ×ˢ s) ≤ ν a (partitionSet n x) := + meas_partitionSet_le_of_fst_le hκν n a x s + by_cases h0 : ν a (partitionSet n x) = 0 · simp [le_antisymm (h_le.trans h0.le) zero_le', h0] - have h0' : ν' a (partitionSet n t) ≠ 0 := + have h0' : ν' a (partitionSet n x) ≠ 0 := fun h ↦ h0 (le_antisymm ((hνν' _ _).trans h.le) zero_le') rw [ENNReal.toReal_le_toReal] · gcongr @@ -361,17 +361,17 @@ lemma densityProcess_antitone_kernel_right {ν' : kernel α γ} exact fun h_top ↦ eq_top_mono h_le h_top @[simp] -lemma densityProcess_empty (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (t : γ) : - densityProcess κ ν n a t ∅ = 0 := by +lemma densityProcess_empty (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (x : γ) : + densityProcess κ ν n a x ∅ = 0 := by simp [densityProcess] lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (γ × β)) (ν : kernel α γ) - [IsFiniteKernel κ] (n : ℕ) (a : α) (t : γ) + [IsFiniteKernel κ] (n : ℕ) (a : α) (x : γ) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : - Tendsto (fun m ↦ densityProcess κ ν n a t (s m)) atTop (𝓝 (densityProcess κ ν n a t ∅)) := by + Tendsto (fun m ↦ densityProcess κ ν n a x (s m)) atTop (𝓝 (densityProcess κ ν n a x ∅)) := by simp_rw [densityProcess] - by_cases h0 : ν a (partitionSet n t) = 0 + by_cases h0 : ν a (partitionSet n x) = 0 · simp_rw [h0, ENNReal.toReal_div] simp refine (ENNReal.tendsto_toReal ?_).comp ?_ @@ -379,8 +379,7 @@ lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (γ × β)) push_neg simp refine ENNReal.Tendsto.div_const ?_ ?_ - · have h := tendsto_measure_iInter (μ := κ a) - (s := fun m ↦ partitionSet n t ×ˢ s m) ?_ ?_ ?_ + · have h := tendsto_measure_iInter (μ := κ a) (s := fun m ↦ partitionSet n x ×ˢ s m) ?_ ?_ ?_ · convert h rw [← prod_iInter, hs_iInter] · exact fun n ↦ MeasurableSet.prod (measurableSet_partitionSet _ _) (hs_meas n) @@ -392,18 +391,18 @@ lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (γ × β)) not_false_iff] lemma tendsto_densityProcess_atTop_of_antitone (κ : kernel α (γ × β)) (ν : kernel α γ) - [IsFiniteKernel κ] (n : ℕ) (a : α) (t : γ) + [IsFiniteKernel κ] (n : ℕ) (a : α) (x : γ) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : - Tendsto (fun m ↦ densityProcess κ ν n a t (s m)) atTop (𝓝 0) := by - rw [← densityProcess_empty κ ν n a t] - exact tendsto_densityProcess_atTop_empty_of_antitone κ ν n a t s hs hs_iInter hs_meas + Tendsto (fun m ↦ densityProcess κ ν n a x (s m)) atTop (𝓝 0) := by + rw [← densityProcess_empty κ ν n a x] + exact tendsto_densityProcess_atTop_empty_of_antitone κ ν n a x s hs hs_iInter hs_meas lemma tendsto_densityProcess_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - ∀ᵐ t ∂(ν a), Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop + ∀ᵐ x ∂(ν a), Tendsto (fun n ↦ densityProcess κ ν n a x s) atTop (𝓝 ((partitionFiltration γ).limitProcess - (fun n t ↦ densityProcess κ ν n a t s) (ν a) t)) := by + (fun n x ↦ densityProcess κ ν n a x s) (ν a) x)) := by refine Submartingale.ae_tendsto_limitProcess (martingale_densityProcess hκν a hs).submartingale (R := (ν a univ).toNNReal) (fun n ↦ ?_) refine (snorm_densityProcess_le hκν n a s).trans_eq ?_ @@ -413,7 +412,7 @@ lemma tendsto_densityProcess_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKerne lemma limitProcess_densityProcess_mem_L1 (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : Memℒp ((partitionFiltration γ).limitProcess - (fun n t ↦ densityProcess κ ν n a t s) (ν a)) 1 (ν a) := by + (fun n x ↦ densityProcess κ ν n a x s) (ν a)) 1 (ν a) := by refine Submartingale.memℒp_limitProcess (martingale_densityProcess hκν a hs).submartingale (R := (ν a univ).toNNReal) (fun n ↦ ?_) refine (snorm_densityProcess_le hκν n a s).trans_eq ?_ @@ -422,8 +421,8 @@ lemma limitProcess_densityProcess_mem_L1 (hκν : fst κ ≤ ν) [IsFiniteKernel lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Tendsto (fun n ↦ snorm ((fun t ↦ densityProcess κ ν n a t s) - - (partitionFiltration γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) + Tendsto (fun n ↦ snorm ((fun x ↦ densityProcess κ ν n a x s) + - (partitionFiltration γ).limitProcess (fun n x ↦ densityProcess κ ν n a x s) (ν a)) 1 (ν a)) atTop (𝓝 0) := by refine Submartingale.tendsto_snorm_one_limitProcess ?_ ?_ · exact (martingale_densityProcess hκν a hs).submartingale @@ -440,8 +439,8 @@ lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : fst κ ≤ ν) [IsF lemma tendsto_snorm_one_restrict_densityProcess_limitProcess [IsFiniteKernel κ] [IsFiniteKernel ν] (hκν : fst κ ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : - Tendsto (fun n ↦ snorm ((fun t ↦ densityProcess κ ν n a t s) - - (partitionFiltration γ).limitProcess (fun n t ↦ densityProcess κ ν n a t s) (ν a)) + Tendsto (fun n ↦ snorm ((fun x ↦ densityProcess κ ν n a x s) + - (partitionFiltration γ).limitProcess (fun n x ↦ densityProcess κ ν n a x s) (ν a)) 1 ((ν a).restrict A)) atTop (𝓝 0) := tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) (fun _ ↦ zero_le') @@ -453,22 +452,22 @@ section Density /-- Density of the kernel `κ` with respect to `ν`. This is a function `α → γ → Set β → ℝ` which is measurable on `α × γ` for all measurable sets `s : Set β` and satisfies that -`∫ t in A, density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal` for all measurable `A : Set γ`. -/ +`∫ x in A, density κ ν a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal` for all measurable `A : Set γ`. -/ noncomputable -def density (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) (t : γ) (s : Set β) : ℝ := - limsup (fun n ↦ densityProcess κ ν n a t s) atTop +def density (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) (x : γ) (s : Set β) : ℝ := + limsup (fun n ↦ densityProcess κ ν n a x s) atTop lemma density_ae_eq_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - (fun t ↦ density κ ν a t s) + (fun x ↦ density κ ν a x s) =ᵐ[ν a] (partitionFiltration γ).limitProcess - (fun n t ↦ densityProcess κ ν n a t s) (ν a) := by + (fun n x ↦ densityProcess κ ν n a x s) (ν a) := by filter_upwards [tendsto_densityProcess_limitProcess hκν a hs] with t ht using ht.limsup_eq lemma tendsto_m_density (hκν : fst κ ≤ ν) (a : α) [IsFiniteKernel κ] [IsFiniteKernel ν] {s : Set β} (hs : MeasurableSet s) : - ∀ᵐ t ∂(ν a), - Tendsto (fun n ↦ densityProcess κ ν n a t s) atTop (𝓝 (density κ ν a t s)) := by + ∀ᵐ x ∂(ν a), + Tendsto (fun n ↦ densityProcess κ ν n a x s) atTop (𝓝 (density κ ν a x s)) := by filter_upwards [tendsto_densityProcess_limitProcess hκν a hs, density_ae_eq_limitProcess hκν a hs] with t h1 h2 using h2 ▸ h1 @@ -477,33 +476,33 @@ lemma measurable_density (κ : kernel α (γ × β)) (ν : kernel α γ) Measurable (fun (p : α × γ) ↦ density κ ν p.1 p.2 s) := measurable_limsup (fun n ↦ measurable_densityProcess κ ν n hs) -lemma measurable_density_left (κ : kernel α (γ × β)) (ν : kernel α γ) (t : γ) +lemma measurable_density_left (κ : kernel α (γ × β)) (ν : kernel α γ) (x : γ) {s : Set β} (hs : MeasurableSet s) : - Measurable (fun a ↦ density κ ν a t s) := by - change Measurable ((fun (p : α × γ) ↦ density κ ν p.1 p.2 s) ∘ (fun a ↦ (a, t))) + Measurable (fun a ↦ density κ ν a x s) := by + change Measurable ((fun (p : α × γ) ↦ density κ ν p.1 p.2 s) ∘ (fun a ↦ (a, x))) exact (measurable_density κ ν hs).comp measurable_prod_mk_right lemma measurable_density_right (κ : kernel α (γ × β)) (ν : kernel α γ) {s : Set β} (hs : MeasurableSet s) (a : α) : - Measurable (fun t ↦ density κ ν a t s) := by - change Measurable ((fun (p : α × γ) ↦ density κ ν p.1 p.2 s) ∘ (fun t ↦ (a, t))) + Measurable (fun x ↦ density κ ν a x s) := by + change Measurable ((fun (p : α × γ) ↦ density κ ν p.1 p.2 s) ∘ (fun x ↦ (a, x))) exact (measurable_density κ ν hs).comp measurable_prod_mk_left -lemma density_mono_set (hκν : fst κ ≤ ν) (a : α) (t : γ) {s s' : Set β} (h : s ⊆ s') : - density κ ν a t s ≤ density κ ν a t s' := by +lemma density_mono_set (hκν : fst κ ≤ ν) (a : α) (x : γ) {s s' : Set β} (h : s ⊆ s') : + density κ ν a x s ≤ density κ ν a x s' := by refine limsup_le_limsup ?_ ?_ ?_ - · exact eventually_of_forall (fun n ↦ densityProcess_mono_set hκν n a t h) + · exact eventually_of_forall (fun n ↦ densityProcess_mono_set hκν n a x h) · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ -lemma density_nonneg (hκν : fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : - 0 ≤ density κ ν a t s := by +lemma density_nonneg (hκν : fst κ ≤ ν) (a : α) (x : γ) (s : Set β) : + 0 ≤ density κ ν a x s := by refine le_limsup_of_frequently_le ?_ ?_ · exact frequently_of_forall (fun n ↦ densityProcess_nonneg _ _ _ _ _ _) · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ -lemma density_le_one (hκν : fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : - density κ ν a t s ≤ 1 := by +lemma density_le_one (hκν : fst κ ≤ ν) (a : α) (x : γ) (s : Set β) : + density κ ν a x s ≤ 1 := by refine limsup_le_of_le ?_ ?_ · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) · exact eventually_of_forall (fun n ↦ densityProcess_le_one hκν _ _ _ _) @@ -511,7 +510,7 @@ lemma density_le_one (hκν : fst κ ≤ ν) (a : α) (t : γ) (s : Set β) : section Integral lemma snorm_density_le (hκν : fst κ ≤ ν) (a : α) (s : Set β) : - snorm (fun t ↦ density κ ν a t s) 1 (ν a) ≤ ν a univ := by + snorm (fun x ↦ density κ ν a x s) 1 (ν a) ≤ ν a univ := by refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ · simp only [Real.norm_eq_abs, abs_of_nonneg (density_nonneg hκν a t s), density_le_one hκν a t s] @@ -519,7 +518,7 @@ lemma snorm_density_le (hκν : fst κ ≤ ν) (a : α) (s : Set β) : lemma integrable_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Integrable (fun t ↦ density κ ν a t s) (ν a) := by + Integrable (fun x ↦ density κ ν a x s) (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ · exact measurable_density_right κ ν hs a @@ -529,8 +528,8 @@ lemma tendsto_set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKerne [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : Tendsto (fun i ↦ ∫ x in A, densityProcess κ ν i a x s ∂(ν a)) atTop (𝓝 (∫ x in A, density κ ν a x s ∂(ν a))) := by - refine tendsto_set_integral_of_L1' (μ := ν a) (fun t ↦ density κ ν a t s) - (integrable_density hκν a hs) (F := fun i t ↦ densityProcess κ ν i a t s) (l := atTop) + refine tendsto_set_integral_of_L1' (μ := ν a) (fun x ↦ density κ ν a x s) + (integrable_density hκν a hs) (F := fun i x ↦ densityProcess κ ν i a x s) (l := atTop) (eventually_of_forall (fun n ↦ integrable_densityProcess hκν _ _ hs)) ?_ A refine (tendsto_congr fun n ↦ ?_).mp (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) refine snorm_congr_ae ?_ @@ -540,12 +539,12 @@ lemma tendsto_set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKerne lemma set_integral_density_of_measurableSet (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet[partitionFiltration γ n] A) : - ∫ t in A, density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - suffices ∫ t in A, density κ ν a t s ∂(ν a) = ∫ t in A, densityProcess κ ν n a t s ∂(ν a) by + ∫ x in A, density κ ν a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + suffices ∫ x in A, density κ ν a x s ∂(ν a) = ∫ x in A, densityProcess κ ν n a x s ∂(ν a) by exact this ▸ set_integral_densityProcess hκν _ _ hs hA - suffices ∫ t in A, density κ ν a t s ∂(ν a) - = limsup (fun i ↦ ∫ t in A, densityProcess κ ν i a t s ∂(ν a)) atTop by - rw [this, ← limsup_const (α := ℕ) (f := atTop) (∫ t in A, densityProcess κ ν n a t s ∂(ν a)), + suffices ∫ x in A, density κ ν a x s ∂(ν a) + = limsup (fun i ↦ ∫ x in A, densityProcess κ ν i a x s ∂(ν a)) atTop by + rw [this, ← limsup_const (α := ℕ) (f := atTop) (∫ x in A, densityProcess κ ν n a x s ∂(ν a)), limsup_congr] simp only [eventually_atTop] refine ⟨n, fun m hnm ↦ ?_⟩ @@ -557,15 +556,15 @@ lemma set_integral_density_of_measurableSet (hκν : fst κ ≤ ν) [IsFiniteKer lemma integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - ∫ t, density κ ν a t s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by + ∫ x, density κ ν a x s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by rw [← integral_univ, set_integral_density_of_measurableSet hκν 0 a hs MeasurableSet.univ] lemma set_integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet A) : - ∫ t in A, density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + ∫ x in A, density κ ν a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by have hA' : MeasurableSet[⨆ n, partitionFiltration γ n] A := by rwa [iSup_partitionFiltration] refine induction_on_inter (m := ⨆ n, partitionFiltration γ n) - (C := fun A ↦ ∫ t in A, density κ ν a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) + (C := fun A ↦ ∫ x in A, density κ ν a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal) (measurableSpace_iSup_eq (partitionFiltration γ)) ?_ ?_ ?_ ?_ ?_ hA' · rintro s ⟨n, hs⟩ t ⟨m, ht⟩ _ exact ⟨max n m, ((partitionFiltration γ).mono (le_max_left n m) _ hs).inter @@ -601,7 +600,7 @@ lemma set_integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFinite lemma set_lintegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet A) : - ∫⁻ t in A, ENNReal.ofReal (density κ ν a t s) ∂(ν a) = κ a (A ×ˢ s) := by + ∫⁻ x in A, ENNReal.ofReal (density κ ν a x s) ∂(ν a) = κ a (A ×ˢ s) := by rw [← ofReal_integral_eq_lintegral_ofReal] · rw [set_integral_density hκν a hs hA, ENNReal.ofReal_toReal (measure_ne_top _ _)] @@ -610,7 +609,7 @@ lemma set_lintegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFinit lemma lintegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ t, ENNReal.ofReal (density κ ν a t s) ∂(ν a) = κ a (univ ×ˢ s) := by + ∫⁻ x, ENNReal.ofReal (density κ ν a x s) ∂(ν a) = κ a (univ ×ˢ s) := by rw [← set_lintegral_univ] exact set_lintegral_density hκν a hs MeasurableSet.univ @@ -619,7 +618,7 @@ end Integral lemma tendsto_integral_density_of_monotone (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : - Tendsto (fun m ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop (𝓝 (κ a univ).toReal) := by + Tendsto (fun m ↦ ∫ x, density κ ν a x (s m) ∂(ν a)) atTop (𝓝 (κ a univ).toReal) := by simp_rw [integral_density hκν a (hs_meas _)] have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := κ a univ) ?_ swap @@ -638,7 +637,7 @@ lemma tendsto_integral_density_of_monotone (hκν : fst κ ≤ ν) [IsFiniteKern lemma tendsto_integral_density_of_antitone (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : - Tendsto (fun m ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop (𝓝 0) := by + Tendsto (fun m ↦ ∫ x, density κ ν a x (s m) ∂(ν a)) atTop (𝓝 0) := by simp_rw [integral_density hκν a (hs_meas _)] rw [← ENNReal.zero_toReal] have h_cont := ENNReal.continuousOn_toReal.continuousAt (x := 0) ?_ @@ -660,29 +659,29 @@ lemma tendsto_integral_density_of_antitone (hκν : fst κ ≤ ν) [IsFiniteKern lemma tendsto_density_atTop_ae_of_antitone (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : - ∀ᵐ t ∂(ν a), Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 0) := by - have h_anti : ∀ t, Antitone (fun m ↦ density κ ν a t (s m)) := - fun t n m hnm ↦ density_mono_set hκν a t (hs hnm) - have h_le_one : ∀ m t, density κ ν a t (s m) ≤ 1 := fun m t ↦ density_le_one hκν a t (s m) - -- for all `t`, `fun m ↦ density κ a (s m) t` has a limit - have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := by - intro t - have h_tendsto : Tendsto (fun m ↦ density κ ν a t (s m)) atTop atBot ∨ - ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := - tendsto_of_antitone (h_anti t) + ∀ᵐ x ∂(ν a), Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 0) := by + have h_anti : ∀ x, Antitone (fun m ↦ density κ ν a x (s m)) := + fun x n m hnm ↦ density_mono_set hκν a x (hs hnm) + have h_le_one : ∀ m x, density κ ν a x (s m) ≤ 1 := fun m x ↦ density_le_one hκν a x (s m) + -- for all `x`, `fun m ↦ density κ a (s m) x` has a limit + have h_exists : ∀ x, ∃ l, Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 l) := by + intro x + have h_tendsto : Tendsto (fun m ↦ density κ ν a x (s m)) atTop atBot ∨ + ∃ l, Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 l) := + tendsto_of_antitone (h_anti x) cases' h_tendsto with h_absurd h_tendsto - · rw [tendsto_atTop_atBot_iff_of_antitone (h_anti t)] at h_absurd + · rw [tendsto_atTop_atBot_iff_of_antitone (h_anti x)] at h_absurd obtain ⟨r, hr⟩ := h_absurd (-1) - have h_nonneg := density_nonneg hκν a t (s r) + have h_nonneg := density_nonneg hκν a x (s r) linarith · exact h_tendsto -- let `F` be the pointwise limit of `fun m ↦ density κ a (s m) t` for all `t` - let F : γ → ℝ := fun t ↦ (h_exists t).choose - have hF_tendsto : ∀ t, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 (F t)) := - fun t ↦ (h_exists t).choose_spec - have hF_nonneg : ∀ t, 0 ≤ F t := - fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ density_nonneg hκν a t (s m)) - have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) + let F : γ → ℝ := fun x ↦ (h_exists x).choose + have hF_tendsto : ∀ x, Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 (F x)) := + fun x ↦ (h_exists x).choose_spec + have hF_nonneg : ∀ x, 0 ≤ F x := + fun x ↦ ge_of_tendsto' (hF_tendsto x) (fun m ↦ density_nonneg hκν a x (s m)) + have hF_le_one : ∀ x, F x ≤ 1 := fun x ↦ le_of_tendsto' (hF_tendsto x) (fun m ↦ h_le_one m x) have hF_int : Integrable F (ν a) := by rw [← memℒp_one_iff_integrable] refine ⟨?_, ?_⟩ @@ -699,20 +698,20 @@ lemma tendsto_density_atTop_ae_of_antitone (hκν : fst κ ≤ ν) [IsFiniteKern exact measure_lt_top _ _ -- it suffices to show that the limit `F` is 0 a.e. suffices F=ᵐ[ν a] 0 by - filter_upwards [this] with t ht_eq - simp only [Pi.zero_apply] at ht_eq - rw [← ht_eq] - exact hF_tendsto t + filter_upwards [this] with x hx_eq + simp only [Pi.zero_apply] at hx_eq + rw [← hx_eq] + exact hF_tendsto x -- since `F` is nonnegative, proving that its integral is 0 is sufficient to get that -- `F` is 0 a.e. rw [← integral_eq_zero_iff_of_nonneg hF_nonneg hF_int] - have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop - (𝓝 (∫ t, F t ∂(ν a))) := by + have h_integral : Tendsto (fun m : ℕ ↦ ∫ x, density κ ν a x (s m) ∂(ν a)) atTop + (𝓝 (∫ x, F x ∂(ν a))) := by refine integral_tendsto_of_tendsto_of_antitone ?_ hF_int ?_ ?_ · exact fun n ↦ integrable_density hκν _ (hs_meas n) · exact ae_of_all _ h_anti · exact ae_of_all _ hF_tendsto - have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop + have h_integral' : Tendsto (fun m : ℕ ↦ ∫ x, density κ ν a x (s m) ∂(ν a)) atTop (𝓝 0) := by exact tendsto_integral_density_of_antitone hκν a s hs hs_iInter hs_meas exact tendsto_nhds_unique h_integral h_integral' @@ -721,18 +720,18 @@ section UnivFst /-! We specialize to `ν = fst κ`, for which `density κ (fst κ) a t univ = 1` almost everywhere. -/ -lemma densityProcess_univ [IsFiniteKernel κ] (n : ℕ) (a : α) (t : γ) : - densityProcess κ (fst κ) n a t univ = if fst κ a (partitionSet n t) = 0 then 0 else 1 := by +lemma densityProcess_univ [IsFiniteKernel κ] (n : ℕ) (a : α) (x : γ) : + densityProcess κ (fst κ) n a x univ = if fst κ a (partitionSet n x) = 0 then 0 else 1 := by rw [densityProcess] - by_cases h : fst κ a (partitionSet n t) = 0 + by_cases h : fst κ a (partitionSet n x) = 0 · simp [h] - by_cases h' : κ a (partitionSet n t ×ˢ univ) = 0 + by_cases h' : κ a (partitionSet n x ×ˢ univ) = 0 · simp [h'] · rw [ENNReal.div_zero h'] simp · simp only [h, ite_false] rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] - have : partitionSet n t ×ˢ univ = {p : γ × β | p.1 ∈ partitionSet n t} := by + have : partitionSet n x ×ˢ univ = {p : γ × β | p.1 ∈ partitionSet n x} := by ext x simp rw [this, ENNReal.div_self] @@ -741,16 +740,16 @@ lemma densityProcess_univ [IsFiniteKernel κ] (n : ℕ) (a : α) (t : γ) : · exact measure_ne_top _ _ lemma densityProcess_univ_ae (κ : kernel α (γ × β)) [IsFiniteKernel κ] (n : ℕ) (a : α) : - ∀ᵐ t ∂(fst κ a), densityProcess κ (fst κ) n a t univ = 1 := by + ∀ᵐ x ∂(fst κ a), densityProcess κ (fst κ) n a x univ = 1 := by rw [ae_iff] - have : {t | ¬ densityProcess κ (fst κ) n a t univ = 1} - ⊆ {t | fst κ a (partitionSet n t) = 0} := by - intro t ht - simp only [mem_setOf_eq] at ht ⊢ - rw [densityProcess_univ] at ht - simpa using ht + have : {x | ¬ densityProcess κ (fst κ) n a x univ = 1} + ⊆ {x | fst κ a (partitionSet n x) = 0} := by + intro x hx + simp only [mem_setOf_eq] at hx ⊢ + rw [densityProcess_univ] at hx + simpa using hx refine measure_mono_null this ?_ - have : {t | fst κ a (partitionSet n t) = 0} + have : {x | fst κ a (partitionSet n x) = 0} ⊆ ⋃ (u) (_ : u ∈ countablePartition γ n) (_ : fst κ a u = 0), u := by intro t ht simp only [mem_setOf_eq, mem_iUnion, exists_prop] at ht ⊢ @@ -768,9 +767,9 @@ lemma densityProcess_univ_ae (κ : kernel α (γ × β)) [IsFiniteKernel κ] (n · simp [h] lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (γ × β)) - (n : ℕ) (a : α) (t : γ) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) : - Tendsto (fun m ↦ densityProcess κ (fst κ) n a t (s m)) atTop - (𝓝 (densityProcess κ (fst κ) n a t univ)) := by + (n : ℕ) (a : α) (x : γ) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) : + Tendsto (fun m ↦ densityProcess κ (fst κ) n a x (s m)) atTop + (𝓝 (densityProcess κ (fst κ) n a x univ)) := by simp_rw [densityProcess] refine (ENNReal.tendsto_toReal ?_).comp ?_ · rw [ne_eq, ENNReal.div_eq_top] @@ -783,11 +782,11 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (γ × β)) · refine fun h_top ↦ eq_top_mono (measure_mono (fun x ↦ ?_)) h_top simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h - by_cases h0 : fst κ a (partitionSet n t) = 0 + by_cases h0 : fst κ a (partitionSet n x) = 0 · rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] at h0 ⊢ - suffices ∀ m, κ a (partitionSet n t ×ˢ s m) = 0 by + suffices ∀ m, κ a (partitionSet n x ×ˢ s m) = 0 by simp only [this, h0, ENNReal.zero_div, tendsto_const_nhds_iff] - suffices κ a (partitionSet n t ×ˢ univ) = 0 by + suffices κ a (partitionSet n x ×ˢ univ) = 0 by simp only [this, ENNReal.zero_div] convert h0 ext x @@ -797,7 +796,7 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (γ × β)) exact fun h _ ↦ h refine ENNReal.Tendsto.div_const ?_ ?_ · have h := tendsto_measure_iUnion (μ := κ a) - (s := fun m ↦ partitionSet n t ×ˢ s m) ?_ + (s := fun m ↦ partitionSet n x ×ˢ s m) ?_ swap · intro m m' hmm' simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] @@ -808,45 +807,44 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (γ × β)) lemma tendsto_densityProcess_atTop_ae_of_monotone (κ : kernel α (γ × β)) [IsFiniteKernel κ] (n : ℕ) (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) : - ∀ᵐ t ∂(fst κ a), Tendsto (fun m ↦ densityProcess κ (fst κ) n a t (s m)) atTop (𝓝 1) := by - filter_upwards [densityProcess_univ_ae κ n a] with t ht - rw [← ht] - exact tendsto_densityProcess_atTop_univ_of_monotone κ n a t s hs hs_iUnion + ∀ᵐ x ∂(fst κ a), Tendsto (fun m ↦ densityProcess κ (fst κ) n a x (s m)) atTop (𝓝 1) := by + filter_upwards [densityProcess_univ_ae κ n a] with x hx + rw [← hx] + exact tendsto_densityProcess_atTop_univ_of_monotone κ n a x s hs hs_iUnion lemma density_univ (κ : kernel α (γ × β)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(fst κ a), density κ (fst κ) a t univ = 1 := by + ∀ᵐ x ∂(fst κ a), density κ (fst κ) a x univ = 1 := by have h := fun n ↦ densityProcess_univ_ae κ n a rw [← ae_all_iff] at h - filter_upwards [h] with t ht - simp [density, ht] + filter_upwards [h] with x hx + simp [density, hx] lemma tendsto_density_atTop_ae_of_monotone [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : - ∀ᵐ t ∂(fst κ a), Tendsto (fun m ↦ density κ (fst κ) a t (s m)) atTop (𝓝 1) := by + ∀ᵐ x ∂(fst κ a), Tendsto (fun m ↦ density κ (fst κ) a x (s m)) atTop (𝓝 1) := by let ν := fst κ - have h_mono : ∀ t, Monotone (fun m ↦ density κ (fst κ) a t (s m)) := - fun t n m hnm ↦ density_mono_set le_rfl a t (hs hnm) - have h_le_one : ∀ m t, density κ ν a t (s m) ≤ 1 := - fun m t ↦ density_le_one le_rfl a t (s m) - -- for all `t`, `fun m ↦ density κ a (s m) t` has a limit - have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := by - intro t - have h_tendsto : Tendsto (fun m ↦ density κ ν a t (s m)) atTop atTop ∨ - ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := - tendsto_of_monotone (h_mono t) + have h_mono : ∀ x, Monotone (fun m ↦ density κ (fst κ) a x (s m)) := + fun x n m hnm ↦ density_mono_set le_rfl a x (hs hnm) + have h_le_one : ∀ m x, density κ ν a x (s m) ≤ 1 := fun m x ↦ density_le_one le_rfl a x (s m) + -- for all `x`, `fun m ↦ density κ a (s m) x` has a limit + have h_exists : ∀ x, ∃ l, Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 l) := by + intro x + have h_tendsto : Tendsto (fun m ↦ density κ ν a x (s m)) atTop atTop ∨ + ∃ l, Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 l) := + tendsto_of_monotone (h_mono x) cases' h_tendsto with h_absurd h_tendsto - · rw [tendsto_atTop_atTop_iff_of_monotone (h_mono t)] at h_absurd + · rw [tendsto_atTop_atTop_iff_of_monotone (h_mono x)] at h_absurd obtain ⟨r, hr⟩ := h_absurd 2 - exact absurd (hr.trans (h_le_one r t)) one_lt_two.not_le + exact absurd (hr.trans (h_le_one r x)) one_lt_two.not_le · exact h_tendsto - -- let `F` be the pointwise limit of `fun m ↦ density κ a (s m) t` for all `t` - let F : γ → ℝ := fun t ↦ (h_exists t).choose - have hF_tendsto : ∀ t, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 (F t)) := - fun t ↦ (h_exists t).choose_spec - have hF_nonneg : ∀ t, 0 ≤ F t := - fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ density_nonneg le_rfl a t (s m)) - have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) + -- let `F` be the pointwise limit of `fun m ↦ density κ a (s m) x` for all `x` + let F : γ → ℝ := fun x ↦ (h_exists x).choose + have hF_tendsto : ∀ x, Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 (F x)) := + fun x ↦ (h_exists x).choose_spec + have hF_nonneg : ∀ x, 0 ≤ F x := + fun x ↦ ge_of_tendsto' (hF_tendsto x) (fun m ↦ density_nonneg le_rfl a x (s m)) + have hF_le_one : ∀ x, F x ≤ 1 := fun x ↦ le_of_tendsto' (hF_tendsto x) (fun m ↦ h_le_one m x) have hF_int : Integrable F (ν a) := by rw [← memℒp_one_iff_integrable] constructor @@ -861,19 +859,19 @@ lemma tendsto_density_atTop_ae_of_monotone [IsFiniteKernel κ] _ < ⊤ := by simp only [MeasureTheory.lintegral_const, measure_univ, one_mul, measure_lt_top] -- it suffices to show that the limit `F` is 1 a.e. suffices F =ᵐ[ν a] (fun _ ↦ 1) by - filter_upwards [this] with t ht_eq - rw [← ht_eq] - exact hF_tendsto t + filter_upwards [this] with x hx_eq + rw [← hx_eq] + exact hF_tendsto x -- since `F` is at most 1, proving that its integral is the same as the integral of 1 will tell -- us that `F` is 1 a.e. rw [← integral_eq_iff_of_ae_le hF_int (integrable_const _) (ae_of_all _ hF_le_one)] - have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop - (𝓝 (∫ t, F t ∂(ν a))) := by + have h_integral : Tendsto (fun m : ℕ ↦ ∫ x, density κ ν a x (s m) ∂(ν a)) atTop + (𝓝 (∫ x, F x ∂(ν a))) := by refine integral_tendsto_of_tendsto_of_monotone ?_ hF_int ?_ ?_ · exact fun n ↦ integrable_density le_rfl _ (hs_meas n) · exact ae_of_all _ h_mono · exact ae_of_all _ hF_tendsto - have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop + have h_integral' : Tendsto (fun m : ℕ ↦ ∫ x, density κ ν a x (s m) ∂(ν a)) atTop (𝓝 (∫ _, 1 ∂(ν a))) := by rw [MeasureTheory.integral_const] simp only [smul_eq_mul, mul_one] From e2330f2936970fafddf7494e1d9cf287747fceba Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 26 Feb 2024 17:00:47 +0100 Subject: [PATCH 080/129] docstrings --- .../Kernel/Disintegration/Basic.lean | 18 ++-- .../Kernel/Disintegration/CondCdf.lean | 10 +- .../Kernel/Disintegration/CondKernelCdf.lean | 15 ++- .../Kernel/Disintegration/KernelCdf.lean | 92 +++++++++++-------- 4 files changed, 74 insertions(+), 61 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 0a585325dd6a4..20ab39e944572 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -102,7 +102,7 @@ On `ℝ`, we get disintegration by constructing a map `f` with the property `IsK noncomputable def condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFunction} - (hf : IsKernelCDF f + (hf : IsCondKernelCDF f (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)) (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) @@ -118,7 +118,7 @@ def condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFu instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFunction} - (hf : IsKernelCDF f + (hf : IsCondKernelCDF f (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)) (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) @@ -137,7 +137,7 @@ instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω)) lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω)) [IsFiniteKernel κ] {f : α × β → StieltjesFunction} - (hf : IsKernelCDF f + (hf : IsCondKernelCDF f (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)) (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) @@ -213,21 +213,21 @@ section StandardBorel noncomputable def kernel.condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : kernel (α × γ) ℝ := - (isKernelCDF_kernelCDF κ).toKernel _ + (isCondKernelCDF_condKernelCDF κ).toKernel _ instance (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : IsMarkovKernel (kernel.condKernelReal κ) := by unfold kernel.condKernelReal; infer_instance lemma kernel.eq_compProd_condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : κ = kernel.fst κ ⊗ₖ kernel.condKernelReal κ := - kernel.eq_compProd_toKernel (isKernelCDF_kernelCDF κ) + kernel.eq_compProd_toKernel (isCondKernelCDF_condKernelCDF κ) noncomputable def condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : kernel (α × γ) Ω := let f := measurableEmbedding_real Ω let hf := measurableEmbedding_measurableEmbedding_real Ω let κ' := kernel.map κ (Prod.map (id : γ → γ) f) (measurable_id.prod_map hf.measurable) - condKernelBorelSnd κ (isKernelCDF_kernelCDF κ') + condKernelBorelSnd κ (isCondKernelCDF_condKernelCDF κ') instance instIsMarkovKernel_condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : IsMarkovKernel (condKernelBorel κ) := by @@ -246,7 +246,7 @@ section Real noncomputable def condKernelUnitReal (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : kernel (Unit × α) ℝ := - (isKernelCDF_condCDF (ρ ())).toKernel (fun (p : Unit × α) ↦ condCDF (ρ ()) p.2) + (isCondKernelCDF_condCDF (ρ ())).toKernel (fun (p : Unit × α) ↦ condCDF (ρ ()) p.2) instance (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : IsMarkovKernel (condKernelUnitReal ρ) := by rw [condKernelUnitReal]; infer_instance @@ -254,7 +254,7 @@ instance (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : IsMarkovKernel (co lemma fst_compProd_condKernelUnitReal (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : kernel.fst ρ ⊗ₖ condKernelUnitReal ρ = ρ := by have : ρ = kernel.const Unit (ρ ()) := by ext; simp - conv_rhs => rw [this, kernel.eq_compProd_toKernel (isKernelCDF_condCDF (ρ ()))] + conv_rhs => rw [this, kernel.eq_compProd_toKernel (isCondKernelCDF_condCDF (ρ ()))] end Real @@ -265,7 +265,7 @@ def condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : kern let f := measurableEmbedding_real Ω let hf := measurableEmbedding_measurableEmbedding_real Ω let κ' := kernel.map κ (Prod.map (id : α → α) f) (measurable_id.prod_map hf.measurable) - condKernelBorelSnd κ (isKernelCDF_condCDF (κ' ())) + condKernelBorelSnd κ (isCondKernelCDF_condCDF (κ' ())) instance instIsMarkovKernel_condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : IsMarkovKernel (condKernelUnitBorel κ) := by diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index 28dcc6b63e43d..4f7d98b83afc6 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -425,8 +425,8 @@ theorem integrable_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : refine (set_lintegral_le_lintegral _ _).trans (lintegral_mono_ae ?_) filter_upwards [preCDF_le_one ρ] with a ha using ENNReal.ofReal_toReal_le.trans (ha _) -lemma isRatKernelCDF_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - IsRatKernelCDF (fun p r ↦ (preCDF ρ r p.2).toReal) +lemma isRatCondKernelCDF_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : + IsRatCondKernelCDF (fun p r ↦ (preCDF ρ r p.2).toReal) (kernel.const Unit ρ) (kernel.const Unit ρ.fst) where measurable := measurable_preCDF'.comp measurable_snd isRatStieltjesPoint_ae a := by @@ -492,7 +492,7 @@ theorem measurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : Measurable fun theorem set_lintegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) {s : Set α} (hs : MeasurableSet s) : ∫⁻ a in s, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x) := by - have h := set_lintegral_stieltjesOfMeasurableRat (isRatKernelCDF_preCDF ρ) () x hs + have h := set_lintegral_stieltjesOfMeasurableRat (isRatCondKernelCDF_preCDF ρ) () x hs simp only [kernel.const_apply] at h rw [← h] congr with a @@ -542,8 +542,8 @@ theorem integral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : rw [← set_integral_condCDF ρ _ MeasurableSet.univ, Measure.restrict_univ] #align probability_theory.integral_cond_cdf ProbabilityTheory.integral_condCDF -lemma isKernelCDF_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - IsKernelCDF (fun p : Unit × α ↦ condCDF ρ p.2) (kernel.const Unit ρ) +lemma isCondKernelCDF_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : + IsCondKernelCDF (fun p : Unit × α ↦ condCDF ρ p.2) (kernel.const Unit ρ) (kernel.const Unit ρ.fst) where measurable x := (measurable_condCDF ρ x).comp measurable_snd integrable _ x := integrable_condCDF ρ x diff --git a/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean index b0dfc3ff6e524..02adcb44128e3 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean @@ -205,8 +205,8 @@ lemma isRatStieltjesPoint_densityIic_ae (κ : kernel α (γ × ℝ)) [IsFiniteKe iInf_rat_gt_densityIic_eq le_rfl a] with t ht_top ht_bot ht_iInf exact ⟨monotone_densityIic le_rfl a t, ht_top, ht_bot, ht_iInf⟩ -lemma isRatKernelCDF_densityIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsRatKernelCDF (fun p : α × γ ↦ kernel.densityIic κ (kernel.fst κ) p.1 p.2) κ (kernel.fst κ) +lemma isRatCondKernelCDF_densityIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsRatCondKernelCDF (fun p : α × γ ↦ kernel.densityIic κ (kernel.fst κ) p.1 p.2) κ (kernel.fst κ) where measurable := measurable_densityIic κ (kernel.fst κ) isRatStieltjesPoint_ae := isRatStieltjesPoint_densityIic_ae κ @@ -214,13 +214,12 @@ lemma isRatKernelCDF_densityIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] set_integral := fun _ _ hs _ ↦ set_integral_densityIic le_rfl _ _ hs noncomputable -def kernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - α × γ → StieltjesFunction := +def condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : α × γ → StieltjesFunction := stieltjesOfMeasurableRat (fun p : α × γ ↦ kernel.densityIic κ (kernel.fst κ) p.1 p.2) - (isRatKernelCDF_densityIic κ).measurable + (isRatCondKernelCDF_densityIic κ).measurable -lemma isKernelCDF_kernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsKernelCDF (kernelCDF κ) κ (kernel.fst κ) := - isKernelCDF_stieltjesOfMeasurableRat (isRatKernelCDF_densityIic κ) +lemma isCondKernelCDF_condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsCondKernelCDF (condKernelCDF κ) κ (kernel.fst κ) := + isCondKernelCDF_stieltjesOfMeasurableRat (isRatCondKernelCDF_densityIic κ) end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCdf.lean b/Mathlib/Probability/Kernel/Disintegration/KernelCdf.lean index 841c5b5a5c4f7..1aa10e9065a9f 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/KernelCdf.lean @@ -8,18 +8,32 @@ import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes /-! -# Cumulative distribution function of a Markov kernel +# Conditional cumulative distribution function of a Markov kernel ## Main definitions -* `ProbabilityTheory.IsRatKernelCDF`: -* `ProbabilityTheory.IsKernelCDF`: -* `ProbabilityTheory.IsKernelCDF.toKernel`: +Let `μ : kernel α (β × ℝ)` and `ν : kernel α β`. + +* `ProbabilityTheory.IsCondKernelCDF`: a function `f : α × β → StieltjesFunction` is called + a conditional kernel CDF of `μ` with respect to `ν` if it is measurable, tends to to 0 at -∞ and + to 1 at +∞ for all `p : α × β`, if `fun t ↦ f (a, t) x` is `(ν a)`-integrable for all `a : α` and + `x : ℝ` and for all measurable sets `s : Set β`, + `∫ t in s, f (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal`. +* `ProbabilityTheory.IsCondKernelCDF.toKernel`: from a function `f : α × β → StieltjesFunction` + with the property `hf : IsCondKernelCDF f μ ν`, build a `kernel (α × β) ℝ` such that + `μ = ν ⊗ₖ hf.toKernel f`. +* `ProbabilityTheory.IsRatCondKernelCDF`: a function `f : α × β → ℚ → ℝ` is called a rational + conditional kernel CDF of `μ` with respect to `ν` if is measurable and satisfies the same + integral conditions as in the definition of `IsCondKernelCDF`, and the `ℚ → ℝ` function `f (a, x)` + satisfies the properties of a Sieltjes function for `(ν a)`-almost all `x : β`. ## Main statements -* `ProbabilityTheory.isKernelCDF_stieltjesOfMeasurableRat`: -* `ProbabilityTheory.kernel.eq_compProd_toKernel`: +* `ProbabilityTheory.isCondKernelCDF_stieltjesOfMeasurableRat`: if `f : α × β → ℚ → ℝ` has the + property `IsRatCondKernelCDF`, then `stieltjesOfMeasurableRat f` is a function + `α × β → StieltjesFunction` with the property `IsCondKernelCDF`. +* `ProbabilityTheory.kernel.eq_compProd_toKernel`: for `hf : IsCondKernelCDF f μ ν`, + `μ = ν ⊗ₖ hf.toKernel f` -/ @@ -36,26 +50,26 @@ section stieltjesOfMeasurableRat variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} {f : α × β → ℚ → ℝ} {μ : kernel α (β × ℝ)} {ν : kernel α β} -structure IsRatKernelCDF (f : α × β → ℚ → ℝ) (μ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := +structure IsRatCondKernelCDF (f : α × β → ℚ → ℝ) (μ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := (measurable : Measurable f) (isRatStieltjesPoint_ae (a : α) : ∀ᵐ t ∂(ν a), IsRatStieltjesPoint f (a, t)) (integrable (a : α) (q : ℚ) : Integrable (fun t ↦ f (a, t) q) (ν a)) (set_integral (a : α) {s : Set β} (_hs : MeasurableSet s) (q : ℚ) : ∫ t in s, f (a, t) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal) -lemma stieltjesOfMeasurableRat_ae_eq (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) : +lemma stieltjesOfMeasurableRat_ae_eq (hf : IsRatCondKernelCDF f μ ν) (a : α) (q : ℚ) : (fun t ↦ stieltjesOfMeasurableRat f hf.measurable (a, t) q) =ᵐ[ν a] fun t ↦ f (a, t) q := by filter_upwards [hf.isRatStieltjesPoint_ae a] with a ha rw [stieltjesOfMeasurableRat_eq, toRatCDF_of_isRatStieltjesPoint ha] -lemma set_integral_stieltjesOfMeasurableRat_rat (hf : IsRatKernelCDF f μ ν) (a : α) (q : ℚ) +lemma set_integral_stieltjesOfMeasurableRat_rat (hf : IsRatCondKernelCDF f μ ν) (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : ∫ t in s, stieltjesOfMeasurableRat f hf.measurable (a, t) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal := by rw [set_integral_congr_ae hs (g := fun t ↦ f (a, t) q) ?_, hf.set_integral a hs] filter_upwards [stieltjesOfMeasurableRat_ae_eq hf a q] with b hb using fun _ ↦ hb -lemma set_lintegral_stieltjesOfMeasurableRat_rat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma set_lintegral_stieltjesOfMeasurableRat_rat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : ∫⁻ t in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) q) ∂(ν a) = μ a (s ×ˢ Iic (q : ℝ)) := by @@ -67,7 +81,7 @@ lemma set_lintegral_stieltjesOfMeasurableRat_rat [IsFiniteKernel μ] (hf : IsRat exact hf.integrable a q · exact ae_of_all _ (fun x ↦ stieltjesOfMeasurableRat_nonneg _ _ _) -lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : ∫⁻ t in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) x) ∂(ν a) = μ a (s ×ˢ Iic x) := by @@ -129,13 +143,13 @@ lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKern exact mod_cast hij · exact ⟨h_nonempty.some, measure_ne_top _ _⟩ -lemma lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) (a : α) (x : ℝ) : ∫⁻ t, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) x) ∂(ν a) = μ a (univ ×ˢ Iic x) := by rw [← set_lintegral_univ, set_lintegral_stieltjesOfMeasurableRat hf _ _ MeasurableSet.univ] -lemma integrable_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma integrable_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) (a : α) (x : ℝ) : Integrable (fun t ↦ stieltjesOfMeasurableRat f hf.measurable (a, t) x) (ν a) := by have : (fun t ↦ stieltjesOfMeasurableRat f hf.measurable (a, t) x) @@ -150,7 +164,7 @@ lemma integrable_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKernelC · rw [lintegral_stieltjesOfMeasurableRat hf] exact measure_ne_top _ _ -lemma set_integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma set_integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : ∫ t in s, stieltjesOfMeasurableRat f hf.measurable (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal := by @@ -163,7 +177,7 @@ lemma set_integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKerne · exact (integrable_stieltjesOfMeasurableRat hf _ _).restrict · exact ae_of_all _ (fun _ ↦ stieltjesOfMeasurableRat_nonneg _ _ _) -lemma integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKernelCDF f μ ν) +lemma integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) (a : α) (x : ℝ) : ∫ t, stieltjesOfMeasurableRat f hf.measurable (a, t) x ∂(ν a) = (μ a (univ ×ˢ Iic x)).toReal := by @@ -171,12 +185,12 @@ lemma integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatKernelCDF end stieltjesOfMeasurableRat -section IsKernelCDF +section IsCondKernelCDF variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} {f : α × β → StieltjesFunction} {μ : kernel α (β × ℝ)} {ν : kernel α β} -structure IsKernelCDF (f : α × β → StieltjesFunction) (μ : kernel α (β × ℝ)) (ν : kernel α β) : +structure IsCondKernelCDF (f : α × β → StieltjesFunction) (μ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := (measurable (x : ℝ) : Measurable fun p ↦ f p x) (integrable (a : α) (x : ℝ) : Integrable (fun t ↦ f (a, t) x) (ν a)) @@ -185,35 +199,35 @@ structure IsKernelCDF (f : α × β → StieltjesFunction) (μ : kernel α (β (set_integral (a : α) {s : Set β} (_hs : MeasurableSet s) (x : ℝ) : ∫ t in s, f (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal) -lemma IsKernelCDF.nonneg (hf : IsKernelCDF f μ ν) (p : α × β) (x : ℝ) : 0 ≤ f p x := +lemma IsCondKernelCDF.nonneg (hf : IsCondKernelCDF f μ ν) (p : α × β) (x : ℝ) : 0 ≤ f p x := Monotone.le_of_tendsto (f p).mono (hf.tendsto_atBot_zero p) x -lemma IsKernelCDF.le_one (hf : IsKernelCDF f μ ν) (p : α × β) (x : ℝ) : f p x ≤ 1 := +lemma IsCondKernelCDF.le_one (hf : IsCondKernelCDF f μ ν) (p : α × β) (x : ℝ) : f p x ≤ 1 := Monotone.ge_of_tendsto (f p).mono (hf.tendsto_atTop_one p) x -lemma IsKernelCDF.set_lintegral [IsFiniteKernel μ] - {f : α × β → StieltjesFunction} (hf : IsKernelCDF f μ ν) +lemma IsCondKernelCDF.set_lintegral [IsFiniteKernel μ] + {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f μ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (x : ℝ) : ∫⁻ t in s, ENNReal.ofReal (f (a, t) x) ∂(ν a) = μ a (s ×ˢ Iic x) := by rw [← ofReal_integral_eq_lintegral_ofReal (hf.integrable a x).restrict (ae_of_all _ (fun _ ↦ hf.nonneg _ _)), hf.set_integral a hs x, ENNReal.ofReal_toReal] exact measure_ne_top _ _ -lemma isKernelCDF_stieltjesOfMeasurableRat {f : α × β → ℚ → ℝ} (hf : IsRatKernelCDF f μ ν) +lemma isCondKernelCDF_stieltjesOfMeasurableRat {f : α × β → ℚ → ℝ} (hf : IsRatCondKernelCDF f μ ν) [IsFiniteKernel μ] : - IsKernelCDF (stieltjesOfMeasurableRat f hf.measurable) μ ν where + IsCondKernelCDF (stieltjesOfMeasurableRat f hf.measurable) μ ν where measurable := measurable_stieltjesOfMeasurableRat hf.measurable integrable := integrable_stieltjesOfMeasurableRat hf tendsto_atTop_one := tendsto_stieltjesOfMeasurableRat_atTop hf.measurable tendsto_atBot_zero := tendsto_stieltjesOfMeasurableRat_atBot hf.measurable set_integral a _ hs x := set_integral_stieltjesOfMeasurableRat hf a x hs -end IsKernelCDF +end IsCondKernelCDF section ToKernel variable {_ : MeasurableSpace β} {f : α × β → StieltjesFunction} - {μ : kernel α (β × ℝ)} {ν : kernel α β} {hf : IsKernelCDF f μ ν} + {μ : kernel α (β × ℝ)} {ν : kernel α β} {hf : IsCondKernelCDF f μ ν} lemma StieltjesFunction.measurable_measure {f : α → StieltjesFunction} (hf : ∀ q, Measurable fun a ↦ f a q) @@ -245,20 +259,20 @@ lemma StieltjesFunction.measurable_measure {f : α → StieltjesFunction} exact Measurable.ennreal_tsum hf_cd_meas noncomputable -def IsKernelCDF.toKernel (f : α × β → StieltjesFunction) (hf : IsKernelCDF f μ ν) : +def IsCondKernelCDF.toKernel (f : α × β → StieltjesFunction) (hf : IsCondKernelCDF f μ ν) : kernel (α × β) ℝ where val p := (f p).measure property := StieltjesFunction.measurable_measure hf.measurable hf.tendsto_atBot_zero hf.tendsto_atTop_one -lemma IsKernelCDF.toKernel_apply (p : α × β) : hf.toKernel f p = (f p).measure := rfl +lemma IsCondKernelCDF.toKernel_apply (p : α × β) : hf.toKernel f p = (f p).measure := rfl instance instIsMarkovKernel_toKernel : IsMarkovKernel (hf.toKernel f) := ⟨fun _ ↦ (f _).isProbabilityMeasure (hf.tendsto_atBot_zero _) (hf.tendsto_atTop_one _)⟩ -lemma IsKernelCDF.toKernel_Iic (p : α × β) (x : ℝ) : +lemma IsCondKernelCDF.toKernel_Iic (p : α × β) (x : ℝ) : hf.toKernel f p (Iic x) = ENNReal.ofReal (f p x) := by - rw [IsKernelCDF.toKernel_apply p, (f p).measure_Iic (hf.tendsto_atBot_zero p)] + rw [IsCondKernelCDF.toKernel_apply p, (f p).measure_Iic (hf.tendsto_atBot_zero p)] simp end ToKernel @@ -267,15 +281,15 @@ section variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} {f : α × β → StieltjesFunction} {μ : kernel α (β × ℝ)} {ν : kernel α β} - {hf : IsKernelCDF f μ ν} + {hf : IsCondKernelCDF f μ ν} -lemma set_lintegral_toKernel_Iic [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) +lemma set_lintegral_toKernel_Iic [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : ∫⁻ t in s, hf.toKernel f (a, t) (Iic x) ∂(ν a) = μ a (s ×ˢ Iic x) := by - simp_rw [IsKernelCDF.toKernel_Iic] + simp_rw [IsCondKernelCDF.toKernel_Iic] exact hf.set_lintegral _ hs _ -lemma set_lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) +lemma set_lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) (a : α) {s : Set β} (hs : MeasurableSet s) : ∫⁻ t in s, hf.toKernel f (a, t) univ ∂(ν a) = μ a (s ×ˢ univ) := by have : ⋃ r : ℚ, Iic (r : ℝ) = univ := by @@ -299,11 +313,11 @@ lemma set_lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) · refine Monotone.directed_le fun i j hij t ↦ measure_mono (Iic_subset_Iic.mpr ?_) exact mod_cast hij -lemma lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) (a : α) : +lemma lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) (a : α) : ∫⁻ t, hf.toKernel f (a, t) univ ∂(ν a) = μ a univ := by rw [← set_lintegral_univ, set_lintegral_toKernel_univ hf a MeasurableSet.univ, univ_prod_univ] -lemma set_lintegral_toKernel_prod [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) +lemma set_lintegral_toKernel_prod [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set ℝ} (ht : MeasurableSet t) : ∫⁻ x in s, hf.toKernel f (a, x) t ∂(ν a) = μ a (s ×ˢ t) := by -- `set_lintegral_toKernel_Iic` gives the result for `t = Iic x`. These sets form a @@ -343,7 +357,7 @@ lemma set_lintegral_toKernel_prod [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) · exact fun i ↦ ((kernel.measurable_coe _ (hf_meas i)).comp measurable_prod_mk_left).aemeasurable.restrict -lemma lintegral_toKernel_mem [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) +lemma lintegral_toKernel_mem [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) (a : α) {s : Set (β × ℝ)} (hs : MeasurableSet s) : ∫⁻ x, hf.toKernel f (a, x) {y | (x, y) ∈ s} ∂(ν a) = μ a s := by -- `set_lintegral_toKernel_prod` gives the result for sets of the form `t₁ × t₂`. These @@ -416,12 +430,12 @@ lemma lintegral_toKernel_mem [IsFiniteKernel μ] (hf : IsKernelCDF f μ ν) _ = μ a (iUnion f') := (measure_iUnion hf_disj hf_meas).symm lemma kernel.eq_compProd_toKernel [IsFiniteKernel μ] [IsSFiniteKernel ν] - (hf : IsKernelCDF f μ ν) : + (hf : IsCondKernelCDF f μ ν) : μ = ν ⊗ₖ hf.toKernel f := by ext a s hs rw [kernel.compProd_apply _ _ _ hs, lintegral_toKernel_mem hf a hs] -lemma ae_toKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsKernelCDF f μ ν) (a : α) +lemma ae_toKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsCondKernelCDF f μ ν) (a : α) {s : Set ℝ} (hs : MeasurableSet s) (hμs : μ a {x | x.snd ∈ sᶜ} = 0) : ∀ᵐ t ∂(ν a), hf.toKernel f (a, t) s = 1 := by have h_eq : μ = ν ⊗ₖ hf.toKernel f := kernel.eq_compProd_toKernel hf @@ -439,7 +453,7 @@ lemma ae_toKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsKernel change hf.toKernel f (a, t) sᶜ = 0 at ht rwa [prob_compl_eq_zero_iff hs] at ht -lemma measurableSet_toKernel_eq_one (hf : IsKernelCDF f μ ν) {s : Set ℝ} (hs : MeasurableSet s) : +lemma measurableSet_toKernel_eq_one (hf : IsCondKernelCDF f μ ν) {s : Set ℝ} (hs : MeasurableSet s) : MeasurableSet {p | hf.toKernel f p s = 1} := (kernel.measurable_coe _ hs) (measurableSet_singleton 1) From d68617d1c23c9152812cd70c733f659c633564ef Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Tue, 27 Feb 2024 10:02:42 +0100 Subject: [PATCH 081/129] file name, docstring --- Mathlib.lean | 2 +- .../{KernelCdf.lean => CdfToKernel.lean} | 11 ++++++++++- .../Probability/Kernel/Disintegration/CondCdf.lean | 2 +- .../Kernel/Disintegration/CondKernelCdf.lean | 2 +- 4 files changed, 13 insertions(+), 4 deletions(-) rename Mathlib/Probability/Kernel/Disintegration/{KernelCdf.lean => CdfToKernel.lean} (97%) diff --git a/Mathlib.lean b/Mathlib.lean index b537073f27a3b..1d3e35187e70e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3138,11 +3138,11 @@ import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Kernel.CondDistrib import Mathlib.Probability.Kernel.Condexp import Mathlib.Probability.Kernel.Disintegration.Basic +import Mathlib.Probability.Kernel.Disintegration.CdfToKernel import Mathlib.Probability.Kernel.Disintegration.CondCdf import Mathlib.Probability.Kernel.Disintegration.CondKernelCdf import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.Disintegration.Integral -import Mathlib.Probability.Kernel.Disintegration.KernelCdf import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes import Mathlib.Probability.Kernel.Disintegration.Unique import Mathlib.Probability.Kernel.IntegralCompProd diff --git a/Mathlib/Probability/Kernel/Disintegration/KernelCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean similarity index 97% rename from Mathlib/Probability/Kernel/Disintegration/KernelCdf.lean rename to Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean index 1aa10e9065a9f..66658d92e6443 100644 --- a/Mathlib/Probability/Kernel/Disintegration/KernelCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean @@ -8,7 +8,16 @@ import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes /-! -# Conditional cumulative distribution function of a Markov kernel +# Building a Markov kernel from a conditional cumulative distribution function + +Let `μ : kernel α (β × ℝ)` and `ν : kernel α β` be two finite kernels. +A function `f : α × β → StieltjesFunction` is called a conditional kernel CDF of `μ` with respect +to `ν` if it is measurable, tends to to 0 at -∞ and to 1 at +∞ for all `p : α × β`, +if `fun t ↦ f (a, t) x` is `(ν a)`-integrable for all `a : α` and `x : ℝ` and for all measurable +sets `s : Set β`, `∫ t in s, f (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal`. + +From such a function with property `hf : IsCondKernelCDF f μ ν`, we can build a `kernel (α × β) ℝ` +denoted by `hf.toKernel f` such that `μ = ν ⊗ₖ hf.toKernel f`. ## Main definitions diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index 4f7d98b83afc6..505fc327cc63b 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.MeasureTheory.Decomposition.RadonNikodym -import Mathlib.Probability.Kernel.Disintegration.KernelCdf +import Mathlib.Probability.Kernel.Disintegration.CdfToKernel import Mathlib.Probability.Kernel.Disintegration.AuxLemmas #align_import probability.kernel.cond_cdf from "leanprover-community/mathlib"@"3b88f4005dc2e28d42f974cc1ce838f0dafb39b8" diff --git a/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean index 02adcb44128e3..ed1f496d54b7d 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.Disintegration.KernelCdf +import Mathlib.Probability.Kernel.Disintegration.CdfToKernel import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.Disintegration.AuxLemmas From a6a0f911cf055a0b76ac34236327149f59dd9a22 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Tue, 27 Feb 2024 13:55:10 +0100 Subject: [PATCH 082/129] mostly variable names --- .../Kernel/Disintegration/Basic.lean | 37 +----- .../Kernel/Disintegration/CdfToKernel.lean | 121 ++++++++++-------- .../Kernel/Disintegration/CondCdf.lean | 67 +++------- 3 files changed, 89 insertions(+), 136 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 20ab39e944572..191968f78f0c7 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -168,7 +168,7 @@ lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω)) [IsFiniteKerne rwa [he.injective he_eq] at h_mem conv_rhs => rw [this] unfold_let κ' - conv_rhs => rw [kernel.eq_compProd_toKernel hf] + conv_rhs => rw [← compProd_toKernel hf] change kernel.fst κ ⊗ₖ condKernelBorelSnd κ hf = kernel.comapRight (kernel.fst κ' ⊗ₖ hf.toKernel f) h_prod_embed ext c t ht : 2 @@ -209,18 +209,7 @@ lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω)) [IsFiniteKerne end BorelSnd -section StandardBorel - -noncomputable -def kernel.condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : kernel (α × γ) ℝ := - (isCondKernelCDF_condKernelCDF κ).toKernel _ - -instance (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : IsMarkovKernel (kernel.condKernelReal κ) := by - unfold kernel.condKernelReal; infer_instance - -lemma kernel.eq_compProd_condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - κ = kernel.fst κ ⊗ₖ kernel.condKernelReal κ := - kernel.eq_compProd_toKernel (isCondKernelCDF_condKernelCDF κ) +section CountablyGenerated noncomputable def condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : kernel (α × γ) Ω := @@ -238,28 +227,10 @@ lemma compProd_fst_condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel kernel.fst κ ⊗ₖ condKernelBorel κ = κ := by rw [condKernelBorel, compProd_fst_condKernelBorelSnd] -end StandardBorel +end CountablyGenerated section Unit -section Real - -noncomputable def condKernelUnitReal (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : - kernel (Unit × α) ℝ := - (isCondKernelCDF_condCDF (ρ ())).toKernel (fun (p : Unit × α) ↦ condCDF (ρ ()) p.2) - -instance (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : IsMarkovKernel (condKernelUnitReal ρ) := by - rw [condKernelUnitReal]; infer_instance - -lemma fst_compProd_condKernelUnitReal (ρ : kernel Unit (α × ℝ)) [IsFiniteKernel ρ] : - kernel.fst ρ ⊗ₖ condKernelUnitReal ρ = ρ := by - have : ρ = kernel.const Unit (ρ ()) := by ext; simp - conv_rhs => rw [this, kernel.eq_compProd_toKernel (isCondKernelCDF_condCDF (ρ ()))] - -end Real - -section BorelSnd - noncomputable def condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : kernel (Unit × α) Ω := let f := measurableEmbedding_real Ω @@ -276,8 +247,6 @@ lemma compProd_fst_condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKe kernel.fst κ ⊗ₖ condKernelUnitBorel κ = κ := by rw [condKernelUnitBorel, compProd_fst_condKernelBorelSnd] -end BorelSnd - end Unit section Measure diff --git a/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean index 66658d92e6443..6c2d4ab72aa67 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean @@ -13,8 +13,8 @@ import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes Let `μ : kernel α (β × ℝ)` and `ν : kernel α β` be two finite kernels. A function `f : α × β → StieltjesFunction` is called a conditional kernel CDF of `μ` with respect to `ν` if it is measurable, tends to to 0 at -∞ and to 1 at +∞ for all `p : α × β`, -if `fun t ↦ f (a, t) x` is `(ν a)`-integrable for all `a : α` and `x : ℝ` and for all measurable -sets `s : Set β`, `∫ t in s, f (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal`. +if `fun b ↦ f (a, b) x` is `(ν a)`-integrable for all `a : α` and `x : ℝ` and for all measurable +sets `s : Set β`, `∫ b in s, f (a, b) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal`. From such a function with property `hf : IsCondKernelCDF f μ ν`, we can build a `kernel (α × β) ℝ` denoted by `hf.toKernel f` such that `μ = ν ⊗ₖ hf.toKernel f`. @@ -25,24 +25,23 @@ Let `μ : kernel α (β × ℝ)` and `ν : kernel α β`. * `ProbabilityTheory.IsCondKernelCDF`: a function `f : α × β → StieltjesFunction` is called a conditional kernel CDF of `μ` with respect to `ν` if it is measurable, tends to to 0 at -∞ and - to 1 at +∞ for all `p : α × β`, if `fun t ↦ f (a, t) x` is `(ν a)`-integrable for all `a : α` and + to 1 at +∞ for all `p : α × β`, if `fun b ↦ f (a, b) x` is `(ν a)`-integrable for all `a : α` and `x : ℝ` and for all measurable sets `s : Set β`, - `∫ t in s, f (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal`. + `∫ b in s, f (a, b) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal`. * `ProbabilityTheory.IsCondKernelCDF.toKernel`: from a function `f : α × β → StieltjesFunction` with the property `hf : IsCondKernelCDF f μ ν`, build a `kernel (α × β) ℝ` such that `μ = ν ⊗ₖ hf.toKernel f`. * `ProbabilityTheory.IsRatCondKernelCDF`: a function `f : α × β → ℚ → ℝ` is called a rational conditional kernel CDF of `μ` with respect to `ν` if is measurable and satisfies the same - integral conditions as in the definition of `IsCondKernelCDF`, and the `ℚ → ℝ` function `f (a, x)` - satisfies the properties of a Sieltjes function for `(ν a)`-almost all `x : β`. + integral conditions as in the definition of `IsCondKernelCDF`, and the `ℚ → ℝ` function `f (a, b)` + satisfies the properties of a Sieltjes function for `(ν a)`-almost all `b : β`. ## Main statements * `ProbabilityTheory.isCondKernelCDF_stieltjesOfMeasurableRat`: if `f : α × β → ℚ → ℝ` has the property `IsRatCondKernelCDF`, then `stieltjesOfMeasurableRat f` is a function `α × β → StieltjesFunction` with the property `IsCondKernelCDF`. -* `ProbabilityTheory.kernel.eq_compProd_toKernel`: for `hf : IsCondKernelCDF f μ ν`, - `μ = ν ⊗ₖ hf.toKernel f` +* `ProbabilityTheory.compProd_toKernel`: for `hf : IsCondKernelCDF f μ ν`, `ν ⊗ₖ hf.toKernel f = μ`. -/ @@ -61,26 +60,26 @@ variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} structure IsRatCondKernelCDF (f : α × β → ℚ → ℝ) (μ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := (measurable : Measurable f) - (isRatStieltjesPoint_ae (a : α) : ∀ᵐ t ∂(ν a), IsRatStieltjesPoint f (a, t)) - (integrable (a : α) (q : ℚ) : Integrable (fun t ↦ f (a, t) q) (ν a)) + (isRatStieltjesPoint_ae (a : α) : ∀ᵐ b ∂(ν a), IsRatStieltjesPoint f (a, b)) + (integrable (a : α) (q : ℚ) : Integrable (fun b ↦ f (a, b) q) (ν a)) (set_integral (a : α) {s : Set β} (_hs : MeasurableSet s) (q : ℚ) : - ∫ t in s, f (a, t) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal) + ∫ b in s, f (a, b) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal) lemma stieltjesOfMeasurableRat_ae_eq (hf : IsRatCondKernelCDF f μ ν) (a : α) (q : ℚ) : - (fun t ↦ stieltjesOfMeasurableRat f hf.measurable (a, t) q) =ᵐ[ν a] fun t ↦ f (a, t) q := by + (fun b ↦ stieltjesOfMeasurableRat f hf.measurable (a, b) q) =ᵐ[ν a] fun b ↦ f (a, b) q := by filter_upwards [hf.isRatStieltjesPoint_ae a] with a ha rw [stieltjesOfMeasurableRat_eq, toRatCDF_of_isRatStieltjesPoint ha] lemma set_integral_stieltjesOfMeasurableRat_rat (hf : IsRatCondKernelCDF f μ ν) (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : - ∫ t in s, stieltjesOfMeasurableRat f hf.measurable (a, t) q ∂(ν a) + ∫ b in s, stieltjesOfMeasurableRat f hf.measurable (a, b) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal := by - rw [set_integral_congr_ae hs (g := fun t ↦ f (a, t) q) ?_, hf.set_integral a hs] + rw [set_integral_congr_ae hs (g := fun b ↦ f (a, b) q) ?_, hf.set_integral a hs] filter_upwards [stieltjesOfMeasurableRat_ae_eq hf a q] with b hb using fun _ ↦ hb lemma set_lintegral_stieltjesOfMeasurableRat_rat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ t in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) q) ∂(ν a) + ∫⁻ b in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, b) q) ∂(ν a) = μ a (s ×ˢ Iic (q : ℝ)) := by rw [← ofReal_integral_eq_lintegral_ofReal] · rw [set_integral_stieltjesOfMeasurableRat_rat hf a q hs, ENNReal.ofReal_toReal] @@ -92,7 +91,7 @@ lemma set_lintegral_stieltjesOfMeasurableRat_rat [IsFiniteKernel μ] (hf : IsRat lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ t in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) x) ∂(ν a) + ∫⁻ b in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, b) x) ∂(ν a) = μ a (s ×ˢ Iic x) := by -- We have the result for `x : ℚ` thanks to `set_lintegral_stieltjesOfMeasurableRat_rat`. -- We use a monotone convergence argument to extend it to the reals. @@ -109,10 +108,10 @@ lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCond simpa [measure_ne_top] using this rw [← hf.set_integral a hs q] simp [hρ_zero] - have h : ∫⁻ t in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) x) ∂(ν a) - = ∫⁻ t in s, ⨅ r : { r' : ℚ // x < r' }, - ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) r) ∂(ν a) := by - congr with t : 1 + have h : ∫⁻ b in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, b) x) ∂(ν a) + = ∫⁻ b in s, ⨅ r : { r' : ℚ // x < r' }, + ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, b) r) ∂(ν a) := by + congr with b : 1 simp_rw [← measure_stieltjesOfMeasurableRat_Iic] rw [← measure_iInter_eq_iInf] · congr with y : 1 @@ -134,7 +133,7 @@ lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCond · intro b rw [set_lintegral_stieltjesOfMeasurableRat_rat hf a _ hs] exact measure_ne_top _ _ - · refine Monotone.directed_ge fun i j hij t ↦ ?_ + · refine Monotone.directed_ge fun i j hij b ↦ ?_ simp_rw [← measure_stieltjesOfMeasurableRat_Iic] refine measure_mono (Iic_subset_Iic.mpr ?_) exact mod_cast hij @@ -154,15 +153,15 @@ lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCond lemma lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) (a : α) (x : ℝ) : - ∫⁻ t, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) x) ∂(ν a) + ∫⁻ b, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, b) x) ∂(ν a) = μ a (univ ×ˢ Iic x) := by rw [← set_lintegral_univ, set_lintegral_stieltjesOfMeasurableRat hf _ _ MeasurableSet.univ] lemma integrable_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) (a : α) (x : ℝ) : - Integrable (fun t ↦ stieltjesOfMeasurableRat f hf.measurable (a, t) x) (ν a) := by - have : (fun t ↦ stieltjesOfMeasurableRat f hf.measurable (a, t) x) - = fun t ↦ (ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, t) x)).toReal := by + Integrable (fun b ↦ stieltjesOfMeasurableRat f hf.measurable (a, b) x) (ν a) := by + have : (fun b ↦ stieltjesOfMeasurableRat f hf.measurable (a, b) x) + = fun b ↦ (ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, b) x)).toReal := by ext t rw [ENNReal.toReal_ofReal] exact stieltjesOfMeasurableRat_nonneg _ _ _ @@ -175,7 +174,7 @@ lemma integrable_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKer lemma set_integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : - ∫ t in s, stieltjesOfMeasurableRat f hf.measurable (a, t) x ∂(ν a) + ∫ b in s, stieltjesOfMeasurableRat f hf.measurable (a, b) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal := by rw [← ENNReal.ofReal_eq_ofReal_iff, ENNReal.ofReal_toReal] rotate_left @@ -188,7 +187,7 @@ lemma set_integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondK lemma integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) (a : α) (x : ℝ) : - ∫ t, stieltjesOfMeasurableRat f hf.measurable (a, t) x ∂(ν a) + ∫ b, stieltjesOfMeasurableRat f hf.measurable (a, b) x ∂(ν a) = (μ a (univ ×ˢ Iic x)).toReal := by rw [← integral_univ, set_integral_stieltjesOfMeasurableRat hf _ _ MeasurableSet.univ] @@ -202,11 +201,11 @@ variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} structure IsCondKernelCDF (f : α × β → StieltjesFunction) (μ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := (measurable (x : ℝ) : Measurable fun p ↦ f p x) - (integrable (a : α) (x : ℝ) : Integrable (fun t ↦ f (a, t) x) (ν a)) + (integrable (a : α) (x : ℝ) : Integrable (fun b ↦ f (a, b) x) (ν a)) (tendsto_atTop_one (p : α × β) : Tendsto (f p) atTop (𝓝 1)) (tendsto_atBot_zero (p : α × β) : Tendsto (f p) atBot (𝓝 0)) (set_integral (a : α) {s : Set β} (_hs : MeasurableSet s) (x : ℝ) : - ∫ t in s, f (a, t) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal) + ∫ b in s, f (a, b) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal) lemma IsCondKernelCDF.nonneg (hf : IsCondKernelCDF f μ ν) (p : α × β) (x : ℝ) : 0 ≤ f p x := Monotone.le_of_tendsto (f p).mono (hf.tendsto_atBot_zero p) x @@ -214,14 +213,24 @@ lemma IsCondKernelCDF.nonneg (hf : IsCondKernelCDF f μ ν) (p : α × β) (x : lemma IsCondKernelCDF.le_one (hf : IsCondKernelCDF f μ ν) (p : α × β) (x : ℝ) : f p x ≤ 1 := Monotone.ge_of_tendsto (f p).mono (hf.tendsto_atTop_one p) x +lemma IsCondKernelCDF.integral [IsFiniteKernel μ] + {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f μ ν) (a : α) (x : ℝ) : + ∫ b, f (a, b) x ∂(ν a) = (μ a (univ ×ˢ Iic x)).toReal := by + rw [← hf.set_integral _ MeasurableSet.univ, Measure.restrict_univ] + lemma IsCondKernelCDF.set_lintegral [IsFiniteKernel μ] {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f μ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (x : ℝ) : - ∫⁻ t in s, ENNReal.ofReal (f (a, t) x) ∂(ν a) = μ a (s ×ˢ Iic x) := by + ∫⁻ b in s, ENNReal.ofReal (f (a, b) x) ∂(ν a) = μ a (s ×ˢ Iic x) := by rw [← ofReal_integral_eq_lintegral_ofReal (hf.integrable a x).restrict (ae_of_all _ (fun _ ↦ hf.nonneg _ _)), hf.set_integral a hs x, ENNReal.ofReal_toReal] exact measure_ne_top _ _ +lemma IsCondKernelCDF.lintegral [IsFiniteKernel μ] + {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f μ ν) (a : α) (x : ℝ) : + ∫⁻ b, ENNReal.ofReal (f (a, b) x) ∂(ν a) = μ a (univ ×ˢ Iic x) := by + rw [← hf.set_lintegral _ MeasurableSet.univ, Measure.restrict_univ] + lemma isCondKernelCDF_stieltjesOfMeasurableRat {f : α × β → ℚ → ℝ} (hf : IsRatCondKernelCDF f μ ν) [IsFiniteKernel μ] : IsCondKernelCDF (stieltjesOfMeasurableRat f hf.measurable) μ ν where @@ -294,13 +303,13 @@ variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} lemma set_lintegral_toKernel_Iic [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ t in s, hf.toKernel f (a, t) (Iic x) ∂(ν a) = μ a (s ×ˢ Iic x) := by + ∫⁻ b in s, hf.toKernel f (a, b) (Iic x) ∂(ν a) = μ a (s ×ˢ Iic x) := by simp_rw [IsCondKernelCDF.toKernel_Iic] exact hf.set_lintegral _ hs _ lemma set_lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) (a : α) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ t in s, hf.toKernel f (a, t) univ ∂(ν a) = μ a (s ×ˢ univ) := by + ∫⁻ b in s, hf.toKernel f (a, b) univ ∂(ν a) = μ a (s ×ˢ univ) := by have : ⋃ r : ℚ, Iic (r : ℝ) = univ := by ext1 x simp only [mem_iUnion, mem_Iic, mem_univ, iff_true_iff] @@ -323,12 +332,12 @@ lemma set_lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ exact mod_cast hij lemma lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) (a : α) : - ∫⁻ t, hf.toKernel f (a, t) univ ∂(ν a) = μ a univ := by + ∫⁻ b, hf.toKernel f (a, b) univ ∂(ν a) = μ a univ := by rw [← set_lintegral_univ, set_lintegral_toKernel_univ hf a MeasurableSet.univ, univ_prod_univ] lemma set_lintegral_toKernel_prod [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set ℝ} (ht : MeasurableSet t) : - ∫⁻ x in s, hf.toKernel f (a, x) t ∂(ν a) = μ a (s ×ˢ t) := by + ∫⁻ b in s, hf.toKernel f (a, b) t ∂(ν a) = μ a (s ×ˢ t) := by -- `set_lintegral_toKernel_Iic` gives the result for `t = Iic x`. These sets form a -- π-system that generates the Borel σ-algebra, hence we can get the same equality for any -- measurable set `t`. @@ -337,11 +346,11 @@ lemma set_lintegral_toKernel_prod [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ · rintro t ⟨q, rfl⟩ exact set_lintegral_toKernel_Iic hf a _ hs · intro t ht ht_lintegral - calc ∫⁻ x in s, hf.toKernel f (a, x) tᶜ ∂(ν a) - = ∫⁻ x in s, hf.toKernel f (a, x) univ - hf.toKernel f (a, x) t ∂(ν a) := by + calc ∫⁻ b in s, hf.toKernel f (a, b) tᶜ ∂(ν a) + = ∫⁻ b in s, hf.toKernel f (a, b) univ - hf.toKernel f (a, b) t ∂(ν a) := by congr with x; rw [measure_compl ht (measure_ne_top (hf.toKernel f (a, x)) _)] - _ = ∫⁻ x in s, hf.toKernel f (a, x) univ ∂(ν a) - - ∫⁻ x in s, hf.toKernel f (a, x) t ∂(ν a) := by + _ = ∫⁻ b in s, hf.toKernel f (a, b) univ ∂(ν a) + - ∫⁻ b in s, hf.toKernel f (a, b) t ∂(ν a) := by rw [lintegral_sub] · exact (kernel.measurable_coe (hf.toKernel f) ht).comp measurable_prod_mk_left · rw [ht_lintegral] @@ -368,7 +377,7 @@ lemma set_lintegral_toKernel_prod [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ lemma lintegral_toKernel_mem [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) (a : α) {s : Set (β × ℝ)} (hs : MeasurableSet s) : - ∫⁻ x, hf.toKernel f (a, x) {y | (x, y) ∈ s} ∂(ν a) = μ a s := by + ∫⁻ b, hf.toKernel f (a, b) {y | (b, y) ∈ s} ∂(ν a) = μ a s := by -- `set_lintegral_toKernel_prod` gives the result for sets of the form `t₁ × t₂`. These -- sets form a π-system that generates the product σ-algebra, hence we can get the same equality -- for any measurable set `s`. @@ -397,10 +406,10 @@ lemma lintegral_toKernel_mem [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) rw [h_eq1, h_eq2, add_zero] exact set_lintegral_toKernel_prod hf a ht₁ ht₂ · intro t ht ht_eq - calc ∫⁻ x, hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ tᶜ} ∂(ν a) - = ∫⁻ x, hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t}ᶜ ∂(ν a) := rfl - _ = ∫⁻ x, hf.toKernel f (a, x) univ - - hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t} ∂(ν a) := by + calc ∫⁻ b, hf.toKernel f (a, b) {y : ℝ | (b, y) ∈ tᶜ} ∂(ν a) + = ∫⁻ b, hf.toKernel f (a, b) {y : ℝ | (b, y) ∈ t}ᶜ ∂(ν a) := rfl + _ = ∫⁻ b, hf.toKernel f (a, b) univ + - hf.toKernel f (a, b) {y : ℝ | (b, y) ∈ t} ∂(ν a) := by congr with x : 1 exact measure_compl (measurable_prod_mk_left ht) (measure_ne_top (hf.toKernel f (a, x)) _) @@ -429,27 +438,26 @@ lemma lintegral_toKernel_mem [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) intro h_mem_both suffices (a, x) ∈ ∅ by rwa [mem_empty_iff_false] at this rwa [← h_disj, mem_inter_iff] - calc ∫⁻ x, hf.toKernel f (a, x) (⋃ i, {y | (x, y) ∈ f' i}) ∂(ν a) - = ∫⁻ x, ∑' i, hf.toKernel f (a, x) {y | (x, y) ∈ f' i} ∂(ν a) := by + calc ∫⁻ b, hf.toKernel f (a, b) (⋃ i, {y | (b, y) ∈ f' i}) ∂(ν a) + = ∫⁻ b, ∑' i, hf.toKernel f (a, b) {y | (b, y) ∈ f' i} ∂(ν a) := by congr with x : 1 rw [measure_iUnion (h_disj x) fun i ↦ measurable_prod_mk_left (hf_meas i)] - _ = ∑' i, ∫⁻ x, hf.toKernel f (a, x) {y | (x, y) ∈ f' i} ∂(ν a) := + _ = ∑' i, ∫⁻ b, hf.toKernel f (a, b) {y | (b, y) ∈ f' i} ∂(ν a) := lintegral_tsum fun i ↦ (kernel.measurable_kernel_prod_mk_left' (hf_meas i) a).aemeasurable _ = ∑' i, μ a (f' i) := by simp_rw [hf_eq] _ = μ a (iUnion f') := (measure_iUnion hf_disj hf_meas).symm -lemma kernel.eq_compProd_toKernel [IsFiniteKernel μ] [IsSFiniteKernel ν] +lemma compProd_toKernel [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsCondKernelCDF f μ ν) : - μ = ν ⊗ₖ hf.toKernel f := by + ν ⊗ₖ hf.toKernel f = μ := by ext a s hs rw [kernel.compProd_apply _ _ _ hs, lintegral_toKernel_mem hf a hs] lemma ae_toKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsCondKernelCDF f μ ν) (a : α) {s : Set ℝ} (hs : MeasurableSet s) (hμs : μ a {x | x.snd ∈ sᶜ} = 0) : - ∀ᵐ t ∂(ν a), hf.toKernel f (a, t) s = 1 := by - have h_eq : μ = ν ⊗ₖ hf.toKernel f := kernel.eq_compProd_toKernel hf - have h : μ a {x | x.snd ∈ sᶜ} = (ν ⊗ₖ hf.toKernel f) a {x | x.snd ∈ sᶜ} := by - rw [← h_eq] + ∀ᵐ b ∂(ν a), hf.toKernel f (a, b) s = 1 := by + have h_eq : ν ⊗ₖ hf.toKernel f = μ := compProd_toKernel hf + have h : μ a {x | x.snd ∈ sᶜ} = (ν ⊗ₖ hf.toKernel f) a {x | x.snd ∈ sᶜ} := by rw [h_eq] rw [hμs, kernel.compProd_apply] at h swap; · exact measurable_snd hs.compl rw [eq_comm, lintegral_eq_zero_iff] at h @@ -458,11 +466,12 @@ lemma ae_toKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsCondKe change Measurable ((fun p ↦ hf.toKernel f p {c | c ∉ s}) ∘ (fun b : β ↦ (a, b))) exact (kernel.measurable_coe _ hs.compl).comp measurable_prod_mk_left simp only [mem_compl_iff, mem_setOf_eq, kernel.prodMkLeft_apply'] at h - filter_upwards [h] with t ht - change hf.toKernel f (a, t) sᶜ = 0 at ht - rwa [prob_compl_eq_zero_iff hs] at ht + filter_upwards [h] with b hb + change hf.toKernel f (a, b) sᶜ = 0 at hb + rwa [prob_compl_eq_zero_iff hs] at hb -lemma measurableSet_toKernel_eq_one (hf : IsCondKernelCDF f μ ν) {s : Set ℝ} (hs : MeasurableSet s) : +lemma measurableSet_toKernel_eq_one (hf : IsCondKernelCDF f μ ν) + {s : Set ℝ} (hs : MeasurableSet s) : MeasurableSet {p | hf.toKernel f p s = 1} := (kernel.measurable_coe _ hs) (measurableSet_singleton 1) diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index 505fc327cc63b..7b9b4c7845085 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -447,6 +447,12 @@ lemma condCDF_eq_stieltjesOfMeasurableRat_unit_prod (ρ : Measure (α × ℝ)) ( ext x rw [condCDF, ← stieltjesOfMeasurableRat_unit_prod] +lemma isCondKernelCDF_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : + IsCondKernelCDF (fun p : Unit × α ↦ condCDF ρ p.2) (kernel.const Unit ρ) + (kernel.const Unit ρ.fst) := by + simp_rw [condCDF_eq_stieltjesOfMeasurableRat_unit_prod ρ] + exact isCondKernelCDF_stieltjesOfMeasurableRat (isRatCondKernelCDF_preCDF ρ) + #noalign probability_theory.cond_cdf_eq_cond_cdf_rat /-- The conditional cdf is non-negative for all `a : α`. -/ @@ -487,70 +493,39 @@ theorem measurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : Measurable fun measurable_stieltjesOfMeasurableRat _ _ #align probability_theory.measurable_cond_cdf ProbabilityTheory.measurable_condCDF +/-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ +theorem stronglyMeasurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : + StronglyMeasurable fun a ↦ condCDF ρ a x := stronglyMeasurable_stieltjesOfMeasurableRat _ _ +#align probability_theory.strongly_measurable_cond_cdf ProbabilityTheory.stronglyMeasurable_condCDF + #noalign probability_theory.set_lintegral_cond_cdf_rat theorem set_lintegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) {s : Set α} (hs : MeasurableSet s) : - ∫⁻ a in s, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x) := by - have h := set_lintegral_stieltjesOfMeasurableRat (isRatCondKernelCDF_preCDF ρ) () x hs - simp only [kernel.const_apply] at h - rw [← h] - congr with a - congr - exact condCDF_eq_stieltjesOfMeasurableRat_unit_prod _ _ + ∫⁻ a in s, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x) := + (isCondKernelCDF_condCDF ρ).set_lintegral () hs x #align probability_theory.set_lintegral_cond_cdf ProbabilityTheory.set_lintegral_condCDF theorem lintegral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) : - ∫⁻ a, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (univ ×ˢ Iic x) := by - rw [← set_lintegral_univ, set_lintegral_condCDF ρ _ MeasurableSet.univ] + ∫⁻ a, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (univ ×ˢ Iic x) := + (isCondKernelCDF_condCDF ρ).lintegral () x #align probability_theory.lintegral_cond_cdf ProbabilityTheory.lintegral_condCDF -/-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ -theorem stronglyMeasurable_condCDF (ρ : Measure (α × ℝ)) (x : ℝ) : - StronglyMeasurable fun a ↦ condCDF ρ a x := stronglyMeasurable_stieltjesOfMeasurableRat _ _ -#align probability_theory.strongly_measurable_cond_cdf ProbabilityTheory.stronglyMeasurable_condCDF - theorem integrable_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) : - Integrable (fun a ↦ condCDF ρ a x) ρ.fst := by - refine integrable_of_forall_fin_meas_le _ (measure_lt_top ρ.fst univ) ?_ fun t _ _ ↦ ?_ - · exact (stronglyMeasurable_condCDF ρ _).aestronglyMeasurable - · have : ∀ y, (‖condCDF ρ y x‖₊ : ℝ≥0∞) ≤ 1 := by - intro y - rw [Real.nnnorm_of_nonneg (condCDF_nonneg _ _ _)] - -- Porting note: was exact_mod_cast condCDF_le_one _ _ _ - simp only [ENNReal.coe_le_one_iff] - exact condCDF_le_one _ _ _ - refine - (set_lintegral_mono (measurable_condCDF _ _).ennnorm measurable_one fun y _ ↦ this y).trans - ?_ - simp only [Pi.one_apply, lintegral_one, Measure.restrict_apply, MeasurableSet.univ, univ_inter] - exact measure_mono (subset_univ _) + Integrable (fun a ↦ condCDF ρ a x) ρ.fst := + (isCondKernelCDF_condCDF ρ).integrable () x #align probability_theory.integrable_cond_cdf ProbabilityTheory.integrable_condCDF theorem set_integral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) {s : Set α} - (hs : MeasurableSet s) : ∫ a in s, condCDF ρ a x ∂ρ.fst = (ρ (s ×ˢ Iic x)).toReal := by - have h := set_lintegral_condCDF ρ x hs - rw [← ofReal_integral_eq_lintegral_ofReal] at h - · rw [← h, ENNReal.toReal_ofReal] - exact integral_nonneg fun _ ↦ condCDF_nonneg _ _ _ - · exact (integrable_condCDF _ _).integrableOn - · exact eventually_of_forall fun _ ↦ condCDF_nonneg _ _ _ + (hs : MeasurableSet s) : ∫ a in s, condCDF ρ a x ∂ρ.fst = (ρ (s ×ˢ Iic x)).toReal := + (isCondKernelCDF_condCDF ρ).set_integral () hs x #align probability_theory.set_integral_cond_cdf ProbabilityTheory.set_integral_condCDF theorem integral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℝ) : - ∫ a, condCDF ρ a x ∂ρ.fst = (ρ (univ ×ˢ Iic x)).toReal := by - rw [← set_integral_condCDF ρ _ MeasurableSet.univ, Measure.restrict_univ] + ∫ a, condCDF ρ a x ∂ρ.fst = (ρ (univ ×ˢ Iic x)).toReal := + (isCondKernelCDF_condCDF ρ).integral () x #align probability_theory.integral_cond_cdf ProbabilityTheory.integral_condCDF -lemma isCondKernelCDF_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - IsCondKernelCDF (fun p : Unit × α ↦ condCDF ρ p.2) (kernel.const Unit ρ) - (kernel.const Unit ρ.fst) where - measurable x := (measurable_condCDF ρ x).comp measurable_snd - integrable _ x := integrable_condCDF ρ x - tendsto_atTop_one p := tendsto_condCDF_atTop ρ p.2 - tendsto_atBot_zero p := tendsto_condCDF_atBot ρ p.2 - set_integral _ _ hs x := set_integral_condCDF ρ x hs - #noalign probability_theory.measure_cond_cdf_Iic #noalign probability_theory.measure_cond_cdf_univ #noalign probability_theory.measurable_measure_cond_cdf From 6ae49323e820dcd99ffe11f244edaaf3e228045f Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Tue, 5 Mar 2024 08:57:38 +0100 Subject: [PATCH 083/129] more work --- Mathlib/MeasureTheory/Integral/Bochner.lean | 78 ++++++ Mathlib/MeasureTheory/Integral/Lebesgue.lean | 36 +++ .../Kernel/Disintegration/CondCdf.lean | 2 +- .../Kernel/Disintegration/CondKernelCdf.lean | 4 +- .../Kernel/Disintegration/Density.lean | 121 +-------- .../Kernel/Disintegration/DensityToCdf.lean | 245 ++++++++++++++++++ 6 files changed, 373 insertions(+), 113 deletions(-) create mode 100644 Mathlib/Probability/Kernel/Disintegration/DensityToCdf.lean diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index d85ddb442a50f..d5db395920d52 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1386,6 +1386,84 @@ lemma integral_tendsto_of_tendsto_of_antitone {μ : Measure α} {f : ℕ → α · filter_upwards [h_mono] with x hx n m hnm using neg_le_neg_iff.mpr <| hx hnm · filter_upwards [h_tendsto] with x hx using hx.neg +lemma tendsto_of_integral_tendsto_of_monotone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} + (hf_int : ∀ n, Integrable (f n) μ) (hF_int : Integrable F μ) + (hf_tendsto : Tendsto (fun i ↦ ∫ a, f i a ∂μ) atTop (𝓝 (∫ a, F a ∂μ))) + (hf_mono : ∀ᵐ a ∂μ, Monotone (fun i ↦ f i a)) + (hf_bound : ∀ᵐ a ∂μ, ∀ i, f i a ≤ F a) : + ∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) atTop (𝓝 (F a)) := by + -- reduce to the `ℝ≥0∞` case + let f' : ℕ → α → ℝ≥0∞ := fun n a ↦ ENNReal.ofReal (f n a - f 0 a) + let F' : α → ℝ≥0∞ := fun a ↦ ENNReal.ofReal (F a - f 0 a) + have hf'_int_eq : ∀ i, ∫⁻ a, f' i a ∂μ = ENNReal.ofReal (∫ a, f i a ∂μ - ∫ a, f 0 a ∂μ) := by + intro i + unfold_let f' + rw [← ofReal_integral_eq_lintegral_ofReal, integral_sub (hf_int i) (hf_int 0)] + · exact (hf_int i).sub (hf_int 0) + · filter_upwards [hf_mono] with a h_mono + simp [h_mono (zero_le i)] + have hF'_int_eq : ∫⁻ a, F' a ∂μ = ENNReal.ofReal (∫ a, F a ∂μ - ∫ a, f 0 a ∂μ) := by + unfold_let F' + rw [← ofReal_integral_eq_lintegral_ofReal, integral_sub hF_int (hf_int 0)] + · exact hF_int.sub (hf_int 0) + · filter_upwards [hf_bound] with a h_bound + simp [h_bound 0] + have h_tendsto : Tendsto (fun i ↦ ∫⁻ a, f' i a ∂μ) atTop (𝓝 (∫⁻ a, F' a ∂μ)) := by + simp_rw [hf'_int_eq, hF'_int_eq] + refine (ENNReal.continuous_ofReal.tendsto _).comp ?_ + rwa [tendsto_sub_const_iff] + have h_mono : ∀ᵐ a ∂μ, Monotone (fun i ↦ f' i a) := by + filter_upwards [hf_mono] with a ha_mono i j hij + refine ENNReal.ofReal_le_ofReal ?_ + simp [ha_mono hij] + have h_bound : ∀ᵐ a ∂μ, ∀ i, f' i a ≤ F' a := by + filter_upwards [hf_bound] with a ha_bound i + refine ENNReal.ofReal_le_ofReal ?_ + simp only [tsub_le_iff_right, sub_add_cancel, ha_bound i] + -- use the corresponding lemma for `ℝ≥0∞` + have h := tendsto_of_lintegral_tendsto_of_monotone ?_ ?_ h_tendsto h_mono h_bound ?_ + rotate_left + · exact fun n ↦ ((hf_int n).1.aemeasurable.sub (hf_int 0).1.aemeasurable).ennreal_ofReal + · exact (hF_int.1.aemeasurable.sub (hf_int 0).1.aemeasurable).ennreal_ofReal + · exact ((lintegral_ofReal_le_lintegral_nnnorm _).trans_lt (hF_int.sub (hf_int 0)).2).ne + filter_upwards [h, hf_mono, hf_bound] with a ha ha_mono ha_bound + have h1 : (fun i ↦ f i a) = fun i ↦ (f' i a).toReal + f 0 a := by + unfold_let f' + ext i + rw [ENNReal.toReal_ofReal] + · abel + · simp [ha_mono (zero_le i)] + have h2 : F a = (F' a).toReal + f 0 a := by + unfold_let F' + rw [ENNReal.toReal_ofReal] + · abel + · simp [ha_bound 0] + rw [h1, h2] + refine Filter.Tendsto.add ?_ tendsto_const_nhds + exact (ENNReal.continuousAt_toReal ENNReal.ofReal_ne_top).tendsto.comp ha + +lemma tendsto_of_integral_tendsto_of_antitone {μ : Measure α} {f : ℕ → α → ℝ} {F : α → ℝ} + (hf_int : ∀ n, Integrable (f n) μ) (hF_int : Integrable F μ) + (hf_tendsto : Tendsto (fun i ↦ ∫ a, f i a ∂μ) atTop (𝓝 (∫ a, F a ∂μ))) + (hf_mono : ∀ᵐ a ∂μ, Antitone (fun i ↦ f i a)) + (hf_bound : ∀ᵐ a ∂μ, ∀ i, F a ≤ f i a) : + ∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) atTop (𝓝 (F a)) := by + let f' : ℕ → α → ℝ := fun i a ↦ - f i a + let F' : α → ℝ := fun a ↦ - F a + suffices ∀ᵐ a ∂μ, Tendsto (fun i ↦ f' i a) atTop (𝓝 (F' a)) by + filter_upwards [this] with a ha_tendsto + convert ha_tendsto.neg + · simp [f'] + · simp [F'] + refine tendsto_of_integral_tendsto_of_monotone (fun n ↦ (hf_int n).neg) hF_int.neg ?_ ?_ ?_ + · convert hf_tendsto.neg + · rw [integral_neg] + · rw [integral_neg] + · filter_upwards [hf_mono] with a ha i j hij + simp [f', ha hij] + · filter_upwards [hf_bound] with a ha i + simp [f', F', ha i] + section NormedAddCommGroup variable {H : Type*} [NormedAddCommGroup H] diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index c75ac1dca4ac7..9bdc2f6c20422 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -1641,6 +1641,42 @@ theorem _root_.IsFiniteMeasure.lintegral_lt_top_of_bounded_to_ennreal {α : Type exact ENNReal.mul_lt_top ENNReal.coe_lt_top.ne μ_fin.measure_univ_lt_top.ne #align is_finite_measure.lintegral_lt_top_of_bounded_to_ennreal IsFiniteMeasure.lintegral_lt_top_of_bounded_to_ennreal +lemma tendsto_of_lintegral_tendsto_of_monotone {α : Type*} {mα : MeasurableSpace α} + {f : ℕ → α → ℝ≥0∞} {F : α → ℝ≥0∞} {μ : Measure α} + (hf_meas : ∀ n, AEMeasurable (f n) μ) (hF_meas : AEMeasurable F μ) + (hf_tendsto : Tendsto (fun i ↦ ∫⁻ a, f i a ∂μ) atTop (𝓝 (∫⁻ a, F a ∂μ))) + (hf_mono : ∀ᵐ a ∂μ, Monotone (fun i ↦ f i a)) + (h_bound : ∀ᵐ a ∂μ, ∀ i, f i a ≤ F a) (h_int_finite : ∫⁻ a, F a ∂μ ≠ ∞) : + ∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) atTop (𝓝 (F a)) := by + have h_bound_finite : ∀ᵐ a ∂μ, F a ≠ ∞ := by + filter_upwards [ae_lt_top' hF_meas h_int_finite] with a ha using ha.ne + have h_exists : ∀ᵐ a ∂μ, ∃ l, Tendsto (fun i ↦ f i a) atTop (𝓝 l) := by + filter_upwards [h_bound, h_bound_finite, hf_mono] with a h_le h_fin h_mono + have h_tendsto : Tendsto (fun i ↦ f i a) atTop atTop ∨ + ∃ l, Tendsto (fun i ↦ f i a) atTop (𝓝 l) := tendsto_of_monotone h_mono + cases' h_tendsto with h_absurd h_tendsto + · rw [tendsto_atTop_atTop_iff_of_monotone h_mono] at h_absurd + obtain ⟨i, hi⟩ := h_absurd (F a + 1) + refine absurd (hi.trans (h_le _)) (not_le.mpr ?_) + exact ENNReal.lt_add_right h_fin one_ne_zero + · exact h_tendsto + classical + let F' : α → ℝ≥0∞ := fun a ↦ if h : ∃ l, Tendsto (fun i ↦ f i a) atTop (𝓝 l) + then h.choose else ∞ + have hF'_tendsto : ∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) atTop (𝓝 (F' a)) := by + filter_upwards [h_exists] with a ha + simp_rw [dif_pos ha] + exact ha.choose_spec + suffices F' =ᵐ[μ] F by + filter_upwards [this, hF'_tendsto] with a h_eq h_tendsto using h_eq ▸ h_tendsto + have hF'_le : F' ≤ᵐ[μ] F := by + filter_upwards [h_bound, hF'_tendsto] with a h_le h_tendsto + exact le_of_tendsto' h_tendsto (fun m ↦ h_le _) + suffices ∫⁻ a, F' a ∂μ = ∫⁻ a, F a ∂μ by + exact ae_eq_of_ae_le_of_lintegral_le hF'_le (this ▸ h_int_finite) hF_meas this.symm.le + refine tendsto_nhds_unique ?_ hf_tendsto + exact lintegral_tendsto_of_tendsto_of_monotone hf_meas hf_mono hF'_tendsto + end Lintegral open MeasureTheory.SimpleFunc diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index 7b9b4c7845085..64038aded22bd 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -45,7 +45,7 @@ namespace MeasureTheory.Measure variable {α β : Type*} {mα : MeasurableSpace α} (ρ : Measure (α × ℝ)) -/-- Measure on `α` such that for a measurable set `s`, `ρ.Iic_snd r s = ρ (s ×ˢ Iic r)`. -/ +/-- Measure on `α` such that for a measurable set `s`, `ρ.IicSnd r s = ρ (s ×ˢ Iic r)`. -/ noncomputable def IicSnd (r : ℝ) : Measure α := (ρ.restrict (univ ×ˢ Iic r)).fst #align measure_theory.measure.Iic_snd MeasureTheory.Measure.IicSnd diff --git a/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean index ed1f496d54b7d..4ec27d8512fae 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean @@ -191,9 +191,7 @@ lemma iInf_rat_gt_densityIic_eq (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel = kernel.densityIic κ ν a t q := by rw [ae_all_iff] refine fun q ↦ ae_eq_of_forall_set_integral_eq_of_sigmaFinite (μ := ν a) ?_ ?_ ?_ - · intro s _ _ - refine Integrable.integrableOn ?_ - exact integrable_iInf_rat_gt_densityIic hκν _ _ + · exact fun s _ _ ↦ (integrable_iInf_rat_gt_densityIic hκν _ _).integrableOn · exact fun s _ _ ↦ (integrable_densityIic hκν a q).integrableOn · intro s hs _ rw [set_integral_densityIic hκν _ _ hs, set_integral_iInf_rat_gt_densityIic hκν _ _ hs] diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 5b0accef399fa..5daedcf568f8f 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -661,61 +661,12 @@ lemma tendsto_density_atTop_ae_of_antitone (hκν : fst κ ≤ ν) [IsFiniteKern [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : ∀ᵐ t ∂(ν a), Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 0) := by - have h_anti : ∀ t, Antitone (fun m ↦ density κ ν a t (s m)) := - fun t n m hnm ↦ density_mono_set hκν a t (hs hnm) - have h_le_one : ∀ m t, density κ ν a t (s m) ≤ 1 := fun m t ↦ density_le_one hκν a t (s m) - -- for all `t`, `fun m ↦ density κ a (s m) t` has a limit - have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := by - intro t - have h_tendsto : Tendsto (fun m ↦ density κ ν a t (s m)) atTop atBot ∨ - ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := - tendsto_of_antitone (h_anti t) - cases' h_tendsto with h_absurd h_tendsto - · rw [tendsto_atTop_atBot_iff_of_antitone (h_anti t)] at h_absurd - obtain ⟨r, hr⟩ := h_absurd (-1) - have h_nonneg := density_nonneg hκν a t (s r) - linarith - · exact h_tendsto - -- let `F` be the pointwise limit of `fun m ↦ density κ a (s m) t` for all `t` - let F : γ → ℝ := fun t ↦ (h_exists t).choose - have hF_tendsto : ∀ t, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 (F t)) := - fun t ↦ (h_exists t).choose_spec - have hF_nonneg : ∀ t, 0 ≤ F t := - fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ density_nonneg hκν a t (s m)) - have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) - have hF_int : Integrable F (ν a) := by - rw [← memℒp_one_iff_integrable] - refine ⟨?_, ?_⟩ - · refine aestronglyMeasurable_of_tendsto_ae atTop (fun n ↦ ?_) (ae_of_all _ hF_tendsto) - exact (measurable_density_right κ ν (hs_meas _) a).aestronglyMeasurable - · rw [snorm_one_eq_lintegral_nnnorm] - calc ∫⁻ x, ‖F x‖₊ ∂(ν a) ≤ ∫⁻ _, 1 ∂(ν a) := by - refine lintegral_mono (fun x ↦ ?_) - rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, ENNReal.ofReal_le_one, - abs_of_nonneg (hF_nonneg _)] - exact hF_le_one _ - _ < ⊤ := by - simp only [MeasureTheory.lintegral_const, one_mul] - exact measure_lt_top _ _ - -- it suffices to show that the limit `F` is 0 a.e. - suffices F=ᵐ[ν a] 0 by - filter_upwards [this] with t ht_eq - simp only [Pi.zero_apply] at ht_eq - rw [← ht_eq] - exact hF_tendsto t - -- since `F` is nonnegative, proving that its integral is 0 is sufficient to get that - -- `F` is 0 a.e. - rw [← integral_eq_zero_iff_of_nonneg hF_nonneg hF_int] - have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop - (𝓝 (∫ t, F t ∂(ν a))) := by - refine integral_tendsto_of_tendsto_of_antitone ?_ hF_int ?_ ?_ - · exact fun n ↦ integrable_density hκν _ (hs_meas n) - · exact ae_of_all _ h_anti - · exact ae_of_all _ hF_tendsto - have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop - (𝓝 0) := by + refine tendsto_of_integral_tendsto_of_antitone ?_ (integrable_const _) ?_ ?_ ?_ + · exact fun n ↦ integrable_density hκν _ (hs_meas n) + · rw [integral_zero] exact tendsto_integral_density_of_antitone hκν a s hs hs_iInter hs_meas - exact tendsto_nhds_unique h_integral h_integral' + · exact ae_of_all _ (fun c n m hnm ↦ density_mono_set hκν a c (hs hnm)) + · exact ae_of_all _ (fun t m ↦ density_nonneg hκν a t (s m)) section UnivFst @@ -824,62 +775,14 @@ lemma tendsto_density_atTop_ae_of_monotone [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : ∀ᵐ t ∂(fst κ a), Tendsto (fun m ↦ density κ (fst κ) a t (s m)) atTop (𝓝 1) := by - let ν := fst κ - have h_mono : ∀ t, Monotone (fun m ↦ density κ (fst κ) a t (s m)) := - fun t n m hnm ↦ density_mono_set le_rfl a t (hs hnm) - have h_le_one : ∀ m t, density κ ν a t (s m) ≤ 1 := - fun m t ↦ density_le_one le_rfl a t (s m) - -- for all `t`, `fun m ↦ density κ a (s m) t` has a limit - have h_exists : ∀ t, ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := by - intro t - have h_tendsto : Tendsto (fun m ↦ density κ ν a t (s m)) atTop atTop ∨ - ∃ l, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 l) := - tendsto_of_monotone (h_mono t) - cases' h_tendsto with h_absurd h_tendsto - · rw [tendsto_atTop_atTop_iff_of_monotone (h_mono t)] at h_absurd - obtain ⟨r, hr⟩ := h_absurd 2 - exact absurd (hr.trans (h_le_one r t)) one_lt_two.not_le - · exact h_tendsto - -- let `F` be the pointwise limit of `fun m ↦ density κ a (s m) t` for all `t` - let F : γ → ℝ := fun t ↦ (h_exists t).choose - have hF_tendsto : ∀ t, Tendsto (fun m ↦ density κ ν a t (s m)) atTop (𝓝 (F t)) := - fun t ↦ (h_exists t).choose_spec - have hF_nonneg : ∀ t, 0 ≤ F t := - fun t ↦ ge_of_tendsto' (hF_tendsto t) (fun m ↦ density_nonneg le_rfl a t (s m)) - have hF_le_one : ∀ t, F t ≤ 1 := fun t ↦ le_of_tendsto' (hF_tendsto t) (fun m ↦ h_le_one m t) - have hF_int : Integrable F (ν a) := by - rw [← memℒp_one_iff_integrable] - constructor - · refine aestronglyMeasurable_of_tendsto_ae atTop (fun n ↦ ?_) (ae_of_all _ hF_tendsto) - exact (measurable_density_right κ ν (hs_meas _) a).aestronglyMeasurable - · rw [snorm_one_eq_lintegral_nnnorm] - calc ∫⁻ x, ‖F x‖₊ ∂(ν a) ≤ ∫⁻ _, 1 ∂(ν a) := by - refine lintegral_mono (fun x ↦ ?_) - rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, ENNReal.ofReal_le_one, - abs_of_nonneg (hF_nonneg _)] - exact hF_le_one _ - _ < ⊤ := by simp only [MeasureTheory.lintegral_const, measure_univ, one_mul, measure_lt_top] - -- it suffices to show that the limit `F` is 1 a.e. - suffices F =ᵐ[ν a] (fun _ ↦ 1) by - filter_upwards [this] with t ht_eq - rw [← ht_eq] - exact hF_tendsto t - -- since `F` is at most 1, proving that its integral is the same as the integral of 1 will tell - -- us that `F` is 1 a.e. - rw [← integral_eq_iff_of_ae_le hF_int (integrable_const _) (ae_of_all _ hF_le_one)] - have h_integral : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop - (𝓝 (∫ t, F t ∂(ν a))) := by - refine integral_tendsto_of_tendsto_of_monotone ?_ hF_int ?_ ?_ - · exact fun n ↦ integrable_density le_rfl _ (hs_meas n) - · exact ae_of_all _ h_mono - · exact ae_of_all _ hF_tendsto - have h_integral' : Tendsto (fun m : ℕ ↦ ∫ t, density κ ν a t (s m) ∂(ν a)) atTop - (𝓝 (∫ _, 1 ∂(ν a))) := by - rw [MeasureTheory.integral_const] - simp only [smul_eq_mul, mul_one] + refine tendsto_of_integral_tendsto_of_monotone ?_ (integrable_const _) ?_ ?_ ?_ + · exact fun n ↦ integrable_density le_rfl _ (hs_meas n) + · rw [MeasureTheory.integral_const, smul_eq_mul, mul_one] + convert tendsto_integral_density_of_monotone (κ := κ) le_rfl a s hs hs_iUnion hs_meas rw [fst_apply' _ _ MeasurableSet.univ] - exact tendsto_integral_density_of_monotone le_rfl a s hs hs_iUnion hs_meas - exact tendsto_nhds_unique h_integral h_integral' + rfl + · exact ae_of_all _ (fun c n m hnm ↦ density_mono_set le_rfl a c (hs hnm)) + · exact ae_of_all _ (fun t m ↦ density_le_one le_rfl a t (s m)) end UnivFst diff --git a/Mathlib/Probability/Kernel/Disintegration/DensityToCdf.lean b/Mathlib/Probability/Kernel/Disintegration/DensityToCdf.lean new file mode 100644 index 0000000000000..91f417b1faa21 --- /dev/null +++ b/Mathlib/Probability/Kernel/Disintegration/DensityToCdf.lean @@ -0,0 +1,245 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Probability.Kernel.Disintegration.CdfToKernel +import Mathlib.Probability.Kernel.Disintegration.Density +import Mathlib.Probability.Kernel.Disintegration.AuxLemmas + +/-! +# Kernel CDF + +We prove IsRatCondKernelCDF from conditions on integrals. + +## Main definitions + +* `FooBar` + +## Main statements + +* `fooBar_unique` + +## Implementation details + + +## References + +-/ + + +open MeasureTheory Set Filter + +open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory + +namespace ProbabilityTheory + +open ProbabilityTheory.kernel + +variable {α γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} + {κ : kernel α (γ × ℝ)} {ν : kernel α γ} + {f : α × γ → ℚ → ℝ} + +structure IsKernelPDF (f : α × γ → ℚ → ℝ) (κ : kernel α (γ × ℝ)) (ν : kernel α γ) : Prop := + (measurable : Measurable f) + (mono' (a : α) {q r : ℚ} (_hqr : q ≤ r) : ∀ᵐ c ∂(ν a), f (a, c) q ≤ f (a, c) r) + (nonneg' (a : α) (q : ℚ) : ∀ᵐ c ∂(ν a), 0 ≤ f (a, c) q) + (le_one' (a : α) (q : ℚ) : ∀ᵐ c ∂(ν a), f (a, c) q ≤ 1) + (tendsto_integral_zero_of_antitone (a : α) (s : ℕ → ℚ) (_hs : Antitone s) + (_hs_tendsto : Tendsto s atTop atBot) : + Tendsto (fun m ↦ ∫ c, f (a, c) (s m) ∂(ν a)) atTop (𝓝 0)) + (tendsto_integral_one_of_monotone (a : α) (s : ℕ → ℚ) (_hs : Monotone s) + (_hs_tendsto : Tendsto s atTop atTop) : + Tendsto (fun m ↦ ∫ c, f (a, c) (s m) ∂(ν a)) atTop (𝓝 (ν a univ).toReal)) + (integrable (a : α) (q : ℚ) : Integrable (fun c ↦ f (a, c) q) (ν a)) + (set_integral (a : α) {A : Set γ} (_hA : MeasurableSet A) (q : ℚ) : + ∫ c in A, f (a, c) q ∂(ν a) = (κ a (A ×ˢ Iic ↑q)).toReal) + +lemma IsKernelPDF.measurable_right (hf : IsKernelPDF f κ ν) (a : α) (q : ℚ) : + Measurable (fun t ↦ f (a, t) q) := by + let h := hf.measurable + rw [measurable_pi_iff] at h + exact (h q).comp measurable_prod_mk_left + +lemma IsKernelPDF.mono (hf : IsKernelPDF f κ ν) (a : α) : + ∀ᵐ c ∂(ν a), ∀ {q r : ℚ} (_ : q ≤ r), f (a, c) q ≤ f (a, c) r := by + simp_rw [ae_all_iff] + intro q r hqr + exact hf.mono' a hqr + +lemma IsKernelPDF.nonneg (hf : IsKernelPDF f κ ν) (a : α) : + ∀ᵐ c ∂(ν a), ∀ (q : ℚ), 0 ≤ f (a, c) q := by + rw [ae_all_iff] + exact hf.nonneg' a + +lemma IsKernelPDF.le_one (hf : IsKernelPDF f κ ν) (a : α) : + ∀ᵐ c ∂(ν a), ∀ (q : ℚ), f (a, c) q ≤ 1 := by + rw [ae_all_iff] + exact hf.le_one' a + +lemma IsKernelPDF.tendsto_zero_of_antitone (hf : IsKernelPDF f κ ν) [IsFiniteKernel ν] (a : α) + (s : ℕ → ℚ) (hs : Antitone s) (hs_tendsto : Tendsto s atTop atBot) : + ∀ᵐ c ∂(ν a), Tendsto (fun m ↦ f (a, c) (s m)) atTop (𝓝 0) := by + refine tendsto_of_integral_tendsto_of_antitone ?_ (integrable_const _) ?_ ?_ ?_ + · exact fun n ↦ hf.integrable a (s n) + · rw [integral_zero] + exact hf.tendsto_integral_zero_of_antitone a s hs hs_tendsto + · filter_upwards [hf.mono a] with t ht + exact fun n m hnm ↦ ht (hs hnm) + · filter_upwards [hf.nonneg a] with c hc using fun i ↦ hc (s i) + +lemma IsKernelPDF.tendsto_one_of_monotone (hf : IsKernelPDF f κ ν) [IsFiniteKernel ν] (a : α) + (s : ℕ → ℚ) (hs : Monotone s) (hs_tendsto : Tendsto s atTop atTop) : + ∀ᵐ c ∂(ν a), Tendsto (fun m ↦ f (a, c) (s m)) atTop (𝓝 1) := by + refine tendsto_of_integral_tendsto_of_monotone ?_ (integrable_const _) ?_ ?_ ?_ + · exact fun n ↦ hf.integrable a (s n) + · rw [MeasureTheory.integral_const, smul_eq_mul, mul_one] + exact hf.tendsto_integral_one_of_monotone a s hs hs_tendsto + · filter_upwards [hf.mono a] with t ht + exact fun n m hnm ↦ ht (hs hnm) + · filter_upwards [hf.le_one a] with c hc using fun i ↦ hc (s i) + +lemma tendsto_atTop_densityIic (hf : IsKernelPDF f κ ν) [IsFiniteKernel ν] (a : α) : + ∀ᵐ t ∂(ν a), Tendsto (fun q : ℚ ↦ f (a, t) q) atTop (𝓝 1) := by + suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ f (a, t) n) atTop (𝓝 1) by + filter_upwards [this, hf.mono a] with t ht h_mono + let f' := fun q : ℚ ↦ f (a, t) ⌊q⌋₊ + let g := fun q : ℚ ↦ f (a, t) ⌈q⌉₊ + have hf_le : ∀ᶠ q in atTop, f' q ≤ f (a, t) q := by + simp only [eventually_atTop, ge_iff_le] + exact ⟨0, fun q hq ↦ h_mono (Nat.floor_le hq)⟩ + have hg_le : ∀ q : ℚ, f (a, t) q ≤ g q := by + exact fun q ↦ h_mono (Nat.le_ceil _) + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ hf_le (eventually_of_forall hg_le) + · exact ht.comp tendsto_nat_floor_atTop + · exact ht.comp tendsto_nat_ceil_atTop + let s : ℕ → ℚ := fun n ↦ n + have hs : Monotone s := fun i j hij ↦ by simp [s, hij] + have hs_tendsto : Tendsto s atTop atTop := by + exact tendsto_nat_cast_atTop_atTop + filter_upwards [hf.tendsto_one_of_monotone a s hs hs_tendsto] with x hx using hx + +lemma tendsto_atBot_densityIic (hf : IsKernelPDF f κ ν) [IsFiniteKernel ν] (a : α) : + ∀ᵐ t ∂(ν a), Tendsto (fun q : ℚ ↦ f (a, t) q) atBot (𝓝 0) := by + suffices ∀ᵐ t ∂(ν a), Tendsto (fun q : ℚ ↦ f (a, t) (-q)) atTop (𝓝 0) by + filter_upwards [this] with t ht + have h_eq_neg : (fun q : ℚ ↦ f (a, t) q) = fun q : ℚ ↦ f (a, t) (- -q) := by + simp_rw [neg_neg] + rw [h_eq_neg] + convert ht.comp tendsto_neg_atBot_atTop + simp + suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ f (a, t) (-n)) atTop (𝓝 0) by + filter_upwards [this, hf.mono a] with t ht h_mono + let f' := fun q : ℚ ↦ f (a, t) (-⌊q⌋₊) + let g := fun q : ℚ ↦ f (a, t) (-⌈q⌉₊) + have hf_le : ∀ᶠ (q : ℚ) in atTop, f (a, t) (-q) ≤ f' q := by + simp only [eventually_atTop, ge_iff_le] + refine ⟨0, fun q hq ↦ ?_⟩ + norm_cast + refine h_mono ?_ + simp only [Rat.ofInt_eq_cast, Int.cast_neg, Int.cast_ofNat, neg_le_neg_iff] + exact Nat.floor_le hq + have hg_le : ∀ q, g q ≤ f (a, t) (-q) := by + refine fun q ↦ ?_ + simp only [g] + norm_cast + refine h_mono ?_ + simp only [Rat.ofInt_eq_cast, Int.cast_neg, Int.cast_ofNat, neg_le_neg_iff] + exact Nat.le_ceil _ + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ (eventually_of_forall hg_le) hf_le + · exact ht.comp tendsto_nat_ceil_atTop + · exact ht.comp tendsto_nat_floor_atTop + let s : ℕ → ℚ := fun n ↦ -n + have hs : Antitone s := fun i j hij ↦ neg_le_neg (by exact mod_cast hij) + have hs_tendsto : Tendsto s atTop atBot := by + simp only [s, tendsto_neg_atBot_iff] + exact tendsto_nat_cast_atTop_atTop + convert hf.tendsto_zero_of_antitone a s hs hs_tendsto with x n + +lemma bddBelow_range_densityIic (hf : IsKernelPDF f κ ν) (a : α) : + ∀ᵐ t ∂(ν a), ∀ q : ℚ, BddBelow (range fun (r : Ioi q) ↦ f (a, t) r) := by + filter_upwards [hf.nonneg a] with c hc + refine fun q ↦ ⟨0, ?_⟩ + rw [mem_lowerBounds] + rintro x ⟨y, rfl⟩ + exact hc y + +lemma integrable_iInf_rat_gt_densityIic (hf : IsKernelPDF f κ ν) [IsFiniteKernel ν] + (a : α) (q : ℚ) : + Integrable (fun t ↦ ⨅ r : Ioi q, f (a, t) r) (ν a) := by + rw [← memℒp_one_iff_integrable] + refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ + · exact measurable_iInf fun i ↦ hf.measurable_right a _ + refine (?_ : _ ≤ (ν a univ : ℝ≥0∞)).trans_lt (measure_lt_top _ _) + refine (snorm_le_of_ae_bound (C := 1) ?_).trans ?_ + · filter_upwards [bddBelow_range_densityIic hf a, hf.nonneg a, hf.le_one a] + with t hbdd_below h_nonneg h_le_one + rw [Real.norm_eq_abs, abs_of_nonneg] + · refine ciInf_le_of_le ?_ ?_ ?_ + · exact hbdd_below _ + · exact ⟨q + 1, by simp⟩ + · exact h_le_one _ + · exact le_ciInf fun r ↦ h_nonneg _ + · simp + +lemma set_integral_iInf_rat_gt_densityIic (hf : IsKernelPDF f κ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) (q : ℚ) {A : Set γ} (hA : MeasurableSet A) : + ∫ t in A, ⨅ r : Ioi q, f (a, t) r ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by + refine le_antisymm ?_ ?_ + · have h : ∀ r : Ioi q, ∫ t in A, ⨅ r' : Ioi q, f (a, t) r' ∂(ν a) + ≤ (κ a (A ×ˢ Iic (r : ℝ))).toReal := by + intro r + rw [← hf.set_integral a hA] + refine set_integral_mono_ae ?_ ?_ ?_ + · exact (integrable_iInf_rat_gt_densityIic hf _ _).integrableOn + · exact (hf.integrable _ _).integrableOn + · filter_upwards [bddBelow_range_densityIic hf a] with t ht + exact ciInf_le (ht _) r + calc ∫ t in A, ⨅ r : Ioi q, f (a, t) r ∂(ν a) + ≤ ⨅ r : Ioi q, (κ a (A ×ˢ Iic (r : ℝ))).toReal := le_ciInf h + _ = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by + rw [← Measure.iInf_rat_gt_prod_Iic hA q] + exact (ENNReal.toReal_iInf (fun r ↦ measure_ne_top _ _)).symm + · rw [← hf.set_integral a hA] + refine set_integral_mono_ae ?_ ?_ ?_ + · exact (hf.integrable _ _).integrableOn + · exact (integrable_iInf_rat_gt_densityIic hf _ _).integrableOn + · filter_upwards [hf.mono a] with c h_mono + exact le_ciInf (fun r ↦ h_mono (le_of_lt r.prop)) + +lemma iInf_rat_gt_densityIic_eq (hf : IsKernelPDF f κ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] + (a : α) : + ∀ᵐ t ∂(ν a), ∀ q : ℚ, ⨅ r : Ioi q, f (a, t) r = f (a, t) q := by + rw [ae_all_iff] + refine fun q ↦ ae_eq_of_forall_set_integral_eq_of_sigmaFinite (μ := ν a) ?_ ?_ ?_ + · intro s _ _ + refine Integrable.integrableOn ?_ + exact integrable_iInf_rat_gt_densityIic hf _ _ + · exact fun s _ _ ↦ (hf.integrable a _).integrableOn + · intro s hs _ + rw [hf.set_integral _ hs, set_integral_iInf_rat_gt_densityIic hf _ _ hs] + +lemma isRatStieltjesPoint_densityIic_ae (hf : IsKernelPDF f κ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) : + ∀ᵐ t ∂(ν a), IsRatStieltjesPoint f (a, t) := by + filter_upwards [tendsto_atTop_densityIic hf a, tendsto_atBot_densityIic hf a, + iInf_rat_gt_densityIic_eq hf a, hf.mono a] with t ht_top ht_bot ht_iInf h_mono + exact ⟨fun _ _ ↦ h_mono, ht_top, ht_bot, ht_iInf⟩ + +lemma isRatCondKernelCDF_densityIic (hf : IsKernelPDF f κ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] : + IsRatCondKernelCDF f κ ν where + measurable := hf.measurable + isRatStieltjesPoint_ae := isRatStieltjesPoint_densityIic_ae hf + integrable := hf.integrable + set_integral := hf.set_integral + +noncomputable +def condKernelCDF (hf : IsKernelPDF f κ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] : + α × γ → StieltjesFunction := + stieltjesOfMeasurableRat f (isRatCondKernelCDF_densityIic hf).measurable + +lemma isCondKernelCDF_condKernelCDF (hf : IsKernelPDF f κ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] : + IsCondKernelCDF (condKernelCDF hf) κ ν := + isCondKernelCDF_stieltjesOfMeasurableRat (isRatCondKernelCDF_densityIic hf) + +end ProbabilityTheory From b61e02499927476f37b152542667eb2318e0024b Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 6 Mar 2024 12:53:19 +0100 Subject: [PATCH 084/129] refactor --- .../Function/AEMeasurableSequence.lean | 9 + Mathlib/MeasureTheory/Integral/Lebesgue.lean | 131 ++++++- .../Kernel/Disintegration/CdfToKernel.lean | 349 ++++++++++++++---- .../Kernel/Disintegration/CondCdf.lean | 266 +++---------- .../Kernel/Disintegration/CondKernelCdf.lean | 215 +++-------- .../Kernel/Disintegration/DensityToCdf.lean | 245 ------------ .../Algebra/Order/MonotoneConvergence.lean | 10 + 7 files changed, 520 insertions(+), 705 deletions(-) delete mode 100644 Mathlib/Probability/Kernel/Disintegration/DensityToCdf.lean diff --git a/Mathlib/MeasureTheory/Function/AEMeasurableSequence.lean b/Mathlib/MeasureTheory/Function/AEMeasurableSequence.lean index cb3dd8bff1ff0..4bb4b4a792f5e 100644 --- a/Mathlib/MeasureTheory/Function/AEMeasurableSequence.lean +++ b/Mathlib/MeasureTheory/Function/AEMeasurableSequence.lean @@ -137,4 +137,13 @@ theorem iSup [CompleteLattice β] [Countable ι] (hf : ∀ i, AEMeasurable (f i) exact measure_mono_null (Set.compl_subset_compl.mpr h_ss) (measure_compl_aeSeqSet_eq_zero hf hp) #align ae_seq.supr aeSeq.iSup +theorem iInf [CompleteLattice β] [Countable ι] (hf : ∀ i, AEMeasurable (f i) μ) + (hp : ∀ᵐ x ∂μ, p x fun n ↦ f n x) : ⨅ n, aeSeq hf p n =ᵐ[μ] ⨅ n, f n := by + simp_rw [Filter.EventuallyEq, ae_iff, iInf_apply] + have h_ss : aeSeqSet hf p ⊆ { a : α | ⨅ i : ι, aeSeq hf p i a = ⨅ i : ι, f i a } := by + intro x hx + congr + exact funext fun i ↦ aeSeq_eq_fun_of_mem_aeSeqSet hf hx i + exact measure_mono_null (Set.compl_subset_compl.mpr h_ss) (measure_compl_aeSeqSet_eq_zero hf hp) + end aeSeq diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index 9bdc2f6c20422..286f67bea7e2a 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -1198,6 +1198,41 @@ theorem tendsto_lintegral_filter_of_dominated_convergence {ι} {l : Filter ι} assumption #align measure_theory.tendsto_lintegral_filter_of_dominated_convergence MeasureTheory.tendsto_lintegral_filter_of_dominated_convergence +theorem lintegral_iInf' {f : ℕ → α → ℝ≥0∞} + (h_meas : ∀ n, AEMeasurable (f n) μ) (h_anti : ∀ᵐ a ∂μ, Antitone (fun i ↦ f i a)) + (h_fin : ∫⁻ a, f 0 a ∂μ ≠ ∞) : + ∫⁻ a, ⨅ n, f n a ∂μ = ⨅ n, ∫⁻ a, f n a ∂μ := by + simp_rw [← iInf_apply] + let p : α → (ℕ → ℝ≥0∞) → Prop := fun _ f' => Antitone f' + have hp : ∀ᵐ x ∂μ, p x fun i => f i x := h_anti + have h_ae_seq_mono : Antitone (aeSeq h_meas p) := by + intro n m hnm x + by_cases hx : x ∈ aeSeqSet h_meas p + · exact aeSeq.prop_of_mem_aeSeqSet h_meas hx hnm + · simp only [aeSeq, hx, if_false] + exact le_rfl + rw [lintegral_congr_ae (aeSeq.iInf h_meas hp).symm] + simp_rw [iInf_apply] + rw [lintegral_iInf (aeSeq.measurable h_meas p) h_ae_seq_mono] + · congr + exact funext fun n ↦ lintegral_congr_ae (aeSeq.aeSeq_n_eq_fun_n_ae h_meas hp n) + · rwa [lintegral_congr_ae (aeSeq.aeSeq_n_eq_fun_n_ae h_meas hp 0)] + +theorem lintegral_tendsto_of_tendsto_of_antitone {f : ℕ → α → ℝ≥0∞} {F : α → ℝ≥0∞} + (hf : ∀ n, AEMeasurable (f n) μ) (h_anti : ∀ᵐ x ∂μ, Antitone fun n ↦ f n x) + (h0 : ∫⁻ a, f 0 a ∂μ ≠ ∞) + (h_tendsto : ∀ᵐ x ∂μ, Tendsto (fun n ↦ f n x) atTop (𝓝 (F x))) : + Tendsto (fun n ↦ ∫⁻ x, f n x ∂μ) atTop (𝓝 (∫⁻ x, F x ∂μ)) := by + have : Antitone fun n ↦ ∫⁻ x, f n x ∂μ := fun i j hij ↦ + lintegral_mono_ae (h_anti.mono fun x hx ↦ hx hij) + suffices key : ∫⁻ x, F x ∂μ = ⨅ n, ∫⁻ x, f n x ∂μ by + rw [key] + exact tendsto_atTop_iInf this + rw [← lintegral_iInf' hf h_anti h0] + refine lintegral_congr_ae ?_ + filter_upwards [h_anti, h_tendsto] with _ hx_anti hx_tendsto + using tendsto_nhds_unique hx_tendsto (tendsto_atTop_iInf hx_anti) + section open Encodable @@ -1648,18 +1683,13 @@ lemma tendsto_of_lintegral_tendsto_of_monotone {α : Type*} {mα : MeasurableSpa (hf_mono : ∀ᵐ a ∂μ, Monotone (fun i ↦ f i a)) (h_bound : ∀ᵐ a ∂μ, ∀ i, f i a ≤ F a) (h_int_finite : ∫⁻ a, F a ∂μ ≠ ∞) : ∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) atTop (𝓝 (F a)) := by - have h_bound_finite : ∀ᵐ a ∂μ, F a ≠ ∞ := by - filter_upwards [ae_lt_top' hF_meas h_int_finite] with a ha using ha.ne have h_exists : ∀ᵐ a ∂μ, ∃ l, Tendsto (fun i ↦ f i a) atTop (𝓝 l) := by - filter_upwards [h_bound, h_bound_finite, hf_mono] with a h_le h_fin h_mono - have h_tendsto : Tendsto (fun i ↦ f i a) atTop atTop ∨ - ∃ l, Tendsto (fun i ↦ f i a) atTop (𝓝 l) := tendsto_of_monotone h_mono - cases' h_tendsto with h_absurd h_tendsto - · rw [tendsto_atTop_atTop_iff_of_monotone h_mono] at h_absurd - obtain ⟨i, hi⟩ := h_absurd (F a + 1) - refine absurd (hi.trans (h_le _)) (not_le.mpr ?_) - exact ENNReal.lt_add_right h_fin one_ne_zero - · exact h_tendsto + filter_upwards [hf_mono] with a h_mono + rcases tendsto_of_monotone h_mono with h | h + · refine ⟨∞, h.mono_right ?_⟩ + rw [OrderTop.atTop_eq] + exact pure_le_nhds _ + · exact h classical let F' : α → ℝ≥0∞ := fun a ↦ if h : ∃ l, Tendsto (fun i ↦ f i a) atTop (𝓝 l) then h.choose else ∞ @@ -1677,6 +1707,85 @@ lemma tendsto_of_lintegral_tendsto_of_monotone {α : Type*} {mα : MeasurableSpa refine tendsto_nhds_unique ?_ hf_tendsto exact lintegral_tendsto_of_tendsto_of_monotone hf_meas hf_mono hF'_tendsto +/-- A limit (over a general filter) of measurable `ℝ≥0∞` valued functions is measurable. -/ +theorem measurable_of_tendsto_ennreal' {ι} {f : ι → α → ℝ≥0∞} {g : α → ℝ≥0∞} (u : Filter ι) + [NeBot u] [IsCountablyGenerated u] (hf : ∀ i, Measurable (f i)) (lim : Tendsto f u (𝓝 g)) : + Measurable g := by + rcases u.exists_seq_tendsto with ⟨x, hx⟩ + rw [tendsto_pi_nhds] at lim + have : (fun y => liminf (fun n => (f (x n) y : ℝ≥0∞)) atTop) = g := by + ext1 y + exact ((lim y).comp hx).liminf_eq + rw [← this] + show Measurable fun y => liminf (fun n => (f (x n) y : ℝ≥0∞)) atTop + exact measurable_liminf fun n => hf (x n) + +/-- A sequential limit of measurable `ℝ≥0∞` valued functions is measurable. -/ +theorem measurable_of_tendsto_ennreal {f : ℕ → α → ℝ≥0∞} {g : α → ℝ≥0∞} (hf : ∀ i, Measurable (f i)) + (lim : Tendsto f atTop (𝓝 g)) : Measurable g := + measurable_of_tendsto_ennreal' atTop hf lim + +lemma aemeasurable_of_tendsto_ennreal' {f : ι → α → ℝ≥0∞} {g : α → ℝ≥0∞} {μ : Measure α} + (u : Filter ι) [NeBot u] [IsCountablyGenerated u] + (hf : ∀ i, AEMeasurable (f i) μ) (hlim : ∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) u (𝓝 (g a))) : + AEMeasurable g μ := by + rcases u.exists_seq_tendsto with ⟨v, hv⟩ + have h'f : ∀ n, AEMeasurable (f (v n)) μ := fun n ↦ hf (v n) + set p : α → (ℕ → ℝ≥0∞) → Prop := fun x f' ↦ Tendsto f' atTop (𝓝 (g x)) + have hp : ∀ᵐ x ∂μ, p x fun n ↦ f (v n) x := by + filter_upwards [hlim] with x hx using hx.comp hv + set aeSeqLim := fun x ↦ ite (x ∈ aeSeqSet h'f p) (g x) (⟨f (v 0) x⟩ : Nonempty ℝ≥0∞).some + refine ⟨aeSeqLim, measurable_of_tendsto_ennreal' atTop (aeSeq.measurable h'f p) + (tendsto_pi_nhds.mpr fun x ↦ ?_), ?_⟩ + · unfold_let aeSeqLim + simp_rw [aeSeq] + split_ifs with hx + · simp_rw [aeSeq.mk_eq_fun_of_mem_aeSeqSet h'f hx] + exact aeSeq.fun_prop_of_mem_aeSeqSet h'f hx + · exact tendsto_const_nhds + · exact (ite_ae_eq_of_measure_compl_zero g (fun x ↦ (⟨f (v 0) x⟩ : Nonempty ℝ≥0∞).some) + (aeSeqSet h'f p) (aeSeq.measure_compl_aeSeqSet_eq_zero h'f hp)).symm + +lemma aemeasurable_of_tendsto_ennreal {f : ℕ → α → ℝ≥0∞} {g : α → ℝ≥0∞} {μ : Measure α} + (hf : ∀ i, AEMeasurable (f i) μ) (hlim : ∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) atTop (𝓝 (g a))) : + AEMeasurable g μ := + aemeasurable_of_tendsto_ennreal' atTop hf hlim + +lemma tendsto_of_lintegral_tendsto_of_antitone {α : Type*} {mα : MeasurableSpace α} + {f : ℕ → α → ℝ≥0∞} {F : α → ℝ≥0∞} {μ : Measure α} + (hf_meas : ∀ n, AEMeasurable (f n) μ) + (hf_tendsto : Tendsto (fun i ↦ ∫⁻ a, f i a ∂μ) atTop (𝓝 (∫⁻ a, F a ∂μ))) + (hf_mono : ∀ᵐ a ∂μ, Antitone (fun i ↦ f i a)) + (h_bound : ∀ᵐ a ∂μ, ∀ i, F a ≤ f i a) (h0 : ∫⁻ a, f 0 a ∂μ ≠ ∞) : + ∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) atTop (𝓝 (F a)) := by + have h_int_finite : ∫⁻ a, F a ∂μ ≠ ∞ := by + refine ((lintegral_mono_ae ?_).trans_lt h0.lt_top).ne + filter_upwards [h_bound] with a ha using ha 0 + have h_exists : ∀ᵐ a ∂μ, ∃ l, Tendsto (fun i ↦ f i a) atTop (𝓝 l) := by + filter_upwards [hf_mono] with a h_mono + rcases tendsto_of_antitone h_mono with h | h + · refine ⟨0, h.mono_right ?_⟩ + rw [OrderBot.atBot_eq] + exact pure_le_nhds _ + · exact h + classical + let F' : α → ℝ≥0∞ := fun a ↦ if h : ∃ l, Tendsto (fun i ↦ f i a) atTop (𝓝 l) + then h.choose else ∞ + have hF'_tendsto : ∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) atTop (𝓝 (F' a)) := by + filter_upwards [h_exists] with a ha + simp_rw [dif_pos ha] + exact ha.choose_spec + suffices F' =ᵐ[μ] F by + filter_upwards [this, hF'_tendsto] with a h_eq h_tendsto using h_eq ▸ h_tendsto + have hF'_le : F ≤ᵐ[μ] F' := by + filter_upwards [h_bound, hF'_tendsto] with a h_le h_tendsto + exact ge_of_tendsto' h_tendsto (fun m ↦ h_le _) + suffices ∫⁻ a, F' a ∂μ = ∫⁻ a, F a ∂μ by + refine (ae_eq_of_ae_le_of_lintegral_le hF'_le h_int_finite ?_ this.le).symm + exact aemeasurable_of_tendsto_ennreal hf_meas hF'_tendsto + refine tendsto_nhds_unique ?_ hf_tendsto + exact lintegral_tendsto_of_tendsto_of_antitone hf_meas hf_mono h0 hF'_tendsto + end Lintegral open MeasureTheory.SimpleFunc diff --git a/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean index 6c2d4ab72aa67..c7edf2b581f0c 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean @@ -3,36 +3,37 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.MeasureTheory.Integral.SetIntegral +import Mathlib.MeasureTheory.Function.AEEqOfIntegral import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes +import Mathlib.Probability.Kernel.Disintegration.AuxLemmas /-! # Building a Markov kernel from a conditional cumulative distribution function -Let `μ : kernel α (β × ℝ)` and `ν : kernel α β` be two finite kernels. -A function `f : α × β → StieltjesFunction` is called a conditional kernel CDF of `μ` with respect +Let `κ : kernel α (β × ℝ)` and `ν : kernel α β` be two finite kernels. +A function `f : α × β → StieltjesFunction` is called a conditional kernel CDF of `κ` with respect to `ν` if it is measurable, tends to to 0 at -∞ and to 1 at +∞ for all `p : α × β`, if `fun b ↦ f (a, b) x` is `(ν a)`-integrable for all `a : α` and `x : ℝ` and for all measurable -sets `s : Set β`, `∫ b in s, f (a, b) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal`. +sets `s : Set β`, `∫ b in s, f (a, b) x ∂(ν a) = (κ a (s ×ˢ Iic x)).toReal`. -From such a function with property `hf : IsCondKernelCDF f μ ν`, we can build a `kernel (α × β) ℝ` -denoted by `hf.toKernel f` such that `μ = ν ⊗ₖ hf.toKernel f`. +From such a function with property `hf : IsCondKernelCDF f κ ν`, we can build a `kernel (α × β) ℝ` +denoted by `hf.toKernel f` such that `κ = ν ⊗ₖ hf.toKernel f`. ## Main definitions -Let `μ : kernel α (β × ℝ)` and `ν : kernel α β`. +Let `κ : kernel α (β × ℝ)` and `ν : kernel α β`. * `ProbabilityTheory.IsCondKernelCDF`: a function `f : α × β → StieltjesFunction` is called - a conditional kernel CDF of `μ` with respect to `ν` if it is measurable, tends to to 0 at -∞ and + a conditional kernel CDF of `κ` with respect to `ν` if it is measurable, tends to to 0 at -∞ and to 1 at +∞ for all `p : α × β`, if `fun b ↦ f (a, b) x` is `(ν a)`-integrable for all `a : α` and `x : ℝ` and for all measurable sets `s : Set β`, - `∫ b in s, f (a, b) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal`. + `∫ b in s, f (a, b) x ∂(ν a) = (κ a (s ×ˢ Iic x)).toReal`. * `ProbabilityTheory.IsCondKernelCDF.toKernel`: from a function `f : α × β → StieltjesFunction` - with the property `hf : IsCondKernelCDF f μ ν`, build a `kernel (α × β) ℝ` such that - `μ = ν ⊗ₖ hf.toKernel f`. + with the property `hf : IsCondKernelCDF f κ ν`, build a `kernel (α × β) ℝ` such that + `κ = ν ⊗ₖ hf.toKernel f`. * `ProbabilityTheory.IsRatCondKernelCDF`: a function `f : α × β → ℚ → ℝ` is called a rational - conditional kernel CDF of `μ` with respect to `ν` if is measurable and satisfies the same + conditional kernel CDF of `κ` with respect to `ν` if is measurable and satisfies the same integral conditions as in the definition of `IsCondKernelCDF`, and the `ℚ → ℝ` function `f (a, b)` satisfies the properties of a Sieltjes function for `(ν a)`-almost all `b : β`. @@ -41,7 +42,7 @@ Let `μ : kernel α (β × ℝ)` and `ν : kernel α β`. * `ProbabilityTheory.isCondKernelCDF_stieltjesOfMeasurableRat`: if `f : α × β → ℚ → ℝ` has the property `IsRatCondKernelCDF`, then `stieltjesOfMeasurableRat f` is a function `α × β → StieltjesFunction` with the property `IsCondKernelCDF`. -* `ProbabilityTheory.compProd_toKernel`: for `hf : IsCondKernelCDF f μ ν`, `ν ⊗ₖ hf.toKernel f = μ`. +* `ProbabilityTheory.compProd_toKernel`: for `hf : IsCondKernelCDF f κ ν`, `ν ⊗ₖ hf.toKernel f = κ`. -/ @@ -56,31 +57,47 @@ variable {α β : Type*} [MeasurableSpace α] section stieltjesOfMeasurableRat variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} - {f : α × β → ℚ → ℝ} {μ : kernel α (β × ℝ)} {ν : kernel α β} + {f : α × β → ℚ → ℝ} {κ : kernel α (β × ℝ)} {ν : kernel α β} -structure IsRatCondKernelCDF (f : α × β → ℚ → ℝ) (μ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := +structure IsRatCondKernelCDF (f : α × β → ℚ → ℝ) (κ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := (measurable : Measurable f) (isRatStieltjesPoint_ae (a : α) : ∀ᵐ b ∂(ν a), IsRatStieltjesPoint f (a, b)) (integrable (a : α) (q : ℚ) : Integrable (fun b ↦ f (a, b) q) (ν a)) (set_integral (a : α) {s : Set β} (_hs : MeasurableSet s) (q : ℚ) : - ∫ b in s, f (a, b) q ∂(ν a) = (μ a (s ×ˢ Iic (q : ℝ))).toReal) + ∫ b in s, f (a, b) q ∂(ν a) = (κ a (s ×ˢ Iic (q : ℝ))).toReal) -lemma stieltjesOfMeasurableRat_ae_eq (hf : IsRatCondKernelCDF f μ ν) (a : α) (q : ℚ) : +lemma IsRatCondKernelCDF.mono (hf : IsRatCondKernelCDF f κ ν) (a : α) : + ∀ᵐ b ∂(ν a), Monotone (f (a, b)) := by + filter_upwards [hf.isRatStieltjesPoint_ae a] with b hb using hb.mono + +lemma IsRatCondKernelCDF.tendsto_atTop_one (hf : IsRatCondKernelCDF f κ ν) (a : α) : + ∀ᵐ b ∂(ν a), Tendsto (f (a, b)) atTop (𝓝 1) := by + filter_upwards [hf.isRatStieltjesPoint_ae a] with b hb using hb.tendsto_atTop_one + +lemma IsRatCondKernelCDF.tendsto_atBot_zero (hf : IsRatCondKernelCDF f κ ν) (a : α) : + ∀ᵐ b ∂(ν a), Tendsto (f (a, b)) atBot (𝓝 0) := by + filter_upwards [hf.isRatStieltjesPoint_ae a] with b hb using hb.tendsto_atBot_zero + +lemma IsRatCondKernelCDF.iInf_rat_gt_eq (hf : IsRatCondKernelCDF f κ ν) (a : α) : + ∀ᵐ b ∂(ν a), ∀ q, ⨅ r : Ioi q, f (a, b) r = f (a, b) q := by + filter_upwards [hf.isRatStieltjesPoint_ae a] with b hb using hb.iInf_rat_gt_eq + +lemma stieltjesOfMeasurableRat_ae_eq (hf : IsRatCondKernelCDF f κ ν) (a : α) (q : ℚ) : (fun b ↦ stieltjesOfMeasurableRat f hf.measurable (a, b) q) =ᵐ[ν a] fun b ↦ f (a, b) q := by filter_upwards [hf.isRatStieltjesPoint_ae a] with a ha rw [stieltjesOfMeasurableRat_eq, toRatCDF_of_isRatStieltjesPoint ha] -lemma set_integral_stieltjesOfMeasurableRat_rat (hf : IsRatCondKernelCDF f μ ν) (a : α) (q : ℚ) +lemma set_integral_stieltjesOfMeasurableRat_rat (hf : IsRatCondKernelCDF f κ ν) (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : ∫ b in s, stieltjesOfMeasurableRat f hf.measurable (a, b) q ∂(ν a) - = (μ a (s ×ˢ Iic (q : ℝ))).toReal := by + = (κ a (s ×ˢ Iic (q : ℝ))).toReal := by rw [set_integral_congr_ae hs (g := fun b ↦ f (a, b) q) ?_, hf.set_integral a hs] filter_upwards [stieltjesOfMeasurableRat_ae_eq hf a q] with b hb using fun _ ↦ hb -lemma set_lintegral_stieltjesOfMeasurableRat_rat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) +lemma set_lintegral_stieltjesOfMeasurableRat_rat [IsFiniteKernel κ] (hf : IsRatCondKernelCDF f κ ν) (a : α) (q : ℚ) {s : Set β} (hs : MeasurableSet s) : ∫⁻ b in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, b) q) ∂(ν a) - = μ a (s ×ˢ Iic (q : ℝ)) := by + = κ a (s ×ˢ Iic (q : ℝ)) := by rw [← ofReal_integral_eq_lintegral_ofReal] · rw [set_integral_stieltjesOfMeasurableRat_rat hf a q hs, ENNReal.ofReal_toReal] exact measure_ne_top _ _ @@ -89,21 +106,21 @@ lemma set_lintegral_stieltjesOfMeasurableRat_rat [IsFiniteKernel μ] (hf : IsRat exact hf.integrable a q · exact ae_of_all _ (fun x ↦ stieltjesOfMeasurableRat_nonneg _ _ _) -lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) +lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel κ] (hf : IsRatCondKernelCDF f κ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : ∫⁻ b in s, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, b) x) ∂(ν a) - = μ a (s ×ˢ Iic x) := by + = κ a (s ×ˢ Iic x) := by -- We have the result for `x : ℚ` thanks to `set_lintegral_stieltjesOfMeasurableRat_rat`. -- We use a monotone convergence argument to extend it to the reals. by_cases hρ_zero : (ν a).restrict s = 0 · rw [hρ_zero, lintegral_zero_measure] have ⟨q, hq⟩ := exists_rat_gt x - suffices μ a (s ×ˢ Iic (q : ℝ)) = 0 by + suffices κ a (s ×ˢ Iic (q : ℝ)) = 0 by symm refine measure_mono_null (fun p ↦ ?_) this simp only [mem_prod, mem_Iic, and_imp] exact fun h1 h2 ↦ ⟨h1, h2.trans hq.le⟩ - suffices (μ a (s ×ˢ Iic (q : ℝ))).toReal = 0 by + suffices (κ a (s ×ˢ Iic (q : ℝ))).toReal = 0 by rw [ENNReal.toReal_eq_zero_iff] at this simpa [measure_ne_top] using this rw [← hf.set_integral a hs q] @@ -151,13 +168,13 @@ lemma set_lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCond exact mod_cast hij · exact ⟨h_nonempty.some, measure_ne_top _ _⟩ -lemma lintegral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) +lemma lintegral_stieltjesOfMeasurableRat [IsFiniteKernel κ] (hf : IsRatCondKernelCDF f κ ν) (a : α) (x : ℝ) : ∫⁻ b, ENNReal.ofReal (stieltjesOfMeasurableRat f hf.measurable (a, b) x) ∂(ν a) - = μ a (univ ×ˢ Iic x) := by + = κ a (univ ×ˢ Iic x) := by rw [← set_lintegral_univ, set_lintegral_stieltjesOfMeasurableRat hf _ _ MeasurableSet.univ] -lemma integrable_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) +lemma integrable_stieltjesOfMeasurableRat [IsFiniteKernel κ] (hf : IsRatCondKernelCDF f κ ν) (a : α) (x : ℝ) : Integrable (fun b ↦ stieltjesOfMeasurableRat f hf.measurable (a, b) x) (ν a) := by have : (fun b ↦ stieltjesOfMeasurableRat f hf.measurable (a, b) x) @@ -172,10 +189,10 @@ lemma integrable_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKer · rw [lintegral_stieltjesOfMeasurableRat hf] exact measure_ne_top _ _ -lemma set_integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) +lemma set_integral_stieltjesOfMeasurableRat [IsFiniteKernel κ] (hf : IsRatCondKernelCDF f κ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : ∫ b in s, stieltjesOfMeasurableRat f hf.measurable (a, b) x ∂(ν a) - = (μ a (s ×ˢ Iic x)).toReal := by + = (κ a (s ×ˢ Iic x)).toReal := by rw [← ENNReal.ofReal_eq_ofReal_iff, ENNReal.ofReal_toReal] rotate_left · exact measure_ne_top _ _ @@ -185,55 +202,231 @@ lemma set_integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondK · exact (integrable_stieltjesOfMeasurableRat hf _ _).restrict · exact ae_of_all _ (fun _ ↦ stieltjesOfMeasurableRat_nonneg _ _ _) -lemma integral_stieltjesOfMeasurableRat [IsFiniteKernel μ] (hf : IsRatCondKernelCDF f μ ν) +lemma integral_stieltjesOfMeasurableRat [IsFiniteKernel κ] (hf : IsRatCondKernelCDF f κ ν) (a : α) (x : ℝ) : ∫ b, stieltjesOfMeasurableRat f hf.measurable (a, b) x ∂(ν a) - = (μ a (univ ×ˢ Iic x)).toReal := by + = (κ a (univ ×ˢ Iic x)).toReal := by rw [← integral_univ, set_integral_stieltjesOfMeasurableRat hf _ _ MeasurableSet.univ] end stieltjesOfMeasurableRat +section isRatCondKernelCDFAux + +variable {β : Type*} {mβ : MeasurableSpace β} {f : α × β → ℚ → ℝ} + {κ : kernel α (β × ℝ)} {ν : kernel α β} + +structure isRatCondKernelCDFAux (f : α × β → ℚ → ℝ) (κ : kernel α (β × ℝ)) (ν : kernel α β) : + Prop := + (measurable : Measurable f) + (mono' (a : α) {q r : ℚ} (_hqr : q ≤ r) : ∀ᵐ c ∂(ν a), f (a, c) q ≤ f (a, c) r) + (nonneg' (a : α) (q : ℚ) : ∀ᵐ c ∂(ν a), 0 ≤ f (a, c) q) + (le_one' (a : α) (q : ℚ) : ∀ᵐ c ∂(ν a), f (a, c) q ≤ 1) + (tendsto_integral_of_antitone (a : α) (s : ℕ → ℚ) (_hs : Antitone s) + (_hs_tendsto : Tendsto s atTop atBot) : + Tendsto (fun m ↦ ∫ c, f (a, c) (s m) ∂(ν a)) atTop (𝓝 0)) + (tendsto_integral_of_monotone (a : α) (s : ℕ → ℚ) (_hs : Monotone s) + (_hs_tendsto : Tendsto s atTop atTop) : + Tendsto (fun m ↦ ∫ c, f (a, c) (s m) ∂(ν a)) atTop (𝓝 (ν a univ).toReal)) + (integrable (a : α) (q : ℚ) : Integrable (fun c ↦ f (a, c) q) (ν a)) + (set_integral (a : α) {A : Set β} (_hA : MeasurableSet A) (q : ℚ) : + ∫ c in A, f (a, c) q ∂(ν a) = (κ a (A ×ˢ Iic ↑q)).toReal) + +lemma isRatCondKernelCDFAux.measurable_right (hf : isRatCondKernelCDFAux f κ ν) (a : α) (q : ℚ) : + Measurable (fun t ↦ f (a, t) q) := by + let h := hf.measurable + rw [measurable_pi_iff] at h + exact (h q).comp measurable_prod_mk_left + +lemma isRatCondKernelCDFAux.mono (hf : isRatCondKernelCDFAux f κ ν) (a : α) : + ∀ᵐ c ∂(ν a), Monotone (f (a, c)) := by + unfold Monotone + simp_rw [ae_all_iff] + exact fun _ _ hqr ↦ hf.mono' a hqr + +lemma isRatCondKernelCDFAux.nonneg (hf : isRatCondKernelCDFAux f κ ν) (a : α) : + ∀ᵐ c ∂(ν a), ∀ q, 0 ≤ f (a, c) q := ae_all_iff.mpr <| hf.nonneg' a + +lemma isRatCondKernelCDFAux.le_one (hf : isRatCondKernelCDFAux f κ ν) (a : α) : + ∀ᵐ c ∂(ν a), ∀ q, f (a, c) q ≤ 1 := ae_all_iff.mpr <| hf.le_one' a + +lemma isRatCondKernelCDFAux.tendsto_zero_of_antitone (hf : isRatCondKernelCDFAux f κ ν) + [IsFiniteKernel ν] (a : α) (s : ℕ → ℚ) (hs : Antitone s) (hs_tendsto : Tendsto s atTop atBot) : + ∀ᵐ c ∂(ν a), Tendsto (fun m ↦ f (a, c) (s m)) atTop (𝓝 0) := by + refine tendsto_of_integral_tendsto_of_antitone ?_ (integrable_const _) ?_ ?_ ?_ + · exact fun n ↦ hf.integrable a (s n) + · rw [integral_zero] + exact hf.tendsto_integral_of_antitone a s hs hs_tendsto + · filter_upwards [hf.mono a] with t ht + exact fun n m hnm ↦ ht (hs hnm) + · filter_upwards [hf.nonneg a] with c hc using fun i ↦ hc (s i) + +lemma isRatCondKernelCDFAux.tendsto_one_of_monotone (hf : isRatCondKernelCDFAux f κ ν) + [IsFiniteKernel ν] (a : α) (s : ℕ → ℚ) (hs : Monotone s) (hs_tendsto : Tendsto s atTop atTop) : + ∀ᵐ c ∂(ν a), Tendsto (fun m ↦ f (a, c) (s m)) atTop (𝓝 1) := by + refine tendsto_of_integral_tendsto_of_monotone ?_ (integrable_const _) ?_ ?_ ?_ + · exact fun n ↦ hf.integrable a (s n) + · rw [MeasureTheory.integral_const, smul_eq_mul, mul_one] + exact hf.tendsto_integral_of_monotone a s hs hs_tendsto + · filter_upwards [hf.mono a] with t ht + exact fun n m hnm ↦ ht (hs hnm) + · filter_upwards [hf.le_one a] with c hc using fun i ↦ hc (s i) + +lemma isRatCondKernelCDFAux.tendsto_atTop_one (hf : isRatCondKernelCDFAux f κ ν) [IsFiniteKernel ν] + (a : α) : + ∀ᵐ t ∂(ν a), Tendsto (f (a, t)) atTop (𝓝 1) := by + suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ f (a, t) n) atTop (𝓝 1) by + filter_upwards [this, hf.mono a] with t ht h_mono + rw [tendsto_iff_tendsto_subseq_of_monotone h_mono tendsto_nat_cast_atTop_atTop] + exact ht + let s : ℕ → ℚ := fun n ↦ n + have hs : Monotone s := fun i j hij ↦ by simp [s, hij] + have hs_tendsto : Tendsto s atTop atTop := by + exact tendsto_nat_cast_atTop_atTop + filter_upwards [hf.tendsto_one_of_monotone a s hs hs_tendsto] with x hx using hx + +lemma isRatCondKernelCDFAux.tendsto_atBot_zero (hf : isRatCondKernelCDFAux f κ ν) [IsFiniteKernel ν] + (a : α) : + ∀ᵐ t ∂(ν a), Tendsto (f (a, t)) atBot (𝓝 0) := by + suffices ∀ᵐ t ∂(ν a), Tendsto (fun q : ℚ ↦ f (a, t) (-q)) atTop (𝓝 0) by + filter_upwards [this] with t ht + have h_eq_neg : f (a, t) = fun q : ℚ ↦ f (a, t) (- -q) := by + simp_rw [neg_neg] + rw [h_eq_neg] + convert ht.comp tendsto_neg_atBot_atTop + simp + suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ f (a, t) (-n)) atTop (𝓝 0) by + filter_upwards [this, hf.mono a] with t ht h_mono + have h_anti : Antitone (fun q ↦ f (a, t) (-q)) := h_mono.comp_antitone monotone_id.neg + exact (tendsto_iff_tendsto_subseq_of_antitone h_anti tendsto_nat_cast_atTop_atTop).mpr ht + let s : ℕ → ℚ := fun n ↦ -n + have hs : Antitone s := fun i j hij ↦ neg_le_neg (by exact mod_cast hij) + have hs_tendsto : Tendsto s atTop atBot := by + simp only [s, tendsto_neg_atBot_iff] + exact tendsto_nat_cast_atTop_atTop + convert hf.tendsto_zero_of_antitone a s hs hs_tendsto with x n + +lemma isRatCondKernelCDFAux.bddBelow_range (hf : isRatCondKernelCDFAux f κ ν) (a : α) : + ∀ᵐ t ∂(ν a), ∀ q : ℚ, BddBelow (range fun (r : Ioi q) ↦ f (a, t) r) := by + filter_upwards [hf.nonneg a] with c hc + refine fun q ↦ ⟨0, ?_⟩ + rw [mem_lowerBounds] + rintro x ⟨y, rfl⟩ + exact hc y + +lemma isRatCondKernelCDFAux.integrable_iInf_rat_gt (hf : isRatCondKernelCDFAux f κ ν) + [IsFiniteKernel ν] (a : α) (q : ℚ) : + Integrable (fun t ↦ ⨅ r : Ioi q, f (a, t) r) (ν a) := by + rw [← memℒp_one_iff_integrable] + refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ + · exact measurable_iInf fun i ↦ hf.measurable_right a _ + refine (?_ : _ ≤ (ν a univ : ℝ≥0∞)).trans_lt (measure_lt_top _ _) + refine (snorm_le_of_ae_bound (C := 1) ?_).trans ?_ + · filter_upwards [hf.bddBelow_range a, hf.nonneg a, hf.le_one a] + with t hbdd_below h_nonneg h_le_one + rw [Real.norm_eq_abs, abs_of_nonneg] + · refine ciInf_le_of_le ?_ ?_ ?_ + · exact hbdd_below _ + · exact ⟨q + 1, by simp⟩ + · exact h_le_one _ + · exact le_ciInf fun r ↦ h_nonneg _ + · simp + +lemma isRatCondKernelCDFAux.set_integral_iInf_rat_gt (hf : isRatCondKernelCDFAux f κ ν) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (q : ℚ) {A : Set β} (hA : MeasurableSet A) : + ∫ t in A, ⨅ r : Ioi q, f (a, t) r ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by + refine le_antisymm ?_ ?_ + · have h : ∀ r : Ioi q, ∫ t in A, ⨅ r' : Ioi q, f (a, t) r' ∂(ν a) + ≤ (κ a (A ×ˢ Iic (r : ℝ))).toReal := by + intro r + rw [← hf.set_integral a hA] + refine set_integral_mono_ae ?_ ?_ ?_ + · exact (hf.integrable_iInf_rat_gt _ _).integrableOn + · exact (hf.integrable _ _).integrableOn + · filter_upwards [hf.bddBelow_range a] with t ht + exact ciInf_le (ht _) r + calc ∫ t in A, ⨅ r : Ioi q, f (a, t) r ∂(ν a) + ≤ ⨅ r : Ioi q, (κ a (A ×ˢ Iic (r : ℝ))).toReal := le_ciInf h + _ = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by + rw [← Measure.iInf_rat_gt_prod_Iic hA q] + exact (ENNReal.toReal_iInf (fun r ↦ measure_ne_top _ _)).symm + · rw [← hf.set_integral a hA] + refine set_integral_mono_ae ?_ ?_ ?_ + · exact (hf.integrable _ _).integrableOn + · exact (hf.integrable_iInf_rat_gt _ _).integrableOn + · filter_upwards [hf.mono a] with c h_mono + exact le_ciInf (fun r ↦ h_mono (le_of_lt r.prop)) + +lemma isRatCondKernelCDFAux.iInf_rat_gt_eq (hf : isRatCondKernelCDFAux f κ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] (a : α) : + ∀ᵐ t ∂(ν a), ∀ q : ℚ, ⨅ r : Ioi q, f (a, t) r = f (a, t) q := by + rw [ae_all_iff] + refine fun q ↦ ae_eq_of_forall_set_integral_eq_of_sigmaFinite (μ := ν a) ?_ ?_ ?_ + · intro s _ _ + refine Integrable.integrableOn ?_ + exact hf.integrable_iInf_rat_gt _ _ + · exact fun s _ _ ↦ (hf.integrable a _).integrableOn + · intro s hs _ + rw [hf.set_integral _ hs, hf.set_integral_iInf_rat_gt _ _ hs] + +lemma isRatCondKernelCDFAux.isRatStieltjesPoint_ae (hf : isRatCondKernelCDFAux f κ ν) + [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + ∀ᵐ t ∂(ν a), IsRatStieltjesPoint f (a, t) := by + filter_upwards [hf.tendsto_atTop_one a, hf.tendsto_atBot_zero a, + hf.iInf_rat_gt_eq a, hf.mono a] with t ht_top ht_bot ht_iInf h_mono + exact ⟨h_mono, ht_top, ht_bot, ht_iInf⟩ + +lemma isRatCondKernelCDFAux.isRatCondKernelCDF (hf : isRatCondKernelCDFAux f κ ν) [IsFiniteKernel κ] + [IsFiniteKernel ν] : + IsRatCondKernelCDF f κ ν where + measurable := hf.measurable + isRatStieltjesPoint_ae := hf.isRatStieltjesPoint_ae + integrable := hf.integrable + set_integral := hf.set_integral + +end isRatCondKernelCDFAux + + section IsCondKernelCDF variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} - {f : α × β → StieltjesFunction} {μ : kernel α (β × ℝ)} {ν : kernel α β} + {f : α × β → StieltjesFunction} {κ : kernel α (β × ℝ)} {ν : kernel α β} -structure IsCondKernelCDF (f : α × β → StieltjesFunction) (μ : kernel α (β × ℝ)) (ν : kernel α β) : +structure IsCondKernelCDF (f : α × β → StieltjesFunction) (κ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := (measurable (x : ℝ) : Measurable fun p ↦ f p x) (integrable (a : α) (x : ℝ) : Integrable (fun b ↦ f (a, b) x) (ν a)) (tendsto_atTop_one (p : α × β) : Tendsto (f p) atTop (𝓝 1)) (tendsto_atBot_zero (p : α × β) : Tendsto (f p) atBot (𝓝 0)) (set_integral (a : α) {s : Set β} (_hs : MeasurableSet s) (x : ℝ) : - ∫ b in s, f (a, b) x ∂(ν a) = (μ a (s ×ˢ Iic x)).toReal) + ∫ b in s, f (a, b) x ∂(ν a) = (κ a (s ×ˢ Iic x)).toReal) -lemma IsCondKernelCDF.nonneg (hf : IsCondKernelCDF f μ ν) (p : α × β) (x : ℝ) : 0 ≤ f p x := +lemma IsCondKernelCDF.nonneg (hf : IsCondKernelCDF f κ ν) (p : α × β) (x : ℝ) : 0 ≤ f p x := Monotone.le_of_tendsto (f p).mono (hf.tendsto_atBot_zero p) x -lemma IsCondKernelCDF.le_one (hf : IsCondKernelCDF f μ ν) (p : α × β) (x : ℝ) : f p x ≤ 1 := +lemma IsCondKernelCDF.le_one (hf : IsCondKernelCDF f κ ν) (p : α × β) (x : ℝ) : f p x ≤ 1 := Monotone.ge_of_tendsto (f p).mono (hf.tendsto_atTop_one p) x -lemma IsCondKernelCDF.integral [IsFiniteKernel μ] - {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f μ ν) (a : α) (x : ℝ) : - ∫ b, f (a, b) x ∂(ν a) = (μ a (univ ×ˢ Iic x)).toReal := by +lemma IsCondKernelCDF.integral [IsFiniteKernel κ] + {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f κ ν) (a : α) (x : ℝ) : + ∫ b, f (a, b) x ∂(ν a) = (κ a (univ ×ˢ Iic x)).toReal := by rw [← hf.set_integral _ MeasurableSet.univ, Measure.restrict_univ] -lemma IsCondKernelCDF.set_lintegral [IsFiniteKernel μ] - {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f μ ν) +lemma IsCondKernelCDF.set_lintegral [IsFiniteKernel κ] + {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f κ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (x : ℝ) : - ∫⁻ b in s, ENNReal.ofReal (f (a, b) x) ∂(ν a) = μ a (s ×ˢ Iic x) := by + ∫⁻ b in s, ENNReal.ofReal (f (a, b) x) ∂(ν a) = κ a (s ×ˢ Iic x) := by rw [← ofReal_integral_eq_lintegral_ofReal (hf.integrable a x).restrict (ae_of_all _ (fun _ ↦ hf.nonneg _ _)), hf.set_integral a hs x, ENNReal.ofReal_toReal] exact measure_ne_top _ _ -lemma IsCondKernelCDF.lintegral [IsFiniteKernel μ] - {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f μ ν) (a : α) (x : ℝ) : - ∫⁻ b, ENNReal.ofReal (f (a, b) x) ∂(ν a) = μ a (univ ×ˢ Iic x) := by +lemma IsCondKernelCDF.lintegral [IsFiniteKernel κ] + {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f κ ν) (a : α) (x : ℝ) : + ∫⁻ b, ENNReal.ofReal (f (a, b) x) ∂(ν a) = κ a (univ ×ˢ Iic x) := by rw [← hf.set_lintegral _ MeasurableSet.univ, Measure.restrict_univ] -lemma isCondKernelCDF_stieltjesOfMeasurableRat {f : α × β → ℚ → ℝ} (hf : IsRatCondKernelCDF f μ ν) - [IsFiniteKernel μ] : - IsCondKernelCDF (stieltjesOfMeasurableRat f hf.measurable) μ ν where +lemma isCondKernelCDF_stieltjesOfMeasurableRat {f : α × β → ℚ → ℝ} (hf : IsRatCondKernelCDF f κ ν) + [IsFiniteKernel κ] : + IsCondKernelCDF (stieltjesOfMeasurableRat f hf.measurable) κ ν where measurable := measurable_stieltjesOfMeasurableRat hf.measurable integrable := integrable_stieltjesOfMeasurableRat hf tendsto_atTop_one := tendsto_stieltjesOfMeasurableRat_atTop hf.measurable @@ -245,7 +438,7 @@ end IsCondKernelCDF section ToKernel variable {_ : MeasurableSpace β} {f : α × β → StieltjesFunction} - {μ : kernel α (β × ℝ)} {ν : kernel α β} {hf : IsCondKernelCDF f μ ν} + {κ : kernel α (β × ℝ)} {ν : kernel α β} {hf : IsCondKernelCDF f κ ν} lemma StieltjesFunction.measurable_measure {f : α → StieltjesFunction} (hf : ∀ q, Measurable fun a ↦ f a q) @@ -277,7 +470,7 @@ lemma StieltjesFunction.measurable_measure {f : α → StieltjesFunction} exact Measurable.ennreal_tsum hf_cd_meas noncomputable -def IsCondKernelCDF.toKernel (f : α × β → StieltjesFunction) (hf : IsCondKernelCDF f μ ν) : +def IsCondKernelCDF.toKernel (f : α × β → StieltjesFunction) (hf : IsCondKernelCDF f κ ν) : kernel (α × β) ℝ where val p := (f p).measure property := StieltjesFunction.measurable_measure hf.measurable @@ -298,18 +491,18 @@ end ToKernel section variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} - {f : α × β → StieltjesFunction} {μ : kernel α (β × ℝ)} {ν : kernel α β} - {hf : IsCondKernelCDF f μ ν} + {f : α × β → StieltjesFunction} {κ : kernel α (β × ℝ)} {ν : kernel α β} + {hf : IsCondKernelCDF f κ ν} -lemma set_lintegral_toKernel_Iic [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) +lemma set_lintegral_toKernel_Iic [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ b in s, hf.toKernel f (a, b) (Iic x) ∂(ν a) = μ a (s ×ˢ Iic x) := by + ∫⁻ b in s, hf.toKernel f (a, b) (Iic x) ∂(ν a) = κ a (s ×ˢ Iic x) := by simp_rw [IsCondKernelCDF.toKernel_Iic] exact hf.set_lintegral _ hs _ -lemma set_lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) +lemma set_lintegral_toKernel_univ [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) (a : α) {s : Set β} (hs : MeasurableSet s) : - ∫⁻ b in s, hf.toKernel f (a, b) univ ∂(ν a) = μ a (s ×ˢ univ) := by + ∫⁻ b in s, hf.toKernel f (a, b) univ ∂(ν a) = κ a (s ×ˢ univ) := by have : ⋃ r : ℚ, Iic (r : ℝ) = univ := by ext1 x simp only [mem_iUnion, mem_Iic, mem_univ, iff_true_iff] @@ -331,13 +524,13 @@ lemma set_lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ · refine Monotone.directed_le fun i j hij t ↦ measure_mono (Iic_subset_Iic.mpr ?_) exact mod_cast hij -lemma lintegral_toKernel_univ [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) (a : α) : - ∫⁻ b, hf.toKernel f (a, b) univ ∂(ν a) = μ a univ := by +lemma lintegral_toKernel_univ [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) (a : α) : + ∫⁻ b, hf.toKernel f (a, b) univ ∂(ν a) = κ a univ := by rw [← set_lintegral_univ, set_lintegral_toKernel_univ hf a MeasurableSet.univ, univ_prod_univ] -lemma set_lintegral_toKernel_prod [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) +lemma set_lintegral_toKernel_prod [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set ℝ} (ht : MeasurableSet t) : - ∫⁻ b in s, hf.toKernel f (a, b) t ∂(ν a) = μ a (s ×ˢ t) := by + ∫⁻ b in s, hf.toKernel f (a, b) t ∂(ν a) = κ a (s ×ˢ t) := by -- `set_lintegral_toKernel_Iic` gives the result for `t = Iic x`. These sets form a -- π-system that generates the Borel σ-algebra, hence we can get the same equality for any -- measurable set `t`. @@ -356,9 +549,9 @@ lemma set_lintegral_toKernel_prod [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ · rw [ht_lintegral] exact measure_ne_top _ _ · exact eventually_of_forall fun a ↦ measure_mono (subset_univ _) - _ = μ a (s ×ˢ univ) - μ a (s ×ˢ t) := by + _ = κ a (s ×ˢ univ) - κ a (s ×ˢ t) := by rw [set_lintegral_toKernel_univ hf a hs, ht_lintegral] - _ = μ a (s ×ˢ tᶜ) := by + _ = κ a (s ×ˢ tᶜ) := by rw [← measure_diff _ (hs.prod ht) (measure_ne_top _ _)] · rw [prod_diff_prod, compl_eq_univ_diff] simp only [diff_self, empty_prod, union_empty] @@ -375,9 +568,9 @@ lemma set_lintegral_toKernel_prod [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ · exact fun i ↦ ((kernel.measurable_coe _ (hf_meas i)).comp measurable_prod_mk_left).aemeasurable.restrict -lemma lintegral_toKernel_mem [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) +lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) (a : α) {s : Set (β × ℝ)} (hs : MeasurableSet s) : - ∫⁻ b, hf.toKernel f (a, b) {y | (b, y) ∈ s} ∂(ν a) = μ a s := by + ∫⁻ b, hf.toKernel f (a, b) {y | (b, y) ∈ s} ∂(ν a) = κ a s := by -- `set_lintegral_toKernel_prod` gives the result for sets of the form `t₁ × t₂`. These -- sets form a π-system that generates the product σ-algebra, hence we can get the same equality -- for any measurable set `s`. @@ -423,8 +616,8 @@ lemma lintegral_toKernel_mem [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) refine ((lintegral_mono_ae h_le).trans_lt ?_).ne rw [lintegral_toKernel_univ hf] exact measure_lt_top _ univ - _ = μ a univ - μ a t := by rw [ht_eq, lintegral_toKernel_univ hf] - _ = μ a tᶜ := (measure_compl ht (measure_ne_top _ _)).symm + _ = κ a univ - κ a t := by rw [ht_eq, lintegral_toKernel_univ hf] + _ = κ a tᶜ := (measure_compl ht (measure_ne_top _ _)).symm · intro f' hf_disj hf_meas hf_eq have h_eq : ∀ a, {x | (a, x) ∈ ⋃ i, f' i} = ⋃ i, {x | (a, x) ∈ f' i} := by intro a; ext x; simp only [mem_iUnion, mem_setOf_eq] @@ -444,21 +637,21 @@ lemma lintegral_toKernel_mem [IsFiniteKernel μ] (hf : IsCondKernelCDF f μ ν) rw [measure_iUnion (h_disj x) fun i ↦ measurable_prod_mk_left (hf_meas i)] _ = ∑' i, ∫⁻ b, hf.toKernel f (a, b) {y | (b, y) ∈ f' i} ∂(ν a) := lintegral_tsum fun i ↦ (kernel.measurable_kernel_prod_mk_left' (hf_meas i) a).aemeasurable - _ = ∑' i, μ a (f' i) := by simp_rw [hf_eq] - _ = μ a (iUnion f') := (measure_iUnion hf_disj hf_meas).symm + _ = ∑' i, κ a (f' i) := by simp_rw [hf_eq] + _ = κ a (iUnion f') := (measure_iUnion hf_disj hf_meas).symm -lemma compProd_toKernel [IsFiniteKernel μ] [IsSFiniteKernel ν] - (hf : IsCondKernelCDF f μ ν) : - ν ⊗ₖ hf.toKernel f = μ := by +lemma compProd_toKernel [IsFiniteKernel κ] [IsSFiniteKernel ν] + (hf : IsCondKernelCDF f κ ν) : + ν ⊗ₖ hf.toKernel f = κ := by ext a s hs rw [kernel.compProd_apply _ _ _ hs, lintegral_toKernel_mem hf a hs] -lemma ae_toKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsCondKernelCDF f μ ν) (a : α) - {s : Set ℝ} (hs : MeasurableSet s) (hμs : μ a {x | x.snd ∈ sᶜ} = 0) : +lemma ae_toKernel_eq_one [IsFiniteKernel κ] [IsSFiniteKernel ν] (hf : IsCondKernelCDF f κ ν) (a : α) + {s : Set ℝ} (hs : MeasurableSet s) (hκs : κ a {x | x.snd ∈ sᶜ} = 0) : ∀ᵐ b ∂(ν a), hf.toKernel f (a, b) s = 1 := by - have h_eq : ν ⊗ₖ hf.toKernel f = μ := compProd_toKernel hf - have h : μ a {x | x.snd ∈ sᶜ} = (ν ⊗ₖ hf.toKernel f) a {x | x.snd ∈ sᶜ} := by rw [h_eq] - rw [hμs, kernel.compProd_apply] at h + have h_eq : ν ⊗ₖ hf.toKernel f = κ := compProd_toKernel hf + have h : κ a {x | x.snd ∈ sᶜ} = (ν ⊗ₖ hf.toKernel f) a {x | x.snd ∈ sᶜ} := by rw [h_eq] + rw [hκs, kernel.compProd_apply] at h swap; · exact measurable_snd hs.compl rw [eq_comm, lintegral_eq_zero_iff] at h swap @@ -470,7 +663,7 @@ lemma ae_toKernel_eq_one [IsFiniteKernel μ] [IsSFiniteKernel ν] (hf : IsCondKe change hf.toKernel f (a, b) sᶜ = 0 at hb rwa [prob_compl_eq_zero_iff hs] at hb -lemma measurableSet_toKernel_eq_one (hf : IsCondKernelCDF f μ ν) +lemma measurableSet_toKernel_eq_one (hf : IsCondKernelCDF f κ ν) {s : Set ℝ} (hs : MeasurableSet s) : MeasurableSet {p | hf.toKernel f p s = 1} := (kernel.measurable_coe _ hs) (measurableSet_singleton 1) diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index 64038aded22bd..2b8c4ff7ca5de 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -175,6 +175,10 @@ theorem set_lintegral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) {s : Set exacts [measurable_const, rfl] #align probability_theory.set_lintegral_pre_cdf_fst ProbabilityTheory.set_lintegral_preCDF_fst +lemma lintegral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) [IsFiniteMeasure ρ] : + ∫⁻ x, preCDF ρ r x ∂ρ.fst = ρ.IicSnd r univ := by + rw [← set_lintegral_univ, set_lintegral_preCDF_fst ρ r MeasurableSet.univ] + theorem monotone_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : ∀ᵐ a ∂ρ.fst, Monotone fun r ↦ preCDF ρ r a := by simp_rw [Monotone, ae_all_iff] @@ -184,26 +188,6 @@ theorem monotone_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : exact Measure.IicSnd_mono ρ (mod_cast hrr') s #align probability_theory.monotone_pre_cdf ProbabilityTheory.monotone_preCDF -theorem set_lintegral_iInf_gt_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (t : ℚ) {s : Set α} - (hs : MeasurableSet s) : ∫⁻ x in s, ⨅ r : Ioi t, preCDF ρ r x ∂ρ.fst = ρ.IicSnd t s := by - refine le_antisymm ?_ ?_ - · have h : ∀ q : Ioi t, ∫⁻ x in s, ⨅ r : Ioi t, preCDF ρ r x ∂ρ.fst ≤ ρ.IicSnd q s := by - intro q - rw [← set_lintegral_preCDF_fst ρ _ hs] - refine set_lintegral_mono_ae ?_ measurable_preCDF ?_ - · exact measurable_iInf fun _ ↦ measurable_preCDF - · filter_upwards [monotone_preCDF _] with a _ - exact fun _ ↦ iInf_le _ q - calc - ∫⁻ x in s, ⨅ r : Ioi t, preCDF ρ r x ∂ρ.fst ≤ ⨅ q : Ioi t, ρ.IicSnd q s := le_iInf h - _ = ρ.IicSnd t s := Measure.iInf_IicSnd_gt ρ t hs - · rw [(set_lintegral_preCDF_fst ρ t hs).symm] - refine set_lintegral_mono_ae measurable_preCDF ?_ ?_ - · exact measurable_iInf fun _ ↦ measurable_preCDF - · filter_upwards [monotone_preCDF _] with a ha_mono - exact fun _ ↦ le_iInf fun r ↦ ha_mono (le_of_lt r.prop) -#align probability_theory.set_lintegral_infi_gt_pre_cdf ProbabilityTheory.set_lintegral_iInf_gt_preCDF - theorem preCDF_le_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : ∀ᵐ a ∂ρ.fst, ∀ r, preCDF ρ r a ≤ 1 := by rw [ae_all_iff] @@ -224,147 +208,62 @@ theorem set_integral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) {s : Set α filter_upwards [preCDF_le_one ρ] with a ha exact (ha r).trans_lt ENNReal.one_lt_top -theorem tendsto_lintegral_preCDF_atTop (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - Tendsto (fun r ↦ ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (ρ univ)) := by - convert ρ.tendsto_IicSnd_atTop MeasurableSet.univ - · rw [← set_lintegral_univ, set_lintegral_preCDF_fst ρ _ MeasurableSet.univ] - · exact Measure.fst_univ.symm -#align probability_theory.tendsto_lintegral_pre_cdf_at_top ProbabilityTheory.tendsto_lintegral_preCDF_atTop - -theorem tendsto_lintegral_preCDF_atBot (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - Tendsto (fun r ↦ ∫⁻ a, preCDF ρ r a ∂ρ.fst) atBot (𝓝 0) := by - convert ρ.tendsto_IicSnd_atBot MeasurableSet.univ - rw [← set_lintegral_univ, set_lintegral_preCDF_fst ρ _ MeasurableSet.univ] -#align probability_theory.tendsto_lintegral_pre_cdf_at_bot ProbabilityTheory.tendsto_lintegral_preCDF_atBot - -theorem tendsto_preCDF_atTop_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - ∀ᵐ a ∂ρ.fst, Tendsto (fun r ↦ preCDF ρ r a) atTop (𝓝 1) := by - -- We show first that `preCDF` has a limit almost everywhere. That limit has to be at most 1. - -- We then show that the integral of `preCDF` tends to the integral of 1, and that it also tends - -- to the integral of the limit. Since the limit is at most 1 and has same integral as 1, it is - -- equal to 1 a.e. - have h_mono := monotone_preCDF ρ - have h_le_one := preCDF_le_one ρ - -- `preCDF` has a limit a.e. - have h_exists : ∀ᵐ a ∂ρ.fst, ∃ l, Tendsto (fun r ↦ preCDF ρ r a) atTop (𝓝 l) := by - filter_upwards [h_mono] with a ha_mono - exact ⟨_, tendsto_atTop_iSup ha_mono⟩ - classical - -- let `F` be the pointwise limit of `preCDF` where it exists, and 0 elsewhere. - let F : α → ℝ≥0∞ := fun a ↦ - if h : ∃ l, Tendsto (fun r ↦ preCDF ρ r a) atTop (𝓝 l) then h.choose else 0 - have h_tendsto_ℚ : ∀ᵐ a ∂ρ.fst, Tendsto (fun r ↦ preCDF ρ r a) atTop (𝓝 (F a)) := by - filter_upwards [h_exists] with a ha - simp_rw [dif_pos ha] - exact ha.choose_spec - have h_tendsto_ℕ : ∀ᵐ a ∂ρ.fst, Tendsto (fun n : ℕ ↦ preCDF ρ n a) atTop (𝓝 (F a)) := by - filter_upwards [h_tendsto_ℚ] with a ha using ha.comp tendsto_nat_cast_atTop_atTop - have hF_ae_meas : AEMeasurable F ρ.fst := by - refine aemeasurable_of_tendsto_metrizable_ae _ (fun n ↦ ?_) h_tendsto_ℚ - exact measurable_preCDF.aemeasurable - have hF_le_one : ∀ᵐ a ∂ρ.fst, F a ≤ 1 := by - filter_upwards [h_tendsto_ℚ, h_le_one] with a ha ha_le using le_of_tendsto' ha ha_le - -- it suffices to show that the limit `F` is 1 a.e. - suffices ∀ᵐ a ∂ρ.fst, F a = 1 by - filter_upwards [h_tendsto_ℚ, this] with a ha_tendsto ha_eq - rwa [ha_eq] at ha_tendsto - -- since `F` is at most 1, proving that its integral is the same as the integral of 1 will tell - -- us that `F` is 1 a.e. - have h_lintegral_eq : ∫⁻ a, F a ∂ρ.fst = ∫⁻ _, 1 ∂ρ.fst := by - have h_lintegral : - Tendsto (fun r : ℕ ↦ ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (∫⁻ a, F a ∂ρ.fst)) := by - refine lintegral_tendsto_of_tendsto_of_monotone - (fun _ ↦ measurable_preCDF.aemeasurable) ?_ h_tendsto_ℕ - filter_upwards [h_mono] with a ha - refine fun n m hnm ↦ ha ?_ - exact mod_cast hnm - have h_lintegral' : - Tendsto (fun r : ℕ ↦ ∫⁻ a, preCDF ρ r a ∂ρ.fst) atTop (𝓝 (∫⁻ _, 1 ∂ρ.fst)) := by - rw [lintegral_one, Measure.fst_univ] - exact (tendsto_lintegral_preCDF_atTop ρ).comp tendsto_nat_cast_atTop_atTop - exact tendsto_nhds_unique h_lintegral h_lintegral' - have : ∫⁻ a, 1 - F a ∂ρ.fst = 0 := by - rw [lintegral_sub' hF_ae_meas _ hF_le_one, h_lintegral_eq, tsub_self] - calc - ∫⁻ a, F a ∂ρ.fst = ∫⁻ _, 1 ∂ρ.fst := h_lintegral_eq - _ = ρ.fst univ := lintegral_one - _ = ρ univ := Measure.fst_univ - _ ≠ ∞ := measure_ne_top ρ _ - rw [lintegral_eq_zero_iff' (aemeasurable_const.sub hF_ae_meas)] at this - filter_upwards [this, hF_le_one] with ha h_one_sub_eq_zero h_le_one - rw [Pi.zero_apply, tsub_eq_zero_iff_le] at h_one_sub_eq_zero - exact le_antisymm h_le_one h_one_sub_eq_zero -#align probability_theory.tendsto_pre_cdf_at_top_one ProbabilityTheory.tendsto_preCDF_atTop_one - -theorem tendsto_preCDF_atBot_zero (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - ∀ᵐ a ∂ρ.fst, Tendsto (fun r ↦ preCDF ρ r a) atBot (𝓝 0) := by - -- We show first that `preCDF` has a limit in ℝ≥0∞ almost everywhere. - -- We then show that the integral of `preCDF` tends to 0, and that it also tends - -- to the integral of the limit. Since the limit has integral 0, it is equal to 0 a.e. - suffices ∀ᵐ a ∂ρ.fst, Tendsto (fun r ↦ preCDF ρ (-r) a) atTop (𝓝 0) by - filter_upwards [this] with a ha - have h_eq_neg : (fun r : ℚ ↦ preCDF ρ r a) = fun r : ℚ ↦ preCDF ρ (- -r) a := by - simp_rw [neg_neg] - rw [h_eq_neg] - exact ha.comp tendsto_neg_atBot_atTop - have h_exists : ∀ᵐ a ∂ρ.fst, ∃ l, Tendsto (fun r ↦ preCDF ρ (-r) a) atTop (𝓝 l) := by - filter_upwards [monotone_preCDF ρ] with a ha - have h_anti : Antitone fun r ↦ preCDF ρ (-r) a := fun p q hpq ↦ ha (neg_le_neg hpq) - exact ⟨_, tendsto_atTop_iInf h_anti⟩ - classical - let F : α → ℝ≥0∞ := fun a ↦ - if h : ∃ l, Tendsto (fun r ↦ preCDF ρ (-r) a) atTop (𝓝 l) then h.choose else 0 - have h_tendsto : ∀ᵐ a ∂ρ.fst, Tendsto (fun r ↦ preCDF ρ (-r) a) atTop (𝓝 (F a)) := by - filter_upwards [h_exists] with a ha - simp_rw [dif_pos ha] - exact ha.choose_spec - suffices h_lintegral_eq : ∫⁻ a, F a ∂ρ.fst = 0 by - have hF_ae_meas : AEMeasurable F ρ.fst := by - refine aemeasurable_of_tendsto_metrizable_ae _ (fun n ↦ ?_) h_tendsto - exact measurable_preCDF.aemeasurable - rw [lintegral_eq_zero_iff' hF_ae_meas] at h_lintegral_eq - filter_upwards [h_tendsto, h_lintegral_eq] with a ha_tendsto ha_eq - rwa [ha_eq] at ha_tendsto - have h_lintegral : - Tendsto (fun r ↦ ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) atTop (𝓝 (∫⁻ a, F a ∂ρ.fst)) := by - refine tendsto_lintegral_filter_of_dominated_convergence (fun _ ↦ 1) - (eventually_of_forall fun _ ↦ measurable_preCDF) (eventually_of_forall fun _ ↦ ?_) ?_ - h_tendsto - · filter_upwards [preCDF_le_one ρ] with a ha using ha _ - · rw [lintegral_one] - exact measure_ne_top _ _ - have h_lintegral' : Tendsto (fun r ↦ ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) atTop (𝓝 0) := by - have h_lintegral_eq : (fun r ↦ ∫⁻ a, preCDF ρ (-r) a ∂ρ.fst) - = fun r : ℚ ↦ ρ (univ ×ˢ Iic (-r : ℝ)) := by - ext1 n - rw [← set_lintegral_univ, set_lintegral_preCDF_fst ρ _ MeasurableSet.univ, - Measure.IicSnd_univ] - norm_cast - rw [h_lintegral_eq] - have h_zero_eq_measure_iInter : (0 : ℝ≥0∞) = ρ (⋂ r : ℚ, univ ×ˢ Iic (-r : ℝ)) := by - suffices ⋂ r : ℚ, Iic (-(r : ℝ)) = ∅ by rw [← prod_iInter, this, prod_empty, measure_empty] - ext1 x - simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false_iff, not_forall, not_le] - simp_rw [neg_lt] - exact exists_rat_gt _ - rw [h_zero_eq_measure_iInter] - refine tendsto_measure_iInter (fun n ↦ MeasurableSet.univ.prod measurableSet_Iic) - (fun i j hij x ↦ ?_) ⟨0, measure_ne_top ρ _⟩ - simp only [mem_prod, mem_univ, mem_Iic, true_and_iff] - refine fun hxj ↦ hxj.trans (neg_le_neg ?_) - exact mod_cast hij - exact tendsto_nhds_unique h_lintegral h_lintegral' -#align probability_theory.tendsto_pre_cdf_at_bot_zero ProbabilityTheory.tendsto_preCDF_atBot_zero - -theorem inf_gt_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - ∀ᵐ a ∂ρ.fst, ∀ t : ℚ, ⨅ r : Ioi t, preCDF ρ r a = preCDF ρ t a := by - rw [ae_all_iff] - refine fun t ↦ ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite ?_ measurable_preCDF ?_ - · exact measurable_iInf fun i ↦ measurable_preCDF - intro s hs _ - rw [set_lintegral_iInf_gt_preCDF ρ t hs, set_lintegral_preCDF_fst ρ t hs] -#align probability_theory.inf_gt_pre_cdf ProbabilityTheory.inf_gt_preCDF +lemma integral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) [IsFiniteMeasure ρ] : + ∫ x, (preCDF ρ r x).toReal ∂ρ.fst = (ρ.IicSnd r univ).toReal := by + rw [← integral_univ, set_integral_preCDF_fst ρ _ MeasurableSet.univ] + +theorem integrable_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℚ) : + Integrable (fun a ↦ (preCDF ρ x a).toReal) ρ.fst := by + refine integrable_of_forall_fin_meas_le _ (measure_lt_top ρ.fst univ) ?_ fun t _ _ ↦ ?_ + · exact measurable_preCDF.ennreal_toReal.aestronglyMeasurable + · simp_rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_of_nonneg ENNReal.toReal_nonneg] + rw [← lintegral_one] + refine (set_lintegral_le_lintegral _ _).trans (lintegral_mono_ae ?_) + filter_upwards [preCDF_le_one ρ] with a ha using ENNReal.ofReal_toReal_le.trans (ha _) + +lemma isRatCondKernelCDFAux_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : + isRatCondKernelCDFAux (fun p r ↦ (preCDF ρ r p.2).toReal) + (kernel.const Unit ρ) (kernel.const Unit ρ.fst) where + measurable := measurable_preCDF'.comp measurable_snd + mono' a r r' hrr' := by + filter_upwards [monotone_preCDF ρ, preCDF_le_one ρ] with a h1 h2 + have h_ne_top : ∀ r, preCDF ρ r a ≠ ∞ := fun r ↦ ((h2 r).trans_lt ENNReal.one_lt_top).ne + rw [ENNReal.toReal_le_toReal (h_ne_top _) (h_ne_top _)] + exact h1 hrr' + nonneg' _ q := by simp + le_one' a q := by + simp only [kernel.const_apply, forall_const] + filter_upwards [preCDF_le_one ρ] with a ha + refine ENNReal.toReal_le_of_le_ofReal zero_le_one ?_ + simp [ha] + tendsto_integral_of_antitone a s _ hs_tendsto := by + simp_rw [kernel.const_apply, integral_preCDF_fst ρ] + have h := ρ.tendsto_IicSnd_atBot MeasurableSet.univ + rw [← ENNReal.zero_toReal] + have h0 : Tendsto ENNReal.toReal (𝓝 0) (𝓝 0) := + ENNReal.continuousAt_toReal ENNReal.zero_ne_top + exact h0.comp (h.comp hs_tendsto) + tendsto_integral_of_monotone a s _ hs_tendsto := by + simp_rw [kernel.const_apply, integral_preCDF_fst ρ] + have h := ρ.tendsto_IicSnd_atTop MeasurableSet.univ + have h0 : Tendsto ENNReal.toReal (𝓝 (ρ.fst univ)) (𝓝 (ρ.fst univ).toReal) := + ENNReal.continuousAt_toReal (measure_ne_top _ _) + exact h0.comp (h.comp hs_tendsto) + integrable _ q := integrable_preCDF ρ q + set_integral a s hs q := by rw [kernel.const_apply, kernel.const_apply, set_integral_preCDF_fst _ _ hs, + Measure.IicSnd_apply _ _ hs] +lemma isRatCondKernelCDF_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : + IsRatCondKernelCDF (fun p r ↦ (preCDF ρ r p.2).toReal) + (kernel.const Unit ρ) (kernel.const Unit ρ.fst) := + (isRatCondKernelCDFAux_preCDF ρ).isRatCondKernelCDF + +#noalign probability_theory.set_lintegral_infi_gt_pre_cdf +#noalign probability_theory.tendsto_lintegral_pre_cdf_at_top +#noalign probability_theory.tendsto_lintegral_pre_cdf_at_bot +#noalign probability_theory.tendsto_pre_cdf_at_top_one +#noalign probability_theory.tendsto_pre_cdf_at_bot_zero +#noalign probability_theory.inf_gt_pre_cdf #noalign probability_theory.has_cond_cdf #noalign probability_theory.has_cond_cdf_ae #noalign probability_theory.cond_cdf_set @@ -393,49 +292,6 @@ theorem inf_gt_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : /-! ### Conditional cdf -/ -lemma isRatStieltjesPoint_ae (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - ∀ᵐ a ∂ρ.fst, IsRatStieltjesPoint (fun a r ↦ (preCDF ρ r a).toReal) a := by - filter_upwards [monotone_preCDF ρ, preCDF_le_one ρ, tendsto_preCDF_atTop_one ρ, - tendsto_preCDF_atBot_zero ρ, inf_gt_preCDF ρ] with a h1 h2 h3 h4 h5 - constructor - · intro r r' hrr' - have h_ne_top : ∀ r, preCDF ρ r a ≠ ∞ := fun r ↦ ((h2 r).trans_lt ENNReal.one_lt_top).ne - rw [ENNReal.toReal_le_toReal (h_ne_top _) (h_ne_top _)] - exact h1 hrr' - · rw [← ENNReal.one_toReal, ENNReal.tendsto_toReal_iff] - · exact h3 - · exact fun r ↦ ((h2 r).trans_lt ENNReal.one_lt_top).ne - · exact ENNReal.one_ne_top - · rw [← ENNReal.zero_toReal, ENNReal.tendsto_toReal_iff] - · exact h4 - · exact fun r ↦ ((h2 r).trans_lt ENNReal.one_lt_top).ne - · exact ENNReal.zero_ne_top - · intro q - rw [← ENNReal.toReal_iInf] - · suffices ⨅ i : ↥(Ioi q), preCDF ρ (↑i) a = preCDF ρ q a by rw [this] - rw [← h5] - · exact fun r ↦ ((h2 r).trans_lt ENNReal.one_lt_top).ne - -theorem integrable_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℚ) : - Integrable (fun a ↦ (preCDF ρ x a).toReal) ρ.fst := by - refine integrable_of_forall_fin_meas_le _ (measure_lt_top ρ.fst univ) ?_ fun t _ _ ↦ ?_ - · exact measurable_preCDF.ennreal_toReal.aestronglyMeasurable - · simp_rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_of_nonneg ENNReal.toReal_nonneg] - rw [← lintegral_one] - refine (set_lintegral_le_lintegral _ _).trans (lintegral_mono_ae ?_) - filter_upwards [preCDF_le_one ρ] with a ha using ENNReal.ofReal_toReal_le.trans (ha _) - -lemma isRatCondKernelCDF_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : - IsRatCondKernelCDF (fun p r ↦ (preCDF ρ r p.2).toReal) - (kernel.const Unit ρ) (kernel.const Unit ρ.fst) where - measurable := measurable_preCDF'.comp measurable_snd - isRatStieltjesPoint_ae a := by - filter_upwards [isRatStieltjesPoint_ae ρ] with a ha - exact ⟨ha.mono, ha.tendsto_atTop_one, ha.tendsto_atBot_zero, ha.iInf_rat_gt_eq⟩ - integrable _ q := integrable_preCDF ρ q - set_integral a s hs q := by rw [kernel.const_apply, kernel.const_apply, set_integral_preCDF_fst _ _ hs, - Measure.IicSnd_apply _ _ hs] - /-- Conditional cdf of the measure given the value on `α`, as a Stieltjes function. -/ noncomputable def condCDF (ρ : Measure (α × ℝ)) (a : α) : StieltjesFunction := stieltjesOfMeasurableRat (fun a r ↦ (preCDF ρ r a).toReal) measurable_preCDF' a @@ -477,8 +333,8 @@ theorem tendsto_condCDF_atTop (ρ : Measure (α × ℝ)) (a : α) : theorem condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) : (fun a ↦ condCDF ρ a r) =ᵐ[ρ.fst] fun a ↦ (preCDF ρ r a).toReal := by - filter_upwards [isRatStieltjesPoint_ae ρ] with a ha - rw [condCDF, stieltjesOfMeasurableRat_eq, toRatCDF_of_isRatStieltjesPoint ha] + simp_rw [condCDF_eq_stieltjesOfMeasurableRat_unit_prod ρ] + exact stieltjesOfMeasurableRat_ae_eq (isRatCondKernelCDF_preCDF ρ) () r #align probability_theory.cond_cdf_ae_eq ProbabilityTheory.condCDF_ae_eq theorem ofReal_condCDF_ae_eq (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (r : ℚ) : diff --git a/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean index 4ec27d8512fae..0cdb09342eb1e 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean @@ -42,174 +42,57 @@ def kernel.densityIic (κ : kernel α (γ × ℝ)) (ν : kernel α γ) (a : α) kernel.density κ ν a t (Set.Iic q) lemma measurable_densityIic (κ : kernel α (γ × ℝ)) (ν : kernel α γ) : - Measurable (fun p : α × γ ↦ kernel.densityIic κ ν p.1 p.2) := by - rw [measurable_pi_iff] - exact fun _ ↦ measurable_density κ ν measurableSet_Iic - -lemma measurable_densityIic_right (κ : kernel α (γ × ℝ)) (ν : kernel α γ) (a : α) (q : ℚ) : - Measurable (fun t ↦ kernel.densityIic κ ν a t q) := - measurable_density_right _ _ measurableSet_Iic _ - -lemma monotone_densityIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) : - Monotone (kernel.densityIic κ ν a t) := by - intro q r hqr - simp_rw [kernel.densityIic] - refine density_mono_set hκν a t ?_ - refine Iic_subset_Iic.mpr ?_ - exact_mod_cast hqr - -lemma densityIic_nonneg (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (q : ℚ) : - 0 ≤ kernel.densityIic κ ν a t q := - density_nonneg hκν a t _ - -lemma densityIic_le_one (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (q : ℚ) : - kernel.densityIic κ ν a t q ≤ 1 := - density_le_one hκν a t _ - -lemma tendsto_atTop_densityIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), - Tendsto (fun q ↦ kernel.densityIic κ (kernel.fst κ) a t q) atTop (𝓝 1) := by - let ν := kernel.fst κ - suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ kernel.densityIic κ ν a t n) atTop (𝓝 1) by - filter_upwards [this] with t ht - let f := fun q : ℚ ↦ kernel.densityIic κ ν a t ⌊q⌋₊ - let g := fun q : ℚ ↦ kernel.densityIic κ ν a t ⌈q⌉₊ - have hf_le : ∀ᶠ q in atTop, f q ≤ kernel.densityIic κ ν a t q := by - simp only [eventually_atTop, ge_iff_le] - exact ⟨0, fun q hq ↦ monotone_densityIic le_rfl a t (Nat.floor_le hq)⟩ - have hg_le : ∀ q, kernel.densityIic κ ν a t q ≤ g q := - fun q ↦ monotone_densityIic le_rfl a t (Nat.le_ceil _) - refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ hf_le (eventually_of_forall hg_le) - · exact ht.comp tendsto_nat_floor_atTop - · exact ht.comp tendsto_nat_ceil_atTop - let s : ℕ → Set ℝ := fun n ↦ Iic n - have hs : Monotone s := fun i j hij ↦ Iic_subset_Iic.mpr (by exact mod_cast hij) - have hs_iUnion : ⋃ i, s i = univ := by - ext x - simp only [mem_iUnion, mem_Iic, mem_univ, iff_true] - exact ⟨Nat.ceil x, Nat.le_ceil x⟩ - have hs_meas : ∀ n, MeasurableSet (s n) := fun _ ↦ measurableSet_Iic - filter_upwards [tendsto_density_atTop_ae_of_monotone a s hs hs_iUnion hs_meas] - with x hx using hx - -lemma tendsto_atBot_densityIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) : - ∀ᵐ t ∂(ν a), Tendsto (fun q ↦ kernel.densityIic κ ν a t q) atBot (𝓝 0) := by - suffices ∀ᵐ t ∂(ν a), Tendsto (fun q ↦ kernel.densityIic κ ν a t (-q)) atTop (𝓝 0) by - filter_upwards [this] with t ht - have h_eq_neg : (fun q ↦ kernel.densityIic κ ν a t q) - = fun q ↦ kernel.densityIic κ ν a t (- -q) := by - simp_rw [neg_neg] - rw [h_eq_neg] - exact ht.comp tendsto_neg_atBot_atTop - suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ kernel.densityIic κ ν a t (-n)) atTop (𝓝 0) by - filter_upwards [this] with t ht - let f := fun q : ℚ ↦ kernel.densityIic κ ν a t (-⌊q⌋₊) - let g := fun q : ℚ ↦ kernel.densityIic κ ν a t (-⌈q⌉₊) - have hf_le : ∀ᶠ q in atTop, kernel.densityIic κ ν a t (-q) ≤ f q := by - simp only [eventually_atTop, ge_iff_le] - refine ⟨0, fun q hq ↦ monotone_densityIic hκν a t ?_⟩ - rw [neg_le_neg_iff] - exact Nat.floor_le hq - have hg_le : ∀ q, g q ≤ kernel.densityIic κ ν a t (-q) := by - refine fun q ↦ monotone_densityIic hκν a t ?_ - rw [neg_le_neg_iff] - exact Nat.le_ceil _ - refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ (eventually_of_forall hg_le) hf_le - · exact ht.comp tendsto_nat_ceil_atTop - · exact ht.comp tendsto_nat_floor_atTop - let s : ℕ → Set ℝ := fun n ↦ Iic (-n) - have hs : Antitone s := fun i j hij ↦ Iic_subset_Iic.mpr (neg_le_neg (by exact mod_cast hij)) - have hs_iInter : ⋂ i, s i = ∅ := by - ext x - simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, neg_lt] - exact exists_nat_gt (-x) - have hs_meas : ∀ n, MeasurableSet (s n) := fun _ ↦ measurableSet_Iic - convert tendsto_density_atTop_ae_of_antitone hκν a s hs hs_iInter hs_meas with x n - simp [kernel.densityIic] - -lemma set_integral_densityIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) (q : ℚ) {A : Set γ} (hA : MeasurableSet A) : - ∫ t in A, kernel.densityIic κ ν a t q ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := - set_integral_density hκν a measurableSet_Iic hA - -lemma integrable_densityIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] - (a : α) (q : ℚ) : - Integrable (fun t ↦ kernel.densityIic κ ν a t q) (ν a) := - integrable_density hκν _ measurableSet_Iic - -lemma bddBelow_range_densityIic (hκν : kernel.fst κ ≤ ν) (a : α) (t : γ) (q : ℚ) : - BddBelow (range fun (r : Ioi q) ↦ kernel.densityIic κ ν a t r) := by - refine ⟨0, ?_⟩ - rw [mem_lowerBounds] - rintro x ⟨y, rfl⟩ - exact densityIic_nonneg hκν _ _ _ - -lemma integrable_iInf_rat_gt_densityIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel ν] - (a : α) (q : ℚ) : - Integrable (fun t ↦ ⨅ r : Ioi q, kernel.densityIic κ ν a t r) (ν a) := by - rw [← memℒp_one_iff_integrable] - refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ - · exact measurable_iInf fun i ↦ measurable_densityIic_right κ ν a i - refine (?_ : _ ≤ (ν a univ : ℝ≥0∞)).trans_lt (measure_lt_top _ _) - refine (snorm_le_of_ae_bound (C := 1) (ae_of_all _ (fun t ↦ ?_))).trans ?_ - · rw [Real.norm_eq_abs, abs_of_nonneg] - · refine ciInf_le_of_le ?_ ?_ ?_ - · exact bddBelow_range_densityIic hκν a t _ - · exact ⟨q + 1, by simp⟩ - · exact densityIic_le_one hκν _ _ _ - · exact le_ciInf fun r ↦ densityIic_nonneg hκν a t r - · simp - -lemma set_integral_iInf_rat_gt_densityIic (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] - [IsFiniteKernel ν] (a : α) (q : ℚ) {A : Set γ} (hA : MeasurableSet A) : - ∫ t in A, ⨅ r : Ioi q, kernel.densityIic κ ν a t r ∂(ν a) - = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by - refine le_antisymm ?_ ?_ - · have h : ∀ r : Ioi q, ∫ t in A, ⨅ r' : Ioi q, kernel.densityIic κ ν a t r' ∂(ν a) - ≤ (κ a (A ×ˢ Iic (r : ℝ))).toReal := by - intro r - rw [← set_integral_densityIic hκν a r hA] - refine set_integral_mono ?_ ?_ ?_ - · exact (integrable_iInf_rat_gt_densityIic hκν _ _).integrableOn - · exact (integrable_densityIic hκν _ _).integrableOn - · exact fun t ↦ ciInf_le (bddBelow_range_densityIic hκν _ _ _) r - calc ∫ t in A, ⨅ r : Ioi q, kernel.densityIic κ ν a t r ∂(ν a) - ≤ ⨅ r : Ioi q, (κ a (A ×ˢ Iic (r : ℝ))).toReal := le_ciInf h - _ = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by - rw [← Measure.iInf_rat_gt_prod_Iic hA q] - exact (ENNReal.toReal_iInf (fun r ↦ measure_ne_top _ _)).symm - · rw [← set_integral_densityIic hκν a q hA] - refine set_integral_mono ?_ ?_ ?_ - · exact (integrable_densityIic hκν _ _).integrableOn - · exact (integrable_iInf_rat_gt_densityIic hκν _ _).integrableOn - · exact fun t ↦ le_ciInf (fun r ↦ monotone_densityIic hκν _ _ (le_of_lt r.prop)) - -lemma iInf_rat_gt_densityIic_eq (hκν : kernel.fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) : - ∀ᵐ t ∂(ν a), ∀ q : ℚ, ⨅ r : Ioi q, kernel.densityIic κ ν a t r - = kernel.densityIic κ ν a t q := by - rw [ae_all_iff] - refine fun q ↦ ae_eq_of_forall_set_integral_eq_of_sigmaFinite (μ := ν a) ?_ ?_ ?_ - · exact fun s _ _ ↦ (integrable_iInf_rat_gt_densityIic hκν _ _).integrableOn - · exact fun s _ _ ↦ (integrable_densityIic hκν a q).integrableOn - · intro s hs _ - rw [set_integral_densityIic hκν _ _ hs, set_integral_iInf_rat_gt_densityIic hκν _ _ hs] - -lemma isRatStieltjesPoint_densityIic_ae (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ t ∂(kernel.fst κ a), - IsRatStieltjesPoint (fun p q ↦ kernel.densityIic κ (kernel.fst κ) p.1 p.2 q) (a, t) := by - filter_upwards [tendsto_atTop_densityIic κ a, tendsto_atBot_densityIic le_rfl a, - iInf_rat_gt_densityIic_eq le_rfl a] with t ht_top ht_bot ht_iInf - exact ⟨monotone_densityIic le_rfl a t, ht_top, ht_bot, ht_iInf⟩ + Measurable (fun p : α × γ ↦ kernel.densityIic κ ν p.1 p.2) := + measurable_pi_iff.mpr <| fun _ ↦ measurable_density κ ν measurableSet_Iic + +lemma isRatCondKernelCDFAux_densityIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + isRatCondKernelCDFAux (Function.uncurry (kernel.densityIic κ (kernel.fst κ))) κ (kernel.fst κ) + where + measurable := measurable_densityIic _ _ + mono' a q r hqr := + ae_of_all _ fun c ↦ density_mono_set le_rfl a c (Iic_subset_Iic.mpr (by exact_mod_cast hqr)) + nonneg' a q := ae_of_all _ fun c ↦ density_nonneg le_rfl _ _ _ + le_one' a q := ae_of_all _ fun c ↦ density_le_one le_rfl _ _ _ + tendsto_integral_of_antitone a s hs_anti hs_tendsto := by + let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) + have hs'_anti : Antitone s' := by + refine fun i j hij ↦ Iic_subset_Iic.mpr ?_ + exact mod_cast hs_anti hij + have hs'_iInter : ⋂ i, s' i = ∅ := by + ext x + simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, neg_lt] + rw [tendsto_atTop_atBot] at hs_tendsto + have ⟨q, hq⟩ := exists_rat_lt x + obtain ⟨i, hi⟩ := hs_tendsto q + refine ⟨i, lt_of_le_of_lt ?_ hq⟩ + exact mod_cast hi i le_rfl + have hs'_meas : ∀ n, MeasurableSet (s' n) := fun _ ↦ measurableSet_Iic + exact tendsto_integral_density_of_antitone (le_rfl : kernel.fst κ ≤ kernel.fst κ) + a s' hs'_anti hs'_iInter hs'_meas + tendsto_integral_of_monotone a s hs_mono hs_tendsto := by + let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) + have hs'_mono : Monotone s' := fun i j hij ↦ Iic_subset_Iic.mpr (by exact mod_cast hs_mono hij) + have hs'_iUnion : ⋃ i, s' i = univ := by + ext x + simp only [mem_iUnion, mem_Iic, mem_univ, iff_true] + rw [tendsto_atTop_atTop] at hs_tendsto + have ⟨q, hq⟩ := exists_rat_gt x + obtain ⟨i, hi⟩ := hs_tendsto q + refine ⟨i, hq.le.trans ?_⟩ + exact mod_cast hi i le_rfl + have hs'_meas : ∀ n, MeasurableSet (s' n) := fun _ ↦ measurableSet_Iic + have h := tendsto_integral_density_of_monotone (le_rfl : kernel.fst κ ≤ kernel.fst κ) + a s' hs'_mono hs'_iUnion hs'_meas + convert h + rw [kernel.fst_apply' _ _ MeasurableSet.univ] + rfl + integrable a q := integrable_density le_rfl a measurableSet_Iic + set_integral a A hA q := set_integral_density le_rfl a measurableSet_Iic hA lemma isRatCondKernelCDF_densityIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsRatCondKernelCDF (fun p : α × γ ↦ kernel.densityIic κ (kernel.fst κ) p.1 p.2) κ (kernel.fst κ) - where - measurable := measurable_densityIic κ (kernel.fst κ) - isRatStieltjesPoint_ae := isRatStieltjesPoint_densityIic_ae κ - integrable := integrable_densityIic le_rfl - set_integral := fun _ _ hs _ ↦ set_integral_densityIic le_rfl _ _ hs + IsRatCondKernelCDF (fun p : α × γ ↦ kernel.densityIic κ (kernel.fst κ) p.1 p.2) κ + (kernel.fst κ) := + (isRatCondKernelCDFAux_densityIic κ).isRatCondKernelCDF noncomputable def condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : α × γ → StieltjesFunction := diff --git a/Mathlib/Probability/Kernel/Disintegration/DensityToCdf.lean b/Mathlib/Probability/Kernel/Disintegration/DensityToCdf.lean deleted file mode 100644 index 91f417b1faa21..0000000000000 --- a/Mathlib/Probability/Kernel/Disintegration/DensityToCdf.lean +++ /dev/null @@ -1,245 +0,0 @@ -/- -Copyright (c) 2024 Rémy Degenne. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Rémy Degenne --/ -import Mathlib.Probability.Kernel.Disintegration.CdfToKernel -import Mathlib.Probability.Kernel.Disintegration.Density -import Mathlib.Probability.Kernel.Disintegration.AuxLemmas - -/-! -# Kernel CDF - -We prove IsRatCondKernelCDF from conditions on integrals. - -## Main definitions - -* `FooBar` - -## Main statements - -* `fooBar_unique` - -## Implementation details - - -## References - --/ - - -open MeasureTheory Set Filter - -open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory - -namespace ProbabilityTheory - -open ProbabilityTheory.kernel - -variable {α γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} - {κ : kernel α (γ × ℝ)} {ν : kernel α γ} - {f : α × γ → ℚ → ℝ} - -structure IsKernelPDF (f : α × γ → ℚ → ℝ) (κ : kernel α (γ × ℝ)) (ν : kernel α γ) : Prop := - (measurable : Measurable f) - (mono' (a : α) {q r : ℚ} (_hqr : q ≤ r) : ∀ᵐ c ∂(ν a), f (a, c) q ≤ f (a, c) r) - (nonneg' (a : α) (q : ℚ) : ∀ᵐ c ∂(ν a), 0 ≤ f (a, c) q) - (le_one' (a : α) (q : ℚ) : ∀ᵐ c ∂(ν a), f (a, c) q ≤ 1) - (tendsto_integral_zero_of_antitone (a : α) (s : ℕ → ℚ) (_hs : Antitone s) - (_hs_tendsto : Tendsto s atTop atBot) : - Tendsto (fun m ↦ ∫ c, f (a, c) (s m) ∂(ν a)) atTop (𝓝 0)) - (tendsto_integral_one_of_monotone (a : α) (s : ℕ → ℚ) (_hs : Monotone s) - (_hs_tendsto : Tendsto s atTop atTop) : - Tendsto (fun m ↦ ∫ c, f (a, c) (s m) ∂(ν a)) atTop (𝓝 (ν a univ).toReal)) - (integrable (a : α) (q : ℚ) : Integrable (fun c ↦ f (a, c) q) (ν a)) - (set_integral (a : α) {A : Set γ} (_hA : MeasurableSet A) (q : ℚ) : - ∫ c in A, f (a, c) q ∂(ν a) = (κ a (A ×ˢ Iic ↑q)).toReal) - -lemma IsKernelPDF.measurable_right (hf : IsKernelPDF f κ ν) (a : α) (q : ℚ) : - Measurable (fun t ↦ f (a, t) q) := by - let h := hf.measurable - rw [measurable_pi_iff] at h - exact (h q).comp measurable_prod_mk_left - -lemma IsKernelPDF.mono (hf : IsKernelPDF f κ ν) (a : α) : - ∀ᵐ c ∂(ν a), ∀ {q r : ℚ} (_ : q ≤ r), f (a, c) q ≤ f (a, c) r := by - simp_rw [ae_all_iff] - intro q r hqr - exact hf.mono' a hqr - -lemma IsKernelPDF.nonneg (hf : IsKernelPDF f κ ν) (a : α) : - ∀ᵐ c ∂(ν a), ∀ (q : ℚ), 0 ≤ f (a, c) q := by - rw [ae_all_iff] - exact hf.nonneg' a - -lemma IsKernelPDF.le_one (hf : IsKernelPDF f κ ν) (a : α) : - ∀ᵐ c ∂(ν a), ∀ (q : ℚ), f (a, c) q ≤ 1 := by - rw [ae_all_iff] - exact hf.le_one' a - -lemma IsKernelPDF.tendsto_zero_of_antitone (hf : IsKernelPDF f κ ν) [IsFiniteKernel ν] (a : α) - (s : ℕ → ℚ) (hs : Antitone s) (hs_tendsto : Tendsto s atTop atBot) : - ∀ᵐ c ∂(ν a), Tendsto (fun m ↦ f (a, c) (s m)) atTop (𝓝 0) := by - refine tendsto_of_integral_tendsto_of_antitone ?_ (integrable_const _) ?_ ?_ ?_ - · exact fun n ↦ hf.integrable a (s n) - · rw [integral_zero] - exact hf.tendsto_integral_zero_of_antitone a s hs hs_tendsto - · filter_upwards [hf.mono a] with t ht - exact fun n m hnm ↦ ht (hs hnm) - · filter_upwards [hf.nonneg a] with c hc using fun i ↦ hc (s i) - -lemma IsKernelPDF.tendsto_one_of_monotone (hf : IsKernelPDF f κ ν) [IsFiniteKernel ν] (a : α) - (s : ℕ → ℚ) (hs : Monotone s) (hs_tendsto : Tendsto s atTop atTop) : - ∀ᵐ c ∂(ν a), Tendsto (fun m ↦ f (a, c) (s m)) atTop (𝓝 1) := by - refine tendsto_of_integral_tendsto_of_monotone ?_ (integrable_const _) ?_ ?_ ?_ - · exact fun n ↦ hf.integrable a (s n) - · rw [MeasureTheory.integral_const, smul_eq_mul, mul_one] - exact hf.tendsto_integral_one_of_monotone a s hs hs_tendsto - · filter_upwards [hf.mono a] with t ht - exact fun n m hnm ↦ ht (hs hnm) - · filter_upwards [hf.le_one a] with c hc using fun i ↦ hc (s i) - -lemma tendsto_atTop_densityIic (hf : IsKernelPDF f κ ν) [IsFiniteKernel ν] (a : α) : - ∀ᵐ t ∂(ν a), Tendsto (fun q : ℚ ↦ f (a, t) q) atTop (𝓝 1) := by - suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ f (a, t) n) atTop (𝓝 1) by - filter_upwards [this, hf.mono a] with t ht h_mono - let f' := fun q : ℚ ↦ f (a, t) ⌊q⌋₊ - let g := fun q : ℚ ↦ f (a, t) ⌈q⌉₊ - have hf_le : ∀ᶠ q in atTop, f' q ≤ f (a, t) q := by - simp only [eventually_atTop, ge_iff_le] - exact ⟨0, fun q hq ↦ h_mono (Nat.floor_le hq)⟩ - have hg_le : ∀ q : ℚ, f (a, t) q ≤ g q := by - exact fun q ↦ h_mono (Nat.le_ceil _) - refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ hf_le (eventually_of_forall hg_le) - · exact ht.comp tendsto_nat_floor_atTop - · exact ht.comp tendsto_nat_ceil_atTop - let s : ℕ → ℚ := fun n ↦ n - have hs : Monotone s := fun i j hij ↦ by simp [s, hij] - have hs_tendsto : Tendsto s atTop atTop := by - exact tendsto_nat_cast_atTop_atTop - filter_upwards [hf.tendsto_one_of_monotone a s hs hs_tendsto] with x hx using hx - -lemma tendsto_atBot_densityIic (hf : IsKernelPDF f κ ν) [IsFiniteKernel ν] (a : α) : - ∀ᵐ t ∂(ν a), Tendsto (fun q : ℚ ↦ f (a, t) q) atBot (𝓝 0) := by - suffices ∀ᵐ t ∂(ν a), Tendsto (fun q : ℚ ↦ f (a, t) (-q)) atTop (𝓝 0) by - filter_upwards [this] with t ht - have h_eq_neg : (fun q : ℚ ↦ f (a, t) q) = fun q : ℚ ↦ f (a, t) (- -q) := by - simp_rw [neg_neg] - rw [h_eq_neg] - convert ht.comp tendsto_neg_atBot_atTop - simp - suffices ∀ᵐ t ∂(ν a), Tendsto (fun (n : ℕ) ↦ f (a, t) (-n)) atTop (𝓝 0) by - filter_upwards [this, hf.mono a] with t ht h_mono - let f' := fun q : ℚ ↦ f (a, t) (-⌊q⌋₊) - let g := fun q : ℚ ↦ f (a, t) (-⌈q⌉₊) - have hf_le : ∀ᶠ (q : ℚ) in atTop, f (a, t) (-q) ≤ f' q := by - simp only [eventually_atTop, ge_iff_le] - refine ⟨0, fun q hq ↦ ?_⟩ - norm_cast - refine h_mono ?_ - simp only [Rat.ofInt_eq_cast, Int.cast_neg, Int.cast_ofNat, neg_le_neg_iff] - exact Nat.floor_le hq - have hg_le : ∀ q, g q ≤ f (a, t) (-q) := by - refine fun q ↦ ?_ - simp only [g] - norm_cast - refine h_mono ?_ - simp only [Rat.ofInt_eq_cast, Int.cast_neg, Int.cast_ofNat, neg_le_neg_iff] - exact Nat.le_ceil _ - refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ (eventually_of_forall hg_le) hf_le - · exact ht.comp tendsto_nat_ceil_atTop - · exact ht.comp tendsto_nat_floor_atTop - let s : ℕ → ℚ := fun n ↦ -n - have hs : Antitone s := fun i j hij ↦ neg_le_neg (by exact mod_cast hij) - have hs_tendsto : Tendsto s atTop atBot := by - simp only [s, tendsto_neg_atBot_iff] - exact tendsto_nat_cast_atTop_atTop - convert hf.tendsto_zero_of_antitone a s hs hs_tendsto with x n - -lemma bddBelow_range_densityIic (hf : IsKernelPDF f κ ν) (a : α) : - ∀ᵐ t ∂(ν a), ∀ q : ℚ, BddBelow (range fun (r : Ioi q) ↦ f (a, t) r) := by - filter_upwards [hf.nonneg a] with c hc - refine fun q ↦ ⟨0, ?_⟩ - rw [mem_lowerBounds] - rintro x ⟨y, rfl⟩ - exact hc y - -lemma integrable_iInf_rat_gt_densityIic (hf : IsKernelPDF f κ ν) [IsFiniteKernel ν] - (a : α) (q : ℚ) : - Integrable (fun t ↦ ⨅ r : Ioi q, f (a, t) r) (ν a) := by - rw [← memℒp_one_iff_integrable] - refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ - · exact measurable_iInf fun i ↦ hf.measurable_right a _ - refine (?_ : _ ≤ (ν a univ : ℝ≥0∞)).trans_lt (measure_lt_top _ _) - refine (snorm_le_of_ae_bound (C := 1) ?_).trans ?_ - · filter_upwards [bddBelow_range_densityIic hf a, hf.nonneg a, hf.le_one a] - with t hbdd_below h_nonneg h_le_one - rw [Real.norm_eq_abs, abs_of_nonneg] - · refine ciInf_le_of_le ?_ ?_ ?_ - · exact hbdd_below _ - · exact ⟨q + 1, by simp⟩ - · exact h_le_one _ - · exact le_ciInf fun r ↦ h_nonneg _ - · simp - -lemma set_integral_iInf_rat_gt_densityIic (hf : IsKernelPDF f κ ν) [IsFiniteKernel κ] - [IsFiniteKernel ν] (a : α) (q : ℚ) {A : Set γ} (hA : MeasurableSet A) : - ∫ t in A, ⨅ r : Ioi q, f (a, t) r ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by - refine le_antisymm ?_ ?_ - · have h : ∀ r : Ioi q, ∫ t in A, ⨅ r' : Ioi q, f (a, t) r' ∂(ν a) - ≤ (κ a (A ×ˢ Iic (r : ℝ))).toReal := by - intro r - rw [← hf.set_integral a hA] - refine set_integral_mono_ae ?_ ?_ ?_ - · exact (integrable_iInf_rat_gt_densityIic hf _ _).integrableOn - · exact (hf.integrable _ _).integrableOn - · filter_upwards [bddBelow_range_densityIic hf a] with t ht - exact ciInf_le (ht _) r - calc ∫ t in A, ⨅ r : Ioi q, f (a, t) r ∂(ν a) - ≤ ⨅ r : Ioi q, (κ a (A ×ˢ Iic (r : ℝ))).toReal := le_ciInf h - _ = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by - rw [← Measure.iInf_rat_gt_prod_Iic hA q] - exact (ENNReal.toReal_iInf (fun r ↦ measure_ne_top _ _)).symm - · rw [← hf.set_integral a hA] - refine set_integral_mono_ae ?_ ?_ ?_ - · exact (hf.integrable _ _).integrableOn - · exact (integrable_iInf_rat_gt_densityIic hf _ _).integrableOn - · filter_upwards [hf.mono a] with c h_mono - exact le_ciInf (fun r ↦ h_mono (le_of_lt r.prop)) - -lemma iInf_rat_gt_densityIic_eq (hf : IsKernelPDF f κ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] - (a : α) : - ∀ᵐ t ∂(ν a), ∀ q : ℚ, ⨅ r : Ioi q, f (a, t) r = f (a, t) q := by - rw [ae_all_iff] - refine fun q ↦ ae_eq_of_forall_set_integral_eq_of_sigmaFinite (μ := ν a) ?_ ?_ ?_ - · intro s _ _ - refine Integrable.integrableOn ?_ - exact integrable_iInf_rat_gt_densityIic hf _ _ - · exact fun s _ _ ↦ (hf.integrable a _).integrableOn - · intro s hs _ - rw [hf.set_integral _ hs, set_integral_iInf_rat_gt_densityIic hf _ _ hs] - -lemma isRatStieltjesPoint_densityIic_ae (hf : IsKernelPDF f κ ν) [IsFiniteKernel κ] - [IsFiniteKernel ν] (a : α) : - ∀ᵐ t ∂(ν a), IsRatStieltjesPoint f (a, t) := by - filter_upwards [tendsto_atTop_densityIic hf a, tendsto_atBot_densityIic hf a, - iInf_rat_gt_densityIic_eq hf a, hf.mono a] with t ht_top ht_bot ht_iInf h_mono - exact ⟨fun _ _ ↦ h_mono, ht_top, ht_bot, ht_iInf⟩ - -lemma isRatCondKernelCDF_densityIic (hf : IsKernelPDF f κ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] : - IsRatCondKernelCDF f κ ν where - measurable := hf.measurable - isRatStieltjesPoint_ae := isRatStieltjesPoint_densityIic_ae hf - integrable := hf.integrable - set_integral := hf.set_integral - -noncomputable -def condKernelCDF (hf : IsKernelPDF f κ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] : - α × γ → StieltjesFunction := - stieltjesOfMeasurableRat f (isRatCondKernelCDF_densityIic hf).measurable - -lemma isCondKernelCDF_condKernelCDF (hf : IsKernelPDF f κ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] : - IsCondKernelCDF (condKernelCDF hf) κ ν := - isCondKernelCDF_stieltjesOfMeasurableRat (isRatCondKernelCDF_densityIic hf) - -end ProbabilityTheory diff --git a/Mathlib/Topology/Algebra/Order/MonotoneConvergence.lean b/Mathlib/Topology/Algebra/Order/MonotoneConvergence.lean index ae4adcba5e268..cb0fb8cb367d1 100644 --- a/Mathlib/Topology/Algebra/Order/MonotoneConvergence.lean +++ b/Mathlib/Topology/Algebra/Order/MonotoneConvergence.lean @@ -244,6 +244,16 @@ theorem tendsto_iff_tendsto_subseq_of_monotone {ι₁ ι₂ α : Type*} [Semilat · rwa [tendsto_nhds_unique h (hl'.comp hg)] #align tendsto_iff_tendsto_subseq_of_monotone tendsto_iff_tendsto_subseq_of_monotone +theorem tendsto_iff_tendsto_subseq_of_antitone {ι₁ ι₂ α : Type*} [SemilatticeSup ι₁] [Preorder ι₂] + [Nonempty ι₁] [TopologicalSpace α] [ConditionallyCompleteLinearOrder α] [OrderTopology α] + [NoMinOrder α] {f : ι₂ → α} {φ : ι₁ → ι₂} {l : α} (hf : Antitone f) + (hg : Tendsto φ atTop atTop) : Tendsto f atTop (𝓝 l) ↔ Tendsto (f ∘ φ) atTop (𝓝 l) := by + constructor <;> intro h + · exact h.comp hg + · rcases tendsto_of_antitone hf with (h' | ⟨l', hl'⟩) + · exact (not_tendsto_atBot_of_tendsto_nhds h (h'.comp hg)).elim + · rwa [tendsto_nhds_unique h (hl'.comp hg)] + /-! The next family of results, such as `isLUB_of_tendsto_atTop` and `iSup_eq_of_tendsto`, are converses to the standard fact that bounded monotone functions converge. They state, that if a monotone function `f` tends to `a` along `Filter.atTop`, then that value `a` is a least upper bound From b70a68b9c7d81f37eb1ae8e08e9a610acf6290e5 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 6 Mar 2024 13:25:15 +0100 Subject: [PATCH 085/129] fix --- Mathlib/MeasureTheory/Integral/Lebesgue.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index 7f5c696c89a62..20e9d78da3f17 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -1694,7 +1694,7 @@ lemma tendsto_of_lintegral_tendsto_of_monotone {α : Type*} {mα : MeasurableSpa then h.choose else ∞ have hF'_tendsto : ∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) atTop (𝓝 (F' a)) := by filter_upwards [h_exists] with a ha - simp_rw [dif_pos ha] + simp_rw [F', dif_pos ha] exact ha.choose_spec suffices F' =ᵐ[μ] F by filter_upwards [this, hF'_tendsto] with a h_eq h_tendsto using h_eq ▸ h_tendsto @@ -1772,7 +1772,7 @@ lemma tendsto_of_lintegral_tendsto_of_antitone {α : Type*} {mα : MeasurableSpa then h.choose else ∞ have hF'_tendsto : ∀ᵐ a ∂μ, Tendsto (fun i ↦ f i a) atTop (𝓝 (F' a)) := by filter_upwards [h_exists] with a ha - simp_rw [dif_pos ha] + simp_rw [F', dif_pos ha] exact ha.choose_spec suffices F' =ᵐ[μ] F by filter_upwards [this, hF'_tendsto] with a h_eq h_tendsto using h_eq ▸ h_tendsto From b69bd871ee2f15782da8a294c5c154f33849267c Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 6 Mar 2024 14:12:07 +0100 Subject: [PATCH 086/129] fix --- Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean index 0cdb09342eb1e..e8b3f4c7333a7 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean @@ -60,7 +60,7 @@ lemma isRatCondKernelCDFAux_densityIic (κ : kernel α (γ × ℝ)) [IsFiniteKer exact mod_cast hs_anti hij have hs'_iInter : ⋂ i, s' i = ∅ := by ext x - simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, neg_lt] + simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, s'] rw [tendsto_atTop_atBot] at hs_tendsto have ⟨q, hq⟩ := exists_rat_lt x obtain ⟨i, hi⟩ := hs_tendsto q From c72e1d7b33339fd6ee6ae04151879a425bd91ccb Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 6 Mar 2024 14:40:46 +0100 Subject: [PATCH 087/129] delete a file --- .../Kernel/Disintegration/Basic.lean | 55 ++++++++- .../Kernel/Disintegration/CondKernelCdf.lean | 106 ------------------ 2 files changed, 54 insertions(+), 107 deletions(-) delete mode 100644 Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 191968f78f0c7..7f2780b44f7b8 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -5,7 +5,8 @@ Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.MeasureCompProd import Mathlib.Probability.Kernel.Disintegration.CondCdf -import Mathlib.Probability.Kernel.Disintegration.CondKernelCdf +import Mathlib.Probability.Kernel.Disintegration.Density +import Mathlib.Probability.Kernel.Disintegration.CdfToKernel /-! # Disintegration of kernels and measures @@ -211,6 +212,58 @@ end BorelSnd section CountablyGenerated +open ProbabilityTheory.kernel + +lemma isRatCondKernelCDFAux_density_Iic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + isRatCondKernelCDFAux (fun (p : α × γ) q ↦ kernel.density κ (kernel.fst κ) p.1 p.2 (Set.Iic q)) + κ (kernel.fst κ) where + measurable := measurable_pi_iff.mpr fun _ ↦ measurable_density κ (kernel.fst κ) measurableSet_Iic + mono' a q r hqr := + ae_of_all _ fun c ↦ density_mono_set le_rfl a c (Iic_subset_Iic.mpr (by exact_mod_cast hqr)) + nonneg' a q := ae_of_all _ fun c ↦ density_nonneg le_rfl _ _ _ + le_one' a q := ae_of_all _ fun c ↦ density_le_one le_rfl _ _ _ + tendsto_integral_of_antitone a s hs_anti hs_tendsto := by + let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) + refine tendsto_integral_density_of_antitone le_rfl a s' ?_ ?_ (fun _ ↦ measurableSet_Iic) + · refine fun i j hij ↦ Iic_subset_Iic.mpr ?_ + exact mod_cast hs_anti hij + · ext x + simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, s'] + rw [tendsto_atTop_atBot] at hs_tendsto + have ⟨q, hq⟩ := exists_rat_lt x + obtain ⟨i, hi⟩ := hs_tendsto q + refine ⟨i, lt_of_le_of_lt ?_ hq⟩ + exact mod_cast hi i le_rfl + tendsto_integral_of_monotone a s hs_mono hs_tendsto := by + rw [kernel.fst_apply' _ _ MeasurableSet.univ] + let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) + refine tendsto_integral_density_of_monotone (le_rfl : kernel.fst κ ≤ kernel.fst κ) + a s' ?_ ?_ (fun _ ↦ measurableSet_Iic) + · exact fun i j hij ↦ Iic_subset_Iic.mpr (by exact mod_cast hs_mono hij) + · ext x + simp only [mem_iUnion, mem_Iic, mem_univ, iff_true] + rw [tendsto_atTop_atTop] at hs_tendsto + have ⟨q, hq⟩ := exists_rat_gt x + obtain ⟨i, hi⟩ := hs_tendsto q + refine ⟨i, hq.le.trans ?_⟩ + exact mod_cast hi i le_rfl + integrable a q := integrable_density le_rfl a measurableSet_Iic + set_integral a A hA q := set_integral_density le_rfl a measurableSet_Iic hA + +lemma isRatCondKernelCDF_density_Iic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsRatCondKernelCDF (fun (p : α × γ) q ↦ kernel.density κ (kernel.fst κ) p.1 p.2 (Set.Iic q)) κ + (kernel.fst κ) := + (isRatCondKernelCDFAux_density_Iic κ).isRatCondKernelCDF + +noncomputable +def condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : α × γ → StieltjesFunction := + stieltjesOfMeasurableRat (fun (p : α × γ) q ↦ kernel.density κ (kernel.fst κ) p.1 p.2 (Set.Iic q)) + (isRatCondKernelCDF_density_Iic κ).measurable + +lemma isCondKernelCDF_condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsCondKernelCDF (condKernelCDF κ) κ (kernel.fst κ) := + isCondKernelCDF_stieltjesOfMeasurableRat (isRatCondKernelCDF_density_Iic κ) + noncomputable def condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : kernel (α × γ) Ω := let f := measurableEmbedding_real Ω diff --git a/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean deleted file mode 100644 index e8b3f4c7333a7..0000000000000 --- a/Mathlib/Probability/Kernel/Disintegration/CondKernelCdf.lean +++ /dev/null @@ -1,106 +0,0 @@ -/- -Copyright (c) 2024 Rémy Degenne. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Rémy Degenne --/ -import Mathlib.Probability.Kernel.Disintegration.CdfToKernel -import Mathlib.Probability.Kernel.Disintegration.Density -import Mathlib.Probability.Kernel.Disintegration.AuxLemmas - -/-! -# Kernel CDF - -## Main definitions - -* `FooBar` - -## Main statements - -* `fooBar_unique` - -## Implementation details - - -## References - --/ - - -open MeasureTheory Set Filter - -open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory - -namespace ProbabilityTheory - -open ProbabilityTheory.kernel - -variable {α γ : Type*} {mα : MeasurableSpace α} {mγ : MeasurableSpace γ} - [MeasurableSpace.CountablyGenerated γ] {κ : kernel α (γ × ℝ)} {ν : kernel α γ} - -noncomputable -def kernel.densityIic (κ : kernel α (γ × ℝ)) (ν : kernel α γ) (a : α) (t : γ) (q : ℚ) : ℝ := - kernel.density κ ν a t (Set.Iic q) - -lemma measurable_densityIic (κ : kernel α (γ × ℝ)) (ν : kernel α γ) : - Measurable (fun p : α × γ ↦ kernel.densityIic κ ν p.1 p.2) := - measurable_pi_iff.mpr <| fun _ ↦ measurable_density κ ν measurableSet_Iic - -lemma isRatCondKernelCDFAux_densityIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - isRatCondKernelCDFAux (Function.uncurry (kernel.densityIic κ (kernel.fst κ))) κ (kernel.fst κ) - where - measurable := measurable_densityIic _ _ - mono' a q r hqr := - ae_of_all _ fun c ↦ density_mono_set le_rfl a c (Iic_subset_Iic.mpr (by exact_mod_cast hqr)) - nonneg' a q := ae_of_all _ fun c ↦ density_nonneg le_rfl _ _ _ - le_one' a q := ae_of_all _ fun c ↦ density_le_one le_rfl _ _ _ - tendsto_integral_of_antitone a s hs_anti hs_tendsto := by - let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) - have hs'_anti : Antitone s' := by - refine fun i j hij ↦ Iic_subset_Iic.mpr ?_ - exact mod_cast hs_anti hij - have hs'_iInter : ⋂ i, s' i = ∅ := by - ext x - simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, s'] - rw [tendsto_atTop_atBot] at hs_tendsto - have ⟨q, hq⟩ := exists_rat_lt x - obtain ⟨i, hi⟩ := hs_tendsto q - refine ⟨i, lt_of_le_of_lt ?_ hq⟩ - exact mod_cast hi i le_rfl - have hs'_meas : ∀ n, MeasurableSet (s' n) := fun _ ↦ measurableSet_Iic - exact tendsto_integral_density_of_antitone (le_rfl : kernel.fst κ ≤ kernel.fst κ) - a s' hs'_anti hs'_iInter hs'_meas - tendsto_integral_of_monotone a s hs_mono hs_tendsto := by - let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) - have hs'_mono : Monotone s' := fun i j hij ↦ Iic_subset_Iic.mpr (by exact mod_cast hs_mono hij) - have hs'_iUnion : ⋃ i, s' i = univ := by - ext x - simp only [mem_iUnion, mem_Iic, mem_univ, iff_true] - rw [tendsto_atTop_atTop] at hs_tendsto - have ⟨q, hq⟩ := exists_rat_gt x - obtain ⟨i, hi⟩ := hs_tendsto q - refine ⟨i, hq.le.trans ?_⟩ - exact mod_cast hi i le_rfl - have hs'_meas : ∀ n, MeasurableSet (s' n) := fun _ ↦ measurableSet_Iic - have h := tendsto_integral_density_of_monotone (le_rfl : kernel.fst κ ≤ kernel.fst κ) - a s' hs'_mono hs'_iUnion hs'_meas - convert h - rw [kernel.fst_apply' _ _ MeasurableSet.univ] - rfl - integrable a q := integrable_density le_rfl a measurableSet_Iic - set_integral a A hA q := set_integral_density le_rfl a measurableSet_Iic hA - -lemma isRatCondKernelCDF_densityIic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsRatCondKernelCDF (fun p : α × γ ↦ kernel.densityIic κ (kernel.fst κ) p.1 p.2) κ - (kernel.fst κ) := - (isRatCondKernelCDFAux_densityIic κ).isRatCondKernelCDF - -noncomputable -def condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : α × γ → StieltjesFunction := - stieltjesOfMeasurableRat (fun p : α × γ ↦ kernel.densityIic κ (kernel.fst κ) p.1 p.2) - (isRatCondKernelCDF_densityIic κ).measurable - -lemma isCondKernelCDF_condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsCondKernelCDF (condKernelCDF κ) κ (kernel.fst κ) := - isCondKernelCDF_stieltjesOfMeasurableRat (isRatCondKernelCDF_densityIic κ) - -end ProbabilityTheory From 668182e0a8148d30513b3501597ab599db419974 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 6 Mar 2024 14:41:13 +0100 Subject: [PATCH 088/129] Mathlib.lean --- Mathlib.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib.lean b/Mathlib.lean index 6ba2673ba75c3..463a16741acfa 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3175,7 +3175,6 @@ import Mathlib.Probability.Kernel.Condexp import Mathlib.Probability.Kernel.Disintegration.Basic import Mathlib.Probability.Kernel.Disintegration.CdfToKernel import Mathlib.Probability.Kernel.Disintegration.CondCdf -import Mathlib.Probability.Kernel.Disintegration.CondKernelCdf import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.Disintegration.Integral import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes From 829495eebb8197175b2c068d2eb22a439fad25f6 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 6 Mar 2024 17:14:16 +0100 Subject: [PATCH 089/129] docstrings --- .../Kernel/Disintegration/CdfToKernel.lean | 45 +++++++++++-------- .../Kernel/Disintegration/Integral.lean | 2 +- 2 files changed, 27 insertions(+), 20 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean index c7edf2b581f0c..477f3e62f7e45 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean @@ -59,6 +59,11 @@ section stieltjesOfMeasurableRat variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} {f : α × β → ℚ → ℝ} {κ : kernel α (β × ℝ)} {ν : kernel α β} +/-- a function `f : α × β → ℚ → ℝ` is called a rational conditional kernel CDF of `κ` with respect +to `ν` if is measurable, if `fun b ↦ f (a, b) x` is `(ν a)`-integrable for all `a : α` and `x : ℝ` +and for all measurable sets `s : Set β`, `∫ b in s, f (a, b) x ∂(ν a) = (κ a (s ×ˢ Iic x)).toReal`. +Also the `ℚ → ℝ` function `f (a, b)` should satisfy the properties of a Sieltjes function for +`(ν a)`-almost all `b : β`. -/ structure IsRatCondKernelCDF (f : α × β → ℚ → ℝ) (κ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := (measurable : Measurable f) (isRatStieltjesPoint_ae (a : α) : ∀ᵐ b ∂(ν a), IsRatStieltjesPoint f (a, b)) @@ -215,6 +220,9 @@ section isRatCondKernelCDFAux variable {β : Type*} {mβ : MeasurableSpace β} {f : α × β → ℚ → ℝ} {κ : kernel α (β × ℝ)} {ν : kernel α β} +/-- This property implies `IsRatCondKernelCDF`. The measurability, integrability and integral +conditions are the same, but the limit properties of `IsRatCondKernelCDF` are replaced by +limits of integrals. -/ structure isRatCondKernelCDFAux (f : α × β → ℚ → ℝ) (κ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := (measurable : Measurable f) @@ -256,8 +264,7 @@ lemma isRatCondKernelCDFAux.tendsto_zero_of_antitone (hf : isRatCondKernelCDFAux · exact fun n ↦ hf.integrable a (s n) · rw [integral_zero] exact hf.tendsto_integral_of_antitone a s hs hs_tendsto - · filter_upwards [hf.mono a] with t ht - exact fun n m hnm ↦ ht (hs hnm) + · filter_upwards [hf.mono a] with t ht using fun n m hnm ↦ ht (hs hnm) · filter_upwards [hf.nonneg a] with c hc using fun i ↦ hc (s i) lemma isRatCondKernelCDFAux.tendsto_one_of_monotone (hf : isRatCondKernelCDFAux f κ ν) @@ -267,8 +274,7 @@ lemma isRatCondKernelCDFAux.tendsto_one_of_monotone (hf : isRatCondKernelCDFAux · exact fun n ↦ hf.integrable a (s n) · rw [MeasureTheory.integral_const, smul_eq_mul, mul_one] exact hf.tendsto_integral_of_monotone a s hs hs_tendsto - · filter_upwards [hf.mono a] with t ht - exact fun n m hnm ↦ ht (hs hnm) + · filter_upwards [hf.mono a] with t ht using fun n m hnm ↦ ht (hs hnm) · filter_upwards [hf.le_one a] with c hc using fun i ↦ hc (s i) lemma isRatCondKernelCDFAux.tendsto_atTop_one (hf : isRatCondKernelCDFAux f κ ν) [IsFiniteKernel ν] @@ -280,8 +286,7 @@ lemma isRatCondKernelCDFAux.tendsto_atTop_one (hf : isRatCondKernelCDFAux f κ exact ht let s : ℕ → ℚ := fun n ↦ n have hs : Monotone s := fun i j hij ↦ by simp [s, hij] - have hs_tendsto : Tendsto s atTop atTop := by - exact tendsto_nat_cast_atTop_atTop + have hs_tendsto : Tendsto s atTop atTop := tendsto_nat_cast_atTop_atTop filter_upwards [hf.tendsto_one_of_monotone a s hs hs_tendsto] with x hx using hx lemma isRatCondKernelCDFAux.tendsto_atBot_zero (hf : isRatCondKernelCDFAux f κ ν) [IsFiniteKernel ν] @@ -317,10 +322,9 @@ lemma isRatCondKernelCDFAux.integrable_iInf_rat_gt (hf : isRatCondKernelCDFAux f [IsFiniteKernel ν] (a : α) (q : ℚ) : Integrable (fun t ↦ ⨅ r : Ioi q, f (a, t) r) (ν a) := by rw [← memℒp_one_iff_integrable] - refine ⟨Measurable.aestronglyMeasurable ?_, ?_⟩ - · exact measurable_iInf fun i ↦ hf.measurable_right a _ + refine ⟨(measurable_iInf fun i ↦ hf.measurable_right a _).aestronglyMeasurable, ?_⟩ refine (?_ : _ ≤ (ν a univ : ℝ≥0∞)).trans_lt (measure_lt_top _ _) - refine (snorm_le_of_ae_bound (C := 1) ?_).trans ?_ + refine (snorm_le_of_ae_bound (C := 1) ?_).trans (by simp) · filter_upwards [hf.bddBelow_range a, hf.nonneg a, hf.le_one a] with t hbdd_below h_nonneg h_le_one rw [Real.norm_eq_abs, abs_of_nonneg] @@ -329,7 +333,6 @@ lemma isRatCondKernelCDFAux.integrable_iInf_rat_gt (hf : isRatCondKernelCDFAux f · exact ⟨q + 1, by simp⟩ · exact h_le_one _ · exact le_ciInf fun r ↦ h_nonneg _ - · simp lemma isRatCondKernelCDFAux.set_integral_iInf_rat_gt (hf : isRatCondKernelCDFAux f κ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (q : ℚ) {A : Set β} (hA : MeasurableSet A) : @@ -342,8 +345,7 @@ lemma isRatCondKernelCDFAux.set_integral_iInf_rat_gt (hf : isRatCondKernelCDFAux refine set_integral_mono_ae ?_ ?_ ?_ · exact (hf.integrable_iInf_rat_gt _ _).integrableOn · exact (hf.integrable _ _).integrableOn - · filter_upwards [hf.bddBelow_range a] with t ht - exact ciInf_le (ht _) r + · filter_upwards [hf.bddBelow_range a] with t ht using ciInf_le (ht _) r calc ∫ t in A, ⨅ r : Ioi q, f (a, t) r ∂(ν a) ≤ ⨅ r : Ioi q, (κ a (A ×ˢ Iic (r : ℝ))).toReal := le_ciInf h _ = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by @@ -353,18 +355,15 @@ lemma isRatCondKernelCDFAux.set_integral_iInf_rat_gt (hf : isRatCondKernelCDFAux refine set_integral_mono_ae ?_ ?_ ?_ · exact (hf.integrable _ _).integrableOn · exact (hf.integrable_iInf_rat_gt _ _).integrableOn - · filter_upwards [hf.mono a] with c h_mono - exact le_ciInf (fun r ↦ h_mono (le_of_lt r.prop)) + · filter_upwards [hf.mono a] with c h_mono using le_ciInf (fun r ↦ h_mono (le_of_lt r.prop)) lemma isRatCondKernelCDFAux.iInf_rat_gt_eq (hf : isRatCondKernelCDFAux f κ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : ∀ᵐ t ∂(ν a), ∀ q : ℚ, ⨅ r : Ioi q, f (a, t) r = f (a, t) q := by rw [ae_all_iff] refine fun q ↦ ae_eq_of_forall_set_integral_eq_of_sigmaFinite (μ := ν a) ?_ ?_ ?_ - · intro s _ _ - refine Integrable.integrableOn ?_ - exact hf.integrable_iInf_rat_gt _ _ - · exact fun s _ _ ↦ (hf.integrable a _).integrableOn + · exact fun _ _ _ ↦ (hf.integrable_iInf_rat_gt _ _).integrableOn + · exact fun _ _ _ ↦ (hf.integrable a _).integrableOn · intro s hs _ rw [hf.set_integral _ hs, hf.set_integral_iInf_rat_gt _ _ hs] @@ -391,6 +390,10 @@ section IsCondKernelCDF variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} {f : α × β → StieltjesFunction} {κ : kernel α (β × ℝ)} {ν : kernel α β} +/-- A function `f : α × β → StieltjesFunction` is called a conditional kernel CDF of `κ` with +respect to `ν` if it is measurable, tends to to 0 at -∞ and to 1 at +∞ for all `p : α × β`, +if `fun b ↦ f (a, b) x` is `(ν a)`-integrable for all `a : α` and `x : ℝ` and for all +measurable sets `s : Set β`, `∫ b in s, f (a, b) x ∂(ν a) = (κ a (s ×ˢ Iic x)).toReal`. -/ structure IsCondKernelCDF (f : α × β → StieltjesFunction) (κ : kernel α (β × ℝ)) (ν : kernel α β) : Prop := (measurable (x : ℝ) : Measurable fun p ↦ f p x) @@ -406,7 +409,7 @@ lemma IsCondKernelCDF.nonneg (hf : IsCondKernelCDF f κ ν) (p : α × β) (x : lemma IsCondKernelCDF.le_one (hf : IsCondKernelCDF f κ ν) (p : α × β) (x : ℝ) : f p x ≤ 1 := Monotone.ge_of_tendsto (f p).mono (hf.tendsto_atTop_one p) x -lemma IsCondKernelCDF.integral [IsFiniteKernel κ] +lemma IsCondKernelCDF.integral {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f κ ν) (a : α) (x : ℝ) : ∫ b, f (a, b) x ∂(ν a) = (κ a (univ ×ˢ Iic x)).toReal := by rw [← hf.set_integral _ MeasurableSet.univ, Measure.restrict_univ] @@ -440,6 +443,8 @@ section ToKernel variable {_ : MeasurableSpace β} {f : α × β → StieltjesFunction} {κ : kernel α (β × ℝ)} {ν : kernel α β} {hf : IsCondKernelCDF f κ ν} +/-- A measurable function `α → StieltjesFunction` with limits 0 at -∞ and 1 at +∞ gives a measurable +function `α → Measure ℝ` by taking `StieltjesFunction.measure` at each point. -/ lemma StieltjesFunction.measurable_measure {f : α → StieltjesFunction} (hf : ∀ q, Measurable fun a ↦ f a q) (hf_bot : ∀ a, Tendsto (f a) atBot (𝓝 0)) @@ -469,6 +474,8 @@ lemma StieltjesFunction.measurable_measure {f : α → StieltjesFunction} simp_rw [measure_iUnion hf_disj hf_meas] exact Measurable.ennreal_tsum hf_cd_meas +/-- A function `f : α × β → StieltjesFunction` with the property `IsCondKernelCDF f κ ν` gives a +Markov kernel from `α × β` to `ℝ`, by taking for each `p : α × β` the measure defined by `f p`. -/ noncomputable def IsCondKernelCDF.toKernel (f : α × β → StieltjesFunction) (hf : IsCondKernelCDF f κ ν) : kernel (α × β) ℝ where diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index 95ce159581fe3..75e01d37399be 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -8,7 +8,7 @@ import Mathlib.Probability.Kernel.Disintegration.Basic /-! # Lebesgue and Bochner integrals of conditional kernels -Integrals of `ProbabilityTheory.kernel.condKernel` and `MeasureTheory.Measure.condKerenl`. +Integrals of `ProbabilityTheory.kernel.condKernel` and `MeasureTheory.Measure.condKernel`. ## Main statements From cef6d20b0a7c0f2bdcc8eebef9ee8a3f597e916d Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 6 Mar 2024 17:18:56 +0100 Subject: [PATCH 090/129] delete AuxLemmas.lean --- .../Kernel/Disintegration/AuxLemmas.lean | 38 ------------------- .../Kernel/Disintegration/CdfToKernel.lean | 18 ++++++++- .../Kernel/Disintegration/CondCdf.lean | 20 +++++++++- 3 files changed, 36 insertions(+), 40 deletions(-) delete mode 100644 Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean diff --git a/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean b/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean deleted file mode 100644 index 065de63bc988f..0000000000000 --- a/Mathlib/Probability/Kernel/Disintegration/AuxLemmas.lean +++ /dev/null @@ -1,38 +0,0 @@ -import Mathlib.MeasureTheory.Constructions.Polish -import Mathlib.MeasureTheory.Integral.Bochner - -open Filter Set MeasureTheory - -open scoped Topology ENNReal - -variable {α β : Type*} {mα : MeasurableSpace α} - -theorem Real.iUnion_Iic_rat : ⋃ r : ℚ, Iic (r : ℝ) = univ := by - ext1 x - simp only [mem_iUnion, mem_Iic, mem_univ, iff_true_iff] - obtain ⟨r, hr⟩ := exists_rat_gt x - exact ⟨r, hr.le⟩ -#align real.Union_Iic_rat Real.iUnion_Iic_rat - -theorem Real.iInter_Iic_rat : ⋂ r : ℚ, Iic (r : ℝ) = ∅ := by - ext1 x - simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false_iff, not_forall, not_le] - exact exists_rat_lt x -#align real.Inter_Iic_rat Real.iInter_Iic_rat - -lemma Measure.iInf_rat_gt_prod_Iic {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] - {s : Set α} (hs : MeasurableSet s) (t : ℚ) : - ⨅ r : { r' : ℚ // t < r' }, ρ (s ×ˢ Iic (r : ℝ)) = ρ (s ×ˢ Iic (t : ℝ)) := by - rw [← measure_iInter_eq_iInf] - · rw [← prod_iInter] - congr with x : 1 - simp only [mem_iInter, mem_Iic, Subtype.forall, Subtype.coe_mk] - refine ⟨fun h ↦ ?_, fun h a hta ↦ h.trans ?_⟩ - · refine le_of_forall_lt_rat_imp_le fun q htq ↦ h q ?_ - exact mod_cast htq - · exact mod_cast hta.le - · exact fun _ => hs.prod measurableSet_Iic - · refine Monotone.directed_ge fun r r' hrr' ↦ prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, ?_⟩) - refine Iic_subset_Iic.mpr ?_ - exact mod_cast hrr' - · exact ⟨⟨t + 1, lt_add_one _⟩, measure_ne_top ρ _⟩ diff --git a/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean index 477f3e62f7e45..8fb71f71c3978 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean @@ -6,7 +6,6 @@ Authors: Rémy Degenne import Mathlib.MeasureTheory.Function.AEEqOfIntegral import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes -import Mathlib.Probability.Kernel.Disintegration.AuxLemmas /-! # Building a Markov kernel from a conditional cumulative distribution function @@ -334,6 +333,23 @@ lemma isRatCondKernelCDFAux.integrable_iInf_rat_gt (hf : isRatCondKernelCDFAux f · exact h_le_one _ · exact le_ciInf fun r ↦ h_nonneg _ +lemma _root_.MeasureTheory.Measure.iInf_rat_gt_prod_Iic {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] + {s : Set α} (hs : MeasurableSet s) (t : ℚ) : + ⨅ r : { r' : ℚ // t < r' }, ρ (s ×ˢ Iic (r : ℝ)) = ρ (s ×ˢ Iic (t : ℝ)) := by + rw [← measure_iInter_eq_iInf] + · rw [← prod_iInter] + congr with x : 1 + simp only [mem_iInter, mem_Iic, Subtype.forall, Subtype.coe_mk] + refine ⟨fun h ↦ ?_, fun h a hta ↦ h.trans ?_⟩ + · refine le_of_forall_lt_rat_imp_le fun q htq ↦ h q ?_ + exact mod_cast htq + · exact mod_cast hta.le + · exact fun _ => hs.prod measurableSet_Iic + · refine Monotone.directed_ge fun r r' hrr' ↦ prod_subset_prod_iff.mpr (Or.inl ⟨subset_rfl, ?_⟩) + refine Iic_subset_Iic.mpr ?_ + exact mod_cast hrr' + · exact ⟨⟨t + 1, lt_add_one _⟩, measure_ne_top ρ _⟩ + lemma isRatCondKernelCDFAux.set_integral_iInf_rat_gt (hf : isRatCondKernelCDFAux f κ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) (q : ℚ) {A : Set β} (hA : MeasurableSet A) : ∫ t in A, ⨅ r : Ioi q, f (a, t) r ∂(ν a) = (κ a (A ×ˢ Iic (q : ℝ))).toReal := by diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index 2b8c4ff7ca5de..f4eb79d464fc5 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -5,7 +5,6 @@ Authors: Rémy Degenne -/ import Mathlib.MeasureTheory.Decomposition.RadonNikodym import Mathlib.Probability.Kernel.Disintegration.CdfToKernel -import Mathlib.Probability.Kernel.Disintegration.AuxLemmas #align_import probability.kernel.cond_cdf from "leanprover-community/mathlib"@"3b88f4005dc2e28d42f974cc1ce838f0dafb39b8" @@ -41,6 +40,25 @@ open MeasureTheory Set Filter TopologicalSpace open scoped NNReal ENNReal MeasureTheory Topology +section AuxLemmasToBeMoved + +variable {α β ι : Type*} + +theorem Real.iUnion_Iic_rat : ⋃ r : ℚ, Iic (r : ℝ) = univ := by + ext1 x + simp only [mem_iUnion, mem_Iic, mem_univ, iff_true_iff] + obtain ⟨r, hr⟩ := exists_rat_gt x + exact ⟨r, hr.le⟩ +#align real.Union_Iic_rat Real.iUnion_Iic_rat + +theorem Real.iInter_Iic_rat : ⋂ r : ℚ, Iic (r : ℝ) = ∅ := by + ext1 x + simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false_iff, not_forall, not_le] + exact exists_rat_lt x +#align real.Inter_Iic_rat Real.iInter_Iic_rat + +end AuxLemmasToBeMoved + namespace MeasureTheory.Measure variable {α β : Type*} {mα : MeasurableSpace α} (ρ : Measure (α × ℝ)) From bf200d3fb86c4e9b0443654873e662cfb70ce3d1 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 6 Mar 2024 17:27:36 +0100 Subject: [PATCH 091/129] variables --- .../Kernel/Disintegration/CdfToKernel.lean | 20 +++++++------------ 1 file changed, 7 insertions(+), 13 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean index 8fb71f71c3978..15245aa7bfcde 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean @@ -51,12 +51,12 @@ open scoped NNReal ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory -variable {α β : Type*} [MeasurableSpace α] +variable {α β : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + {κ : kernel α (β × ℝ)} {ν : kernel α β} section stieltjesOfMeasurableRat -variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} - {f : α × β → ℚ → ℝ} {κ : kernel α (β × ℝ)} {ν : kernel α β} +variable {f : α × β → ℚ → ℝ} /-- a function `f : α × β → ℚ → ℝ` is called a rational conditional kernel CDF of `κ` with respect to `ν` if is measurable, if `fun b ↦ f (a, b) x` is `(ν a)`-integrable for all `a : α` and `x : ℝ` @@ -216,8 +216,7 @@ end stieltjesOfMeasurableRat section isRatCondKernelCDFAux -variable {β : Type*} {mβ : MeasurableSpace β} {f : α × β → ℚ → ℝ} - {κ : kernel α (β × ℝ)} {ν : kernel α β} +variable {f : α × β → ℚ → ℝ} /-- This property implies `IsRatCondKernelCDF`. The measurability, integrability and integral conditions are the same, but the limit properties of `IsRatCondKernelCDF` are replaced by @@ -400,11 +399,9 @@ lemma isRatCondKernelCDFAux.isRatCondKernelCDF (hf : isRatCondKernelCDFAux f κ end isRatCondKernelCDFAux - section IsCondKernelCDF -variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} - {f : α × β → StieltjesFunction} {κ : kernel α (β × ℝ)} {ν : kernel α β} +variable {f : α × β → StieltjesFunction} /-- A function `f : α × β → StieltjesFunction` is called a conditional kernel CDF of `κ` with respect to `ν` if it is measurable, tends to to 0 at -∞ and to 1 at +∞ for all `p : α × β`, @@ -513,9 +510,7 @@ end ToKernel section -variable {α β : Type*} [MeasurableSpace α] {mβ : MeasurableSpace β} - {f : α × β → StieltjesFunction} {κ : kernel α (β × ℝ)} {ν : kernel α β} - {hf : IsCondKernelCDF f κ ν} +variable {f : α × β → StieltjesFunction} {hf : IsCondKernelCDF f κ ν} lemma set_lintegral_toKernel_Iic [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) (a : α) (x : ℝ) {s : Set β} (hs : MeasurableSet s) : @@ -663,8 +658,7 @@ lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) _ = ∑' i, κ a (f' i) := by simp_rw [hf_eq] _ = κ a (iUnion f') := (measure_iUnion hf_disj hf_meas).symm -lemma compProd_toKernel [IsFiniteKernel κ] [IsSFiniteKernel ν] - (hf : IsCondKernelCDF f κ ν) : +lemma compProd_toKernel [IsFiniteKernel κ] [IsSFiniteKernel ν] (hf : IsCondKernelCDF f κ ν) : ν ⊗ₖ hf.toKernel f = κ := by ext a s hs rw [kernel.compProd_apply _ _ _ hs, lintegral_toKernel_mem hf a hs] From 681d3590bf20cda5e9d3d1e818a432410f2cb273 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 10 Mar 2024 10:04:16 +0100 Subject: [PATCH 092/129] more about RN --- Mathlib/Probability/Kernel/RadonNikodym.lean | 483 +++++++++++++++++++ 1 file changed, 483 insertions(+) diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index 621e5a393cb2c..1142685a9ce75 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -5,6 +5,7 @@ Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.WithDensity +import Mathlib.Probability.Kernel.MeasureCompProd /-! # Radon-Nikodym derivative and Lebesgue decomposition for kernels @@ -431,4 +432,486 @@ lemma measurableSet_mutuallySingular (κ ν : kernel α γ) [IsFiniteKernel κ] exact kernel.measurable_kernel_prod_mk_left (measurableSet_mutuallySingularSet κ ν).compl (measurableSet_singleton 0) +-- ok +lemma Measure.compProd_apply_prod {μ : Measure α} [SFinite μ] [IsSFiniteKernel κ] + {s : Set α} {t : Set γ} (hs : MeasurableSet s) (ht : MeasurableSet t) : + (μ ⊗ₘ κ) (s ×ˢ t) = ∫⁻ a in s, κ a t ∂μ := by + rw [Measure.compProd_apply (hs.prod ht), ← lintegral_indicator _ hs] + congr with a + classical + rw [indicator_apply] + split_ifs with ha <;> simp [ha] + +-- ok +lemma set_lintegral_withDensity_eq_lintegral_mul₀' {μ : Measure α} {f : α → ℝ≥0∞} + (hf : AEMeasurable f μ) {g : α → ℝ≥0∞} (hg : AEMeasurable g (μ.withDensity f)) + {s : Set α} (hs : MeasurableSet s) : + ∫⁻ a in s, g a ∂μ.withDensity f = ∫⁻ a in s, (f * g) a ∂μ := by + rw [restrict_withDensity hs, lintegral_withDensity_eq_lintegral_mul₀' hf.restrict] + rw [← restrict_withDensity hs] + exact hg.restrict + +-- ok +lemma set_lintegral_withDensity_eq_lintegral_mul₀ {μ : Measure α} {f : α → ℝ≥0∞} + (hf : AEMeasurable f μ) {g : α → ℝ≥0∞} (hg : AEMeasurable g μ) + {s : Set α} (hs : MeasurableSet s) : + ∫⁻ a in s, g a ∂μ.withDensity f = ∫⁻ a in s, (f * g) a ∂μ := + set_lintegral_withDensity_eq_lintegral_mul₀' hf + (hg.mono' (MeasureTheory.withDensity_absolutelyContinuous μ f)) hs + +-- ok +lemma set_lintegral_rnDeriv_mul {μ ν : Measure α} [Measure.HaveLebesgueDecomposition μ ν] + (hμν : μ ≪ ν) {f : α → ℝ≥0∞} (hf : AEMeasurable f ν) {s : Set α} (hs : MeasurableSet s) : + ∫⁻ x in s, μ.rnDeriv ν x * f x ∂ν = ∫⁻ x in s, f x ∂μ := by + nth_rw 2 [← Measure.withDensity_rnDeriv_eq μ ν hμν] + rw [set_lintegral_withDensity_eq_lintegral_mul₀ (Measure.measurable_rnDeriv μ ν).aemeasurable hf + hs] + rfl + +-- ok +lemma Measure.absolutelyContinuous_compProd_left {μ ν : Measure α} [SFinite μ] [SFinite ν] + (hμν : μ ≪ ν) (κ : kernel α γ) [IsSFiniteKernel κ] : + μ ⊗ₘ κ ≪ ν ⊗ₘ κ := by + refine Measure.AbsolutelyContinuous.mk fun s hs hs_zero ↦ ?_ + rw [Measure.compProd_apply hs, lintegral_eq_zero_iff (measurable_kernel_prod_mk_left hs)] + at hs_zero ⊢ + exact hμν.ae_eq hs_zero + +-- ok +lemma Measure.absolutelyContinuous_compProd_right (μ : Measure α) {κ η : kernel α γ} + [SFinite μ] [IsSFiniteKernel κ] [IsSFiniteKernel η] + (hκη : ∀ᵐ a ∂μ, κ a ≪ η a) : + μ ⊗ₘ κ ≪ μ ⊗ₘ η := by + refine Measure.AbsolutelyContinuous.mk fun s hs hs_zero ↦ ?_ + rw [Measure.compProd_apply hs, lintegral_eq_zero_iff (measurable_kernel_prod_mk_left hs)] + at hs_zero ⊢ + filter_upwards [hs_zero, hκη] with a ha_zero ha_ac using ha_ac ha_zero + +-- ok +lemma Measure.absolutelyContinuous_compProd {μ ν : Measure α} {κ η : kernel α γ} + [SFinite μ] [SFinite ν] [IsSFiniteKernel κ] [IsSFiniteKernel η] + (hμν : μ ≪ ν) (hκη : ∀ᵐ a ∂ν, κ a ≪ η a) : + μ ⊗ₘ κ ≪ ν ⊗ₘ η := + (Measure.absolutelyContinuous_compProd_left hμν κ).trans + (Measure.absolutelyContinuous_compProd_right ν hκη) + +lemma Measure.mutuallySingular_compProd_left {μ ν : Measure α} [SFinite μ] [SFinite ν] + (hμν : μ ⟂ₘ ν) (κ η : kernel α γ) [IsSFiniteKernel κ] [IsSFiniteKernel η] : + μ ⊗ₘ κ ⟂ₘ ν ⊗ₘ η := by + let s := hμν.nullSet + have hs_prod : MeasurableSet (s ×ˢ (univ : Set γ)) := + hμν.measurableSet_nullSet.prod MeasurableSet.univ + refine ⟨s ×ˢ univ, hs_prod, ?_⟩ + rw [Measure.compProd_apply_prod hμν.measurableSet_nullSet MeasurableSet.univ, compl_prod_eq_union] + simp only [Measure.MutuallySingular.restrict_nullSet, lintegral_zero_measure, compl_univ, + prod_empty, union_empty, true_and] + rw [Measure.compProd_apply_prod hμν.measurableSet_nullSet.compl MeasurableSet.univ] + simp + +lemma Measure.mutuallySingular_compProd_right (μ ν : Measure α) [SFinite μ] [SFinite ν] + {κ η : kernel α γ} [IsFiniteKernel κ] [IsFiniteKernel η] (hκη : ∀ a, κ a ⟂ₘ η a) : + μ ⊗ₘ κ ⟂ₘ ν ⊗ₘ η := by + let s := mutuallySingularSet κ η + have hs : MeasurableSet s := measurableSet_mutuallySingularSet κ η + symm + refine ⟨s, hs, ?_⟩ + rw [Measure.compProd_apply hs, Measure.compProd_apply hs.compl] + have h1 : ∀ a, η a (Prod.mk a ⁻¹' s) = 0 := by + intro a + have : Prod.mk a ⁻¹' s = mutuallySingularSetSlice κ η a := rfl + rw [this, measure_mutuallySingularSetSlice] + have h2 : ∀ a, κ a (Prod.mk a ⁻¹' s)ᶜ = 0 := by + intro a + have : (Prod.mk a ⁻¹' s)ᶜ ⊆ Prod.mk a ⁻¹' sᶜ := by intro; simp + refine measure_mono_null this ?_ + have : Prod.mk a ⁻¹' sᶜ = (mutuallySingularSetSlice κ η a)ᶜ := rfl + rw [this, ← withDensity_rnDeriv_eq_zero_iff_measure_eq_zero κ η a, + withDensity_rnDeriv_eq_zero_iff_mutuallySingular] + exact hκη a + simp [h1, h2] + +-- ok +@[simp] +lemma prodMkLeft_add (β : Type*) [MeasurableSpace β] (κ η : kernel α γ) : + prodMkLeft β (κ + η) = prodMkLeft β κ + prodMkLeft β η := by ext; simp + +-- ok +@[simp] +lemma prodMkRight_add (β : Type*) [MeasurableSpace β] (κ η : kernel α γ) : + prodMkRight β (κ + η) = prodMkRight β κ + prodMkRight β η := by ext; simp + +-- ok +lemma compProd_add_left {β : Type*} {_ : MeasurableSpace β} + (μ κ : kernel α β) [IsSFiniteKernel μ] (η : kernel (α × β) γ) + [IsSFiniteKernel κ] [IsSFiniteKernel η] : + (μ + κ) ⊗ₖ η = μ ⊗ₖ η + κ ⊗ₖ η := by + ext a s hs + rw [compProd_apply _ _ _ hs] + simp only [coeFn_add, Pi.add_apply, lintegral_add_measure, Measure.add_toOuterMeasure, + OuterMeasure.coe_add] + rw [compProd_apply _ _ _ hs, compProd_apply _ _ _ hs] + +-- ok +lemma compProd_add_right {β : Type*} {_ : MeasurableSpace β} + (μ : kernel α β) [IsSFiniteKernel μ] (κ η : kernel (α × β) γ) + [IsSFiniteKernel κ] [IsSFiniteKernel η] : + μ ⊗ₖ (κ + η) = μ ⊗ₖ κ + μ ⊗ₖ η := by + ext a s hs + rw [compProd_apply _ _ _ hs] + simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add] + rw [lintegral_add_left, compProd_apply _ _ _ hs, compProd_apply _ _ _ hs] + exact measurable_kernel_prod_mk_left' hs a + +-- ok +lemma const_add (β : Type*) [MeasurableSpace β] (μ ν : Measure α) : + const β (μ + ν) = const β μ + const β ν := by ext; simp + +-- ok +nonrec lemma Measure.compProd_add_left (μ ν : Measure α) [SFinite μ] [SFinite ν] (κ : kernel α γ) + [IsSFiniteKernel κ] : + (μ + ν) ⊗ₘ κ = μ ⊗ₘ κ + ν ⊗ₘ κ := by + rw [Measure.compProd, const_add, compProd_add_left]; rfl + +-- ok +nonrec lemma Measure.compProd_add_right (μ : Measure α) [SFinite μ] (κ η : kernel α γ) + [IsSFiniteKernel κ] [IsSFiniteKernel η] : + μ ⊗ₘ (κ + η) = μ ⊗ₘ κ + μ ⊗ₘ η := by + rw [Measure.compProd, prodMkLeft_add, compProd_add_right]; rfl + +lemma Measure.fst_map_compProd (μ : Measure α) [SFinite μ] (κ : kernel α γ) + [IsMarkovKernel κ] : + (μ ⊗ₘ κ).map Prod.fst = μ := by + ext s hs + rw [Measure.map_apply measurable_fst hs, Measure.compProd_apply] + swap; · exact measurable_fst hs + have : ∀ a, (κ a) (Prod.mk a ⁻¹' (Prod.fst ⁻¹' s)) = s.indicator (fun a ↦ κ a univ) a := by + intro a + classical + rw [indicator_apply] + split_ifs with ha + · congr with x + simp [ha] + · suffices Prod.mk a ⁻¹' (Prod.fst ⁻¹' s) = ∅ by rw [this]; simp + ext x + simp [ha] + simp_rw [this] + rw [lintegral_indicator _ hs] + simp + +-- ok +lemma ae_compProd_of_ae_ae {β : Type*} {_ : MeasurableSpace β} + {κ : kernel α β} {η : kernel (α × β) γ} [IsSFiniteKernel κ] [IsSFiniteKernel η] {a : α} + {p : β × γ → Prop} (hp : MeasurableSet {x | p x}) (h : ∀ᵐ b ∂κ a, ∀ᵐ c ∂η (a, b), p (b, c)) : + ∀ᵐ bc ∂(κ ⊗ₖ η) a, p bc := by + simp_rw [ae_iff] at h ⊢ + rw [compProd_null] + · exact h + · exact hp.compl + +-- ok +lemma ae_compProd_iff {β : Type*} {_ : MeasurableSpace β} + {κ : kernel α β} {η : kernel (α × β) γ} [IsSFiniteKernel κ] [IsSFiniteKernel η] {a : α} + {p : β × γ → Prop} (hp : MeasurableSet {x | p x}) : + (∀ᵐ bc ∂(κ ⊗ₖ η) a, p bc) ↔ ∀ᵐ b ∂κ a, ∀ᵐ c ∂η (a, b), p (b, c) := + ⟨fun h ↦ ae_ae_of_ae_compProd h, fun h ↦ ae_compProd_of_ae_ae hp h⟩ + +-- ok +lemma Measure.ae_compProd_of_ae_ae {μ : Measure α} + [SFinite μ] [IsSFiniteKernel κ] {p : α × γ → Prop} (hp : MeasurableSet {x | p x}) + (h : ∀ᵐ a ∂μ, ∀ᵐ b ∂(κ a), p (a, b)) : + ∀ᵐ x ∂(μ ⊗ₘ κ), p x := + kernel.ae_compProd_of_ae_ae hp h + +-- ok +lemma Measure.ae_ae_of_ae_compProd {μ : Measure α} [SFinite μ] [IsSFiniteKernel κ] + {p : α × γ → Prop} (h : ∀ᵐ x ∂(μ ⊗ₘ κ), p x) : + ∀ᵐ a ∂μ, ∀ᵐ b ∂(κ a), p (a, b) := + kernel.ae_ae_of_ae_compProd h + +-- ok +lemma Measure.ae_compProd_iff {μ : Measure α} + [SFinite μ] [IsSFiniteKernel κ] {p : α × γ → Prop} (hp : MeasurableSet {x | p x}) : + (∀ᵐ x ∂(μ ⊗ₘ κ), p x) ↔ ∀ᵐ a ∂μ, ∀ᵐ b ∂(κ a), p (a, b) := + kernel.ae_compProd_iff hp + +lemma ae_compProd_of_ae_fst {μ : Measure α} (κ : kernel α γ) + [SFinite μ] [IsSFiniteKernel κ] {p : α → Prop} (hp : MeasurableSet {x | p x}) + (h : ∀ᵐ a ∂μ, p a) : + ∀ᵐ x ∂(μ ⊗ₘ κ), p x.1 := by + refine ae_compProd_of_ae_ae (measurable_fst hp) ?_ + filter_upwards [h] with a ha + simp [ha] + +lemma ae_eq_compProd_of_ae_eq_fst {β : Type*} {_ : MeasurableSpace β} [AddGroup β] + [MeasurableSingletonClass β] [MeasurableSub₂ β] + (μ : Measure α) (κ : kernel α γ) + [SFinite μ] [IsSFiniteKernel κ] {f g : α → β} + (hf : Measurable f) (hg : Measurable g) (h : f =ᵐ[μ] g) : + (fun p ↦ f p.1) =ᵐ[μ ⊗ₘ κ] (fun p ↦ g p.1) := + ae_compProd_of_ae_fst κ (measurableSet_eq_fun hf hg) h + +lemma ae_eq_compProd_of_forall_ae_eq {β : Type*} {_ : MeasurableSpace β} [AddGroup β] + [MeasurableSingletonClass β] [MeasurableSub₂ β] + (μ : Measure α) (κ : kernel α γ) + [SFinite μ] [IsSFiniteKernel κ] {f g : α → γ → β} + (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) + (h : ∀ a, f a =ᵐ[κ a] g a) : + (fun p ↦ f p.1 p.2) =ᵐ[μ ⊗ₘ κ] (fun p ↦ g p.1 p.2) := + ae_compProd_of_ae_ae (measurableSet_eq_fun hf hg) + (ae_of_all _ (fun a ↦ measure_mono_null (fun x ↦ by simp) (h a))) + +-- ok +lemma measurableSet_eq_fun' {β : Type*} [CanonicallyOrderedAddCommMonoid β] [Sub β] [OrderedSub β] + {_ : MeasurableSpace β} [MeasurableSub₂ β] [MeasurableSingletonClass β] + {f g : α → β} (hf : Measurable f) (hg : Measurable g) : + MeasurableSet {x | f x = g x} := by + have : {a | f a = g a} = {a | (f - g) a = 0} ∩ {a | (g - f) a = 0} := by + ext a + simp only [mem_setOf_eq, Pi.sub_apply, tsub_eq_zero_iff_le, mem_inter_iff] + exact ⟨fun h ↦ ⟨h.le, h.symm.le⟩, fun h ↦ le_antisymm h.1 h.2⟩ + rw [this] + exact ((hf.sub hg) (measurableSet_singleton 0)).inter ((hg.sub hf) (measurableSet_singleton 0)) + +lemma ENNReal.ae_eq_compProd_of_ae_eq_fst (μ : Measure α) (κ : kernel α γ) + [SFinite μ] [IsSFiniteKernel κ] {f g : α → ℝ≥0∞} + (hf : Measurable f) (hg : Measurable g) (h : f =ᵐ[μ] g) : + (fun p ↦ f p.1) =ᵐ[μ ⊗ₘ κ] (fun p ↦ g p.1) := + ae_compProd_of_ae_fst κ (measurableSet_eq_fun' hf hg) h + +lemma ENNReal.ae_eq_compProd_of_forall_ae_eq + (μ : Measure α) (κ : kernel α γ) + [SFinite μ] [IsSFiniteKernel κ] {f g : α → γ → ℝ≥0∞} + (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) + (h : ∀ a, f a =ᵐ[κ a] g a) : + (fun p ↦ f p.1 p.2) =ᵐ[μ ⊗ₘ κ] (fun p ↦ g p.1 p.2) := + ae_compProd_of_ae_ae (measurableSet_eq_fun' hf hg) + (ae_of_all _ (fun a ↦ measure_mono_null (fun x ↦ by simp) (h a))) + +section Unique + +variable {ξ : kernel α γ} {f : α → γ → ℝ≥0∞} + +lemma eq_rnDeriv_measure [IsFiniteKernel ν] (h : κ = kernel.withDensity ν f + ξ) + (hf : Measurable (Function.uncurry f)) (hξ : ∀ a, ξ a ⟂ₘ ν a) (a : α) : + f a =ᵐ[ν a] ∂(κ a)/∂(ν a) := by + have : κ a = ξ a + (ν a).withDensity (f a) := by + rw [h, coeFn_add, Pi.add_apply, kernel.withDensity_apply _ hf, add_comm] + exact (κ a).eq_rnDeriv₀ (hf.comp measurable_prod_mk_left).aemeasurable (hξ a) this + +lemma eq_singularPart_measure [IsFiniteKernel κ] [IsFiniteKernel ν] + (h : κ = kernel.withDensity ν f + ξ) + (hf : Measurable (Function.uncurry f)) (hξ : ∀ a, ξ a ⟂ₘ ν a) (a : α) : + ξ a = (κ a).singularPart (ν a) := by + have : κ a = ξ a + (ν a).withDensity (f a) := by + rw [h, coeFn_add, Pi.add_apply, kernel.withDensity_apply _ hf, add_comm] + exact (κ a).eq_singularPart (hf.comp measurable_prod_mk_left) (hξ a) this + +lemma rnDeriv_eq_rnDeriv_measure (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + rnDeriv κ ν a =ᵐ[ν a] ∂(κ a)/∂(ν a) := + eq_rnDeriv_measure (rnDeriv_add_singularPart κ ν).symm (measurable_rnDeriv κ ν) + (mutuallySingular_singularPart κ ν) a + +lemma singularPart_eq_singularPart_measure [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + singularPart κ ν a = (κ a).singularPart (ν a) := + eq_singularPart_measure (rnDeriv_add_singularPart κ ν).symm (measurable_rnDeriv κ ν) + (mutuallySingular_singularPart κ ν) a + +lemma eq_rnDeriv [IsFiniteKernel κ] [IsFiniteKernel ν] (h : κ = kernel.withDensity ν f + ξ) + (hf : Measurable (Function.uncurry f)) (hξ : ∀ a, ξ a ⟂ₘ ν a) (a : α) : + f a =ᵐ[ν a] rnDeriv κ ν a := + (eq_rnDeriv_measure h hf hξ a).trans (rnDeriv_eq_rnDeriv_measure _ _ a).symm + +lemma eq_singularPart [IsFiniteKernel κ] [IsFiniteKernel ν] (h : κ = kernel.withDensity ν f + ξ) + (hf : Measurable (Function.uncurry f)) (hξ : ∀ a, ξ a ⟂ₘ ν a) (a : α) : + ξ a = singularPart κ ν a := + (eq_singularPart_measure h hf hξ a).trans (singularPart_eq_singularPart_measure a).symm + +end Unique + +section MeasureCompProd + +lemma set_lintegral_prod_rnDeriv {μ ν : Measure α} {κ η : kernel α γ} + [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteKernel κ] [IsFiniteKernel η] + (hμν : μ ≪ ν) (hκη : ∀ᵐ a ∂ν, κ a ≪ η a) + {s : Set α} (hs : MeasurableSet s) {t : Set γ} (ht : MeasurableSet t) : + ∫⁻ x in s, ∫⁻ y in t, (∂μ/∂ν) x * rnDeriv κ η x y ∂(η x) ∂ν = (μ ⊗ₘ κ) (s ×ˢ t) := by + have : ∀ᵐ x ∂ν, ∫⁻ y in t, (∂μ/∂ν) x * rnDeriv κ η x y ∂(η x) + = (∂μ/∂ν) x * κ x t := by + filter_upwards [hκη] with x hx + calc ∫⁻ y in t, (∂μ/∂ν) x * rnDeriv κ η x y ∂(η x) + = (∂μ/∂ν) x * ∫⁻ y in t, rnDeriv κ η x y ∂(η x) := by + rw [lintegral_const_mul] + exact measurable_rnDeriv_right κ η x + _ = (∂μ/∂ν) x * ∫⁻ y in t, (∂(κ x)/∂(η x)) y ∂(η x) := by + congr 1 + refine lintegral_congr_ae (ae_restrict_of_ae ?_) + exact rnDeriv_eq_rnDeriv_measure _ _ x + _ = (∂μ/∂ν) x * κ x t := by + congr + rw [Measure.set_lintegral_rnDeriv hx] + rw [lintegral_congr_ae (ae_restrict_of_ae this), + set_lintegral_rnDeriv_mul hμν (kernel.measurable_coe _ ht).aemeasurable hs, + Measure.compProd_apply_prod hs ht] + +lemma rnDeriv_measure_compProd_aux {μ ν : Measure α} {κ η : kernel α γ} + [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteKernel κ] [IsFiniteKernel η] + (hμν : μ ≪ ν) (hκη : ∀ᵐ a ∂ν, κ a ≪ η a) : + ∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] fun p ↦ (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 := by + have h_meas : Measurable fun p : α × γ ↦ (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 := + ((Measure.measurable_rnDeriv _ _).comp measurable_fst).mul (measurable_rnDeriv _ _) + suffices ∀ s, MeasurableSet s → ∫⁻ p in s, (∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η)) p ∂(ν ⊗ₘ η) + = ∫⁻ p in s, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂(ν ⊗ₘ η) from + ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite (Measure.measurable_rnDeriv _ _) h_meas + fun s hs _ ↦ this s hs + have h_left : ∀ s, MeasurableSet s → ∫⁻ p in s, (∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η)) p ∂(ν ⊗ₘ η) + = (μ ⊗ₘ κ) s := by + intro s _ + rw [Measure.set_lintegral_rnDeriv] + exact Measure.absolutelyContinuous_compProd hμν hκη + have : ∀ s t, MeasurableSet s → MeasurableSet t → + ∫⁻ x in s, ∫⁻ y in t, (∂μ/∂ν) x * rnDeriv κ η x y ∂(η x) ∂ν = (μ ⊗ₘ κ) (s ×ˢ t) := + fun _ _ hs ht ↦ set_lintegral_prod_rnDeriv hμν hκη hs ht + intro s hs + apply MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod _ _ _ _ hs + · simp + · rintro _ ⟨t₁, ht₁, t₂, ht₂, rfl⟩ + simp only [mem_setOf_eq] at ht₁ ht₂ + rw [h_left (t₁ ×ˢ t₂) (ht₁.prod ht₂), Measure.set_lintegral_compProd h_meas ht₁ ht₂, + this t₁ t₂ ht₁ ht₂] + · intro t ht ht_eq + calc ∫⁻ p in tᶜ, (∂μ ⊗ₘ κ/∂ν ⊗ₘ η) p ∂ν ⊗ₘ η + = ∫⁻ p, (∂μ ⊗ₘ κ/∂ν ⊗ₘ η) p ∂ν ⊗ₘ η - ∫⁻ p in t, (∂μ ⊗ₘ κ/∂ν ⊗ₘ η) p ∂ν ⊗ₘ η := by + refine (ENNReal.sub_eq_of_add_eq ?_ ?_).symm + · rw [h_left _ ht] + exact measure_ne_top _ _ + · rw [add_comm, lintegral_add_compl _ ht] + _ = ∫⁻ p, (∂μ ⊗ₘ κ/∂ν ⊗ₘ η) p ∂ν ⊗ₘ η + - ∫⁻ p in t, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂ν ⊗ₘ η := by rw [ht_eq] + _ = (μ ⊗ₘ κ) univ + - ∫⁻ p in t, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂ν ⊗ₘ η := by + have h := h_left univ MeasurableSet.univ + rw [set_lintegral_univ] at h + rw [h] + _ = ∫⁻ x, ∫⁻ y, (∂μ/∂ν) x * rnDeriv κ η x y ∂η x ∂ν + - ∫⁻ p in t, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂ν ⊗ₘ η := by + have h := this univ univ MeasurableSet.univ MeasurableSet.univ + simp only [Measure.restrict_univ, univ_prod_univ] at h + rw [← h] + _ = ∫⁻ p, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂ν ⊗ₘ η + - ∫⁻ p in t, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂ν ⊗ₘ η := by + congr + rw [Measure.lintegral_compProd h_meas] + _ = ∫⁻ p in tᶜ, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂ν ⊗ₘ η := by + refine ENNReal.sub_eq_of_add_eq ?_ ?_ + · rw [← ht_eq, h_left _ ht] + exact measure_ne_top _ _ + rw [add_comm, lintegral_add_compl _ ht] + · intro f' hf_disj hf_meas hf_eq + rw [lintegral_iUnion hf_meas hf_disj, lintegral_iUnion hf_meas hf_disj] + congr with i + exact hf_eq i + +instance instIsFiniteKernel_withDensity_rnDeriv [hκ : IsFiniteKernel κ] [IsFiniteKernel ν] : + IsFiniteKernel (withDensity ν (rnDeriv κ ν)) := by + constructor + refine ⟨hκ.bound, hκ.bound_lt_top, fun a ↦ ?_⟩ + rw [kernel.withDensity_apply', set_lintegral_univ] + swap; · exact measurable_rnDeriv κ ν + rw [lintegral_congr_ae (rnDeriv_eq_rnDeriv_measure _ _ a), ← set_lintegral_univ] + exact (Measure.set_lintegral_rnDeriv_le _).trans (measure_le_bound _ _ _) + +instance instIsFiniteKernel_singularPart [hκ : IsFiniteKernel κ] [IsFiniteKernel ν] : + IsFiniteKernel (singularPart κ ν) := by + constructor + refine ⟨hκ.bound, hκ.bound_lt_top, fun a ↦ ?_⟩ + have h : withDensity ν (rnDeriv κ ν) a univ + singularPart κ ν a univ = κ a univ := by + conv_rhs => rw [← rnDeriv_add_singularPart κ ν] + exact (self_le_add_left _ _).trans (h.le.trans (measure_le_bound _ _ _)) + +lemma rnDeriv_add (κ ν η : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] [IsFiniteKernel η] + (a : α) : + rnDeriv (κ + ν) η a =ᵐ[η a] rnDeriv κ η a + rnDeriv ν η a := by + filter_upwards [rnDeriv_eq_rnDeriv_measure (κ + ν) η a, rnDeriv_eq_rnDeriv_measure κ η a, + rnDeriv_eq_rnDeriv_measure ν η a, Measure.rnDeriv_add (κ a) (ν a) (η a)] with x h1 h2 h3 h4 + rw [h1, Pi.add_apply, h2, h3, coeFn_add, Pi.add_apply, h4, Pi.add_apply] + +lemma rnDeriv_singularPart (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : + rnDeriv (singularPart κ ν) ν a =ᵐ[ν a] 0 := by + filter_upwards [rnDeriv_eq_rnDeriv_measure (singularPart κ ν) ν a, + (Measure.rnDeriv_eq_zero _ _).mpr (mutuallySingular_singularPart κ ν a)] with x h1 h2 + rw [h1, h2] + +lemma rnDeriv_measure_compProd (μ ν : Measure α) (κ η : kernel α γ) + [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteKernel κ] [IsFiniteKernel η] : + ∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] fun p ↦ (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 := by + let μ' := ν.withDensity (∂μ/∂ν) + let κ' := withDensity η (rnDeriv κ η) + have h1 : μ = μ' + μ.singularPart ν := by + conv_lhs => rw [μ.haveLebesgueDecomposition_add ν, add_comm] + have h2 : κ = κ' + singularPart κ η := by + conv_lhs => rw [← rnDeriv_add_singularPart κ η] + have h3 : μ ⊗ₘ κ = μ' ⊗ₘ κ' + (μ.singularPart ν) ⊗ₘ κ + μ' ⊗ₘ (singularPart κ η) := by + conv_lhs => rw [h1, Measure.compProd_add_left] + conv_lhs => conv in μ' ⊗ₘ κ => rw [h2] + rw [Measure.compProd_add_right, add_assoc, add_comm (μ' ⊗ₘ singularPart κ η), ← add_assoc] + have h_left : ∂(μ' ⊗ₘ κ')/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] ∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η) := by + rw [h3] + have h_add := Measure.rnDeriv_add (μ' ⊗ₘ κ' + μ.singularPart ν ⊗ₘ κ) + (μ' ⊗ₘ (singularPart κ η)) (ν ⊗ₘ η) + have h_add' := Measure.rnDeriv_add (μ' ⊗ₘ κ') (μ.singularPart ν ⊗ₘ κ) (ν ⊗ₘ η) + have h01 : ∂(μ.singularPart ν ⊗ₘ κ)/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] 0 := by + rw [Measure.rnDeriv_eq_zero] + exact Measure.mutuallySingular_compProd_left (Measure.mutuallySingular_singularPart _ _) κ η + have h02 : ∂(μ' ⊗ₘ (singularPart κ η))/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] 0 := by + rw [Measure.rnDeriv_eq_zero] + exact Measure.mutuallySingular_compProd_right μ' ν (mutuallySingular_singularPart _ _) + filter_upwards [h_add, h_add', h01, h02] with a h_add h_add' h01 h02 + rw [h_add, Pi.add_apply, h_add', Pi.add_apply, h01, h02] + simp + have h_right : (fun p ↦ (∂μ'/∂ν) p.1 * rnDeriv κ' η p.1 p.2) + =ᵐ[ν ⊗ₘ η] (fun p ↦ (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2) := by + refine EventuallyEq.mul ?_ ?_ + · have h := Measure.rnDeriv_withDensity ν (Measure.measurable_rnDeriv μ ν) + rw [EventuallyEq, ae_iff] at h ⊢ + exact ENNReal.ae_eq_compProd_of_ae_eq_fst ν η (Measure.measurable_rnDeriv μ' ν) + (Measure.measurable_rnDeriv μ ν) h + · have : ∀ a, rnDeriv κ' η a =ᵐ[η a] rnDeriv κ η a := by + intro a + rw [h2] + filter_upwards [rnDeriv_add κ' (singularPart κ η) η a, + rnDeriv_singularPart κ η a] with x hx1 hx2 + rw [hx1, Pi.add_apply, hx2, Pi.zero_apply, add_zero] + exact ENNReal.ae_eq_compProd_of_forall_ae_eq _ _ (measurable_rnDeriv κ' η) + (measurable_rnDeriv κ η) this + refine h_left.symm.trans (EventuallyEq.trans ?_ h_right) + refine rnDeriv_measure_compProd_aux ?_ ?_ + · exact MeasureTheory.withDensity_absolutelyContinuous _ _ + · exact ae_of_all _ (fun _ ↦ withDensity_absolutelyContinuous _ _) + +lemma rnDeriv_self (κ : kernel α γ) [IsFiniteKernel κ] (a : α) : + rnDeriv κ κ a =ᵐ[κ a] 1 := + (rnDeriv_eq_rnDeriv_measure κ κ a).trans (Measure.rnDeriv_self _) + +lemma rnDeriv_measure_compProd_left (μ ν : Measure α) (κ : kernel α γ) + [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteKernel κ] : + ∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ κ) =ᵐ[ν ⊗ₘ κ] fun p ↦ (∂μ/∂ν) p.1 := by + have h : (fun p ↦ rnDeriv κ κ p.1 p.2) =ᵐ[ν ⊗ₘ κ] (fun p ↦ (1 : α → γ → ℝ≥0∞) p.1 p.2) := + ENNReal.ae_eq_compProd_of_forall_ae_eq ν κ (measurable_rnDeriv _ _) measurable_const + (rnDeriv_self κ) + filter_upwards [rnDeriv_measure_compProd μ ν κ κ, h] with p hp1 hp2 + rw [hp1, hp2] + simp + +lemma rnDeriv_measure_compProd_right (μ : Measure α) (κ η : kernel α γ) + [IsFiniteMeasure μ] [IsFiniteKernel κ] [IsFiniteKernel η] : + ∂(μ ⊗ₘ κ)/∂(μ ⊗ₘ η) =ᵐ[μ ⊗ₘ η] fun p ↦ rnDeriv κ η p.1 p.2 := by + have h : (fun p ↦ (∂μ/∂μ) p.1) =ᵐ[μ ⊗ₘ η] (fun p ↦ (1 : α → ℝ≥0∞) p.1) := + ENNReal.ae_eq_compProd_of_ae_eq_fst μ η (Measure.measurable_rnDeriv _ _) + measurable_const (Measure.rnDeriv_self _) + filter_upwards [rnDeriv_measure_compProd μ μ κ η, h] with p hp1 hp2 + rw [hp1, hp2] + simp + +end MeasureCompProd + end ProbabilityTheory.kernel From e5841102152e012dae17967b9958f7c7f4e23c31 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 10 Mar 2024 10:16:43 +0100 Subject: [PATCH 093/129] remove duplicates --- Mathlib/Probability/Kernel/RadonNikodym.lean | 58 -------------------- 1 file changed, 58 deletions(-) diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index 1142685a9ce75..1e6977bdab7b3 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -432,16 +432,6 @@ lemma measurableSet_mutuallySingular (κ ν : kernel α γ) [IsFiniteKernel κ] exact kernel.measurable_kernel_prod_mk_left (measurableSet_mutuallySingularSet κ ν).compl (measurableSet_singleton 0) --- ok -lemma Measure.compProd_apply_prod {μ : Measure α} [SFinite μ] [IsSFiniteKernel κ] - {s : Set α} {t : Set γ} (hs : MeasurableSet s) (ht : MeasurableSet t) : - (μ ⊗ₘ κ) (s ×ˢ t) = ∫⁻ a in s, κ a t ∂μ := by - rw [Measure.compProd_apply (hs.prod ht), ← lintegral_indicator _ hs] - congr with a - classical - rw [indicator_apply] - split_ifs with ha <;> simp [ha] - -- ok lemma set_lintegral_withDensity_eq_lintegral_mul₀' {μ : Measure α} {f : α → ℝ≥0∞} (hf : AEMeasurable f μ) {g : α → ℝ≥0∞} (hg : AEMeasurable g (μ.withDensity f)) @@ -530,54 +520,6 @@ lemma Measure.mutuallySingular_compProd_right (μ ν : Measure α) [SFinite μ] exact hκη a simp [h1, h2] --- ok -@[simp] -lemma prodMkLeft_add (β : Type*) [MeasurableSpace β] (κ η : kernel α γ) : - prodMkLeft β (κ + η) = prodMkLeft β κ + prodMkLeft β η := by ext; simp - --- ok -@[simp] -lemma prodMkRight_add (β : Type*) [MeasurableSpace β] (κ η : kernel α γ) : - prodMkRight β (κ + η) = prodMkRight β κ + prodMkRight β η := by ext; simp - --- ok -lemma compProd_add_left {β : Type*} {_ : MeasurableSpace β} - (μ κ : kernel α β) [IsSFiniteKernel μ] (η : kernel (α × β) γ) - [IsSFiniteKernel κ] [IsSFiniteKernel η] : - (μ + κ) ⊗ₖ η = μ ⊗ₖ η + κ ⊗ₖ η := by - ext a s hs - rw [compProd_apply _ _ _ hs] - simp only [coeFn_add, Pi.add_apply, lintegral_add_measure, Measure.add_toOuterMeasure, - OuterMeasure.coe_add] - rw [compProd_apply _ _ _ hs, compProd_apply _ _ _ hs] - --- ok -lemma compProd_add_right {β : Type*} {_ : MeasurableSpace β} - (μ : kernel α β) [IsSFiniteKernel μ] (κ η : kernel (α × β) γ) - [IsSFiniteKernel κ] [IsSFiniteKernel η] : - μ ⊗ₖ (κ + η) = μ ⊗ₖ κ + μ ⊗ₖ η := by - ext a s hs - rw [compProd_apply _ _ _ hs] - simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add] - rw [lintegral_add_left, compProd_apply _ _ _ hs, compProd_apply _ _ _ hs] - exact measurable_kernel_prod_mk_left' hs a - --- ok -lemma const_add (β : Type*) [MeasurableSpace β] (μ ν : Measure α) : - const β (μ + ν) = const β μ + const β ν := by ext; simp - --- ok -nonrec lemma Measure.compProd_add_left (μ ν : Measure α) [SFinite μ] [SFinite ν] (κ : kernel α γ) - [IsSFiniteKernel κ] : - (μ + ν) ⊗ₘ κ = μ ⊗ₘ κ + ν ⊗ₘ κ := by - rw [Measure.compProd, const_add, compProd_add_left]; rfl - --- ok -nonrec lemma Measure.compProd_add_right (μ : Measure α) [SFinite μ] (κ η : kernel α γ) - [IsSFiniteKernel κ] [IsSFiniteKernel η] : - μ ⊗ₘ (κ + η) = μ ⊗ₘ κ + μ ⊗ₘ η := by - rw [Measure.compProd, prodMkLeft_add, compProd_add_right]; rfl - lemma Measure.fst_map_compProd (μ : Measure α) [SFinite μ] (κ : kernel α γ) [IsMarkovKernel κ] : (μ ⊗ₘ κ).map Prod.fst = μ := by From 5b08df8bbd52e687a660dfe47711bc613d212551 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 13 Mar 2024 17:37:04 +0100 Subject: [PATCH 094/129] remove duplicate file --- .../Kernel/Disintegration/Density.lean | 2 +- .../Process/CountablyGenerated.lean | 84 ------------------- 2 files changed, 1 insertion(+), 85 deletions(-) delete mode 100644 Mathlib/Probability/Process/CountablyGenerated.lean diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 424bd6de20a93..fe77d6c6a682b 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -5,7 +5,7 @@ Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Martingale.Convergence -import Mathlib.Probability.Process.CountablyGenerated +import Mathlib.Probability.Process.PartitionFiltration /-! # Kernel density diff --git a/Mathlib/Probability/Process/CountablyGenerated.lean b/Mathlib/Probability/Process/CountablyGenerated.lean deleted file mode 100644 index 6b77075a26429..0000000000000 --- a/Mathlib/Probability/Process/CountablyGenerated.lean +++ /dev/null @@ -1,84 +0,0 @@ -/- -Copyright (c) 2024 Rémy Degenne. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Rémy Degenne --/ -import Mathlib.MeasureTheory.MeasurableSpace.CountablyGenerated -import Mathlib.Probability.Process.Filtration - -/-! -# Filtration built from the finite partitions of a countably generated measurable space - -In a countably generated measurable space `α`, we can build a sequence of finer and finer finite -measurable partitions of the space such that the measurable space is generated by the union of all -partitions. -This sequence of partitions is defined in `MeasureTheory.MeasurableSpace.CountablyGenerated`. - -Here, we build the filtration of the measurable spaces generated by `countablePartition α n` for all -`n : ℕ`, which we call `partitionFiltration α`. -Since each measurable space in the filtration is finite, we can easily build measurable functions on -those spaces. By building a martingale with respect to `partitionFiltration α` and using the -martingale convergence theorems, we can define a measurable function on `α`. - -## Main definitions - -* `ProbabilityTheory.partitionFiltration`: A filtration built from the measurable spaces generated - by `countablePartition α n` for all `n : ℕ`. - -## Main statements - -* `ProbabilityTheory.iSup_partitionFiltration`: `⨆ n, partitionFiltration α n` is the measurable - space on `α`. - --/ - -open MeasureTheory MeasurableSpace - -namespace ProbabilityTheory - -variable {α : Type*} [MeasurableSpace α] [CountablyGenerated α] - -/-- A filtration built from the measurable spaces generated by `countablePartition α n` for -all `n : ℕ`. -/ -def partitionFiltration (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] : - Filtration ℕ m where - seq n := generateFrom (countablePartition α n) - mono' := monotone_nat_of_le_succ (generateFrom_countablePartition_le_succ _) - le' := generateFrom_countablePartition_le α - -lemma measurableSet_partitionFiltration_of_mem_countablePartition (n : ℕ) {s : Set α} - (hs : s ∈ countablePartition α n) : - MeasurableSet[partitionFiltration α n] s := - measurableSet_generateFrom hs - -lemma measurableSet_partitionFiltration_partitionSet (n : ℕ) (t : α) : - MeasurableSet[partitionFiltration α n] (partitionSet n t) := - measurableSet_partitionFiltration_of_mem_countablePartition n (partitionSet_mem n t) - -lemma measurable_partitionSet_aux (n : ℕ) (m : MeasurableSpace (countablePartition α n)) : - @Measurable α (countablePartition α n) (partitionFiltration α n) m - (fun c : α ↦ ⟨partitionSet n c, partitionSet_mem n c⟩) := by - refine @measurable_to_countable' (countablePartition α n) α m _ - (partitionFiltration α n) _ (fun t ↦ ?_) - rcases t with ⟨t, ht⟩ - suffices MeasurableSet[partitionFiltration α n] {x | partitionSet n x = t} by - convert this - ext x - simp - simp_rw [partitionSet_eq_iff _ ht] - exact measurableSet_partitionFiltration_of_mem_countablePartition _ ht - -lemma measurable_partitionFiltration_partitionSet (α : Type*) - [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : - Measurable[partitionFiltration α n] (partitionSet n) := - measurable_subtype_coe.comp (measurable_partitionSet_aux _ _) - -lemma measurable_partitionSet (α : Type*) [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : - Measurable (partitionSet (α := α) n) := - (measurable_partitionFiltration_partitionSet α n).mono ((partitionFiltration α).le n) le_rfl - -lemma iSup_partitionFiltration (α : Type*) [m : MeasurableSpace α] [CountablyGenerated α] : - ⨆ n, partitionFiltration α n = m := by - conv_rhs => rw [← generateFrom_iUnion_countablePartition α, ← iSup_generateFrom] - -end ProbabilityTheory From 0da240a36f391c08ad4719d3f175c5acf8f7f42d Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 13 Mar 2024 17:41:11 +0100 Subject: [PATCH 095/129] fix --- .../Kernel/Disintegration/Density.lean | 188 +++++++++--------- 1 file changed, 94 insertions(+), 94 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index fe77d6c6a682b..8e95e0acfdf67 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -28,18 +28,18 @@ In order to obtain measurability through countability, we use the fact that the is countably generated. For each `n : ℕ`, we define a finite partition of `γ`, such that those partitions are finer as `n` grows, and the σ-algebra generated by the union of all partitions is the σ-algebra of `γ`. -For `x : γ`, let `partitionSet n x` be the set in the partition such that `x ∈ partitionSet n x`. +For `x : γ`, let `countablePartitionSet n x` be the set in the partition such that `x ∈ countablePartitionSet n x`. For a given `n`, the function `densityProcess κ ν n : α → γ → Set β → ℝ` defined by -`fun a x s ↦ (κ a (partitionSet n x ×ˢ s) / ν a (partitionSet n x)).toReal` has the desired +`fun a x s ↦ (κ a (countablePartitionSet n x ×ˢ s) / ν a (countablePartitionSet n x)).toReal` has the desired property that `∫ x in A, densityProcess κ ν n a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal` for all `A` in the σ-algebra generated by the partition at scale `n` and is measurable in `(a, x)`. -Let `partitionFiltration γ` be the filtration of those σ-algebras for all `n : ℕ`. +Let `countableFiltration γ` be the filtration of those σ-algebras for all `n : ℕ`. The functions `densityProcess κ ν n` described here are a bounded `ν`-martingale for the filtration -`partitionFiltration γ`. By Doob's martingale L1 convergence theorem, that martingale converges to +`countableFiltration γ`. By Doob's martingale L1 convergence theorem, that martingale converges to a limit, which has a product-measurable version and satisfies the integral equality for all `A` in -`⨆ n, partitionFiltration γ n`. Finally, the partitions were chosen such that that supremum is equal +`⨆ n, countableFiltration γ n`. Finally, the partitions were chosen such that that supremum is equal to the σ-algebra on `γ`, hence the equality holds for all measurable sets. We have obtained the desired density function. @@ -79,19 +79,19 @@ a density for those kernels for all measurable sets. -/ noncomputable def densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (x : γ) (s : Set β) : ℝ := - (κ a (partitionSet n x ×ˢ s) / ν a (partitionSet n x)).toReal + (κ a (countablePartitionSet n x ×ˢ s) / ν a (countablePartitionSet n x)).toReal lemma densityProcess_def (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (s : Set β) : (fun t ↦ densityProcess κ ν n a t s) - = fun t ↦ (κ a (partitionSet n t ×ˢ s) / ν a (partitionSet n t)).toReal := rfl + = fun t ↦ (κ a (countablePartitionSet n t ×ˢ s) / ν a (countablePartitionSet n t)).toReal := rfl -lemma measurable_densityProcess_partitionFiltration_aux (κ : kernel α (γ × β)) (ν : kernel α γ) +lemma measurable_densityProcess_countableFiltration_aux (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) {s : Set β} (hs : MeasurableSet s) : - Measurable[mα.prod (partitionFiltration γ n)] (fun (p : α × γ) ↦ - κ p.1 (partitionSet n p.2 ×ˢ s) / ν p.1 (partitionSet n p.2)) := by - change Measurable[mα.prod (partitionFiltration γ n)] + Measurable[mα.prod (countableFiltration γ n)] (fun (p : α × γ) ↦ + κ p.1 (countablePartitionSet n p.2 ×ˢ s) / ν p.1 (countablePartitionSet n p.2)) := by + change Measurable[mα.prod (countableFiltration γ n)] ((fun (p : α × countablePartition γ n) ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) - ∘ (fun (p : α × γ) ↦ (p.1, ⟨partitionSet n p.2, partitionSet_mem n p.2⟩))) + ∘ (fun (p : α × γ) ↦ (p.1, ⟨countablePartitionSet n p.2, countablePartitionSet_mem n p.2⟩))) have h1 : @Measurable _ _ (mα.prod ⊤) _ (fun p : α × countablePartition γ n ↦ κ p.1 (↑p.2 ×ˢ s) / ν p.1 p.2) := by refine Measurable.div ?_ ?_ @@ -101,16 +101,16 @@ lemma measurable_densityProcess_partitionFiltration_aux (κ : kernel α (γ × rintro ⟨t, ht⟩ exact kernel.measurable_coe _ (measurableSet_countablePartition _ ht) refine h1.comp (measurable_fst.prod_mk ?_) - change @Measurable (α × γ) (countablePartition γ n) (mα.prod (partitionFiltration γ n)) ⊤ - ((fun c : γ ↦ ⟨partitionSet n c, partitionSet_mem n c⟩) ∘ (fun p : α × γ ↦ p.2)) - exact (measurable_partitionSet_aux n ⊤).comp measurable_snd + change @Measurable (α × γ) (countablePartition γ n) (mα.prod (countableFiltration γ n)) ⊤ + ((fun c : γ ↦ ⟨countablePartitionSet n c, countablePartitionSet_mem n c⟩) ∘ (fun p : α × γ ↦ p.2)) + exact (measurable_countablePartitionSet_subtype n ⊤).comp measurable_snd lemma measurable_densityProcess_aux (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) {s : Set β} (hs : MeasurableSet s) : Measurable (fun (p : α × γ) ↦ - κ p.1 (partitionSet n p.2 ×ˢ s) / ν p.1 (partitionSet n p.2)) := by - refine Measurable.mono (measurable_densityProcess_partitionFiltration_aux κ ν n hs) ?_ le_rfl - exact sup_le_sup le_rfl (comap_mono ((partitionFiltration γ).le _)) + κ p.1 (countablePartitionSet n p.2 ×ˢ s) / ν p.1 (countablePartitionSet n p.2)) := by + refine Measurable.mono (measurable_densityProcess_countableFiltration_aux κ ν n hs) ?_ le_rfl + exact sup_le_sup le_rfl (comap_mono ((countableFiltration γ).le _)) lemma measurable_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) {s : Set β} (hs : MeasurableSet s) : @@ -127,42 +127,42 @@ lemma measurable_densityProcess_right (κ : kernel α (γ × β)) (ν : kernel Measurable (fun x ↦ densityProcess κ ν n a x s) := (measurable_densityProcess κ ν n hs).comp (measurable_const.prod_mk measurable_id) -lemma measurable_partitionFiltration_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) +lemma measurable_countableFiltration_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : - Measurable[partitionFiltration γ n] (fun x ↦ densityProcess κ ν n a x s) := by - refine @Measurable.ennreal_toReal _ (partitionFiltration γ n) _ ?_ - exact (measurable_densityProcess_partitionFiltration_aux κ ν n hs).comp measurable_prod_mk_left + Measurable[countableFiltration γ n] (fun x ↦ densityProcess κ ν n a x s) := by + refine @Measurable.ennreal_toReal _ (countableFiltration γ n) _ ?_ + exact (measurable_densityProcess_countableFiltration_aux κ ν n hs).comp measurable_prod_mk_left -lemma stronglyMeasurable_partitionFiltration_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) +lemma stronglyMeasurable_countableFiltration_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : - StronglyMeasurable[partitionFiltration γ n] (fun x ↦ densityProcess κ ν n a x s) := - (measurable_partitionFiltration_densityProcess κ ν n a hs).stronglyMeasurable + StronglyMeasurable[countableFiltration γ n] (fun x ↦ densityProcess κ ν n a x s) := + (measurable_countableFiltration_densityProcess κ ν n a hs).stronglyMeasurable lemma adapted_densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) {s : Set β} (hs : MeasurableSet s) : - Adapted (partitionFiltration γ) (fun n x ↦ densityProcess κ ν n a x s) := - fun n ↦ stronglyMeasurable_partitionFiltration_densityProcess κ ν n a hs + Adapted (countableFiltration γ) (fun n x ↦ densityProcess κ ν n a x s) := + fun n ↦ stronglyMeasurable_countableFiltration_densityProcess κ ν n a hs lemma densityProcess_nonneg (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (x : γ) (s : Set β) : 0 ≤ densityProcess κ ν n a x s := ENNReal.toReal_nonneg -lemma meas_partitionSet_le_of_fst_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : - κ a (partitionSet n x ×ˢ s) ≤ ν a (partitionSet n x) := by - calc κ a (partitionSet n x ×ˢ s) - ≤ fst κ a (partitionSet n x) := by - rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] +lemma meas_countablePartitionSet_le_of_fst_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : + κ a (countablePartitionSet n x ×ˢ s) ≤ ν a (countablePartitionSet n x) := by + calc κ a (countablePartitionSet n x ×ˢ s) + ≤ fst κ a (countablePartitionSet n x) := by + rw [fst_apply' _ _ (measurableSet_countablePartitionSet _ _)] refine measure_mono (fun x ↦ ?_) simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h - _ ≤ ν a (partitionSet n x) := hκν a _ + _ ≤ ν a (countablePartitionSet n x) := hκν a _ lemma densityProcess_le_one (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : densityProcess κ ν n a x s ≤ 1 := by refine ENNReal.toReal_le_of_le_ofReal zero_le_one (ENNReal.div_le_of_le_mul ?_) rw [ENNReal.ofReal_one, one_mul] - exact meas_partitionSet_le_of_fst_le hκν n a x s + exact meas_countablePartitionSet_le_of_fst_le hκν n a x s lemma snorm_densityProcess_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (s : Set β) : snorm (fun x ↦ densityProcess κ ν n a x s) 1 (ν a) ≤ ν a univ := by @@ -188,24 +188,24 @@ lemma set_integral_densityProcess_I (hκν : fst κ ≤ ν) [IsFiniteKernel κ] rw [integral_toReal] rotate_left · refine Measurable.aemeasurable ?_ - change Measurable ((fun (p : α × _) ↦ κ p.1 (partitionSet n p.2 ×ˢ s) - / ν p.1 (partitionSet n p.2)) ∘ (fun x ↦ (a, x))) + change Measurable ((fun (p : α × _) ↦ κ p.1 (countablePartitionSet n p.2 ×ˢ s) + / ν p.1 (countablePartitionSet n p.2)) ∘ (fun x ↦ (a, x))) exact (measurable_densityProcess_aux κ ν n hs).comp measurable_prod_mk_left · refine ae_of_all _ (fun x ↦ ?_) - by_cases h0 : ν a (partitionSet n x) = 0 - · suffices κ a (partitionSet n x ×ˢ s) = 0 by simp [h0, this] - have h0' : fst κ a (partitionSet n x) = 0 := + by_cases h0 : ν a (countablePartitionSet n x) = 0 + · suffices κ a (countablePartitionSet n x ×ˢ s) = 0 by simp [h0, this] + have h0' : fst κ a (countablePartitionSet n x) = 0 := le_antisymm ((hκν a _).trans h0.le) zero_le' - rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] at h0' + rw [fst_apply' _ _ (measurableSet_countablePartitionSet _ _)] at h0' refine measure_mono_null (fun x ↦ ?_) h0' simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h · exact ENNReal.div_lt_top (measure_ne_top _ _) h0 congr - have : ∫⁻ x in u, κ a (partitionSet n x ×ˢ s) / ν a (partitionSet n x) ∂(ν a) + have : ∫⁻ x in u, κ a (countablePartitionSet n x ×ˢ s) / ν a (countablePartitionSet n x) ∂(ν a) = ∫⁻ _ in u, κ a (u ×ˢ s) / ν a u ∂(ν a) := by refine set_lintegral_congr_fun hu_meas (ae_of_all _ (fun t ht ↦ ?_)) - rw [partitionSet_of_mem hu ht] + rw [countablePartitionSet_of_mem hu ht] rw [this] simp only [MeasureTheory.lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter] by_cases h0 : ν a u = 0 @@ -244,9 +244,9 @@ lemma integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFin lemma set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} - (hA : MeasurableSet[partitionFiltration γ n] A) : + (hA : MeasurableSet[countableFiltration γ n] A) : ∫ x in A, densityProcess κ ν n a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - refine induction_on_inter (m := partitionFiltration γ n) (s := countablePartition γ n) + refine induction_on_inter (m := countableFiltration γ n) (s := countablePartition γ n) (C := fun A ↦ ∫ t in A, densityProcess κ ν n a t s ∂(ν a) = (κ a (A ×ˢ s)).toReal) rfl ?_ ?_ ?_ ?_ ?_ hA · rintro s hs t ht hst @@ -258,7 +258,7 @@ lemma set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [I · rintro u hu rw [set_integral_densityProcess_I hκν _ _ hs hu] · intro A hA hA_eq - have hA' : MeasurableSet A := (partitionFiltration γ).le _ _ hA + have hA' : MeasurableSet A := (countableFiltration γ).le _ _ hA have h := integral_add_compl hA' (integrable_densityProcess hκν n a hs) rw [hA_eq, integral_densityProcess hκν n a hs] at h have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by rw [prod_diff_prod, compl_eq_univ_diff]; simp @@ -269,11 +269,11 @@ lemma set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [I · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] exact measure_mono (by intro x; simp) · intro f hf_disj hf h_eq - rw [integral_iUnion (fun i ↦ (partitionFiltration γ).le n _ (hf i)) hf_disj + rw [integral_iUnion (fun i ↦ (countableFiltration γ).le n _ (hf i)) hf_disj (integrable_densityProcess hκν _ _ hs).integrableOn] simp_rw [h_eq] rw [iUnion_prod_const, - measure_iUnion _ (fun i ↦ ((partitionFiltration γ).le n _ (hf i)).prod hs)] + measure_iUnion _ (fun i ↦ ((countableFiltration γ).le n _ (hf i)).prod hs)] · rw [ENNReal.tsum_toReal_eq] exact fun _ ↦ measure_ne_top _ _ · intro i j hij @@ -282,13 +282,13 @@ lemma set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [I lemma set_integral_densityProcess_of_le (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] {n m : ℕ} (hnm : n ≤ m) (a : α) {s : Set β} (hs : MeasurableSet s) - {A : Set γ} (hA : MeasurableSet[partitionFiltration γ n] A) : + {A : Set γ} (hA : MeasurableSet[countableFiltration γ n] A) : ∫ x in A, densityProcess κ ν m a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal := - set_integral_densityProcess hκν m a hs ((partitionFiltration γ).mono hnm A hA) + set_integral_densityProcess hκν m a hs ((countableFiltration γ).mono hnm A hA) lemma condexp_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] {i j : ℕ} (hij : i ≤ j) (a : α) {s : Set β} (hs : MeasurableSet s) : - (ν a)[fun x ↦ densityProcess κ ν j a x s | partitionFiltration γ i] + (ν a)[fun x ↦ densityProcess κ ν j a x s | countableFiltration γ i] =ᵐ[ν a] fun x ↦ densityProcess κ ν i a x s := by refine (ae_eq_condexp_of_forall_set_integral_eq ?_ ?_ ?_ ?_ ?_).symm · exact integrable_densityProcess hκν j a hs @@ -297,26 +297,26 @@ lemma condexp_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFini rw [set_integral_densityProcess hκν i a hs hx, set_integral_densityProcess_of_le hκν hij a hs hx] · exact StronglyMeasurable.aeStronglyMeasurable' - (stronglyMeasurable_partitionFiltration_densityProcess κ ν i a hs) + (stronglyMeasurable_countableFiltration_densityProcess κ ν i a hs) lemma martingale_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Martingale (fun n x ↦ densityProcess κ ν n a x s) (partitionFiltration γ) (ν a) := + Martingale (fun n x ↦ densityProcess κ ν n a x s) (countableFiltration γ) (ν a) := ⟨adapted_densityProcess κ ν a hs, fun _ _ h ↦ condexp_densityProcess hκν h a hs⟩ lemma densityProcess_mono_set (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) {s s' : Set β} (h : s ⊆ s') : densityProcess κ ν n a x s ≤ densityProcess κ ν n a x s' := by unfold densityProcess - by_cases h0 : ν a (partitionSet n x) = 0 + by_cases h0 : ν a (countablePartitionSet n x) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] simp - have h_ne_top : ∀ s, κ a (partitionSet n x ×ˢ s) / ν a (partitionSet n x) ≠ ⊤ := by + have h_ne_top : ∀ s, κ a (countablePartitionSet n x ×ˢ s) / ν a (countablePartitionSet n x) ≠ ⊤ := by intro s rw [ne_eq, ENNReal.div_eq_top] simp only [ne_eq, h0, and_false, false_or, not_and, not_not] refine fun h_top ↦ eq_top_mono ?_ h_top - exact meas_partitionSet_le_of_fst_le hκν n a x s + exact meas_countablePartitionSet_le_of_fst_le hκν n a x s rw [ENNReal.toReal_le_toReal (h_ne_top s) (h_ne_top s')] gcongr simp [prod_subset_prod_iff, subset_rfl, h] @@ -325,11 +325,11 @@ lemma densityProcess_mono_kernel_left {κ' : kernel α (γ × β)} (hκκ' : κ (hκ'ν : fst κ' ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : densityProcess κ ν n a x s ≤ densityProcess κ' ν n a x s := by unfold densityProcess - by_cases h0 : ν a (partitionSet n x) = 0 + by_cases h0 : ν a (countablePartitionSet n x) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] simp - have h_le : κ' a (partitionSet n x ×ˢ s) ≤ ν a (partitionSet n x) := - meas_partitionSet_le_of_fst_le hκ'ν n a x s + have h_le : κ' a (countablePartitionSet n x ×ˢ s) ≤ ν a (countablePartitionSet n x) := + meas_countablePartitionSet_le_of_fst_le hκ'ν n a x s rw [ENNReal.toReal_le_toReal] · gcongr exact hκκ' _ _ @@ -345,11 +345,11 @@ lemma densityProcess_antitone_kernel_right {ν' : kernel α γ} (hνν' : ν ≤ ν') (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : densityProcess κ ν' n a x s ≤ densityProcess κ ν n a x s := by unfold densityProcess - have h_le : κ a (partitionSet n x ×ˢ s) ≤ ν a (partitionSet n x) := - meas_partitionSet_le_of_fst_le hκν n a x s - by_cases h0 : ν a (partitionSet n x) = 0 + have h_le : κ a (countablePartitionSet n x ×ˢ s) ≤ ν a (countablePartitionSet n x) := + meas_countablePartitionSet_le_of_fst_le hκν n a x s + by_cases h0 : ν a (countablePartitionSet n x) = 0 · simp [le_antisymm (h_le.trans h0.le) zero_le', h0] - have h0' : ν' a (partitionSet n x) ≠ 0 := + have h0' : ν' a (countablePartitionSet n x) ≠ 0 := fun h ↦ h0 (le_antisymm ((hνν' _ _).trans h.le) zero_le') rw [ENNReal.toReal_le_toReal] · gcongr @@ -371,7 +371,7 @@ lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (γ × β)) (hs_meas : ∀ n, MeasurableSet (s n)) : Tendsto (fun m ↦ densityProcess κ ν n a x (s m)) atTop (𝓝 (densityProcess κ ν n a x ∅)) := by simp_rw [densityProcess] - by_cases h0 : ν a (partitionSet n x) = 0 + by_cases h0 : ν a (countablePartitionSet n x) = 0 · simp_rw [h0, ENNReal.toReal_div] simp refine (ENNReal.tendsto_toReal ?_).comp ?_ @@ -379,10 +379,10 @@ lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (γ × β)) push_neg simp refine ENNReal.Tendsto.div_const ?_ ?_ - · have h := tendsto_measure_iInter (μ := κ a) (s := fun m ↦ partitionSet n x ×ˢ s m) ?_ ?_ ?_ + · have h := tendsto_measure_iInter (μ := κ a) (s := fun m ↦ countablePartitionSet n x ×ˢ s m) ?_ ?_ ?_ · convert h rw [← prod_iInter, hs_iInter] - · exact fun n ↦ MeasurableSet.prod (measurableSet_partitionSet _ _) (hs_meas n) + · exact fun n ↦ MeasurableSet.prod (measurableSet_countablePartitionSet _ _) (hs_meas n) · intro m m' hmm' simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] exact Or.inl <| hs hmm' @@ -401,7 +401,7 @@ lemma tendsto_densityProcess_atTop_of_antitone (κ : kernel α (γ × β)) (ν : lemma tendsto_densityProcess_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : ∀ᵐ x ∂(ν a), Tendsto (fun n ↦ densityProcess κ ν n a x s) atTop - (𝓝 ((partitionFiltration γ).limitProcess + (𝓝 ((countableFiltration γ).limitProcess (fun n x ↦ densityProcess κ ν n a x s) (ν a) x)) := by refine Submartingale.ae_tendsto_limitProcess (martingale_densityProcess hκν a hs).submartingale (R := (ν a univ).toNNReal) (fun n ↦ ?_) @@ -411,7 +411,7 @@ lemma tendsto_densityProcess_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKerne lemma limitProcess_densityProcess_mem_L1 (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : - Memℒp ((partitionFiltration γ).limitProcess + Memℒp ((countableFiltration γ).limitProcess (fun n x ↦ densityProcess κ ν n a x s) (ν a)) 1 (ν a) := by refine Submartingale.memℒp_limitProcess (martingale_densityProcess hκν a hs).submartingale (R := (ν a univ).toNNReal) (fun n ↦ ?_) @@ -422,7 +422,7 @@ lemma limitProcess_densityProcess_mem_L1 (hκν : fst κ ≤ ν) [IsFiniteKernel lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : Tendsto (fun n ↦ snorm ((fun x ↦ densityProcess κ ν n a x s) - - (partitionFiltration γ).limitProcess (fun n x ↦ densityProcess κ ν n a x s) (ν a)) + - (countableFiltration γ).limitProcess (fun n x ↦ densityProcess κ ν n a x s) (ν a)) 1 (ν a)) atTop (𝓝 0) := by refine Submartingale.tendsto_snorm_one_limitProcess ?_ ?_ · exact (martingale_densityProcess hκν a hs).submartingale @@ -440,7 +440,7 @@ lemma tendsto_snorm_one_densityProcess_limitProcess (hκν : fst κ ≤ ν) [IsF lemma tendsto_snorm_one_restrict_densityProcess_limitProcess [IsFiniteKernel κ] [IsFiniteKernel ν] (hκν : fst κ ≤ ν) (a : α) {s : Set β} (hs : MeasurableSet s) (A : Set γ) : Tendsto (fun n ↦ snorm ((fun x ↦ densityProcess κ ν n a x s) - - (partitionFiltration γ).limitProcess (fun n x ↦ densityProcess κ ν n a x s) (ν a)) + - (countableFiltration γ).limitProcess (fun n x ↦ densityProcess κ ν n a x s) (ν a)) 1 ((ν a).restrict A)) atTop (𝓝 0) := tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds (tendsto_snorm_one_densityProcess_limitProcess hκν a hs) (fun _ ↦ zero_le') @@ -460,7 +460,7 @@ def density (κ : kernel α (γ × β)) (ν : kernel α γ) (a : α) (x : γ) (s lemma density_ae_eq_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : (fun x ↦ density κ ν a x s) - =ᵐ[ν a] (partitionFiltration γ).limitProcess + =ᵐ[ν a] (countableFiltration γ).limitProcess (fun n x ↦ densityProcess κ ν n a x s) (ν a) := by filter_upwards [tendsto_densityProcess_limitProcess hκν a hs] with t ht using ht.limsup_eq @@ -538,7 +538,7 @@ lemma tendsto_set_integral_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKerne /-- Auxiliary lemma for `set_integral_density`. -/ lemma set_integral_density_of_measurableSet (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} - (hA : MeasurableSet[partitionFiltration γ n] A) : + (hA : MeasurableSet[countableFiltration γ n] A) : ∫ x in A, density κ ν a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by suffices ∫ x in A, density κ ν a x s ∂(ν a) = ∫ x in A, densityProcess κ ν n a x s ∂(ν a) by exact this ▸ set_integral_densityProcess hκν _ _ hs hA @@ -562,18 +562,18 @@ lemma integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKern lemma set_integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet A) : ∫ x in A, density κ ν a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by - have hA' : MeasurableSet[⨆ n, partitionFiltration γ n] A := by rwa [iSup_partitionFiltration] - refine induction_on_inter (m := ⨆ n, partitionFiltration γ n) + have hA' : MeasurableSet[⨆ n, countableFiltration γ n] A := by rwa [iSup_countableFiltration] + refine induction_on_inter (m := ⨆ n, countableFiltration γ n) (C := fun A ↦ ∫ x in A, density κ ν a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal) - (measurableSpace_iSup_eq (partitionFiltration γ)) ?_ ?_ ?_ ?_ ?_ hA' + (measurableSpace_iSup_eq (countableFiltration γ)) ?_ ?_ ?_ ?_ ?_ hA' · rintro s ⟨n, hs⟩ t ⟨m, ht⟩ _ - exact ⟨max n m, ((partitionFiltration γ).mono (le_max_left n m) _ hs).inter - ((partitionFiltration γ).mono (le_max_right n m) _ ht)⟩ + exact ⟨max n m, ((countableFiltration γ).mono (le_max_left n m) _ hs).inter + ((countableFiltration γ).mono (le_max_right n m) _ ht)⟩ · simp · intro A ⟨n, hA⟩ exact set_integral_density_of_measurableSet hκν n a hs hA · intro A hA hA_eq - rw [iSup_partitionFiltration] at hA + rw [iSup_countableFiltration] at hA have h := integral_add_compl hA (integrable_density hκν a hs) rw [hA_eq, integral_density hκν a hs] at h have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by @@ -594,9 +594,9 @@ lemma set_integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFinite · intro i j hij rw [Function.onFun, Set.disjoint_prod] exact Or.inl (hf_disj hij) - · rw [iSup_partitionFiltration] at hf + · rw [iSup_countableFiltration] at hf exact fun i ↦ (hf i).prod hs - · rwa [iSup_partitionFiltration] at hf + · rwa [iSup_countableFiltration] at hf lemma set_lintegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet A) : @@ -721,39 +721,39 @@ section UnivFst /-! We specialize to `ν = fst κ`, for which `density κ (fst κ) a t univ = 1` almost everywhere. -/ lemma densityProcess_univ [IsFiniteKernel κ] (n : ℕ) (a : α) (x : γ) : - densityProcess κ (fst κ) n a x univ = if fst κ a (partitionSet n x) = 0 then 0 else 1 := by + densityProcess κ (fst κ) n a x univ = if fst κ a (countablePartitionSet n x) = 0 then 0 else 1 := by rw [densityProcess] - by_cases h : fst κ a (partitionSet n x) = 0 + by_cases h : fst κ a (countablePartitionSet n x) = 0 · simp [h] - by_cases h' : κ a (partitionSet n x ×ˢ univ) = 0 + by_cases h' : κ a (countablePartitionSet n x ×ˢ univ) = 0 · simp [h'] · rw [ENNReal.div_zero h'] simp · simp only [h, ite_false] - rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] - have : partitionSet n x ×ˢ univ = {p : γ × β | p.1 ∈ partitionSet n x} := by + rw [fst_apply' _ _ (measurableSet_countablePartitionSet _ _)] + have : countablePartitionSet n x ×ˢ univ = {p : γ × β | p.1 ∈ countablePartitionSet n x} := by ext x simp rw [this, ENNReal.div_self] · simp - · rwa [fst_apply' _ _ (measurableSet_partitionSet _ _)] at h + · rwa [fst_apply' _ _ (measurableSet_countablePartitionSet _ _)] at h · exact measure_ne_top _ _ lemma densityProcess_univ_ae (κ : kernel α (γ × β)) [IsFiniteKernel κ] (n : ℕ) (a : α) : ∀ᵐ x ∂(fst κ a), densityProcess κ (fst κ) n a x univ = 1 := by rw [ae_iff] have : {x | ¬ densityProcess κ (fst κ) n a x univ = 1} - ⊆ {x | fst κ a (partitionSet n x) = 0} := by + ⊆ {x | fst κ a (countablePartitionSet n x) = 0} := by intro x hx simp only [mem_setOf_eq] at hx ⊢ rw [densityProcess_univ] at hx simpa using hx refine measure_mono_null this ?_ - have : {x | fst κ a (partitionSet n x) = 0} + have : {x | fst κ a (countablePartitionSet n x) = 0} ⊆ ⋃ (u) (_ : u ∈ countablePartition γ n) (_ : fst κ a u = 0), u := by intro t ht simp only [mem_setOf_eq, mem_iUnion, exists_prop] at ht ⊢ - exact ⟨partitionSet n t, partitionSet_mem _ _, ht, mem_partitionSet _ _⟩ + exact ⟨countablePartitionSet n t, countablePartitionSet_mem _ _, ht, mem_countablePartitionSet _ _⟩ refine measure_mono_null this ?_ rw [measure_biUnion] · simp @@ -774,7 +774,7 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (γ × β)) refine (ENNReal.tendsto_toReal ?_).comp ?_ · rw [ne_eq, ENNReal.div_eq_top] push_neg - simp_rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] + simp_rw [fst_apply' _ _ (measurableSet_countablePartitionSet _ _)] constructor · refine fun h h0 ↦ h (measure_mono_null (fun x ↦ ?_) h0) simp only [mem_prod, mem_setOf_eq, and_imp] @@ -782,11 +782,11 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (γ × β)) · refine fun h_top ↦ eq_top_mono (measure_mono (fun x ↦ ?_)) h_top simp only [mem_prod, mem_setOf_eq, and_imp] exact fun h _ ↦ h - by_cases h0 : fst κ a (partitionSet n x) = 0 - · rw [fst_apply' _ _ (measurableSet_partitionSet _ _)] at h0 ⊢ - suffices ∀ m, κ a (partitionSet n x ×ˢ s m) = 0 by + by_cases h0 : fst κ a (countablePartitionSet n x) = 0 + · rw [fst_apply' _ _ (measurableSet_countablePartitionSet _ _)] at h0 ⊢ + suffices ∀ m, κ a (countablePartitionSet n x ×ˢ s m) = 0 by simp only [this, h0, ENNReal.zero_div, tendsto_const_nhds_iff] - suffices κ a (partitionSet n x ×ˢ univ) = 0 by + suffices κ a (countablePartitionSet n x ×ˢ univ) = 0 by simp only [this, ENNReal.zero_div] convert h0 ext x @@ -796,7 +796,7 @@ lemma tendsto_densityProcess_atTop_univ_of_monotone (κ : kernel α (γ × β)) exact fun h _ ↦ h refine ENNReal.Tendsto.div_const ?_ ?_ · have h := tendsto_measure_iUnion (μ := κ a) - (s := fun m ↦ partitionSet n x ×ˢ s m) ?_ + (s := fun m ↦ countablePartitionSet n x ×ˢ s m) ?_ swap · intro m m' hmm' simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] From 0ef67e73d27d26f859d4ddf2db9172c50ea0e5b8 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 13 Mar 2024 17:41:58 +0100 Subject: [PATCH 096/129] fix Mathlib.lean --- Mathlib.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib.lean b/Mathlib.lean index 468e0e6ae2e7d..ba007720fa340 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3207,7 +3207,6 @@ import Mathlib.Probability.ProbabilityMassFunction.Constructions import Mathlib.Probability.ProbabilityMassFunction.Integrals import Mathlib.Probability.ProbabilityMassFunction.Monad import Mathlib.Probability.Process.Adapted -import Mathlib.Probability.Process.CountablyGenerated import Mathlib.Probability.Process.Filtration import Mathlib.Probability.Process.HittingTime import Mathlib.Probability.Process.PartitionFiltration From 5a4a32d1fdc0bec81c35a7132b2215d7156f57fb Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 13 Mar 2024 17:52:47 +0100 Subject: [PATCH 097/129] golf --- .../Kernel/Disintegration/Density.lean | 143 ++++-------------- 1 file changed, 26 insertions(+), 117 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 8e95e0acfdf67..e523268e6d899 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -28,12 +28,13 @@ In order to obtain measurability through countability, we use the fact that the is countably generated. For each `n : ℕ`, we define a finite partition of `γ`, such that those partitions are finer as `n` grows, and the σ-algebra generated by the union of all partitions is the σ-algebra of `γ`. -For `x : γ`, let `countablePartitionSet n x` be the set in the partition such that `x ∈ countablePartitionSet n x`. +For `x : γ`, let `countablePartitionSet n x` be the set in the partition such that +`x ∈ countablePartitionSet n x`. For a given `n`, the function `densityProcess κ ν n : α → γ → Set β → ℝ` defined by -`fun a x s ↦ (κ a (countablePartitionSet n x ×ˢ s) / ν a (countablePartitionSet n x)).toReal` has the desired -property that `∫ x in A, densityProcess κ ν n a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal` for all `A` -in the σ-algebra generated by the partition at scale `n` and is measurable in `(a, x)`. +`fun a x s ↦ (κ a (countablePartitionSet n x ×ˢ s) / ν a (countablePartitionSet n x)).toReal` has +the desired property that `∫ x in A, densityProcess κ ν n a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal` for +all `A` in the σ-algebra generated by the partition at scale `n` and is measurable in `(a, x)`. Let `countableFiltration γ` be the filtration of those σ-algebras for all `n : ℕ`. The functions `densityProcess κ ν n` described here are a bounded `ν`-martingale for the filtration @@ -83,7 +84,8 @@ def densityProcess (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a lemma densityProcess_def (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) (a : α) (s : Set β) : (fun t ↦ densityProcess κ ν n a t s) - = fun t ↦ (κ a (countablePartitionSet n t ×ˢ s) / ν a (countablePartitionSet n t)).toReal := rfl + = fun t ↦ (κ a (countablePartitionSet n t ×ˢ s) / ν a (countablePartitionSet n t)).toReal := + rfl lemma measurable_densityProcess_countableFiltration_aux (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) {s : Set β} (hs : MeasurableSet s) : @@ -102,7 +104,7 @@ lemma measurable_densityProcess_countableFiltration_aux (κ : kernel α (γ × exact kernel.measurable_coe _ (measurableSet_countablePartition _ ht) refine h1.comp (measurable_fst.prod_mk ?_) change @Measurable (α × γ) (countablePartition γ n) (mα.prod (countableFiltration γ n)) ⊤ - ((fun c : γ ↦ ⟨countablePartitionSet n c, countablePartitionSet_mem n c⟩) ∘ (fun p : α × γ ↦ p.2)) + ((fun c ↦ ⟨countablePartitionSet n c, countablePartitionSet_mem n c⟩) ∘ (fun p : α × γ ↦ p.2)) exact (measurable_countablePartitionSet_subtype n ⊤).comp measurable_snd lemma measurable_densityProcess_aux (κ : kernel α (γ × β)) (ν : kernel α γ) (n : ℕ) @@ -148,7 +150,8 @@ lemma densityProcess_nonneg (κ : kernel α (γ × β)) (ν : kernel α γ) (n : 0 ≤ densityProcess κ ν n a x s := ENNReal.toReal_nonneg -lemma meas_countablePartitionSet_le_of_fst_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : +lemma meas_countablePartitionSet_le_of_fst_le (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) + (s : Set β) : κ a (countablePartitionSet n x ×ˢ s) ≤ ν a (countablePartitionSet n x) := by calc κ a (countablePartitionSet n x ×ˢ s) ≤ fst κ a (countablePartitionSet n x) := by @@ -311,7 +314,8 @@ lemma densityProcess_mono_set (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ by_cases h0 : ν a (countablePartitionSet n x) = 0 · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] simp - have h_ne_top : ∀ s, κ a (countablePartitionSet n x ×ˢ s) / ν a (countablePartitionSet n x) ≠ ⊤ := by + have h_ne_top : ∀ s, + κ a (countablePartitionSet n x ×ˢ s) / ν a (countablePartitionSet n x) ≠ ⊤ := by intro s rw [ne_eq, ENNReal.div_eq_top] simp only [ne_eq, h0, and_false, false_or, not_and, not_not] @@ -379,7 +383,8 @@ lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (γ × β)) push_neg simp refine ENNReal.Tendsto.div_const ?_ ?_ - · have h := tendsto_measure_iInter (μ := κ a) (s := fun m ↦ countablePartitionSet n x ×ˢ s m) ?_ ?_ ?_ + · have h := tendsto_measure_iInter (μ := κ a) (s := fun m ↦ countablePartitionSet n x ×ˢ s m) + ?_ ?_ ?_ · convert h rw [← prod_iInter, hs_iInter] · exact fun n ↦ MeasurableSet.prod (measurableSet_countablePartitionSet _ _) (hs_meas n) @@ -660,61 +665,12 @@ lemma tendsto_density_atTop_ae_of_antitone (hκν : fst κ ≤ ν) [IsFiniteKern [IsFiniteKernel ν] (a : α) (s : ℕ → Set β) (hs : Antitone s) (hs_iInter : ⋂ i, s i = ∅) (hs_meas : ∀ n, MeasurableSet (s n)) : ∀ᵐ x ∂(ν a), Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 0) := by - have h_anti : ∀ x, Antitone (fun m ↦ density κ ν a x (s m)) := - fun x n m hnm ↦ density_mono_set hκν a x (hs hnm) - have h_le_one : ∀ m x, density κ ν a x (s m) ≤ 1 := fun m x ↦ density_le_one hκν a x (s m) - -- for all `x`, `fun m ↦ density κ a (s m) x` has a limit - have h_exists : ∀ x, ∃ l, Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 l) := by - intro x - have h_tendsto : Tendsto (fun m ↦ density κ ν a x (s m)) atTop atBot ∨ - ∃ l, Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 l) := - tendsto_of_antitone (h_anti x) - cases' h_tendsto with h_absurd h_tendsto - · rw [tendsto_atTop_atBot_iff_of_antitone (h_anti x)] at h_absurd - obtain ⟨r, hr⟩ := h_absurd (-1) - have h_nonneg := density_nonneg hκν a x (s r) - linarith - · exact h_tendsto - -- let `F` be the pointwise limit of `fun m ↦ density κ a (s m) t` for all `t` - let F : γ → ℝ := fun x ↦ (h_exists x).choose - have hF_tendsto : ∀ x, Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 (F x)) := - fun x ↦ (h_exists x).choose_spec - have hF_nonneg : ∀ x, 0 ≤ F x := - fun x ↦ ge_of_tendsto' (hF_tendsto x) (fun m ↦ density_nonneg hκν a x (s m)) - have hF_le_one : ∀ x, F x ≤ 1 := fun x ↦ le_of_tendsto' (hF_tendsto x) (fun m ↦ h_le_one m x) - have hF_int : Integrable F (ν a) := by - rw [← memℒp_one_iff_integrable] - refine ⟨?_, ?_⟩ - · refine aestronglyMeasurable_of_tendsto_ae atTop (fun n ↦ ?_) (ae_of_all _ hF_tendsto) - exact (measurable_density_right κ ν (hs_meas _) a).aestronglyMeasurable - · rw [snorm_one_eq_lintegral_nnnorm] - calc ∫⁻ x, ‖F x‖₊ ∂(ν a) ≤ ∫⁻ _, 1 ∂(ν a) := by - refine lintegral_mono (fun x ↦ ?_) - rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, ENNReal.ofReal_le_one, - abs_of_nonneg (hF_nonneg _)] - exact hF_le_one _ - _ < ⊤ := by - simp only [MeasureTheory.lintegral_const, one_mul] - exact measure_lt_top _ _ - -- it suffices to show that the limit `F` is 0 a.e. - suffices F=ᵐ[ν a] 0 by - filter_upwards [this] with x hx_eq - simp only [Pi.zero_apply] at hx_eq - rw [← hx_eq] - exact hF_tendsto x - -- since `F` is nonnegative, proving that its integral is 0 is sufficient to get that - -- `F` is 0 a.e. - rw [← integral_eq_zero_iff_of_nonneg hF_nonneg hF_int] - have h_integral : Tendsto (fun m : ℕ ↦ ∫ x, density κ ν a x (s m) ∂(ν a)) atTop - (𝓝 (∫ x, F x ∂(ν a))) := by - refine integral_tendsto_of_tendsto_of_antitone ?_ hF_int ?_ ?_ - · exact fun n ↦ integrable_density hκν _ (hs_meas n) - · exact ae_of_all _ h_anti - · exact ae_of_all _ hF_tendsto - have h_integral' : Tendsto (fun m : ℕ ↦ ∫ x, density κ ν a x (s m) ∂(ν a)) atTop - (𝓝 0) := by + refine tendsto_of_integral_tendsto_of_antitone ?_ (integrable_const _) ?_ ?_ ?_ + · exact fun n ↦ integrable_density hκν _ (hs_meas n) + · rw [integral_zero] exact tendsto_integral_density_of_antitone hκν a s hs hs_iInter hs_meas - exact tendsto_nhds_unique h_integral h_integral' + · exact ae_of_all _ (fun c n m hnm ↦ density_mono_set hκν a c (hs hnm)) + · exact ae_of_all _ (fun x m ↦ density_nonneg hκν a x (s m)) section UnivFst @@ -823,61 +779,14 @@ lemma tendsto_density_atTop_ae_of_monotone [IsFiniteKernel κ] (a : α) (s : ℕ → Set β) (hs : Monotone s) (hs_iUnion : ⋃ i, s i = univ) (hs_meas : ∀ n, MeasurableSet (s n)) : ∀ᵐ x ∂(fst κ a), Tendsto (fun m ↦ density κ (fst κ) a x (s m)) atTop (𝓝 1) := by - let ν := fst κ - have h_mono : ∀ x, Monotone (fun m ↦ density κ (fst κ) a x (s m)) := - fun x n m hnm ↦ density_mono_set le_rfl a x (hs hnm) - have h_le_one : ∀ m x, density κ ν a x (s m) ≤ 1 := fun m x ↦ density_le_one le_rfl a x (s m) - -- for all `x`, `fun m ↦ density κ a (s m) x` has a limit - have h_exists : ∀ x, ∃ l, Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 l) := by - intro x - have h_tendsto : Tendsto (fun m ↦ density κ ν a x (s m)) atTop atTop ∨ - ∃ l, Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 l) := - tendsto_of_monotone (h_mono x) - cases' h_tendsto with h_absurd h_tendsto - · rw [tendsto_atTop_atTop_iff_of_monotone (h_mono x)] at h_absurd - obtain ⟨r, hr⟩ := h_absurd 2 - exact absurd (hr.trans (h_le_one r x)) one_lt_two.not_le - · exact h_tendsto - -- let `F` be the pointwise limit of `fun m ↦ density κ a (s m) x` for all `x` - let F : γ → ℝ := fun x ↦ (h_exists x).choose - have hF_tendsto : ∀ x, Tendsto (fun m ↦ density κ ν a x (s m)) atTop (𝓝 (F x)) := - fun x ↦ (h_exists x).choose_spec - have hF_nonneg : ∀ x, 0 ≤ F x := - fun x ↦ ge_of_tendsto' (hF_tendsto x) (fun m ↦ density_nonneg le_rfl a x (s m)) - have hF_le_one : ∀ x, F x ≤ 1 := fun x ↦ le_of_tendsto' (hF_tendsto x) (fun m ↦ h_le_one m x) - have hF_int : Integrable F (ν a) := by - rw [← memℒp_one_iff_integrable] - constructor - · refine aestronglyMeasurable_of_tendsto_ae atTop (fun n ↦ ?_) (ae_of_all _ hF_tendsto) - exact (measurable_density_right κ ν (hs_meas _) a).aestronglyMeasurable - · rw [snorm_one_eq_lintegral_nnnorm] - calc ∫⁻ x, ‖F x‖₊ ∂(ν a) ≤ ∫⁻ _, 1 ∂(ν a) := by - refine lintegral_mono (fun x ↦ ?_) - rw [← ofReal_norm_eq_coe_nnnorm, Real.norm_eq_abs, ENNReal.ofReal_le_one, - abs_of_nonneg (hF_nonneg _)] - exact hF_le_one _ - _ < ⊤ := by simp only [MeasureTheory.lintegral_const, measure_univ, one_mul, measure_lt_top] - -- it suffices to show that the limit `F` is 1 a.e. - suffices F =ᵐ[ν a] (fun _ ↦ 1) by - filter_upwards [this] with x hx_eq - rw [← hx_eq] - exact hF_tendsto x - -- since `F` is at most 1, proving that its integral is the same as the integral of 1 will tell - -- us that `F` is 1 a.e. - rw [← integral_eq_iff_of_ae_le hF_int (integrable_const _) (ae_of_all _ hF_le_one)] - have h_integral : Tendsto (fun m : ℕ ↦ ∫ x, density κ ν a x (s m) ∂(ν a)) atTop - (𝓝 (∫ x, F x ∂(ν a))) := by - refine integral_tendsto_of_tendsto_of_monotone ?_ hF_int ?_ ?_ - · exact fun n ↦ integrable_density le_rfl _ (hs_meas n) - · exact ae_of_all _ h_mono - · exact ae_of_all _ hF_tendsto - have h_integral' : Tendsto (fun m : ℕ ↦ ∫ x, density κ ν a x (s m) ∂(ν a)) atTop - (𝓝 (∫ _, 1 ∂(ν a))) := by - rw [MeasureTheory.integral_const] - simp only [smul_eq_mul, mul_one] + refine tendsto_of_integral_tendsto_of_monotone ?_ (integrable_const _) ?_ ?_ ?_ + · exact fun n ↦ integrable_density le_rfl _ (hs_meas n) + · rw [MeasureTheory.integral_const, smul_eq_mul, mul_one] + convert tendsto_integral_density_of_monotone (κ := κ) le_rfl a s hs hs_iUnion hs_meas rw [fst_apply' _ _ MeasurableSet.univ] - exact tendsto_integral_density_of_monotone le_rfl a s hs hs_iUnion hs_meas - exact tendsto_nhds_unique h_integral h_integral' + rfl + · exact ae_of_all _ (fun c n m hnm ↦ density_mono_set le_rfl a c (hs hnm)) + · exact ae_of_all _ (fun x m ↦ density_le_one le_rfl a x (s m)) end UnivFst From 044808fcade9c1f557e918f5f2d83844234e0fdf Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 13 Mar 2024 17:57:41 +0100 Subject: [PATCH 098/129] docstring --- .../Probability/Kernel/Disintegration/Density.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index e523268e6d899..2110132f1c676 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -25,10 +25,10 @@ def f (κ : kernel α (γ × β)) (ν : kernel a γ) (a : α) (x : γ) (s : Set However, we can't turn those functions for each `a` into a measurable function of the pair `(a, x)`. In order to obtain measurability through countability, we use the fact that the measurable space `γ` -is countably generated. For each `n : ℕ`, we define a finite partition of `γ`, such that those -partitions are finer as `n` grows, and the σ-algebra generated by the union of all partitions is -the σ-algebra of `γ`. -For `x : γ`, let `countablePartitionSet n x` be the set in the partition such that +is countably generated. For each `n : ℕ`, we define (in the file +Probability/Process/PartitionFiltration) a finite partition of `γ`, such that those partitions are +finer as `n` grows, and the σ-algebra generated by the union of all partitions is the σ-algebra +of `γ`. For `x : γ`, `countablePartitionSet n x` denotes the set in the partition such that `x ∈ countablePartitionSet n x`. For a given `n`, the function `densityProcess κ ν n : α → γ → Set β → ℝ` defined by @@ -36,7 +36,7 @@ For a given `n`, the function `densityProcess κ ν n : α → γ → Set β → the desired property that `∫ x in A, densityProcess κ ν n a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal` for all `A` in the σ-algebra generated by the partition at scale `n` and is measurable in `(a, x)`. -Let `countableFiltration γ` be the filtration of those σ-algebras for all `n : ℕ`. +`countableFiltration γ` is the filtration of those σ-algebras for all `n : ℕ`. The functions `densityProcess κ ν n` described here are a bounded `ν`-martingale for the filtration `countableFiltration γ`. By Doob's martingale L1 convergence theorem, that martingale converges to a limit, which has a product-measurable version and satisfies the integral equality for all `A` in @@ -46,7 +46,7 @@ We have obtained the desired density function. ## Main definitions -* `ProbabilityTheory.kernel.density`: for `k : kernel α (γ × β)` and `ν : kernel α γ` two finite +* `ProbabilityTheory.kernel.density`: for `κ : kernel α (γ × β)` and `ν : kernel α γ` two finite kernels, `kernel.density κ ν` is a function `α → γ → Set β → ℝ`. ## Main statements From a045edee66a7387d41041c8c291951dd41a175c6 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 13 Mar 2024 17:58:38 +0100 Subject: [PATCH 099/129] line length --- Mathlib/Probability/Kernel/Disintegration/Density.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 2110132f1c676..e4f8607d8cd65 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -677,7 +677,8 @@ section UnivFst /-! We specialize to `ν = fst κ`, for which `density κ (fst κ) a t univ = 1` almost everywhere. -/ lemma densityProcess_univ [IsFiniteKernel κ] (n : ℕ) (a : α) (x : γ) : - densityProcess κ (fst κ) n a x univ = if fst κ a (countablePartitionSet n x) = 0 then 0 else 1 := by + densityProcess κ (fst κ) n a x univ + = if fst κ a (countablePartitionSet n x) = 0 then 0 else 1 := by rw [densityProcess] by_cases h : fst κ a (countablePartitionSet n x) = 0 · simp [h] @@ -709,7 +710,8 @@ lemma densityProcess_univ_ae (κ : kernel α (γ × β)) [IsFiniteKernel κ] (n ⊆ ⋃ (u) (_ : u ∈ countablePartition γ n) (_ : fst κ a u = 0), u := by intro t ht simp only [mem_setOf_eq, mem_iUnion, exists_prop] at ht ⊢ - exact ⟨countablePartitionSet n t, countablePartitionSet_mem _ _, ht, mem_countablePartitionSet _ _⟩ + exact ⟨countablePartitionSet n t, countablePartitionSet_mem _ _, ht, + mem_countablePartitionSet _ _⟩ refine measure_mono_null this ?_ rw [measure_biUnion] · simp From ae811167b92f93272d3707f0b01022a7b5f04367 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 13 Mar 2024 18:12:53 +0100 Subject: [PATCH 100/129] remove duplicates --- Mathlib/Probability/Kernel/RadonNikodym.lean | 74 -------------------- 1 file changed, 74 deletions(-) diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index 1e6977bdab7b3..7fe973d263596 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -432,32 +432,6 @@ lemma measurableSet_mutuallySingular (κ ν : kernel α γ) [IsFiniteKernel κ] exact kernel.measurable_kernel_prod_mk_left (measurableSet_mutuallySingularSet κ ν).compl (measurableSet_singleton 0) --- ok -lemma set_lintegral_withDensity_eq_lintegral_mul₀' {μ : Measure α} {f : α → ℝ≥0∞} - (hf : AEMeasurable f μ) {g : α → ℝ≥0∞} (hg : AEMeasurable g (μ.withDensity f)) - {s : Set α} (hs : MeasurableSet s) : - ∫⁻ a in s, g a ∂μ.withDensity f = ∫⁻ a in s, (f * g) a ∂μ := by - rw [restrict_withDensity hs, lintegral_withDensity_eq_lintegral_mul₀' hf.restrict] - rw [← restrict_withDensity hs] - exact hg.restrict - --- ok -lemma set_lintegral_withDensity_eq_lintegral_mul₀ {μ : Measure α} {f : α → ℝ≥0∞} - (hf : AEMeasurable f μ) {g : α → ℝ≥0∞} (hg : AEMeasurable g μ) - {s : Set α} (hs : MeasurableSet s) : - ∫⁻ a in s, g a ∂μ.withDensity f = ∫⁻ a in s, (f * g) a ∂μ := - set_lintegral_withDensity_eq_lintegral_mul₀' hf - (hg.mono' (MeasureTheory.withDensity_absolutelyContinuous μ f)) hs - --- ok -lemma set_lintegral_rnDeriv_mul {μ ν : Measure α} [Measure.HaveLebesgueDecomposition μ ν] - (hμν : μ ≪ ν) {f : α → ℝ≥0∞} (hf : AEMeasurable f ν) {s : Set α} (hs : MeasurableSet s) : - ∫⁻ x in s, μ.rnDeriv ν x * f x ∂ν = ∫⁻ x in s, f x ∂μ := by - nth_rw 2 [← Measure.withDensity_rnDeriv_eq μ ν hμν] - rw [set_lintegral_withDensity_eq_lintegral_mul₀ (Measure.measurable_rnDeriv μ ν).aemeasurable hf - hs] - rfl - -- ok lemma Measure.absolutelyContinuous_compProd_left {μ ν : Measure α} [SFinite μ] [SFinite ν] (hμν : μ ≪ ν) (κ : kernel α γ) [IsSFiniteKernel κ] : @@ -540,42 +514,6 @@ lemma Measure.fst_map_compProd (μ : Measure α) [SFinite μ] (κ : kernel α γ rw [lintegral_indicator _ hs] simp --- ok -lemma ae_compProd_of_ae_ae {β : Type*} {_ : MeasurableSpace β} - {κ : kernel α β} {η : kernel (α × β) γ} [IsSFiniteKernel κ] [IsSFiniteKernel η] {a : α} - {p : β × γ → Prop} (hp : MeasurableSet {x | p x}) (h : ∀ᵐ b ∂κ a, ∀ᵐ c ∂η (a, b), p (b, c)) : - ∀ᵐ bc ∂(κ ⊗ₖ η) a, p bc := by - simp_rw [ae_iff] at h ⊢ - rw [compProd_null] - · exact h - · exact hp.compl - --- ok -lemma ae_compProd_iff {β : Type*} {_ : MeasurableSpace β} - {κ : kernel α β} {η : kernel (α × β) γ} [IsSFiniteKernel κ] [IsSFiniteKernel η] {a : α} - {p : β × γ → Prop} (hp : MeasurableSet {x | p x}) : - (∀ᵐ bc ∂(κ ⊗ₖ η) a, p bc) ↔ ∀ᵐ b ∂κ a, ∀ᵐ c ∂η (a, b), p (b, c) := - ⟨fun h ↦ ae_ae_of_ae_compProd h, fun h ↦ ae_compProd_of_ae_ae hp h⟩ - --- ok -lemma Measure.ae_compProd_of_ae_ae {μ : Measure α} - [SFinite μ] [IsSFiniteKernel κ] {p : α × γ → Prop} (hp : MeasurableSet {x | p x}) - (h : ∀ᵐ a ∂μ, ∀ᵐ b ∂(κ a), p (a, b)) : - ∀ᵐ x ∂(μ ⊗ₘ κ), p x := - kernel.ae_compProd_of_ae_ae hp h - --- ok -lemma Measure.ae_ae_of_ae_compProd {μ : Measure α} [SFinite μ] [IsSFiniteKernel κ] - {p : α × γ → Prop} (h : ∀ᵐ x ∂(μ ⊗ₘ κ), p x) : - ∀ᵐ a ∂μ, ∀ᵐ b ∂(κ a), p (a, b) := - kernel.ae_ae_of_ae_compProd h - --- ok -lemma Measure.ae_compProd_iff {μ : Measure α} - [SFinite μ] [IsSFiniteKernel κ] {p : α × γ → Prop} (hp : MeasurableSet {x | p x}) : - (∀ᵐ x ∂(μ ⊗ₘ κ), p x) ↔ ∀ᵐ a ∂μ, ∀ᵐ b ∂(κ a), p (a, b) := - kernel.ae_compProd_iff hp - lemma ae_compProd_of_ae_fst {μ : Measure α} (κ : kernel α γ) [SFinite μ] [IsSFiniteKernel κ] {p : α → Prop} (hp : MeasurableSet {x | p x}) (h : ∀ᵐ a ∂μ, p a) : @@ -602,18 +540,6 @@ lemma ae_eq_compProd_of_forall_ae_eq {β : Type*} {_ : MeasurableSpace β} [AddG ae_compProd_of_ae_ae (measurableSet_eq_fun hf hg) (ae_of_all _ (fun a ↦ measure_mono_null (fun x ↦ by simp) (h a))) --- ok -lemma measurableSet_eq_fun' {β : Type*} [CanonicallyOrderedAddCommMonoid β] [Sub β] [OrderedSub β] - {_ : MeasurableSpace β} [MeasurableSub₂ β] [MeasurableSingletonClass β] - {f g : α → β} (hf : Measurable f) (hg : Measurable g) : - MeasurableSet {x | f x = g x} := by - have : {a | f a = g a} = {a | (f - g) a = 0} ∩ {a | (g - f) a = 0} := by - ext a - simp only [mem_setOf_eq, Pi.sub_apply, tsub_eq_zero_iff_le, mem_inter_iff] - exact ⟨fun h ↦ ⟨h.le, h.symm.le⟩, fun h ↦ le_antisymm h.1 h.2⟩ - rw [this] - exact ((hf.sub hg) (measurableSet_singleton 0)).inter ((hg.sub hf) (measurableSet_singleton 0)) - lemma ENNReal.ae_eq_compProd_of_ae_eq_fst (μ : Measure α) (κ : kernel α γ) [SFinite μ] [IsSFiniteKernel κ] {f g : α → ℝ≥0∞} (hf : Measurable f) (hg : Measurable g) (h : f =ᵐ[μ] g) : From 6b696344f99b906a898cd394b8b64e472fb5aba5 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 13 Mar 2024 18:24:14 +0100 Subject: [PATCH 101/129] refactor, with todo names --- Mathlib/Probability/Kernel/RadonNikodym.lean | 80 +++++++++++--------- 1 file changed, 45 insertions(+), 35 deletions(-) diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index 7fe973d263596..dd5b87e623635 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -709,49 +709,59 @@ lemma rnDeriv_singularPart (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteK (Measure.rnDeriv_eq_zero _ _).mpr (mutuallySingular_singularPart κ ν a)] with x h1 h2 rw [h1, h2] -lemma rnDeriv_measure_compProd (μ ν : Measure α) (κ η : kernel α γ) +lemma todo1 (μ ν : Measure α) (κ η : kernel α γ) [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteKernel κ] [IsFiniteKernel η] : - ∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] fun p ↦ (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 := by + ∂(ν.withDensity (∂μ/∂ν) ⊗ₘ withDensity η (rnDeriv κ η))/∂(ν ⊗ₘ η) + =ᵐ[ν ⊗ₘ η] ∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η) := by let μ' := ν.withDensity (∂μ/∂ν) let κ' := withDensity η (rnDeriv κ η) have h1 : μ = μ' + μ.singularPart ν := by conv_lhs => rw [μ.haveLebesgueDecomposition_add ν, add_comm] - have h2 : κ = κ' + singularPart κ η := by - conv_lhs => rw [← rnDeriv_add_singularPart κ η] - have h3 : μ ⊗ₘ κ = μ' ⊗ₘ κ' + (μ.singularPart ν) ⊗ₘ κ + μ' ⊗ₘ (singularPart κ η) := by + have h2 : μ ⊗ₘ κ = μ' ⊗ₘ κ' + (μ.singularPart ν) ⊗ₘ κ + μ' ⊗ₘ (singularPart κ η) := by conv_lhs => rw [h1, Measure.compProd_add_left] - conv_lhs => conv in μ' ⊗ₘ κ => rw [h2] + conv_lhs => conv in μ' ⊗ₘ κ => rw [← rnDeriv_add_singularPart κ η] rw [Measure.compProd_add_right, add_assoc, add_comm (μ' ⊗ₘ singularPart κ η), ← add_assoc] - have h_left : ∂(μ' ⊗ₘ κ')/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] ∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η) := by - rw [h3] - have h_add := Measure.rnDeriv_add (μ' ⊗ₘ κ' + μ.singularPart ν ⊗ₘ κ) - (μ' ⊗ₘ (singularPart κ η)) (ν ⊗ₘ η) - have h_add' := Measure.rnDeriv_add (μ' ⊗ₘ κ') (μ.singularPart ν ⊗ₘ κ) (ν ⊗ₘ η) - have h01 : ∂(μ.singularPart ν ⊗ₘ κ)/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] 0 := by - rw [Measure.rnDeriv_eq_zero] - exact Measure.mutuallySingular_compProd_left (Measure.mutuallySingular_singularPart _ _) κ η - have h02 : ∂(μ' ⊗ₘ (singularPart κ η))/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] 0 := by - rw [Measure.rnDeriv_eq_zero] - exact Measure.mutuallySingular_compProd_right μ' ν (mutuallySingular_singularPart _ _) - filter_upwards [h_add, h_add', h01, h02] with a h_add h_add' h01 h02 - rw [h_add, Pi.add_apply, h_add', Pi.add_apply, h01, h02] - simp - have h_right : (fun p ↦ (∂μ'/∂ν) p.1 * rnDeriv κ' η p.1 p.2) + rw [h2] + have h_add := Measure.rnDeriv_add (μ' ⊗ₘ κ' + μ.singularPart ν ⊗ₘ κ) + (μ' ⊗ₘ (singularPart κ η)) (ν ⊗ₘ η) + have h_add' := Measure.rnDeriv_add (μ' ⊗ₘ κ') (μ.singularPart ν ⊗ₘ κ) (ν ⊗ₘ η) + have h01 : ∂(μ.singularPart ν ⊗ₘ κ)/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] 0 := by + rw [Measure.rnDeriv_eq_zero] + exact Measure.mutuallySingular_compProd_left (Measure.mutuallySingular_singularPart _ _) κ η + have h02 : ∂(μ' ⊗ₘ (singularPart κ η))/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] 0 := by + rw [Measure.rnDeriv_eq_zero] + exact Measure.mutuallySingular_compProd_right μ' ν (mutuallySingular_singularPart _ _) + filter_upwards [h_add, h_add', h01, h02] with a h_add h_add' h01 h02 + rw [h_add, Pi.add_apply, h_add', Pi.add_apply, h01, h02] + simp + +lemma todo2 (μ ν : Measure α) (κ η : kernel α γ) + [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteKernel κ] [IsFiniteKernel η] : + (fun p ↦ (∂(ν.withDensity (∂μ/∂ν))/∂ν) p.1 * rnDeriv (withDensity η (rnDeriv κ η)) η p.1 p.2) =ᵐ[ν ⊗ₘ η] (fun p ↦ (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2) := by - refine EventuallyEq.mul ?_ ?_ - · have h := Measure.rnDeriv_withDensity ν (Measure.measurable_rnDeriv μ ν) - rw [EventuallyEq, ae_iff] at h ⊢ - exact ENNReal.ae_eq_compProd_of_ae_eq_fst ν η (Measure.measurable_rnDeriv μ' ν) - (Measure.measurable_rnDeriv μ ν) h - · have : ∀ a, rnDeriv κ' η a =ᵐ[η a] rnDeriv κ η a := by - intro a - rw [h2] - filter_upwards [rnDeriv_add κ' (singularPart κ η) η a, - rnDeriv_singularPart κ η a] with x hx1 hx2 - rw [hx1, Pi.add_apply, hx2, Pi.zero_apply, add_zero] - exact ENNReal.ae_eq_compProd_of_forall_ae_eq _ _ (measurable_rnDeriv κ' η) - (measurable_rnDeriv κ η) this - refine h_left.symm.trans (EventuallyEq.trans ?_ h_right) + let μ' := ν.withDensity (∂μ/∂ν) + let κ' := withDensity η (rnDeriv κ η) + refine EventuallyEq.mul ?_ ?_ + · have h := Measure.rnDeriv_withDensity ν (Measure.measurable_rnDeriv μ ν) + rw [EventuallyEq, ae_iff] at h ⊢ + exact ENNReal.ae_eq_compProd_of_ae_eq_fst ν η (Measure.measurable_rnDeriv μ' ν) + (Measure.measurable_rnDeriv μ ν) h + · have : ∀ a, rnDeriv κ' η a =ᵐ[η a] rnDeriv κ η a := by + intro a + rw [← rnDeriv_add_singularPart κ η] + filter_upwards [rnDeriv_add κ' (singularPart κ η) η a, + rnDeriv_singularPart κ η a] with x hx1 hx2 + rw [hx1, Pi.add_apply, hx2, Pi.zero_apply, add_zero] + exact ENNReal.ae_eq_compProd_of_forall_ae_eq _ _ (measurable_rnDeriv κ' η) + (measurable_rnDeriv κ η) this + +lemma rnDeriv_measure_compProd (μ ν : Measure α) (κ η : kernel α γ) + [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteKernel κ] [IsFiniteKernel η] : + ∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] fun p ↦ (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 := by + let μ' := ν.withDensity (∂μ/∂ν) + let κ' := withDensity η (rnDeriv κ η) + suffices ∂(μ' ⊗ₘ κ')/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] fun p ↦ (∂μ'/∂ν) p.1 * rnDeriv κ' η p.1 p.2 from + (todo1 μ ν κ η).symm.trans (this.trans (todo2 μ ν κ η)) refine rnDeriv_measure_compProd_aux ?_ ?_ · exact MeasureTheory.withDensity_absolutelyContinuous _ _ · exact ae_of_all _ (fun _ ↦ withDensity_absolutelyContinuous _ _) From 297ade43cf20ec47a91828ca98e6a1200000894b Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 20 Mar 2024 10:35:29 +0100 Subject: [PATCH 102/129] minor changes in CondCdf --- Mathlib/Probability/Kernel/Disintegration/CondCdf.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index f4eb79d464fc5..bdc0be09b1f0a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -173,7 +173,7 @@ theorem measurable_preCDF {ρ : Measure (α × ℝ)} {r : ℚ} : Measurable (pre #align probability_theory.measurable_pre_cdf ProbabilityTheory.measurable_preCDF lemma measurable_preCDF' {ρ : Measure (α × ℝ)} : - Measurable fun a r ↦ ENNReal.toReal (preCDF ρ r a) := by + Measurable fun a r ↦ (preCDF ρ r a).toReal := by rw [measurable_pi_iff] exact fun _ ↦ measurable_preCDF.ennreal_toReal @@ -216,7 +216,7 @@ theorem preCDF_le_one (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : exact Measure.IicSnd_le_fst ρ r s #align probability_theory.pre_cdf_le_one ProbabilityTheory.preCDF_le_one -theorem set_integral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) {s : Set α} (hs : MeasurableSet s) +lemma set_integral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) {s : Set α} (hs : MeasurableSet s) [IsFiniteMeasure ρ] : ∫ x in s, (preCDF ρ r x).toReal ∂ρ.fst = (ρ.IicSnd r s).toReal := by rw [integral_toReal] @@ -230,7 +230,7 @@ lemma integral_preCDF_fst (ρ : Measure (α × ℝ)) (r : ℚ) [IsFiniteMeasure ∫ x, (preCDF ρ r x).toReal ∂ρ.fst = (ρ.IicSnd r univ).toReal := by rw [← integral_univ, set_integral_preCDF_fst ρ _ MeasurableSet.univ] -theorem integrable_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℚ) : +lemma integrable_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : ℚ) : Integrable (fun a ↦ (preCDF ρ x a).toReal) ρ.fst := by refine integrable_of_forall_fin_meas_le _ (measure_lt_top ρ.fst univ) ?_ fun t _ _ ↦ ?_ · exact measurable_preCDF.ennreal_toReal.aestronglyMeasurable @@ -268,8 +268,8 @@ lemma isRatCondKernelCDFAux_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ENNReal.continuousAt_toReal (measure_ne_top _ _) exact h0.comp (h.comp hs_tendsto) integrable _ q := integrable_preCDF ρ q - set_integral a s hs q := by rw [kernel.const_apply, kernel.const_apply, set_integral_preCDF_fst _ _ hs, - Measure.IicSnd_apply _ _ hs] + set_integral a s hs q := by rw [kernel.const_apply, kernel.const_apply, + set_integral_preCDF_fst _ _ hs, Measure.IicSnd_apply _ _ hs] lemma isRatCondKernelCDF_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : IsRatCondKernelCDF (fun p r ↦ (preCDF ρ r p.2).toReal) From 8edbc010a63a41a2c0399783ec8a995885bb088d Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 27 Mar 2024 16:55:52 +0100 Subject: [PATCH 103/129] fix merge --- Mathlib/Probability/Kernel/Disintegration/Density.lean | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index bc34d8ea28f56..b34cc4335b51f 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -452,7 +452,10 @@ lemma density_ae_eq_limitProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] filter_upwards [tendsto_densityProcess_limitProcess hκν a hs] with t ht using ht.limsup_eq lemma tendsto_m_density (hκν : fst κ ≤ ν) (a : α) [IsFiniteKernel ν] + {s : Set β} (hs : MeasurableSet s) : + ∀ᵐ x ∂(ν a), Tendsto (fun n ↦ densityProcess κ ν n a x s) atTop (𝓝 (density κ ν a x s)) := by + filter_upwards [tendsto_densityProcess_limitProcess hκν a hs, density_ae_eq_limitProcess hκν a hs] with t h1 h2 using h2 ▸ h1 lemma measurable_density (κ : kernel α (γ × β)) (ν : kernel α γ) @@ -543,12 +546,14 @@ lemma integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] ∫ x, density κ ν a x s ∂(ν a) = (κ a (univ ×ˢ s)).toReal := by rw [← integral_univ, set_integral_density_of_measurableSet hκν 0 a hs MeasurableSet.univ] -lemma set_integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel κ] [IsFiniteKernel ν] +lemma set_integral_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet A) : ∫ x in A, density κ ν a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by + have : IsFiniteKernel κ := isFiniteKernel_of_isFiniteKernel_fst (h := isFiniteKernel_of_le hκν) have hA' : MeasurableSet[⨆ n, countableFiltration γ n] A := by rwa [iSup_countableFiltration] refine induction_on_inter (m := ⨆ n, countableFiltration γ n) (C := fun A ↦ ∫ x in A, density κ ν a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal) + (measurableSpace_iSup_eq (countableFiltration γ)) ?_ ?_ ?_ ?_ ?_ hA' · rintro s ⟨n, hs⟩ t ⟨m, ht⟩ _ exact ⟨max n m, ((countableFiltration γ).mono (le_max_left n m) _ hs).inter ((countableFiltration γ).mono (le_max_right n m) _ ht)⟩ From 6029651b6b80717defcc7245bb28b139ac74f621 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 27 Mar 2024 16:57:28 +0100 Subject: [PATCH 104/129] remove duplicate --- Mathlib/Probability/Kernel/WithDensity.lean | 15 --------------- 1 file changed, 15 deletions(-) diff --git a/Mathlib/Probability/Kernel/WithDensity.lean b/Mathlib/Probability/Kernel/WithDensity.lean index 474b5da1faeb1..6c694f321027d 100644 --- a/Mathlib/Probability/Kernel/WithDensity.lean +++ b/Mathlib/Probability/Kernel/WithDensity.lean @@ -186,21 +186,6 @@ theorem withDensity_tsum [Countable ι] (κ : kernel α β) [IsSFiniteKernel κ] rw [kernel.withDensity_apply' _ (hf n) a s] #align probability_theory.kernel.with_density_tsum ProbabilityTheory.kernel.withDensity_tsum -lemma withDensity_sub_add [IsSFiniteKernel κ] {f g : α → β → ℝ≥0∞} - (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) - (hg_int : ∀ a, ∫⁻ x, g a x ∂(κ a) ≠ ∞) (hfg : ∀ a, g a ≤ᵐ[κ a] f a) : - withDensity κ (fun a x ↦ f a x - g a x) + withDensity κ g = withDensity κ f := by - ext a s - simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add] - rw [kernel.withDensity_apply' _ hf, kernel.withDensity_apply' _ hg, kernel.withDensity_apply'] - swap; · exact hf.sub hg - rw [lintegral_sub] - · rw [tsub_add_cancel_iff_le] - exact lintegral_mono_ae (ae_restrict_of_ae (hfg a)) - · exact hg.comp measurable_prod_mk_left - · exact ((set_lintegral_le_lintegral _ _).trans_lt (hg_int a).lt_top).ne - · exact ae_restrict_of_ae (hfg a) - /-- If a kernel `κ` is finite and a function `f : α → β → ℝ≥0∞` is bounded, then `withDensity κ f` is finite. -/ theorem isFiniteKernel_withDensity_of_bounded (κ : kernel α β) [IsFiniteKernel κ] {B : ℝ≥0∞} From 54dd5ee92716c1026f988ba7aa43e2e99a35b14b Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Wed, 27 Mar 2024 18:33:09 +0100 Subject: [PATCH 105/129] fix --- Mathlib/Probability/Kernel/Disintegration/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 7f2780b44f7b8..f494e5168e24e 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -215,7 +215,7 @@ section CountablyGenerated open ProbabilityTheory.kernel lemma isRatCondKernelCDFAux_density_Iic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - isRatCondKernelCDFAux (fun (p : α × γ) q ↦ kernel.density κ (kernel.fst κ) p.1 p.2 (Set.Iic q)) + IsRatCondKernelCDFAux (fun (p : α × γ) q ↦ kernel.density κ (kernel.fst κ) p.1 p.2 (Set.Iic q)) κ (kernel.fst κ) where measurable := measurable_pi_iff.mpr fun _ ↦ measurable_density κ (kernel.fst κ) measurableSet_Iic mono' a q r hqr := From 99f0bc0878f78e0e7ac4adf27522b45a74fe345b Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 29 Mar 2024 08:19:40 +0100 Subject: [PATCH 106/129] remove unrelated material --- Mathlib/Probability/Kernel/RadonNikodym.lean | 360 ------------------- 1 file changed, 360 deletions(-) diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index 2c9c16a5c11ba..519d8c8a5d80f 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -436,364 +436,4 @@ lemma measurableSet_mutuallySingular (κ η : kernel α γ) [IsFiniteKernel κ] exact measurable_kernel_prod_mk_left (measurableSet_mutuallySingularSet κ η).compl (measurableSet_singleton 0) --- ok -lemma Measure.absolutelyContinuous_compProd_left {μ ν : Measure α} [SFinite μ] [SFinite ν] - (hμν : μ ≪ ν) (κ : kernel α γ) [IsSFiniteKernel κ] : - μ ⊗ₘ κ ≪ ν ⊗ₘ κ := by - refine Measure.AbsolutelyContinuous.mk fun s hs hs_zero ↦ ?_ - rw [Measure.compProd_apply hs, lintegral_eq_zero_iff (measurable_kernel_prod_mk_left hs)] - at hs_zero ⊢ - exact hμν.ae_eq hs_zero - --- ok -lemma Measure.absolutelyContinuous_compProd_right (μ : Measure α) {κ η : kernel α γ} - [SFinite μ] [IsSFiniteKernel κ] [IsSFiniteKernel η] - (hκη : ∀ᵐ a ∂μ, κ a ≪ η a) : - μ ⊗ₘ κ ≪ μ ⊗ₘ η := by - refine Measure.AbsolutelyContinuous.mk fun s hs hs_zero ↦ ?_ - rw [Measure.compProd_apply hs, lintegral_eq_zero_iff (measurable_kernel_prod_mk_left hs)] - at hs_zero ⊢ - filter_upwards [hs_zero, hκη] with a ha_zero ha_ac using ha_ac ha_zero - --- ok -lemma Measure.absolutelyContinuous_compProd {μ ν : Measure α} {κ η : kernel α γ} - [SFinite μ] [SFinite ν] [IsSFiniteKernel κ] [IsSFiniteKernel η] - (hμν : μ ≪ ν) (hκη : ∀ᵐ a ∂ν, κ a ≪ η a) : - μ ⊗ₘ κ ≪ ν ⊗ₘ η := - (Measure.absolutelyContinuous_compProd_left hμν κ).trans - (Measure.absolutelyContinuous_compProd_right ν hκη) - -lemma Measure.mutuallySingular_compProd_left {μ ν : Measure α} [SFinite μ] [SFinite ν] - (hμν : μ ⟂ₘ ν) (κ η : kernel α γ) [IsSFiniteKernel κ] [IsSFiniteKernel η] : - μ ⊗ₘ κ ⟂ₘ ν ⊗ₘ η := by - let s := hμν.nullSet - have hs_prod : MeasurableSet (s ×ˢ (univ : Set γ)) := - hμν.measurableSet_nullSet.prod MeasurableSet.univ - refine ⟨s ×ˢ univ, hs_prod, ?_⟩ - rw [Measure.compProd_apply_prod hμν.measurableSet_nullSet MeasurableSet.univ, compl_prod_eq_union] - simp only [Measure.MutuallySingular.restrict_nullSet, lintegral_zero_measure, compl_univ, - prod_empty, union_empty, true_and] - rw [Measure.compProd_apply_prod hμν.measurableSet_nullSet.compl MeasurableSet.univ] - simp - -lemma Measure.mutuallySingular_compProd_right (μ ν : Measure α) [SFinite μ] [SFinite ν] - {κ η : kernel α γ} [IsFiniteKernel κ] [IsFiniteKernel η] (hκη : ∀ a, κ a ⟂ₘ η a) : - μ ⊗ₘ κ ⟂ₘ ν ⊗ₘ η := by - let s := mutuallySingularSet κ η - have hs : MeasurableSet s := measurableSet_mutuallySingularSet κ η - symm - refine ⟨s, hs, ?_⟩ - rw [Measure.compProd_apply hs, Measure.compProd_apply hs.compl] - have h1 : ∀ a, η a (Prod.mk a ⁻¹' s) = 0 := by - intro a - have : Prod.mk a ⁻¹' s = mutuallySingularSetSlice κ η a := rfl - rw [this, measure_mutuallySingularSetSlice] - have h2 : ∀ a, κ a (Prod.mk a ⁻¹' s)ᶜ = 0 := by - intro a - have : (Prod.mk a ⁻¹' s)ᶜ ⊆ Prod.mk a ⁻¹' sᶜ := by intro; simp - refine measure_mono_null this ?_ - have : Prod.mk a ⁻¹' sᶜ = (mutuallySingularSetSlice κ η a)ᶜ := rfl - rw [this, ← withDensity_rnDeriv_eq_zero_iff_measure_eq_zero κ η a, - withDensity_rnDeriv_eq_zero_iff_mutuallySingular] - exact hκη a - simp [h1, h2] - -lemma Measure.fst_map_compProd (μ : Measure α) [SFinite μ] (κ : kernel α γ) - [IsMarkovKernel κ] : - (μ ⊗ₘ κ).map Prod.fst = μ := by - ext s hs - rw [Measure.map_apply measurable_fst hs, Measure.compProd_apply] - swap; · exact measurable_fst hs - have : ∀ a, (κ a) (Prod.mk a ⁻¹' (Prod.fst ⁻¹' s)) = s.indicator (fun a ↦ κ a univ) a := by - intro a - classical - rw [indicator_apply] - split_ifs with ha - · congr with x - simp [ha] - · suffices Prod.mk a ⁻¹' (Prod.fst ⁻¹' s) = ∅ by rw [this]; simp - ext x - simp [ha] - simp_rw [this] - rw [lintegral_indicator _ hs] - simp - -lemma ae_compProd_of_ae_fst {μ : Measure α} (κ : kernel α γ) - [SFinite μ] [IsSFiniteKernel κ] {p : α → Prop} (hp : MeasurableSet {x | p x}) - (h : ∀ᵐ a ∂μ, p a) : - ∀ᵐ x ∂(μ ⊗ₘ κ), p x.1 := by - refine ae_compProd_of_ae_ae (measurable_fst hp) ?_ - filter_upwards [h] with a ha - simp [ha] - -lemma ae_eq_compProd_of_ae_eq_fst {β : Type*} {_ : MeasurableSpace β} [AddGroup β] - [MeasurableSingletonClass β] [MeasurableSub₂ β] - (μ : Measure α) (κ : kernel α γ) - [SFinite μ] [IsSFiniteKernel κ] {f g : α → β} - (hf : Measurable f) (hg : Measurable g) (h : f =ᵐ[μ] g) : - (fun p ↦ f p.1) =ᵐ[μ ⊗ₘ κ] (fun p ↦ g p.1) := - ae_compProd_of_ae_fst κ (measurableSet_eq_fun hf hg) h - -lemma ae_eq_compProd_of_forall_ae_eq {β : Type*} {_ : MeasurableSpace β} [AddGroup β] - [MeasurableSingletonClass β] [MeasurableSub₂ β] - (μ : Measure α) (κ : kernel α γ) - [SFinite μ] [IsSFiniteKernel κ] {f g : α → γ → β} - (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) - (h : ∀ a, f a =ᵐ[κ a] g a) : - (fun p ↦ f p.1 p.2) =ᵐ[μ ⊗ₘ κ] (fun p ↦ g p.1 p.2) := - ae_compProd_of_ae_ae (measurableSet_eq_fun hf hg) - (ae_of_all _ (fun a ↦ measure_mono_null (fun x ↦ by simp) (h a))) - -lemma ENNReal.ae_eq_compProd_of_ae_eq_fst (μ : Measure α) (κ : kernel α γ) - [SFinite μ] [IsSFiniteKernel κ] {f g : α → ℝ≥0∞} - (hf : Measurable f) (hg : Measurable g) (h : f =ᵐ[μ] g) : - (fun p ↦ f p.1) =ᵐ[μ ⊗ₘ κ] (fun p ↦ g p.1) := - ae_compProd_of_ae_fst κ (measurableSet_eq_fun' hf hg) h - -lemma ENNReal.ae_eq_compProd_of_forall_ae_eq - (μ : Measure α) (κ : kernel α γ) - [SFinite μ] [IsSFiniteKernel κ] {f g : α → γ → ℝ≥0∞} - (hf : Measurable (Function.uncurry f)) (hg : Measurable (Function.uncurry g)) - (h : ∀ a, f a =ᵐ[κ a] g a) : - (fun p ↦ f p.1 p.2) =ᵐ[μ ⊗ₘ κ] (fun p ↦ g p.1 p.2) := - ae_compProd_of_ae_ae (measurableSet_eq_fun' hf hg) - (ae_of_all _ (fun a ↦ measure_mono_null (fun x ↦ by simp) (h a))) - -section Unique - -variable {ξ : kernel α γ} {f : α → γ → ℝ≥0∞} - -lemma eq_rnDeriv_measure [IsFiniteKernel ν] (h : κ = kernel.withDensity ν f + ξ) - (hf : Measurable (Function.uncurry f)) (hξ : ∀ a, ξ a ⟂ₘ ν a) (a : α) : - f a =ᵐ[ν a] ∂(κ a)/∂(ν a) := by - have : κ a = ξ a + (ν a).withDensity (f a) := by - rw [h, coeFn_add, Pi.add_apply, kernel.withDensity_apply _ hf, add_comm] - exact (κ a).eq_rnDeriv₀ (hf.comp measurable_prod_mk_left).aemeasurable (hξ a) this - -lemma eq_singularPart_measure [IsFiniteKernel κ] [IsFiniteKernel ν] - (h : κ = kernel.withDensity ν f + ξ) - (hf : Measurable (Function.uncurry f)) (hξ : ∀ a, ξ a ⟂ₘ ν a) (a : α) : - ξ a = (κ a).singularPart (ν a) := by - have : κ a = ξ a + (ν a).withDensity (f a) := by - rw [h, coeFn_add, Pi.add_apply, kernel.withDensity_apply _ hf, add_comm] - exact (κ a).eq_singularPart (hf.comp measurable_prod_mk_left) (hξ a) this - -lemma rnDeriv_eq_rnDeriv_measure (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : - rnDeriv κ ν a =ᵐ[ν a] ∂(κ a)/∂(ν a) := - eq_rnDeriv_measure (rnDeriv_add_singularPart κ ν).symm (measurable_rnDeriv κ ν) - (mutuallySingular_singularPart κ ν) a - -lemma singularPart_eq_singularPart_measure [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : - singularPart κ ν a = (κ a).singularPart (ν a) := - eq_singularPart_measure (rnDeriv_add_singularPart κ ν).symm (measurable_rnDeriv κ ν) - (mutuallySingular_singularPart κ ν) a - -lemma eq_rnDeriv [IsFiniteKernel κ] [IsFiniteKernel ν] (h : κ = kernel.withDensity ν f + ξ) - (hf : Measurable (Function.uncurry f)) (hξ : ∀ a, ξ a ⟂ₘ ν a) (a : α) : - f a =ᵐ[ν a] rnDeriv κ ν a := - (eq_rnDeriv_measure h hf hξ a).trans (rnDeriv_eq_rnDeriv_measure _ _ a).symm - -lemma eq_singularPart [IsFiniteKernel κ] [IsFiniteKernel ν] (h : κ = kernel.withDensity ν f + ξ) - (hf : Measurable (Function.uncurry f)) (hξ : ∀ a, ξ a ⟂ₘ ν a) (a : α) : - ξ a = singularPart κ ν a := - (eq_singularPart_measure h hf hξ a).trans (singularPart_eq_singularPart_measure a).symm - -end Unique - -section MeasureCompProd - -lemma set_lintegral_prod_rnDeriv {μ ν : Measure α} {κ η : kernel α γ} - [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteKernel κ] [IsFiniteKernel η] - (hμν : μ ≪ ν) (hκη : ∀ᵐ a ∂ν, κ a ≪ η a) - {s : Set α} (hs : MeasurableSet s) {t : Set γ} (ht : MeasurableSet t) : - ∫⁻ x in s, ∫⁻ y in t, (∂μ/∂ν) x * rnDeriv κ η x y ∂(η x) ∂ν = (μ ⊗ₘ κ) (s ×ˢ t) := by - have : ∀ᵐ x ∂ν, ∫⁻ y in t, (∂μ/∂ν) x * rnDeriv κ η x y ∂(η x) - = (∂μ/∂ν) x * κ x t := by - filter_upwards [hκη] with x hx - calc ∫⁻ y in t, (∂μ/∂ν) x * rnDeriv κ η x y ∂(η x) - = (∂μ/∂ν) x * ∫⁻ y in t, rnDeriv κ η x y ∂(η x) := by - rw [lintegral_const_mul] - exact measurable_rnDeriv_right κ η x - _ = (∂μ/∂ν) x * ∫⁻ y in t, (∂(κ x)/∂(η x)) y ∂(η x) := by - congr 1 - refine lintegral_congr_ae (ae_restrict_of_ae ?_) - exact rnDeriv_eq_rnDeriv_measure _ _ x - _ = (∂μ/∂ν) x * κ x t := by - congr - rw [Measure.set_lintegral_rnDeriv hx] - rw [lintegral_congr_ae (ae_restrict_of_ae this), - set_lintegral_rnDeriv_mul hμν (kernel.measurable_coe _ ht).aemeasurable hs, - Measure.compProd_apply_prod hs ht] - -lemma rnDeriv_measure_compProd_aux {μ ν : Measure α} {κ η : kernel α γ} - [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteKernel κ] [IsFiniteKernel η] - (hμν : μ ≪ ν) (hκη : ∀ᵐ a ∂ν, κ a ≪ η a) : - ∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] fun p ↦ (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 := by - have h_meas : Measurable fun p : α × γ ↦ (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 := - ((Measure.measurable_rnDeriv _ _).comp measurable_fst).mul (measurable_rnDeriv _ _) - suffices ∀ s, MeasurableSet s → ∫⁻ p in s, (∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η)) p ∂(ν ⊗ₘ η) - = ∫⁻ p in s, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂(ν ⊗ₘ η) from - ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite (Measure.measurable_rnDeriv _ _) h_meas - fun s hs _ ↦ this s hs - have h_left : ∀ s, MeasurableSet s → ∫⁻ p in s, (∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η)) p ∂(ν ⊗ₘ η) - = (μ ⊗ₘ κ) s := by - intro s _ - rw [Measure.set_lintegral_rnDeriv] - exact Measure.absolutelyContinuous_compProd hμν hκη - have : ∀ s t, MeasurableSet s → MeasurableSet t → - ∫⁻ x in s, ∫⁻ y in t, (∂μ/∂ν) x * rnDeriv κ η x y ∂(η x) ∂ν = (μ ⊗ₘ κ) (s ×ˢ t) := - fun _ _ hs ht ↦ set_lintegral_prod_rnDeriv hμν hκη hs ht - intro s hs - apply MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod _ _ _ _ hs - · simp - · rintro _ ⟨t₁, ht₁, t₂, ht₂, rfl⟩ - simp only [mem_setOf_eq] at ht₁ ht₂ - rw [h_left (t₁ ×ˢ t₂) (ht₁.prod ht₂), Measure.set_lintegral_compProd h_meas ht₁ ht₂, - this t₁ t₂ ht₁ ht₂] - · intro t ht ht_eq - calc ∫⁻ p in tᶜ, (∂μ ⊗ₘ κ/∂ν ⊗ₘ η) p ∂ν ⊗ₘ η - = ∫⁻ p, (∂μ ⊗ₘ κ/∂ν ⊗ₘ η) p ∂ν ⊗ₘ η - ∫⁻ p in t, (∂μ ⊗ₘ κ/∂ν ⊗ₘ η) p ∂ν ⊗ₘ η := by - refine (ENNReal.sub_eq_of_add_eq ?_ ?_).symm - · rw [h_left _ ht] - exact measure_ne_top _ _ - · rw [add_comm, lintegral_add_compl _ ht] - _ = ∫⁻ p, (∂μ ⊗ₘ κ/∂ν ⊗ₘ η) p ∂ν ⊗ₘ η - - ∫⁻ p in t, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂ν ⊗ₘ η := by rw [ht_eq] - _ = (μ ⊗ₘ κ) univ - - ∫⁻ p in t, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂ν ⊗ₘ η := by - have h := h_left univ MeasurableSet.univ - rw [set_lintegral_univ] at h - rw [h] - _ = ∫⁻ x, ∫⁻ y, (∂μ/∂ν) x * rnDeriv κ η x y ∂η x ∂ν - - ∫⁻ p in t, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂ν ⊗ₘ η := by - have h := this univ univ MeasurableSet.univ MeasurableSet.univ - simp only [Measure.restrict_univ, univ_prod_univ] at h - rw [← h] - _ = ∫⁻ p, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂ν ⊗ₘ η - - ∫⁻ p in t, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂ν ⊗ₘ η := by - congr - rw [Measure.lintegral_compProd h_meas] - _ = ∫⁻ p in tᶜ, (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 ∂ν ⊗ₘ η := by - refine ENNReal.sub_eq_of_add_eq ?_ ?_ - · rw [← ht_eq, h_left _ ht] - exact measure_ne_top _ _ - rw [add_comm, lintegral_add_compl _ ht] - · intro f' hf_disj hf_meas hf_eq - rw [lintegral_iUnion hf_meas hf_disj, lintegral_iUnion hf_meas hf_disj] - congr with i - exact hf_eq i - -instance instIsFiniteKernel_withDensity_rnDeriv [hκ : IsFiniteKernel κ] [IsFiniteKernel ν] : - IsFiniteKernel (withDensity ν (rnDeriv κ ν)) := by - constructor - refine ⟨hκ.bound, hκ.bound_lt_top, fun a ↦ ?_⟩ - rw [kernel.withDensity_apply', set_lintegral_univ] - swap; · exact measurable_rnDeriv κ ν - rw [lintegral_congr_ae (rnDeriv_eq_rnDeriv_measure _ _ a), ← set_lintegral_univ] - exact (Measure.set_lintegral_rnDeriv_le _).trans (measure_le_bound _ _ _) - -instance instIsFiniteKernel_singularPart [hκ : IsFiniteKernel κ] [IsFiniteKernel ν] : - IsFiniteKernel (singularPart κ ν) := by - constructor - refine ⟨hκ.bound, hκ.bound_lt_top, fun a ↦ ?_⟩ - have h : withDensity ν (rnDeriv κ ν) a univ + singularPart κ ν a univ = κ a univ := by - conv_rhs => rw [← rnDeriv_add_singularPart κ ν] - exact (self_le_add_left _ _).trans (h.le.trans (measure_le_bound _ _ _)) - -lemma rnDeriv_add (κ ν η : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] [IsFiniteKernel η] - (a : α) : - rnDeriv (κ + ν) η a =ᵐ[η a] rnDeriv κ η a + rnDeriv ν η a := by - filter_upwards [rnDeriv_eq_rnDeriv_measure (κ + ν) η a, rnDeriv_eq_rnDeriv_measure κ η a, - rnDeriv_eq_rnDeriv_measure ν η a, Measure.rnDeriv_add (κ a) (ν a) (η a)] with x h1 h2 h3 h4 - rw [h1, Pi.add_apply, h2, h3, coeFn_add, Pi.add_apply, h4, Pi.add_apply] - -lemma rnDeriv_singularPart (κ ν : kernel α γ) [IsFiniteKernel κ] [IsFiniteKernel ν] (a : α) : - rnDeriv (singularPart κ ν) ν a =ᵐ[ν a] 0 := by - filter_upwards [rnDeriv_eq_rnDeriv_measure (singularPart κ ν) ν a, - (Measure.rnDeriv_eq_zero _ _).mpr (mutuallySingular_singularPart κ ν a)] with x h1 h2 - rw [h1, h2] - -lemma todo1 (μ ν : Measure α) (κ η : kernel α γ) - [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteKernel κ] [IsFiniteKernel η] : - ∂(ν.withDensity (∂μ/∂ν) ⊗ₘ withDensity η (rnDeriv κ η))/∂(ν ⊗ₘ η) - =ᵐ[ν ⊗ₘ η] ∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η) := by - let μ' := ν.withDensity (∂μ/∂ν) - let κ' := withDensity η (rnDeriv κ η) - have h1 : μ = μ' + μ.singularPart ν := by - conv_lhs => rw [μ.haveLebesgueDecomposition_add ν, add_comm] - have h2 : μ ⊗ₘ κ = μ' ⊗ₘ κ' + (μ.singularPart ν) ⊗ₘ κ + μ' ⊗ₘ (singularPart κ η) := by - conv_lhs => rw [h1, Measure.compProd_add_left] - conv_lhs => conv in μ' ⊗ₘ κ => rw [← rnDeriv_add_singularPart κ η] - rw [Measure.compProd_add_right, add_assoc, add_comm (μ' ⊗ₘ singularPart κ η), ← add_assoc] - rw [h2] - have h_add := Measure.rnDeriv_add (μ' ⊗ₘ κ' + μ.singularPart ν ⊗ₘ κ) - (μ' ⊗ₘ (singularPart κ η)) (ν ⊗ₘ η) - have h_add' := Measure.rnDeriv_add (μ' ⊗ₘ κ') (μ.singularPart ν ⊗ₘ κ) (ν ⊗ₘ η) - have h01 : ∂(μ.singularPart ν ⊗ₘ κ)/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] 0 := by - rw [Measure.rnDeriv_eq_zero] - exact Measure.mutuallySingular_compProd_left (Measure.mutuallySingular_singularPart _ _) κ η - have h02 : ∂(μ' ⊗ₘ (singularPart κ η))/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] 0 := by - rw [Measure.rnDeriv_eq_zero] - exact Measure.mutuallySingular_compProd_right μ' ν (mutuallySingular_singularPart _ _) - filter_upwards [h_add, h_add', h01, h02] with a h_add h_add' h01 h02 - rw [h_add, Pi.add_apply, h_add', Pi.add_apply, h01, h02] - simp - -lemma todo2 (μ ν : Measure α) (κ η : kernel α γ) - [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteKernel κ] [IsFiniteKernel η] : - (fun p ↦ (∂(ν.withDensity (∂μ/∂ν))/∂ν) p.1 * rnDeriv (withDensity η (rnDeriv κ η)) η p.1 p.2) - =ᵐ[ν ⊗ₘ η] (fun p ↦ (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2) := by - let μ' := ν.withDensity (∂μ/∂ν) - let κ' := withDensity η (rnDeriv κ η) - refine EventuallyEq.mul ?_ ?_ - · have h := Measure.rnDeriv_withDensity ν (Measure.measurable_rnDeriv μ ν) - rw [EventuallyEq, ae_iff] at h ⊢ - exact ENNReal.ae_eq_compProd_of_ae_eq_fst ν η (Measure.measurable_rnDeriv μ' ν) - (Measure.measurable_rnDeriv μ ν) h - · have : ∀ a, rnDeriv κ' η a =ᵐ[η a] rnDeriv κ η a := by - intro a - rw [← rnDeriv_add_singularPart κ η] - filter_upwards [rnDeriv_add κ' (singularPart κ η) η a, - rnDeriv_singularPart κ η a] with x hx1 hx2 - rw [hx1, Pi.add_apply, hx2, Pi.zero_apply, add_zero] - exact ENNReal.ae_eq_compProd_of_forall_ae_eq _ _ (measurable_rnDeriv κ' η) - (measurable_rnDeriv κ η) this - -lemma rnDeriv_measure_compProd (μ ν : Measure α) (κ η : kernel α γ) - [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteKernel κ] [IsFiniteKernel η] : - ∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] fun p ↦ (∂μ/∂ν) p.1 * rnDeriv κ η p.1 p.2 := by - let μ' := ν.withDensity (∂μ/∂ν) - let κ' := withDensity η (rnDeriv κ η) - suffices ∂(μ' ⊗ₘ κ')/∂(ν ⊗ₘ η) =ᵐ[ν ⊗ₘ η] fun p ↦ (∂μ'/∂ν) p.1 * rnDeriv κ' η p.1 p.2 from - (todo1 μ ν κ η).symm.trans (this.trans (todo2 μ ν κ η)) - refine rnDeriv_measure_compProd_aux ?_ ?_ - · exact MeasureTheory.withDensity_absolutelyContinuous _ _ - · exact ae_of_all _ (fun _ ↦ withDensity_absolutelyContinuous _ _) - -lemma rnDeriv_self (κ : kernel α γ) [IsFiniteKernel κ] (a : α) : - rnDeriv κ κ a =ᵐ[κ a] 1 := - (rnDeriv_eq_rnDeriv_measure κ κ a).trans (Measure.rnDeriv_self _) - -lemma rnDeriv_measure_compProd_left (μ ν : Measure α) (κ : kernel α γ) - [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteKernel κ] : - ∂(μ ⊗ₘ κ)/∂(ν ⊗ₘ κ) =ᵐ[ν ⊗ₘ κ] fun p ↦ (∂μ/∂ν) p.1 := by - have h : (fun p ↦ rnDeriv κ κ p.1 p.2) =ᵐ[ν ⊗ₘ κ] (fun p ↦ (1 : α → γ → ℝ≥0∞) p.1 p.2) := - ENNReal.ae_eq_compProd_of_forall_ae_eq ν κ (measurable_rnDeriv _ _) measurable_const - (rnDeriv_self κ) - filter_upwards [rnDeriv_measure_compProd μ ν κ κ, h] with p hp1 hp2 - rw [hp1, hp2] - simp - -lemma rnDeriv_measure_compProd_right (μ : Measure α) (κ η : kernel α γ) - [IsFiniteMeasure μ] [IsFiniteKernel κ] [IsFiniteKernel η] : - ∂(μ ⊗ₘ κ)/∂(μ ⊗ₘ η) =ᵐ[μ ⊗ₘ η] fun p ↦ rnDeriv κ η p.1 p.2 := by - have h : (fun p ↦ (∂μ/∂μ) p.1) =ᵐ[μ ⊗ₘ η] (fun p ↦ (1 : α → ℝ≥0∞) p.1) := - ENNReal.ae_eq_compProd_of_ae_eq_fst μ η (Measure.measurable_rnDeriv _ _) - measurable_const (Measure.rnDeriv_self _) - filter_upwards [rnDeriv_measure_compProd μ μ κ η, h] with p hp1 hp2 - rw [hp1, hp2] - simp - -end MeasureCompProd - end ProbabilityTheory.kernel From cdaeb0ae7aad79aedea7551af9733c5288b527c3 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 29 Mar 2024 08:20:31 +0100 Subject: [PATCH 107/129] remove import --- Mathlib/Probability/Kernel/RadonNikodym.lean | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index 519d8c8a5d80f..9806bcd609578 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -5,7 +5,6 @@ Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.WithDensity -import Mathlib.Probability.Kernel.MeasureCompProd /-! # Radon-Nikodym derivative and Lebesgue decomposition for kernels @@ -60,6 +59,12 @@ everywhere (although `rnDeriv κ η` has that value almost everywhere). See the Theorem 1.28 in [O. Kallenberg, Random Measures, Theory and Applications][kallenberg2017]. +## TODO + +* prove uniqueness results. +* link kernel Radon-Nikodym derivative and Radon-Nikodym derivative of measures, and similarly for + singular parts. + -/ open MeasureTheory Set Filter From d264e3ab202661a1c820c70fedf4a05e69b494eb Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 29 Mar 2024 08:24:09 +0100 Subject: [PATCH 108/129] rename embeddingReal --- .../MeasureTheory/Constructions/Polish.lean | 7 ++- .../Kernel/Disintegration/Basic.lean | 44 +++++++++---------- 2 files changed, 25 insertions(+), 26 deletions(-) diff --git a/Mathlib/MeasureTheory/Constructions/Polish.lean b/Mathlib/MeasureTheory/Constructions/Polish.lean index c1e5d4d9cb332..b10d46e2bc098 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish.lean @@ -1083,12 +1083,11 @@ theorem exists_measurableEmbedding_real : ∃ f : α → ℝ, MeasurableEmbeddin /-- A measurable embedding of a standard Borel space into `ℝ`. -/ noncomputable -def measurableEmbedding_real (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : Ω → ℝ := +def embeddingReal (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : Ω → ℝ := (exists_measurableEmbedding_real Ω).choose -lemma measurableEmbedding_measurableEmbedding_real - (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : - MeasurableEmbedding (measurableEmbedding_real Ω) := +lemma measurableEmbedding_embeddingReal (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : + MeasurableEmbedding (embeddingReal Ω) := (exists_measurableEmbedding_real Ω).choose_spec end MeasureTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index f494e5168e24e..6a4e553a7b3bf 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -104,13 +104,13 @@ On `ℝ`, we get disintegration by constructing a map `f` with the property `IsK noncomputable def condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f - (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) - (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)) - (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) - (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)))) : + (kernel.map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) + (kernel.fst (kernel.map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)))) : kernel (α × β) Ω := - let e := measurableEmbedding_real Ω - let he := measurableEmbedding_measurableEmbedding_real Ω + let e := embeddingReal Ω + have he := measurableEmbedding_embeddingReal Ω let x₀ := (range_nonempty e).choose kernel.comapRight (kernel.piecewise (measurableSet_toKernel_eq_one hf he.measurableSet_range) @@ -120,10 +120,10 @@ def condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFu instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f - (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) - (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)) - (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) - (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)))) : + (kernel.map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) + (kernel.fst (kernel.map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)))) : IsMarkovKernel (condKernelBorelSnd κ hf) := by rw [condKernelBorelSnd] refine kernel.IsMarkovKernel.comapRight _ _ fun a ↦ ?_ @@ -132,20 +132,20 @@ instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω)) · exact h_mem · classical rw [kernel.deterministic_apply' _ _ - (measurableEmbedding_measurableEmbedding_real Ω).measurableSet_range, + (measurableEmbedding_embeddingReal Ω).measurableSet_range, Set.indicator_apply, if_pos] - exact (range_nonempty (measurableEmbedding_real Ω)).choose_spec + exact (range_nonempty (embeddingReal Ω)).choose_spec lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω)) [IsFiniteKernel κ] {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f - (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) - (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)) - (kernel.fst (kernel.map κ (Prod.map (id : β → β) (measurableEmbedding_real Ω)) - (measurable_id.prod_map (measurableEmbedding_measurableEmbedding_real Ω).measurable)))) : + (kernel.map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) + (kernel.fst (kernel.map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)))) : kernel.fst κ ⊗ₖ condKernelBorelSnd κ hf = κ := by - let e := measurableEmbedding_real Ω - let he := measurableEmbedding_measurableEmbedding_real Ω + let e := embeddingReal Ω + let he := measurableEmbedding_embeddingReal Ω let κ' := kernel.map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable) have h_prod_embed : MeasurableEmbedding (Prod.map (id : β → β) e) := MeasurableEmbedding.id.prod_mk he @@ -266,8 +266,8 @@ lemma isCondKernelCDF_condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel noncomputable def condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : kernel (α × γ) Ω := - let f := measurableEmbedding_real Ω - let hf := measurableEmbedding_measurableEmbedding_real Ω + let f := embeddingReal Ω + let hf := measurableEmbedding_embeddingReal Ω let κ' := kernel.map κ (Prod.map (id : γ → γ) f) (measurable_id.prod_map hf.measurable) condKernelBorelSnd κ (isCondKernelCDF_condKernelCDF κ') @@ -286,8 +286,8 @@ section Unit noncomputable def condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : kernel (Unit × α) Ω := - let f := measurableEmbedding_real Ω - let hf := measurableEmbedding_measurableEmbedding_real Ω + let f := embeddingReal Ω + let hf := measurableEmbedding_embeddingReal Ω let κ' := kernel.map κ (Prod.map (id : α → α) f) (measurable_id.prod_map hf.measurable) condKernelBorelSnd κ (isCondKernelCDF_condCDF (κ' ())) From 635060845ad14ce8f95a24e77a3b241cfc12fb0d Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 29 Mar 2024 08:26:29 +0100 Subject: [PATCH 109/129] make condKernel irreducible --- Mathlib/Probability/Kernel/Disintegration/Basic.lean | 6 +++--- .../Probability/Kernel/Disintegration/Unique.lean | 12 ++++++------ 2 files changed, 9 insertions(+), 9 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 6a4e553a7b3bf..364a49900be7a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -424,7 +424,7 @@ instance instCountableOrCountablyGenerated_of_standardBorelSpace open Classical in noncomputable -def kernel.condKernel [h : CountableOrCountablyGenerated α β] +irreducible_def kernel.condKernel [h : CountableOrCountablyGenerated α β] (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : kernel (α × β) Ω' := if hα : Countable α ∧ MeasurableSingletonClass α then @@ -435,13 +435,13 @@ def kernel.condKernel [h : CountableOrCountablyGenerated α β] instance kernel.instIsMarkovKernel_condKernel [CountableOrCountablyGenerated α β] (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : IsMarkovKernel (kernel.condKernel κ) := by - unfold kernel.condKernel + rw [kernel.condKernel_def] split_ifs <;> infer_instance lemma kernel.compProd_fst_condKernel [hαβ : CountableOrCountablyGenerated α β] (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : kernel.fst κ ⊗ₖ kernel.condKernel κ = κ := by - unfold kernel.condKernel + rw [kernel.condKernel_def] split_ifs with h · have := h.1 have := h.2 diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index 16d8c718a3f74..7085af4fad8c0 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -77,10 +77,10 @@ disintegration kernel. -/ theorem eq_condKernel_of_kernel_eq_compProd (κ : kernel (α × β) Ω) [IsFiniteKernel κ] (hκ : ρ = kernel.fst ρ ⊗ₖ κ) (a : α) : ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) = kernel.condKernel ρ (a, x) := by --- The idea is to transporting the question to `ℝ` from `Ω` using `exists_measurableEmbedding_real` +-- The idea is to transporting the question to `ℝ` from `Ω` using `exists_embeddingReal` -- and then constructing a measure on `α × ℝ` using the obtained measurable embedding - let f := measurableEmbedding_real Ω - let hf := measurableEmbedding_measurableEmbedding_real Ω + let f := embeddingReal Ω + let hf := measurableEmbedding_embeddingReal Ω set ρ' : kernel α (β × ℝ) := kernel.map ρ (Prod.map id f) (measurable_id.prod_map hf.measurable) with hρ'def have hρ' : kernel.fst ρ' = kernel.fst ρ := by @@ -176,10 +176,10 @@ theorem Measure.eq_condKernel_of_measure_eq_compProd (κ : kernel α Ω) [IsFini (hκ : ρ = ρ.fst ⊗ₘ κ) : ∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by -- The idea is to transporting the question to `ℝ` from `Ω` using - -- `exists_measurableEmbedding_real` and then constructing a measure on `α × ℝ` using + -- `exists_embeddingReal` and then constructing a measure on `α × ℝ` using -- the obtained measurable embedding - let f := measurableEmbedding_real Ω - let hf := measurableEmbedding_measurableEmbedding_real Ω + let f := embeddingReal Ω + let hf := measurableEmbedding_embeddingReal Ω set ρ' : Measure (α × ℝ) := ρ.map (Prod.map id f) with hρ'def have hρ' : ρ'.fst = ρ.fst := by ext s hs From cea55478b06e9c1a97bc81d3ea8dc8cb8501a81a Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 29 Mar 2024 08:27:28 +0100 Subject: [PATCH 110/129] don't use Omega' --- Mathlib/Probability/Kernel/Disintegration/Basic.lean | 11 +++++------ 1 file changed, 5 insertions(+), 6 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 364a49900be7a..486175b034dab 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -87,10 +87,9 @@ open scoped ENNReal MeasureTheory Topology ProbabilityTheory namespace ProbabilityTheory -variable {α β γ Ω Ω': Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} +variable {α β γ Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} [MeasurableSpace.CountablyGenerated γ] [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] - [MeasurableSpace Ω'] [StandardBorelSpace Ω'] [Nonempty Ω'] section BorelSnd @@ -425,21 +424,21 @@ instance instCountableOrCountablyGenerated_of_standardBorelSpace open Classical in noncomputable irreducible_def kernel.condKernel [h : CountableOrCountablyGenerated α β] - (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : - kernel (α × β) Ω' := + (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : + kernel (α × β) Ω := if hα : Countable α ∧ MeasurableSingletonClass α then letI := hα.1; letI := hα.2; condKernelCountable κ else letI := h.countableOrCountablyGenerated.resolve_left hα; condKernelBorel κ instance kernel.instIsMarkovKernel_condKernel [CountableOrCountablyGenerated α β] - (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : + (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : IsMarkovKernel (kernel.condKernel κ) := by rw [kernel.condKernel_def] split_ifs <;> infer_instance lemma kernel.compProd_fst_condKernel [hαβ : CountableOrCountablyGenerated α β] - (κ : kernel α (β × Ω')) [IsFiniteKernel κ] : + (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : kernel.fst κ ⊗ₖ kernel.condKernel κ = κ := by rw [kernel.condKernel_def] split_ifs with h From e2dc65a3b1681f9bd95638dbcbd5cd07c17fbfb1 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 29 Mar 2024 08:50:05 +0100 Subject: [PATCH 111/129] move CountableOrCountablyGenerated --- .../MeasurableSpace/CountablyGenerated.lean | 20 +++++++++++++++++++ .../Kernel/Disintegration/Basic.lean | 14 +------------ .../Kernel/Disintegration/Integral.lean | 2 +- .../Kernel/Disintegration/Unique.lean | 2 +- 4 files changed, 23 insertions(+), 15 deletions(-) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index f5d0dca71cfe3..4c5e20deb0c75 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -459,4 +459,24 @@ lemma measurableSet_countablePartitionSet (n : ℕ) (a : α) : MeasurableSet (countablePartitionSet n a) := measurableSet_countablePartition n (countablePartitionSet_mem n a) +section CountableOrCountablyGenerated + +variable [MeasurableSpace β] + +/-- A class registering that either `α` is countable with measurable singletons or `β` is a +countably generated measurable space. -/ +class CountableOrCountablyGenerated (α β : Type*) [MeasurableSpace α] [MeasurableSpace β] : Prop := + (countableOrCountablyGenerated : + (Countable α ∧ MeasurableSingletonClass α) ∨ MeasurableSpace.CountablyGenerated β) + +instance instCountableOrCountablyGeneratedOfCountable + [h1 : Countable α] [h2 : MeasurableSingletonClass α] : + CountableOrCountablyGenerated α β := ⟨Or.inl ⟨h1, h2⟩⟩ + +instance instCountableOrCountablyGeneratedOfCountablyGenerated + [h : MeasurableSpace.CountablyGenerated β] : + CountableOrCountablyGenerated α β := ⟨Or.inr h⟩ + +end CountableOrCountablyGenerated + end MeasurableSpace diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 486175b034dab..7e87c97bcc54a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -63,7 +63,7 @@ file `Unique.lean`. #align_import probability.kernel.disintegration from "leanprover-community/mathlib"@"6315581f5650ffa2fbdbbbedc41243c8d7070981" -open MeasureTheory Set Filter +open MeasureTheory Set Filter MeasurableSpace open scoped ENNReal MeasureTheory Topology ProbabilityTheory @@ -409,18 +409,6 @@ end Countable section CountableOrCountablyGenerated -class CountableOrCountablyGenerated (α β : Type*) [MeasurableSpace α] [MeasurableSpace β] : Prop := - (countableOrCountablyGenerated : - (Countable α ∧ MeasurableSingletonClass α) ∨ MeasurableSpace.CountablyGenerated β) - -instance instCountableOrCountablyGenerated_of_countable - [h1 : Countable α] [h2 : MeasurableSingletonClass α] : - CountableOrCountablyGenerated α β := ⟨Or.inl ⟨h1, h2⟩⟩ - -instance instCountableOrCountablyGenerated_of_standardBorelSpace - [h : MeasurableSpace.CountablyGenerated β] : - CountableOrCountablyGenerated α β := ⟨Or.inr h⟩ - open Classical in noncomputable irreducible_def kernel.condKernel [h : CountableOrCountablyGenerated α β] diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index 75e01d37399be..81eef60f18cfc 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -23,7 +23,7 @@ provided. -/ -open MeasureTheory ProbabilityTheory +open MeasureTheory ProbabilityTheory MeasurableSpace open scoped ENNReal diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index 7085af4fad8c0..3b99e514257fe 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -15,7 +15,7 @@ import Mathlib.Probability.Kernel.Disintegration.Integral -/ -open MeasureTheory Set Filter +open MeasureTheory Set Filter MeasurableSpace open scoped ENNReal MeasureTheory Topology ProbabilityTheory From 1529974b2d8839c8af9e73d6c8f4d0ea081252a4 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Fri, 29 Mar 2024 08:51:16 +0100 Subject: [PATCH 112/129] undo changes to condCdf --- .../Kernel/Disintegration/CondCdf.lean | 47 +++++++++++++++++-- 1 file changed, 44 insertions(+), 3 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index 3aadea1fad010..ddc2ba44f06ea 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -383,8 +383,49 @@ theorem integral_condCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] (x : (isCondKernelCDF_condCDF ρ).integral () x #align probability_theory.integral_cond_cdf ProbabilityTheory.integral_condCDF -#noalign probability_theory.measure_cond_cdf_Iic -#noalign probability_theory.measure_cond_cdf_univ -#noalign probability_theory.measurable_measure_cond_cdf +section Measure + +theorem measure_condCDF_Iic (ρ : Measure (α × ℝ)) (a : α) (x : ℝ) : + (condCDF ρ a).measure (Iic x) = ENNReal.ofReal (condCDF ρ a x) := by + rw [← sub_zero (condCDF ρ a x)] + exact (condCDF ρ a).measure_Iic (tendsto_condCDF_atBot ρ a) _ +#align probability_theory.measure_cond_cdf_Iic ProbabilityTheory.measure_condCDF_Iic + +theorem measure_condCDF_univ (ρ : Measure (α × ℝ)) (a : α) : (condCDF ρ a).measure univ = 1 := by + rw [← ENNReal.ofReal_one, ← sub_zero (1 : ℝ)] + exact StieltjesFunction.measure_univ _ (tendsto_condCDF_atBot ρ a) (tendsto_condCDF_atTop ρ a) +#align probability_theory.measure_cond_cdf_univ ProbabilityTheory.measure_condCDF_univ + +instance instIsProbabilityMeasureCondCDF (ρ : Measure (α × ℝ)) (a : α) : + IsProbabilityMeasure (condCDF ρ a).measure := + ⟨measure_condCDF_univ ρ a⟩ + +/-- The function `a ↦ (condCDF ρ a).measure` is measurable. -/ +theorem measurable_measure_condCDF (ρ : Measure (α × ℝ)) : + Measurable fun a => (condCDF ρ a).measure := by + rw [Measure.measurable_measure] + refine' fun s hs => ?_ + -- Porting note: supplied `C` + refine' MeasurableSpace.induction_on_inter + (C := fun s => Measurable fun b ↦ StieltjesFunction.measure (condCDF ρ b) s) + (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic _ _ _ _ hs + · simp only [measure_empty, measurable_const] + · rintro S ⟨u, rfl⟩ + simp_rw [measure_condCDF_Iic ρ _ u] + exact (measurable_condCDF ρ u).ennreal_ofReal + · intro t ht ht_cd_meas + have : + (fun a => (condCDF ρ a).measure tᶜ) = + (fun a => (condCDF ρ a).measure univ) - fun a => (condCDF ρ a).measure t := by + ext1 a + rw [measure_compl ht (measure_ne_top (condCDF ρ a).measure _), Pi.sub_apply] + simp_rw [this, measure_condCDF_univ ρ] + exact Measurable.sub measurable_const ht_cd_meas + · intro f hf_disj hf_meas hf_cd_meas + simp_rw [measure_iUnion hf_disj hf_meas] + exact Measurable.ennreal_tsum hf_cd_meas +#align probability_theory.measurable_measure_cond_cdf ProbabilityTheory.measurable_measure_condCDF + +end Measure end ProbabilityTheory From a0899530d137b419e38b81c48dac3f9dc4dce552 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 31 Mar 2024 19:39:17 +0200 Subject: [PATCH 113/129] add docstrings to Basic --- .../Kernel/Disintegration/Basic.lean | 194 ++++++++++-------- 1 file changed, 105 insertions(+), 89 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 7e87c97bcc54a..b9e359cabade8 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -12,9 +12,9 @@ import Mathlib.Probability.Kernel.Disintegration.CdfToKernel # Disintegration of kernels and measures Let `κ : kernel α (β × Ω)` be a finite kernel, where `Ω` is a standard Borel space. Then if `α` is -countable or `β` has a countably generated σ-algebra (for example if it is standard Borel), -there exists a `kernel (α × β) Ω`, called conditional kernel and denoted by `kernel.condKernel κ` -such that `κ = kernel.fst κ ⊗ₖ kernel.condKernel κ`. +countable or `β` has a countably generated σ-algebra (for example if it is standard Borel), then +there exists a `kernel (α × β) Ω` called conditional kernel and denoted by `condKernel κ` such that +`κ = fst κ ⊗ₖ condKernel κ`. We also define a conditional kernel for a measure `ρ : Measure (β × Ω)`, where `Ω` is a standard Borel space. This is a `kernel β Ω` denoted by `ρ.condKernel` such that `ρ = ρ.fst ⊗ₘ ρ.condKernel`. @@ -24,11 +24,11 @@ measurably into `ℝ`: it then suffices to define a suitable kernel for `Ω = For `κ : kernel α (β × ℝ)`, the construction of the conditional kernel proceeds as follows: * Build a measurable function `f : (α × β) → ℚ → ℝ` such that for all measurable sets `s` and all `q : ℚ`, `∫ x in s, f (a, x) q ∂(kernel.fst κ a) = (κ a (s ×ˢ Iic (q : ℝ))).toReal`. - We restrict to `Q` here to be able to prove the measurability. + We restrict to `ℚ` here to be able to prove the measurability. * Extend that function to `(α × β) → StieltjesFunction`. See the file `MeasurableStieltjes.lean`. * Finally obtain from the measurable Stieltjes function a measure on `ℝ` for each element of `α × β` in a measurable way: we have obtained a `kernel (α × β) ℝ`. - See the file `KernelCdf.lean` for that step. + See the file `CdfToKernel.lean` for that step. The first step (building the measurable function on `ℚ`) is done differently depending on whether `α` is countable or not. @@ -39,15 +39,15 @@ The first step (building the measurable function on `ℚ`) is done differently d * If `α` is not countable, we can't proceed separately for each `a : α` and have to build a function `f : α × β → ℚ → ℝ` which is measurable on the product. We are able to do so if `β` has a countably generated σ-algebra (this is the case in particular for standard Borel spaces). - See the files `Density.lean` and `CondKernelCdf.lean`. + See the file `Density.lean`. -We define a class `CountableOrCountablyGenerated α β` which encodes the property -`(Countable α ∧ MeasurableSingletonClass α) ∨ CountablyGenerated β`. The conditional kernel is -defined under that assumption. +The conditional kernel is defined under the typeclass assumption +`CountableOrCountablyGenerated α β`, which encodes the property +`(Countable α ∧ MeasurableSingletonClass α) ∨ CountablyGenerated β`. -Properties of integrals involving `kernel.condKernel` are collated in the file `Integral.lean`. -The conditional kernel is unique (almost everywhere w.r.t. `kernel.fst κ`): this is proved in the -file `Unique.lean`. +Properties of integrals involving `condKernel` are collated in the file `Integral.lean`. +The conditional kernel is unique (almost everywhere w.r.t. `fst κ`): this is proved in the file +`Unique.lean`. ## Main definitions @@ -56,7 +56,7 @@ file `Unique.lean`. ## Main statements -* `ProbabilityTheory.kernel.compProd_fst_condKernel`: `kernel.fst κ ⊗ₖ kernel.condKernel κ = κ` +* `ProbabilityTheory.kernel.compProd_fst_condKernel`: `fst κ ⊗ₖ condKernel κ = κ` * `MeasureTheory.Measure.compProd_fst_condKernel`: `ρ.fst ⊗ₘ ρ.condKernel = ρ` -/ @@ -67,7 +67,6 @@ open MeasureTheory Set Filter MeasurableSpace open scoped ENNReal MeasureTheory Topology ProbabilityTheory - #noalign probability_theory.cond_kernel_real #noalign probability_theory.cond_kernel_real_Iic #noalign probability_theory.set_lintegral_cond_kernel_real_Iic @@ -84,8 +83,7 @@ open scoped ENNReal MeasureTheory Topology ProbabilityTheory #noalign probability_theory.kernel.const_unit_eq_comp_prod #noalign probability_theory.kernel.const_eq_comp_prod - -namespace ProbabilityTheory +namespace ProbabilityTheory.kernel variable {α β γ Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} [MeasurableSpace.CountablyGenerated γ] @@ -93,72 +91,75 @@ variable {α β γ Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace section BorelSnd -/-! ### Disintegration of kenrels on standard Borel spaces +/-! ### Disintegration of kernels on standard Borel spaces Since every standard Borel space embeds measurably into `ℝ`, we can generalize a disintegration -property on `ℝ` to all these spaces. +property on `ℝ` to all these spaces. -/ -On `ℝ`, we get disintegration by constructing a map `f` with the property `IsKernelCDF f`. -/ +/-- Auxiliary definition for `ProbabilityTheory.kernel.condKernel`. +For `κ` a kernel from `α` to `β × Ω` where `Ω` is standard borel, and `f` a conditional kernel +CDF of the map of `κ` by an embedding of `Ω` into `ℝ` (a kernel from `α` to `β × ℝ`), we build a +`kernel (α × β) Ω`. +`f` gives a Markov kernel from `α × β` to `ℝ`: we can't take directly the comap of that kernel by +the embedding `embeddingReal Ω`, because it would not be a Markov kernel (its value is a probability +measure only almost everywhere). We thus take the comap of a slight modification. -/ noncomputable def condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f - (kernel.map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (map κ (Prod.map (id : β → β) (embeddingReal Ω)) (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) - (kernel.fst (kernel.map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)))) : kernel (α × β) Ω := - let e := embeddingReal Ω have he := measurableEmbedding_embeddingReal Ω - let x₀ := (range_nonempty e).choose - kernel.comapRight - (kernel.piecewise (measurableSet_toKernel_eq_one hf he.measurableSet_range) - (hf.toKernel f) (kernel.deterministic (fun _ ↦ x₀) measurable_const)) - he + let x₀ := (range_nonempty (embeddingReal Ω)).choose + comapRight (piecewise (measurableSet_toKernel_eq_one hf he.measurableSet_range) + (hf.toKernel f) (deterministic (fun _ ↦ x₀) measurable_const)) he +/-- `condKernelBorelSnd` is a Markov kernel. -/ instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f - (kernel.map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (map κ (Prod.map (id : β → β) (embeddingReal Ω)) (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) - (kernel.fst (kernel.map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)))) : IsMarkovKernel (condKernelBorelSnd κ hf) := by rw [condKernelBorelSnd] - refine kernel.IsMarkovKernel.comapRight _ _ fun a ↦ ?_ - rw [kernel.piecewise_apply'] + refine IsMarkovKernel.comapRight _ _ fun a ↦ ?_ + rw [piecewise_apply'] split_ifs with h_mem · exact h_mem · classical - rw [kernel.deterministic_apply' _ _ - (measurableEmbedding_embeddingReal Ω).measurableSet_range, - Set.indicator_apply, if_pos] + rw [deterministic_apply' _ _ (measurableEmbedding_embeddingReal Ω).measurableSet_range, + indicator_apply, if_pos] exact (range_nonempty (embeddingReal Ω)).choose_spec +/-- `fst κ ⊗ₖ condKernelBorelSnd κ hf = κ`. -/ lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω)) [IsFiniteKernel κ] {f : α × β → StieltjesFunction} (hf : IsCondKernelCDF f - (kernel.map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (map κ (Prod.map (id : β → β) (embeddingReal Ω)) (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) - (kernel.fst (kernel.map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)))) : - kernel.fst κ ⊗ₖ condKernelBorelSnd κ hf = κ := by + fst κ ⊗ₖ condKernelBorelSnd κ hf = κ := by let e := embeddingReal Ω let he := measurableEmbedding_embeddingReal Ω - let κ' := kernel.map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable) + let κ' := map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable) have h_prod_embed : MeasurableEmbedding (Prod.map (id : β → β) e) := MeasurableEmbedding.id.prod_mk he - have h_fst : kernel.fst κ' = kernel.fst κ := by + have h_fst : fst κ' = fst κ := by ext a u hu unfold_let κ' - rw [kernel.fst_apply' _ _ hu, kernel.fst_apply' _ _ hu, - kernel.map_apply' κ h_prod_embed.measurable] + rw [fst_apply' _ _ hu, fst_apply' _ _ hu, map_apply' κ h_prod_embed.measurable] · rfl · exact measurable_fst hu - have : κ = kernel.comapRight κ' h_prod_embed := by + have : κ = comapRight κ' h_prod_embed := by ext c t ht : 2 unfold_let κ' - rw [kernel.comapRight_apply' _ _ _ ht, kernel.map_apply' κ h_prod_embed.measurable + rw [comapRight_apply' _ _ _ ht, map_apply' κ h_prod_embed.measurable _ (h_prod_embed.measurableSet_image.mpr ht)] congr with x : 1 rw [← @Prod.mk.eta _ _ x] @@ -169,20 +170,18 @@ lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω)) [IsFiniteKerne conv_rhs => rw [this] unfold_let κ' conv_rhs => rw [← compProd_toKernel hf] - change kernel.fst κ ⊗ₖ condKernelBorelSnd κ hf - = kernel.comapRight (kernel.fst κ' ⊗ₖ hf.toKernel f) h_prod_embed + change fst κ ⊗ₖ condKernelBorelSnd κ hf = comapRight (fst κ' ⊗ₖ hf.toKernel f) h_prod_embed ext c t ht : 2 - rw [kernel.comapRight_apply' _ _ _ ht, - kernel.compProd_apply _ _ _ (h_prod_embed.measurableSet_image.mpr ht)] - simp_rw [h_fst, kernel.compProd_apply _ _ _ ht] + rw [comapRight_apply' _ _ _ ht, compProd_apply _ _ _ (h_prod_embed.measurableSet_image.mpr ht)] + simp_rw [h_fst, compProd_apply _ _ _ ht] refine lintegral_congr_ae ?_ let ρ_set := {p : α × β | hf.toKernel f p (range e) = 1} - have h_ae : ∀ a, ∀ᵐ t ∂(kernel.fst κ a), (a, t) ∈ ρ_set := by + have h_ae : ∀ a, ∀ᵐ t ∂(fst κ a), (a, t) ∈ ρ_set := by intro a rw [← h_fst] refine ae_toKernel_eq_one hf a he.measurableSet_range ?_ simp only [mem_compl_iff, mem_range, not_exists] - rw [kernel.map_apply'] + rw [map_apply'] · have h_empty : {a : β × Ω | ∀ (x : Ω), ¬e x = e a.2} = ∅ := by ext x simp only [mem_setOf_eq, mem_empty_iff_false, iff_false, not_forall, not_not, @@ -194,7 +193,7 @@ lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω)) [IsFiniteKerne rw [this] exact MeasurableSet.univ.prod he.measurableSet_range.compl filter_upwards [h_ae c] with a ha - rw [condKernelBorelSnd, kernel.comapRight_apply'] + rw [condKernelBorelSnd, comapRight_apply'] swap; · exact measurable_prod_mk_left ht have h1 : {c : ℝ | (a, c) ∈ Prod.map id e '' t} = e '' {c : Ω | (a, c) ∈ t} := by ext1 x @@ -204,7 +203,7 @@ lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω)) [IsFiniteKerne exact ⟨b, h_mem, hf_eq⟩ · rintro ⟨b, h_mem, hf_eq⟩ exact ⟨a, b, h_mem, rfl, hf_eq⟩ - rw [h1, kernel.piecewise_apply, if_pos ha] + rw [h1, piecewise_apply, if_pos ha] rfl end BorelSnd @@ -214,9 +213,9 @@ section CountablyGenerated open ProbabilityTheory.kernel lemma isRatCondKernelCDFAux_density_Iic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsRatCondKernelCDFAux (fun (p : α × γ) q ↦ kernel.density κ (kernel.fst κ) p.1 p.2 (Set.Iic q)) - κ (kernel.fst κ) where - measurable := measurable_pi_iff.mpr fun _ ↦ measurable_density κ (kernel.fst κ) measurableSet_Iic + IsRatCondKernelCDFAux (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) κ (fst κ) + where + measurable := measurable_pi_iff.mpr fun _ ↦ measurable_density κ (fst κ) measurableSet_Iic mono' a q r hqr := ae_of_all _ fun c ↦ density_mono_set le_rfl a c (Iic_subset_Iic.mpr (by exact_mod_cast hqr)) nonneg' a q := ae_of_all _ fun c ↦ density_nonneg le_rfl _ _ _ @@ -234,9 +233,9 @@ lemma isRatCondKernelCDFAux_density_Iic (κ : kernel α (γ × ℝ)) [IsFiniteKe refine ⟨i, lt_of_le_of_lt ?_ hq⟩ exact mod_cast hi i le_rfl tendsto_integral_of_monotone a s hs_mono hs_tendsto := by - rw [kernel.fst_apply' _ _ MeasurableSet.univ] + rw [fst_apply' _ _ MeasurableSet.univ] let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) - refine tendsto_integral_density_of_monotone (le_rfl : kernel.fst κ ≤ kernel.fst κ) + refine tendsto_integral_density_of_monotone (le_rfl : fst κ ≤ fst κ) a s' ?_ ?_ (fun _ ↦ measurableSet_Iic) · exact fun i j hij ↦ Iic_subset_Iic.mpr (by exact mod_cast hs_mono hij) · ext x @@ -249,25 +248,31 @@ lemma isRatCondKernelCDFAux_density_Iic (κ : kernel α (γ × ℝ)) [IsFiniteKe integrable a q := integrable_density le_rfl a measurableSet_Iic set_integral a A hA q := set_integral_density le_rfl a measurableSet_Iic hA +/-- Taking the kernel density of intervals `Iic q` for `q : ℚ` gives a function with the property +`isRatCondKernelCDF`. -/ lemma isRatCondKernelCDF_density_Iic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsRatCondKernelCDF (fun (p : α × γ) q ↦ kernel.density κ (kernel.fst κ) p.1 p.2 (Set.Iic q)) κ - (kernel.fst κ) := + IsRatCondKernelCDF (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) κ (fst κ) := (isRatCondKernelCDFAux_density_Iic κ).isRatCondKernelCDF +/-- The conditional kernel CDF of a kernel `κ : kernel α (γ × ℝ)`, where `γ` is countably generated. +-/ noncomputable def condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : α × γ → StieltjesFunction := - stieltjesOfMeasurableRat (fun (p : α × γ) q ↦ kernel.density κ (kernel.fst κ) p.1 p.2 (Set.Iic q)) + stieltjesOfMeasurableRat (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) (isRatCondKernelCDF_density_Iic κ).measurable lemma isCondKernelCDF_condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsCondKernelCDF (condKernelCDF κ) κ (kernel.fst κ) := + IsCondKernelCDF (condKernelCDF κ) κ (fst κ) := isCondKernelCDF_stieltjesOfMeasurableRat (isRatCondKernelCDF_density_Iic κ) +/-- Auxiliary definition for `ProbabilityTheory.kernel.condKernel`. +A conditional kernel for `κ : kernel α (γ × Ω)` where `γ` is countably generated and `Ω` is +standard Borel. -/ noncomputable def condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : kernel (α × γ) Ω := let f := embeddingReal Ω let hf := measurableEmbedding_embeddingReal Ω - let κ' := kernel.map κ (Prod.map (id : γ → γ) f) (measurable_id.prod_map hf.measurable) + let κ' := map κ (Prod.map (id : γ → γ) f) (measurable_id.prod_map hf.measurable) condKernelBorelSnd κ (isCondKernelCDF_condKernelCDF κ') instance instIsMarkovKernel_condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : @@ -276,18 +281,21 @@ instance instIsMarkovKernel_condKernelBorel (κ : kernel α (γ × Ω)) [IsFinit infer_instance lemma compProd_fst_condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : - kernel.fst κ ⊗ₖ condKernelBorel κ = κ := by + fst κ ⊗ₖ condKernelBorel κ = κ := by rw [condKernelBorel, compProd_fst_condKernelBorelSnd] end CountablyGenerated section Unit +/-- Auxiliary definition for `MeasureTheory.Measure.condKernel` and +`ProbabilityTheory.kernel.condKernel`. +A conditional kernel for `κ : kernel Unit (α × Ω)` where `Ω` is standard Borel. -/ noncomputable def condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : kernel (Unit × α) Ω := let f := embeddingReal Ω let hf := measurableEmbedding_embeddingReal Ω - let κ' := kernel.map κ (Prod.map (id : α → α) f) (measurable_id.prod_map hf.measurable) + let κ' := map κ (Prod.map (id : α → α) f) (measurable_id.prod_map hf.measurable) condKernelBorelSnd κ (isCondKernelCDF_condCDF (κ' ())) instance instIsMarkovKernel_condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : @@ -296,7 +304,7 @@ instance instIsMarkovKernel_condKernelUnitBorel (κ : kernel Unit (α × Ω)) [I infer_instance lemma compProd_fst_condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : - kernel.fst κ ⊗ₖ condKernelUnitBorel κ = κ := by + fst κ ⊗ₖ condKernelUnitBorel κ = κ := by rw [condKernelUnitBorel, compProd_fst_condKernelBorelSnd] end Unit @@ -310,12 +318,12 @@ variable {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] noncomputable def _root_.MeasureTheory.Measure.condKernel (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : kernel α Ω := - kernel.comap (condKernelUnitBorel (kernel.const Unit ρ)) (fun a ↦ ((), a)) measurable_prod_mk_left + comap (condKernelUnitBorel (const Unit ρ)) (fun a ↦ ((), a)) measurable_prod_mk_left #align measure_theory.measure.cond_kernel MeasureTheory.Measure.condKernel lemma _root_.MeasureTheory.Measure.condKernel_apply (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] (a : α) : - ρ.condKernel a = condKernelUnitBorel (kernel.const Unit ρ) ((), a) := rfl + ρ.condKernel a = condKernelUnitBorel (const Unit ρ) ((), a) := rfl instance _root_.MeasureTheory.Measure.instIsMarkovKernel_condKernel (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : @@ -329,13 +337,13 @@ a Markov kernel from `α` to `Ω`. We call that Markov kernel `ρ.condKernel`. - lemma _root_.MeasureTheory.Measure.compProd_fst_condKernel (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : ρ.fst ⊗ₘ ρ.condKernel = ρ := by - have h1 : kernel.const Unit (Measure.fst ρ) = kernel.fst (kernel.const Unit ρ) := by + have h1 : const Unit (Measure.fst ρ) = fst (const Unit ρ) := by ext - simp only [kernel.fst_apply, Measure.fst, kernel.const_apply] - have h2 : kernel.prodMkLeft Unit (Measure.condKernel ρ) - = condKernelUnitBorel (kernel.const Unit ρ) := by + simp only [fst_apply, Measure.fst, const_apply] + have h2 : prodMkLeft Unit (Measure.condKernel ρ) + = condKernelUnitBorel (const Unit ρ) := by ext - simp only [kernel.prodMkLeft_apply, Measure.condKernel_apply] + simp only [prodMkLeft_apply, Measure.condKernel_apply] rw [Measure.compProd, h1, h2, compProd_fst_condKernelUnitBorel] simp #align probability_theory.measure_eq_comp_prod MeasureTheory.Measure.compProd_fst_condKernel @@ -351,11 +359,10 @@ lemma _root_.MeasureTheory.Measure.condKernel_apply_of_ne_zero_of_measurableSet = ({x} : Set α).indicator (fun a ↦ ρ.condKernel a s) a := by intro a by_cases hax : a = x - · simp only [hax, Set.singleton_prod, Set.mem_singleton_iff, Set.indicator_of_mem] + · simp only [hax, singleton_prod, mem_singleton_iff, indicator_of_mem] congr with y simp - · simp only [Set.singleton_prod, Set.mem_singleton_iff, hax, not_false_eq_true, - Set.indicator_of_not_mem] + · simp only [singleton_prod, mem_singleton_iff, hax, not_false_eq_true, indicator_of_not_mem] have : Prod.mk a ⁻¹' (Prod.mk x '' s) = ∅ := by ext y simp [Ne.symm hax] @@ -382,6 +389,9 @@ section Countable variable [MeasurableSingletonClass α] [Countable α] +/-- Auxiliary definition for `ProbabilityTheory.kernel.condKernel`. +A conditional kernel for `κ : kernel α (β × Ω)` where `α` is countable with measurable singletons +and `Ω` is standard Borel. -/ noncomputable def condKernelCountable (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : kernel (α × β) Ω where val p := (κ p.1).condKernel p.2 @@ -398,11 +408,11 @@ instance instIsMarkovKernel_condKernelCountable (κ : kernel α (β × Ω)) [IsF ⟨fun p ↦ (Measure.instIsMarkovKernel_condKernel (κ p.1)).isProbabilityMeasure p.2⟩ lemma compProd_fst_condKernelCountable (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : - kernel.fst κ ⊗ₖ condKernelCountable κ = κ := by + fst κ ⊗ₖ condKernelCountable κ = κ := by ext a s hs have h := (κ a).compProd_fst_condKernel conv_rhs => rw [← h] - simp_rw [kernel.compProd_apply _ _ _ hs, condKernelCountable_apply, Measure.compProd_apply hs] + simp_rw [compProd_apply _ _ _ hs, condKernelCountable_apply, Measure.compProd_apply hs] congr end Countable @@ -410,32 +420,38 @@ end Countable section CountableOrCountablyGenerated open Classical in + +/-- Conditional kernel of a kernel `κ : kernel α (β × Ω)`: a Markov kernel such that +`fst κ ⊗ₖ condKernel κ = κ` (see `MeasureTheory.Measure.compProd_fst_condKernel`). +It exists whenever `Ω` is standard Borel and either `α` is countable with measurable singletons +or `β` is countably generated. -/ noncomputable -irreducible_def kernel.condKernel [h : CountableOrCountablyGenerated α β] +irreducible_def condKernel [h : CountableOrCountablyGenerated α β] (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : kernel (α × β) Ω := if hα : Countable α ∧ MeasurableSingletonClass α then letI := hα.1; letI := hα.2; condKernelCountable κ - else - letI := h.countableOrCountablyGenerated.resolve_left hα; condKernelBorel κ + else letI := h.countableOrCountablyGenerated.resolve_left hα; condKernelBorel κ -instance kernel.instIsMarkovKernel_condKernel [CountableOrCountablyGenerated α β] +/-- `condKernel κ` is a Markov kernel. -/ +instance instIsMarkovKernel_condKernel [CountableOrCountablyGenerated α β] (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : - IsMarkovKernel (kernel.condKernel κ) := by - rw [kernel.condKernel_def] + IsMarkovKernel (condKernel κ) := by + rw [condKernel_def] split_ifs <;> infer_instance -lemma kernel.compProd_fst_condKernel [hαβ : CountableOrCountablyGenerated α β] +/-- **Disintegration** of finite kernels. +The composition-product of `fst κ` and `condKernel κ` is equal to `κ`. -/ +lemma compProd_fst_condKernel [hαβ : CountableOrCountablyGenerated α β] (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : - kernel.fst κ ⊗ₖ kernel.condKernel κ = κ := by - rw [kernel.condKernel_def] + fst κ ⊗ₖ condKernel κ = κ := by + rw [condKernel_def] split_ifs with h - · have := h.1 - have := h.2 + · cases h exact compProd_fst_condKernelCountable κ · have := hαβ.countableOrCountablyGenerated.resolve_left h exact compProd_fst_condKernelBorel κ end CountableOrCountablyGenerated -end ProbabilityTheory +end ProbabilityTheory.kernel From 00b0862fd97924719d131832e5c7adc5741f2fb3 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 31 Mar 2024 20:00:21 +0200 Subject: [PATCH 114/129] add docstrings in Unique.lean --- .../Kernel/Disintegration/Basic.lean | 1 - .../Kernel/Disintegration/Integral.lean | 1 - .../Kernel/Disintegration/Unique.lean | 69 +++++++++++-------- 3 files changed, 39 insertions(+), 32 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index b9e359cabade8..f1d5af01d1d08 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -58,7 +58,6 @@ The conditional kernel is unique (almost everywhere w.r.t. `fst κ`): this is pr * `ProbabilityTheory.kernel.compProd_fst_condKernel`: `fst κ ⊗ₖ condKernel κ = κ` * `MeasureTheory.Measure.compProd_fst_condKernel`: `ρ.fst ⊗ₘ ρ.condKernel = ρ` - -/ #align_import probability.kernel.disintegration from "leanprover-community/mathlib"@"6315581f5650ffa2fbdbbbedc41243c8d7070981" diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index 81eef60f18cfc..4644193c0fa57 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -20,7 +20,6 @@ Integrals of `ProbabilityTheory.kernel.condKernel` and `MeasureTheory.Measure.co Corresponding statements for the Lebesgue integral and/or without the sets `s` and `t` are also provided. - -/ open MeasureTheory ProbabilityTheory MeasurableSpace diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index 3b99e514257fe..f9c36fc2fb810 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -6,15 +6,21 @@ Authors: Kexing Ying, Rémy Degenne import Mathlib.Probability.Kernel.Disintegration.Integral /-! -# Disintegration of measures on product spaces +# Uniqueness of the conditional kernel -## Main statements +We prove that the conditional kernels `ProbabilityTheory.kernel.condKernel` and +`MeasureTheory.Measure.condKernel` are almost everywhere unique. -* `ProbabilityTheory.eq_condKernel_of_measure_eq_compProd`: a.e. uniqueness of `Measure.condKernel` +## Main statements +* `ProbabilityTheory.eq_condKernel_of_kernel_eq_compProd`: a.e. uniqueness of + `ProbabilityTheory.kernel.condKernel` +* `ProbabilityTheory.eq_condKernel_of_measure_eq_compProd`: a.e. uniqueness of + `MeasureTheory.Measure.condKernel` +* `ProbabilityTheory.kernel.condKernel_apply_eq_condKernel`: the kernel `condKernel` is almost + everywhere equal to the measure `condKernel`. -/ - open MeasureTheory Set Filter MeasurableSpace open scoped ENNReal MeasureTheory Topology ProbabilityTheory @@ -28,17 +34,17 @@ section Kernel variable [CountableOrCountablyGenerated α β] {ρ : kernel α (β × Ω)} [IsFiniteKernel ρ] -/-! ### Uniqueness +/-! ### Uniqueness of `kernel.condKernel` -This section will prove that the conditional kernel is unique almost everywhere. -/ +The conditional kernel is unique almost everywhere. -/ -/-- A s-finite kernel which satisfy the disintegration property of the given measure `ρ` is almost +/-- A s-finite kernel which satisfy the disintegration property of a kernel `ρ` is almost everywhere equal to the disintegration kernel of `ρ` when evaluated on a measurable set. -This theorem in the case of finite kernels is weaker than `eq_condKernel_of_measure_eq_compProd` +This theorem in the case of finite kernels is weaker than `eq_condKernel_of_kernel_eq_compProd` which asserts that the kernels are equal almost everywhere and not just on a given measurable set. -/ -theorem eq_condKernel_of_measure_eq_compProd' {κ : kernel (α × β) Ω} [IsSFiniteKernel κ] +theorem eq_condKernel_of_kernel_eq_compProd' {κ : kernel (α × β) Ω} [IsSFiniteKernel κ] (hκ : ρ = kernel.fst ρ ⊗ₖ κ) {s : Set Ω} (hs : MeasurableSet s) (a : α) : ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) s = kernel.condKernel ρ (a, x) s := by refine ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite ?_ ?_ ?_ @@ -51,12 +57,13 @@ theorem eq_condKernel_of_measure_eq_compProd' {κ : kernel (α × β) Ω} [IsSFi by_cases hx : x ∈ t all_goals simp [hx] --- This lemma establishes uniqueness of the disintegration kernel on ℝ -lemma eq_condKernel_of_measure_eq_compProd_real {ρ : kernel α (β × ℝ)} [IsFiniteKernel ρ] +/-- Auxiliary lemma for `eq_condKernel_of_kernel_eq_compProd`. +Uniqueness of the disintegration kernel on ℝ. -/ +lemma eq_condKernel_of_kernel_eq_compProd_real {ρ : kernel α (β × ℝ)} [IsFiniteKernel ρ] {κ : kernel (α × β) ℝ} [IsFiniteKernel κ] (hκ : ρ = kernel.fst ρ ⊗ₖ κ) (a : α) : ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) = kernel.condKernel ρ (a, x) := by have huniv : ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) Set.univ = kernel.condKernel ρ (a, x) Set.univ := - eq_condKernel_of_measure_eq_compProd' hκ MeasurableSet.univ a + eq_condKernel_of_kernel_eq_compProd' hκ MeasurableSet.univ a suffices ∀ᵐ x ∂(kernel.fst ρ a), ∀ ⦃t⦄, MeasurableSet t → κ (a, x) t = kernel.condKernel ρ (a, x) t by filter_upwards [this] with x hx @@ -65,7 +72,7 @@ lemma eq_condKernel_of_measure_eq_compProd_real {ρ : kernel α (β × ℝ)} [Is Real.isPiSystem_Iic_rat · simp only [OuterMeasure.empty', Filter.eventually_true] · simp only [iUnion_singleton_eq_range, mem_range, forall_exists_index, forall_apply_eq_imp_iff] - exact ae_all_iff.2 fun q => eq_condKernel_of_measure_eq_compProd' hκ measurableSet_Iic a + exact ae_all_iff.2 fun q => eq_condKernel_of_kernel_eq_compProd' hκ measurableSet_Iic a · filter_upwards [huniv] with x hxuniv t ht heq rw [measure_compl ht <| measure_ne_top _ _, heq, hxuniv, measure_compl ht <| measure_ne_top _ _] · refine ae_of_all _ (fun x f hdisj hf heq ↦ ?_) @@ -77,10 +84,10 @@ disintegration kernel. -/ theorem eq_condKernel_of_kernel_eq_compProd (κ : kernel (α × β) Ω) [IsFiniteKernel κ] (hκ : ρ = kernel.fst ρ ⊗ₖ κ) (a : α) : ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) = kernel.condKernel ρ (a, x) := by --- The idea is to transporting the question to `ℝ` from `Ω` using `exists_embeddingReal` --- and then constructing a measure on `α × ℝ` using the obtained measurable embedding +-- The idea is to transport the question to `ℝ` from `Ω` using `embeddingReal Ω` +-- and then construct a measure on `α × ℝ` let f := embeddingReal Ω - let hf := measurableEmbedding_embeddingReal Ω + have hf := measurableEmbedding_embeddingReal Ω set ρ' : kernel α (β × ℝ) := kernel.map ρ (Prod.map id f) (measurable_id.prod_map hf.measurable) with hρ'def have hρ' : kernel.fst ρ' = kernel.fst ρ := by @@ -90,7 +97,7 @@ theorem eq_condKernel_of_kernel_eq_compProd (κ : kernel (α × β) Ω) [IsFinit have hρ'' : ∀ᵐ x ∂(kernel.fst ρ a), kernel.map κ f hf.measurable (a, x) = kernel.condKernel ρ' (a, x) := by rw [← hρ'] - refine eq_condKernel_of_measure_eq_compProd_real ?_ a + refine eq_condKernel_of_kernel_eq_compProd_real ?_ a ext a s hs conv_lhs => rw [hρ'def, hκ] rw [kernel.map_apply' _ _ _ hs, kernel.compProd_apply _ _ _ hs, kernel.compProd_apply] @@ -111,7 +118,7 @@ theorem eq_condKernel_of_kernel_eq_compProd (κ : kernel (α × β) Ω) [IsFinit h _ <| hf.measurableSet_image.2 hs] suffices ρ' = (kernel.fst ρ ⊗ₖ (kernel.map (kernel.condKernel ρ) f hf.measurable)) by rw [← hρ'] at this - have heq := eq_condKernel_of_measure_eq_compProd_real this a + have heq := eq_condKernel_of_kernel_eq_compProd_real this a rw [hρ'] at heq filter_upwards [heq] with x hx s hs rw [← hx, kernel.map_apply, Measure.map_apply hf.measurable hs] @@ -129,9 +136,9 @@ section Measure variable {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] -/-! ### Uniqueness +/-! ### Uniqueness of `Measure.condKernel` -This section will prove that the conditional kernel is unique almost everywhere. -/ +The conditional kernel of a measure is unique almost everywhere. -/ /-- A s-finite kernel which satisfy the disintegration property of the given measure `ρ` is almost everywhere equal to the disintegration kernel of `ρ` when evaluated on a measurable set. @@ -139,7 +146,7 @@ everywhere equal to the disintegration kernel of `ρ` when evaluated on a measur This theorem in the case of finite kernels is weaker than `eq_condKernel_of_measure_eq_compProd` which asserts that the kernels are equal almost everywhere and not just on a given measurable set. -/ -theorem Measure.eq_condKernel_of_measure_eq_compProd' (κ : kernel α Ω) [IsSFiniteKernel κ] +theorem eq_condKernel_of_measure_eq_compProd' (κ : kernel α Ω) [IsSFiniteKernel κ] (hκ : ρ = ρ.fst ⊗ₘ κ) {s : Set Ω} (hs : MeasurableSet s) : ∀ᵐ x ∂ρ.fst, κ x s = ρ.condKernel x s := by refine ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite @@ -150,8 +157,9 @@ theorem Measure.eq_condKernel_of_measure_eq_compProd' (κ : kernel α Ω) [IsSFi by_cases hx : x ∈ t all_goals simp [hx] --- This lemma establishes uniqueness of the disintegration kernel on ℝ -lemma Measure.eq_condKernel_of_measure_eq_compProd_real {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] +/-- Auxiliary lemma for `eq_condKernel_of_measure_eq_compProd`. +Uniqueness of the disintegration kernel on ℝ. -/ +lemma eq_condKernel_of_measure_eq_compProd_real {ρ : Measure (α × ℝ)} [IsFiniteMeasure ρ] (κ : kernel α ℝ) [IsFiniteKernel κ] (hκ : ρ = ρ.fst ⊗ₘ κ) : ∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by have huniv : ∀ᵐ x ∂ρ.fst, κ x Set.univ = ρ.condKernel x Set.univ := @@ -172,14 +180,13 @@ lemma Measure.eq_condKernel_of_measure_eq_compProd_real {ρ : Measure (α × ℝ /-- A finite kernel which satisfies the disintegration property is almost everywhere equal to the disintegration kernel. -/ -theorem Measure.eq_condKernel_of_measure_eq_compProd (κ : kernel α Ω) [IsFiniteKernel κ] +theorem eq_condKernel_of_measure_eq_compProd (κ : kernel α Ω) [IsFiniteKernel κ] (hκ : ρ = ρ.fst ⊗ₘ κ) : ∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by - -- The idea is to transporting the question to `ℝ` from `Ω` using - -- `exists_embeddingReal` and then constructing a measure on `α × ℝ` using - -- the obtained measurable embedding + -- The idea is to transport the question to `ℝ` from `Ω` using `embeddingReal` + -- and then construct a measure on `α × ℝ` let f := embeddingReal Ω - let hf := measurableEmbedding_embeddingReal Ω + have hf := measurableEmbedding_embeddingReal Ω set ρ' : Measure (α × ℝ) := ρ.map (Prod.map id f) with hρ'def have hρ' : ρ'.fst = ρ.fst := by ext s hs @@ -222,16 +229,18 @@ end Measure section KernelAndMeasure +/-- For `fst κ a`-almost all `b`, the conditional kernel `kernel.condKernel κ` applied to `(a, b)` +is equal to the conditional kernel of the measure `κ a` applied to `b`. -/ lemma kernel.condKernel_apply_eq_condKernel [CountableOrCountablyGenerated α β] (κ : kernel α (β × Ω)) [IsFiniteKernel κ] (a : α) : - ∀ᵐ b ∂(kernel.fst κ a), kernel.condKernel κ (a, b) = (κ a).condKernel b := by + (fun b ↦ kernel.condKernel κ (a, b)) =ᵐ[kernel.fst κ a] (κ a).condKernel := by have : κ a = (κ a).fst ⊗ₘ kernel.comap (kernel.condKernel κ) (fun b ↦ (a, b)) measurable_prod_mk_left := by ext s hs conv_lhs => rw [← compProd_fst_condKernel κ] rw [Measure.compProd_apply hs, kernel.compProd_apply _ _ _ hs] rfl - have h := Measure.eq_condKernel_of_measure_eq_compProd _ this + have h := eq_condKernel_of_measure_eq_compProd _ this rw [kernel.fst_apply] filter_upwards [h] with b hb rw [← hb, kernel.comap_apply] From f5704dc948685d126507d2fddde4ff27db2d2407 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 31 Mar 2024 20:03:19 +0200 Subject: [PATCH 115/129] fix --- Mathlib/Probability/Kernel/CondDistrib.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Probability/Kernel/CondDistrib.lean b/Mathlib/Probability/Kernel/CondDistrib.lean index 9f9582fcb2e03..d856907837bac 100644 --- a/Mathlib/Probability/Kernel/CondDistrib.lean +++ b/Mathlib/Probability/Kernel/CondDistrib.lean @@ -125,7 +125,7 @@ theorem condDistrib_ae_eq_of_measure_eq_compProd (hX : Measurable X) (hY : Measu rw [Measure.map_apply hX hs, Measure.fst_apply hs, Measure.map_apply] exacts [rfl, Measurable.prod hX hY, measurable_fst hs] rw [heq, condDistrib] - refine Measure.eq_condKernel_of_measure_eq_compProd _ ?_ + refine eq_condKernel_of_measure_eq_compProd _ ?_ convert hκ exact heq.symm From 3ddfbd8ea8ffb7bf93864569468984297796181b Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 1 Apr 2024 08:52:24 +0200 Subject: [PATCH 116/129] golf Unique --- .../Kernel/Disintegration/Unique.lean | 143 ++++-------------- 1 file changed, 32 insertions(+), 111 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index f9c36fc2fb810..23039ac5421d5 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -30,108 +30,6 @@ namespace ProbabilityTheory variable {α β Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] -section Kernel - -variable [CountableOrCountablyGenerated α β] {ρ : kernel α (β × Ω)} [IsFiniteKernel ρ] - -/-! ### Uniqueness of `kernel.condKernel` - -The conditional kernel is unique almost everywhere. -/ - -/-- A s-finite kernel which satisfy the disintegration property of a kernel `ρ` is almost -everywhere equal to the disintegration kernel of `ρ` when evaluated on a measurable set. - -This theorem in the case of finite kernels is weaker than `eq_condKernel_of_kernel_eq_compProd` -which asserts that the kernels are equal almost everywhere and not just on a given measurable -set. -/ -theorem eq_condKernel_of_kernel_eq_compProd' {κ : kernel (α × β) Ω} [IsSFiniteKernel κ] - (hκ : ρ = kernel.fst ρ ⊗ₖ κ) {s : Set Ω} (hs : MeasurableSet s) (a : α) : - ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) s = kernel.condKernel ρ (a, x) s := by - refine ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite ?_ ?_ ?_ - · exact (kernel.measurable_coe κ hs).comp measurable_prod_mk_left - · exact (kernel.measurable_coe _ hs).comp measurable_prod_mk_left - intros t ht _ - conv_rhs => rw [set_lintegral_condKernel_eq_measure_prod _ ht hs, hκ] - simp only [kernel.compProd_apply _ _ _ (ht.prod hs), Set.mem_prod, ← lintegral_indicator _ ht] - congr with x - by_cases hx : x ∈ t - all_goals simp [hx] - -/-- Auxiliary lemma for `eq_condKernel_of_kernel_eq_compProd`. -Uniqueness of the disintegration kernel on ℝ. -/ -lemma eq_condKernel_of_kernel_eq_compProd_real {ρ : kernel α (β × ℝ)} [IsFiniteKernel ρ] - {κ : kernel (α × β) ℝ} [IsFiniteKernel κ] (hκ : ρ = kernel.fst ρ ⊗ₖ κ) (a : α) : - ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) = kernel.condKernel ρ (a, x) := by - have huniv : ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) Set.univ = kernel.condKernel ρ (a, x) Set.univ := - eq_condKernel_of_kernel_eq_compProd' hκ MeasurableSet.univ a - suffices ∀ᵐ x ∂(kernel.fst ρ a), - ∀ ⦃t⦄, MeasurableSet t → κ (a, x) t = kernel.condKernel ρ (a, x) t by - filter_upwards [this] with x hx - ext t ht; exact hx ht - apply MeasurableSpace.ae_induction_on_inter Real.borel_eq_generateFrom_Iic_rat - Real.isPiSystem_Iic_rat - · simp only [OuterMeasure.empty', Filter.eventually_true] - · simp only [iUnion_singleton_eq_range, mem_range, forall_exists_index, forall_apply_eq_imp_iff] - exact ae_all_iff.2 fun q => eq_condKernel_of_kernel_eq_compProd' hκ measurableSet_Iic a - · filter_upwards [huniv] with x hxuniv t ht heq - rw [measure_compl ht <| measure_ne_top _ _, heq, hxuniv, measure_compl ht <| measure_ne_top _ _] - · refine ae_of_all _ (fun x f hdisj hf heq ↦ ?_) - rw [measure_iUnion hdisj hf, measure_iUnion hdisj hf] - exact tsum_congr heq - -/-- A finite kernel which satisfies the disintegration property is almost everywhere equal to the -disintegration kernel. -/ -theorem eq_condKernel_of_kernel_eq_compProd (κ : kernel (α × β) Ω) [IsFiniteKernel κ] - (hκ : ρ = kernel.fst ρ ⊗ₖ κ) (a : α) : - ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) = kernel.condKernel ρ (a, x) := by --- The idea is to transport the question to `ℝ` from `Ω` using `embeddingReal Ω` --- and then construct a measure on `α × ℝ` - let f := embeddingReal Ω - have hf := measurableEmbedding_embeddingReal Ω - set ρ' : kernel α (β × ℝ) := kernel.map ρ (Prod.map id f) - (measurable_id.prod_map hf.measurable) with hρ'def - have hρ' : kernel.fst ρ' = kernel.fst ρ := by - ext a s hs - rw [hρ'def, kernel.fst_apply' _ _ hs, kernel.fst_apply' _ _ hs, kernel.map_apply'] - exacts [rfl, measurable_fst hs] - have hρ'' : ∀ᵐ x ∂(kernel.fst ρ a), - kernel.map κ f hf.measurable (a, x) = kernel.condKernel ρ' (a, x) := by - rw [← hρ'] - refine eq_condKernel_of_kernel_eq_compProd_real ?_ a - ext a s hs - conv_lhs => rw [hρ'def, hκ] - rw [kernel.map_apply' _ _ _ hs, kernel.compProd_apply _ _ _ hs, kernel.compProd_apply] - swap; · exact measurable_id.prod_map hf.measurable hs - congr with b - · congr with a t ht - rw [kernel.fst_apply' _ _ ht, kernel.fst_apply' _ _ ht, kernel.map_apply'] - exacts [rfl, measurable_fst ht] - · rw [kernel.map_apply'] - exacts [rfl, measurable_prod_mk_left hs] - suffices ∀ᵐ x ∂(kernel.fst ρ a), ∀ s, MeasurableSet s → - kernel.condKernel ρ' (a, x) s = kernel.condKernel ρ (a, x) (f ⁻¹' s) by - filter_upwards [hρ'', this] with x hx h - rw [kernel.map_apply] at hx - ext s hs - rw [← Set.preimage_image_eq s hf.injective, - ← Measure.map_apply hf.measurable <| hf.measurableSet_image.2 hs, hx, - h _ <| hf.measurableSet_image.2 hs] - suffices ρ' = (kernel.fst ρ ⊗ₖ (kernel.map (kernel.condKernel ρ) f hf.measurable)) by - rw [← hρ'] at this - have heq := eq_condKernel_of_kernel_eq_compProd_real this a - rw [hρ'] at heq - filter_upwards [heq] with x hx s hs - rw [← hx, kernel.map_apply, Measure.map_apply hf.measurable hs] - ext a s hs - rw [hρ'def, kernel.map_apply' _ _ _ hs, kernel.compProd_apply _ _ _ hs, - ← lintegral_condKernel_mem] - swap; · exact measurable_id.prod_map hf.measurable hs - congr with b - rw [kernel.map_apply'] - exacts [rfl, measurable_prod_mk_left hs] - -end Kernel - section Measure variable {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] @@ -229,15 +127,13 @@ end Measure section KernelAndMeasure -/-- For `fst κ a`-almost all `b`, the conditional kernel `kernel.condKernel κ` applied to `(a, b)` -is equal to the conditional kernel of the measure `κ a` applied to `b`. -/ -lemma kernel.condKernel_apply_eq_condKernel [CountableOrCountablyGenerated α β] - (κ : kernel α (β × Ω)) [IsFiniteKernel κ] (a : α) : - (fun b ↦ kernel.condKernel κ (a, b)) =ᵐ[kernel.fst κ a] (κ a).condKernel := by - have : κ a = (κ a).fst - ⊗ₘ kernel.comap (kernel.condKernel κ) (fun b ↦ (a, b)) measurable_prod_mk_left := by +lemma kernel.apply_eq_measure_condKernel_of_compProd_eq [CountableOrCountablyGenerated α β] + {ρ : kernel α (β × Ω)} [IsFiniteKernel ρ] {κ : kernel (α × β) Ω} [IsFiniteKernel κ] + (hκ : kernel.fst ρ ⊗ₖ κ = ρ) (a : α) : + (fun b ↦ κ (a, b)) =ᵐ[kernel.fst ρ a] (ρ a).condKernel := by + have : ρ a = (ρ a).fst ⊗ₘ kernel.comap κ (fun b ↦ (a, b)) measurable_prod_mk_left := by ext s hs - conv_lhs => rw [← compProd_fst_condKernel κ] + conv_lhs => rw [← hκ] rw [Measure.compProd_apply hs, kernel.compProd_apply _ _ _ hs] rfl have h := eq_condKernel_of_measure_eq_compProd _ this @@ -245,13 +141,38 @@ lemma kernel.condKernel_apply_eq_condKernel [CountableOrCountablyGenerated α β filter_upwards [h] with b hb rw [← hb, kernel.comap_apply] +/-- For `fst κ a`-almost all `b`, the conditional kernel `kernel.condKernel κ` applied to `(a, b)` +is equal to the conditional kernel of the measure `κ a` applied to `b`. -/ +lemma kernel.condKernel_apply_eq_condKernel [CountableOrCountablyGenerated α β] + (κ : kernel α (β × Ω)) [IsFiniteKernel κ] (a : α) : + (fun b ↦ kernel.condKernel κ (a, b)) =ᵐ[kernel.fst κ a] (κ a).condKernel := + kernel.apply_eq_measure_condKernel_of_compProd_eq (compProd_fst_condKernel κ) a + lemma condKernel_const [CountableOrCountablyGenerated α β] (ρ : Measure (β × Ω)) [IsFiniteMeasure ρ] (a : α) : - ∀ᵐ b ∂ρ.fst, kernel.condKernel (kernel.const α ρ) (a, b) = ρ.condKernel b := by + (fun b ↦ kernel.condKernel (kernel.const α ρ) (a, b)) =ᵐ[ρ.fst] ρ.condKernel := by have h := kernel.condKernel_apply_eq_condKernel (kernel.const α ρ) a simp_rw [kernel.fst_apply, kernel.const_apply] at h filter_upwards [h] with b hb using hb end KernelAndMeasure +section Kernel + +/-! ### Uniqueness of `kernel.condKernel` + +The conditional kernel is unique almost everywhere. -/ + +/-- A finite kernel which satisfies the disintegration property is almost everywhere equal to the +disintegration kernel. -/ +theorem eq_condKernel_of_kernel_eq_compProd [CountableOrCountablyGenerated α β] + {ρ : kernel α (β × Ω)} [IsFiniteKernel ρ] {κ : kernel (α × β) Ω} [IsFiniteKernel κ] + (hκ : kernel.fst ρ ⊗ₖ κ = ρ) (a : α) : + ∀ᵐ x ∂(kernel.fst ρ a), κ (a, x) = kernel.condKernel ρ (a, x) := by + filter_upwards [kernel.condKernel_apply_eq_condKernel ρ a, + kernel.apply_eq_measure_condKernel_of_compProd_eq hκ a] with a h1 h2 + rw [h1, h2] + +end Kernel + end ProbabilityTheory From b5563248c77cd52c0c860a51cdb4cc4bbfb7d8e0 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Mon, 1 Apr 2024 09:09:08 +0200 Subject: [PATCH 117/129] lint --- Mathlib/Probability/Kernel/Disintegration/Unique.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index 23039ac5421d5..16ee273fbe27a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -127,7 +127,7 @@ end Measure section KernelAndMeasure -lemma kernel.apply_eq_measure_condKernel_of_compProd_eq [CountableOrCountablyGenerated α β] +lemma kernel.apply_eq_measure_condKernel_of_compProd_eq {ρ : kernel α (β × Ω)} [IsFiniteKernel ρ] {κ : kernel (α × β) Ω} [IsFiniteKernel κ] (hκ : kernel.fst ρ ⊗ₖ κ = ρ) (a : α) : (fun b ↦ κ (a, b)) =ᵐ[kernel.fst ρ a] (ρ a).condKernel := by From c36174275fa5849d3cfd45e586abf5f7b1a76745 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 3 Apr 2024 17:05:08 +0200 Subject: [PATCH 118/129] remove assumption to see what breaks --- .../MeasurableSpace/CountablyGenerated.lean | 10 ++++------ 1 file changed, 4 insertions(+), 6 deletions(-) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index 4c5e20deb0c75..d282a538f71d8 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -466,16 +466,14 @@ variable [MeasurableSpace β] /-- A class registering that either `α` is countable with measurable singletons or `β` is a countably generated measurable space. -/ class CountableOrCountablyGenerated (α β : Type*) [MeasurableSpace α] [MeasurableSpace β] : Prop := - (countableOrCountablyGenerated : - (Countable α ∧ MeasurableSingletonClass α) ∨ MeasurableSpace.CountablyGenerated β) + countableOrCountablyGenerated : Countable α ∨ MeasurableSpace.CountablyGenerated β -instance instCountableOrCountablyGeneratedOfCountable - [h1 : Countable α] [h2 : MeasurableSingletonClass α] : - CountableOrCountablyGenerated α β := ⟨Or.inl ⟨h1, h2⟩⟩ +instance instCountableOrCountablyGeneratedOfCountable [h1 : Countable α] : + CountableOrCountablyGenerated α β := ⟨Or.inl h1⟩ instance instCountableOrCountablyGeneratedOfCountablyGenerated [h : MeasurableSpace.CountablyGenerated β] : - CountableOrCountablyGenerated α β := ⟨Or.inr h⟩ + CountableOrCountablyGenerated α β := ⟨Or.inr h⟩ end CountableOrCountablyGenerated From 2e9b90b48a481f7463a301a0e7d199b9c740f6dd Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 5 Apr 2024 14:40:48 +0200 Subject: [PATCH 119/129] measurable atoms --- .../MeasureTheory/MeasurableSpace/Basic.lean | 60 +++++++++++++++++-- 1 file changed, 55 insertions(+), 5 deletions(-) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index b3ffa3930443d..afc5b2fc25648 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -815,14 +815,64 @@ instance Prod.instMeasurableSingletonClass ⟨fun ⟨a, b⟩ => @singleton_prod_singleton _ _ a b ▸ .prod (.singleton a) (.singleton b)⟩ #align prod.measurable_singleton_class Prod.instMeasurableSingletonClass -theorem measurable_from_prod_countable [Countable β] [MeasurableSingletonClass β] - {_ : MeasurableSpace γ} {f : α × β → γ} (hf : ∀ y, Measurable fun x => f (x, y)) : +def measurableAtom {β : Type*} [MeasurableSpace β] (x : β) : + Set β := ⋂ (s : Set β) (_h's : x ∈ s) (_hs : MeasurableSet s), s + +@[simp] lemma mem_measurableAtom_self {β : Type*} [MeasurableSpace β] (x : β) : + x ∈ measurableAtom x := by + simp (config := {contextual := true}) [measurableAtom] + +lemma mem_of_mem_measurableAtom {β : Type*} [MeasurableSpace β] + {x y : β} (h : y ∈ measurableAtom x) {s : Set β} + (hs : MeasurableSet s) (hxs : x ∈ s) : y ∈ s := by + simp only [measurableAtom, mem_iInter] at h + exact h s hxs hs + +lemma measurableAtom_subset {β : Type*} [MeasurableSpace β] {s : Set β} {x : β} + (hs : MeasurableSet s) (hx : x ∈ s) : measurableAtom x ⊆ s := + iInter₂_subset_of_subset s hx fun ⦃a⦄ ↦ (by simp [hs]) + +@[simp] lemma measurableAtom_of_measurableSingletonClass + {β : Type*} [MeasurableSpace β] [MeasurableSingletonClass β] (x : β) : + measurableAtom x = {x} := + Subset.antisymm (measurableAtom_subset (measurableSet_singleton x) rfl) (by simp) + +lemma MeasurableSet.measurableAtom_of_countable {β : Type*} [MeasurableSpace β] [Countable β] (x : β) : + MeasurableSet (measurableAtom x) := by + have : ∀ (y : β), y ∉ measurableAtom x → ∃ s, MeasurableSet s ∧ x ∈ s ∧ y ∉ s := + fun y hy ↦ by simpa [measurableAtom] using hy + choose! s hs using this + have : measurableAtom x = ⋂ (y ∈ (measurableAtom x)ᶜ), s y := by + apply Subset.antisymm + · intro z hz + simp only [mem_iInter, mem_compl_iff] + intro i hi + show z ∈ s i + exact mem_of_mem_measurableAtom hz (hs i hi).1 (hs i hi).2.1 + · apply compl_subset_compl.1 + intro z hz + simp only [compl_iInter, mem_iUnion, mem_compl_iff, exists_prop] + exact ⟨z, hz, (hs z hz).2.2⟩ + rw [this] + exact MeasurableSet.biInter (to_countable (measurableAtom x)ᶜ) (fun i hi ↦ (hs i hi).1) + +theorem measurable_from_prod_countable' [Countable β] + {_ : MeasurableSpace γ} {f : α × β → γ} (hf : ∀ y, Measurable fun x => f (x, y)) + (h'f : ∀ y y' x, y' ∈ measurableAtom y → f (x, y') = f (x, y)) : Measurable f := fun s hs => by - have : f ⁻¹' s = ⋃ y, ((fun x => f (x, y)) ⁻¹' s) ×ˢ ({y} : Set β) := by + have : f ⁻¹' s = ⋃ y, ((fun x => f (x, y)) ⁻¹' s) ×ˢ (measurableAtom y : Set β) := by ext1 ⟨x, y⟩ - simp [and_assoc, and_left_comm] + simp only [mem_preimage, mem_iUnion, mem_prod] + refine ⟨fun h ↦ ⟨y, h, mem_measurableAtom_self y⟩, ?_⟩ + rintro ⟨y', hy's, hy'⟩ + rwa [h'f y' y x hy'] rw [this] - exact .iUnion fun y => (hf y hs).prod (.singleton y) + exact .iUnion (fun y ↦ (hf y hs).prod (.measurableAtom_of_countable y)) + +theorem measurable_from_prod_countable [Countable β] [MeasurableSingletonClass β] + {_ : MeasurableSpace γ} {f : α × β → γ} (hf : ∀ y, Measurable fun x => f (x, y)) : + Measurable f := + measurable_from_prod_countable' hf (by simp (config := {contextual := true})) #align measurable_from_prod_countable measurable_from_prod_countable /-- A piecewise function on countably many pieces is measurable if all the data is measurable. -/ From 9325b6201973926f929bac29406d66d0c19d5363 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 5 Apr 2024 15:02:29 +0200 Subject: [PATCH 120/129] start fixing --- .../Kernel/Disintegration/Basic.lean | 23 +++++++++++-------- 1 file changed, 13 insertions(+), 10 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index f1d5af01d1d08..1738b9f832f88 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -386,18 +386,23 @@ end Measure section Countable -variable [MeasurableSingletonClass α] [Countable α] +variable [Countable α] /-- Auxiliary definition for `ProbabilityTheory.kernel.condKernel`. -A conditional kernel for `κ : kernel α (β × Ω)` where `α` is countable with measurable singletons -and `Ω` is standard Borel. -/ +A conditional kernel for `κ : kernel α (β × Ω)` where `α` is countable and `Ω` is standard Borel. -/ noncomputable def condKernelCountable (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : kernel (α × β) Ω where val p := (κ p.1).condKernel p.2 property := by change Measurable ((fun q : β × α ↦ (κ q.2).condKernel q.1) ∘ Prod.swap) - refine (measurable_from_prod_countable (fun a ↦ ?_)).comp measurable_swap - exact kernel.measurable (κ a).condKernel + refine (measurable_from_prod_countable' (fun a ↦ ?_) ?_).comp measurable_swap + · exact kernel.measurable (κ a).condKernel + · intro y y' x hy' + have : κ y' = κ y := by + ext s hs + exact mem_of_mem_measurableAtom hy' + (kernel.measurable_coe κ hs (measurableSet_singleton (κ y s))) rfl + simp [this] lemma condKernelCountable_apply (κ : kernel α (β × Ω)) [IsFiniteKernel κ] (p : α × β) : condKernelCountable κ p = (κ p.1).condKernel p.2 := rfl @@ -422,14 +427,13 @@ open Classical in /-- Conditional kernel of a kernel `κ : kernel α (β × Ω)`: a Markov kernel such that `fst κ ⊗ₖ condKernel κ = κ` (see `MeasureTheory.Measure.compProd_fst_condKernel`). -It exists whenever `Ω` is standard Borel and either `α` is countable with measurable singletons +It exists whenever `Ω` is standard Borel and either `α` is countable or `β` is countably generated. -/ noncomputable irreducible_def condKernel [h : CountableOrCountablyGenerated α β] (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : kernel (α × β) Ω := - if hα : Countable α ∧ MeasurableSingletonClass α then - letI := hα.1; letI := hα.2; condKernelCountable κ + if hα : Countable α then condKernelCountable κ else letI := h.countableOrCountablyGenerated.resolve_left hα; condKernelBorel κ /-- `condKernel κ` is a Markov kernel. -/ @@ -446,8 +450,7 @@ lemma compProd_fst_condKernel [hαβ : CountableOrCountablyGenerated α β] fst κ ⊗ₖ condKernel κ = κ := by rw [condKernel_def] split_ifs with h - · cases h - exact compProd_fst_condKernelCountable κ + · exact compProd_fst_condKernelCountable κ · have := hαβ.countableOrCountablyGenerated.resolve_left h exact compProd_fst_condKernelBorel κ From 0c3e29ed4139b40172c947c003344a5e43d84a39 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 5 Apr 2024 15:30:52 +0200 Subject: [PATCH 121/129] move to better place, lint, cleanup --- .../MeasureTheory/MeasurableSpace/Basic.lean | 88 ++++++++++--------- .../MeasurableSpace/CountablyGenerated.lean | 20 ++++- 2 files changed, 64 insertions(+), 44 deletions(-) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index afc5b2fc25648..0334114775217 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -691,6 +691,53 @@ theorem measurable_of_measurable_on_compl_singleton [MeasurableSingletonClass α end Subtype +section Atoms + +variable [MeasurableSpace β] + +/-- The *measurable atom* of `x` is the intersection of all the measurable sets countaining `x`. +It is measurable when the space is countable (or more generally when the measurable space is +countably generated). -/ +def measurableAtom (x : β) : Set β := + ⋂ (s : Set β) (_h's : x ∈ s) (_hs : MeasurableSet s), s + +@[simp] lemma mem_measurableAtom_self (x : β) : x ∈ measurableAtom x := by + simp (config := {contextual := true}) [measurableAtom] + +lemma mem_of_mem_measurableAtom {x y : β} (h : y ∈ measurableAtom x) {s : Set β} + (hs : MeasurableSet s) (hxs : x ∈ s) : y ∈ s := by + simp only [measurableAtom, mem_iInter] at h + exact h s hxs hs + +lemma measurableAtom_subset {s : Set β} {x : β} (hs : MeasurableSet s) (hx : x ∈ s) : + measurableAtom x ⊆ s := + iInter₂_subset_of_subset s hx fun ⦃a⦄ ↦ (by simp [hs]) + +@[simp] lemma measurableAtom_of_measurableSingletonClass [MeasurableSingletonClass β] (x : β) : + measurableAtom x = {x} := + Subset.antisymm (measurableAtom_subset (measurableSet_singleton x) rfl) (by simp) + +lemma MeasurableSet.measurableAtom_of_countable [Countable β] (x : β) : + MeasurableSet (measurableAtom x) := by + have : ∀ (y : β), y ∉ measurableAtom x → ∃ s, MeasurableSet s ∧ x ∈ s ∧ y ∉ s := + fun y hy ↦ by simpa [measurableAtom] using hy + choose! s hs using this + have : measurableAtom x = ⋂ (y ∈ (measurableAtom x)ᶜ), s y := by + apply Subset.antisymm + · intro z hz + simp only [mem_iInter, mem_compl_iff] + intro i hi + show z ∈ s i + exact mem_of_mem_measurableAtom hz (hs i hi).1 (hs i hi).2.1 + · apply compl_subset_compl.1 + intro z hz + simp only [compl_iInter, mem_iUnion, mem_compl_iff, exists_prop] + exact ⟨z, hz, (hs z hz).2.2⟩ + rw [this] + exact MeasurableSet.biInter (to_countable (measurableAtom x)ᶜ) (fun i hi ↦ (hs i hi).1) + +end Atoms + section Prod /-- A `MeasurableSpace` structure on the product of two measurable spaces. -/ @@ -815,47 +862,6 @@ instance Prod.instMeasurableSingletonClass ⟨fun ⟨a, b⟩ => @singleton_prod_singleton _ _ a b ▸ .prod (.singleton a) (.singleton b)⟩ #align prod.measurable_singleton_class Prod.instMeasurableSingletonClass -def measurableAtom {β : Type*} [MeasurableSpace β] (x : β) : - Set β := ⋂ (s : Set β) (_h's : x ∈ s) (_hs : MeasurableSet s), s - -@[simp] lemma mem_measurableAtom_self {β : Type*} [MeasurableSpace β] (x : β) : - x ∈ measurableAtom x := by - simp (config := {contextual := true}) [measurableAtom] - -lemma mem_of_mem_measurableAtom {β : Type*} [MeasurableSpace β] - {x y : β} (h : y ∈ measurableAtom x) {s : Set β} - (hs : MeasurableSet s) (hxs : x ∈ s) : y ∈ s := by - simp only [measurableAtom, mem_iInter] at h - exact h s hxs hs - -lemma measurableAtom_subset {β : Type*} [MeasurableSpace β] {s : Set β} {x : β} - (hs : MeasurableSet s) (hx : x ∈ s) : measurableAtom x ⊆ s := - iInter₂_subset_of_subset s hx fun ⦃a⦄ ↦ (by simp [hs]) - -@[simp] lemma measurableAtom_of_measurableSingletonClass - {β : Type*} [MeasurableSpace β] [MeasurableSingletonClass β] (x : β) : - measurableAtom x = {x} := - Subset.antisymm (measurableAtom_subset (measurableSet_singleton x) rfl) (by simp) - -lemma MeasurableSet.measurableAtom_of_countable {β : Type*} [MeasurableSpace β] [Countable β] (x : β) : - MeasurableSet (measurableAtom x) := by - have : ∀ (y : β), y ∉ measurableAtom x → ∃ s, MeasurableSet s ∧ x ∈ s ∧ y ∉ s := - fun y hy ↦ by simpa [measurableAtom] using hy - choose! s hs using this - have : measurableAtom x = ⋂ (y ∈ (measurableAtom x)ᶜ), s y := by - apply Subset.antisymm - · intro z hz - simp only [mem_iInter, mem_compl_iff] - intro i hi - show z ∈ s i - exact mem_of_mem_measurableAtom hz (hs i hi).1 (hs i hi).2.1 - · apply compl_subset_compl.1 - intro z hz - simp only [compl_iInter, mem_iUnion, mem_compl_iff, exists_prop] - exact ⟨z, hz, (hs z hz).2.2⟩ - rw [this] - exact MeasurableSet.biInter (to_countable (measurableAtom x)ᶜ) (fun i hi ↦ (hs i hi).1) - theorem measurable_from_prod_countable' [Countable β] {_ : MeasurableSpace γ} {f : α × β → γ} (hf : ∀ y, Measurable fun x => f (x, y)) (h'f : ∀ y y' x, y' ∈ measurableAtom y → f (x, y') = f (x, y)) : diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index d282a538f71d8..1fc60788b555b 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -106,9 +106,23 @@ theorem CountablyGenerated.sup {m₁ m₂ : MeasurableSpace β} (h₁ : @Countab rcases h₂ with ⟨⟨b₂, hb₂c, rfl⟩⟩ exact @mk _ (_ ⊔ _) ⟨_, hb₁c.union hb₂c, generateFrom_sup_generateFrom⟩ -instance (priority := 100) [MeasurableSpace α] [Finite α] : CountablyGenerated α where - isCountablyGenerated := - ⟨{s | MeasurableSet s}, Set.to_countable _, generateFrom_measurableSet.symm⟩ +/-- Any measurable space structure on a countable space is countably generated. -/ +instance (priority := 100) [MeasurableSpace α] [Countable α] : CountablyGenerated α where + isCountablyGenerated := by + refine ⟨⋃ y, {measurableAtom y}, countable_iUnion (fun i ↦ countable_singleton _), ?_⟩ + refine le_antisymm ?_ (generateFrom_le (by simp [MeasurableSet.measurableAtom_of_countable])) + intro s hs + have : s = ⋃ y ∈ s, measurableAtom y := by + apply Subset.antisymm + · intro x hx + simpa using ⟨x, hx, by simp⟩ + · simp only [iUnion_subset_iff] + intro x hx + exact measurableAtom_subset hs hx + rw [this] + apply MeasurableSet.biUnion (to_countable s) (fun x _hx ↦ ?_) + apply measurableSet_generateFrom + simp instance [MeasurableSpace α] [CountablyGenerated α] {p : α → Prop} : CountablyGenerated { x // p x } := .comap _ From 486309df1f938fdc9ee187bcf9171eff2c4027c6 Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sat, 6 Apr 2024 09:40:42 +0200 Subject: [PATCH 122/129] fix docstring --- Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index 1fc60788b555b..a39718deddf8d 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -477,8 +477,8 @@ section CountableOrCountablyGenerated variable [MeasurableSpace β] -/-- A class registering that either `α` is countable with measurable singletons or `β` is a -countably generated measurable space. -/ +/-- A class registering that either `α` is countable or `β` is a countably generated +measurable space. -/ class CountableOrCountablyGenerated (α β : Type*) [MeasurableSpace α] [MeasurableSpace β] : Prop := countableOrCountablyGenerated : Countable α ∨ MeasurableSpace.CountablyGenerated β From cb7621508d6e69ce42fc15b40213f3a425c29f2d Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 7 Apr 2024 09:10:22 +0200 Subject: [PATCH 123/129] refactor real to Borel --- Mathlib/Probability/Kernel/Composition.lean | 19 ++ .../Kernel/Disintegration/Basic.lean | 292 +++++++++++------- .../Kernel/Disintegration/CdfToKernel.lean | 22 -- 3 files changed, 195 insertions(+), 138 deletions(-) diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index bdf97e6f347d4..2f9050bc491c1 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -583,6 +583,25 @@ lemma compProd_add_right (μ : kernel α β) (κ η : kernel (α × β) γ) rw [lintegral_add_left] exact measurable_kernel_prod_mk_left' hs a +lemma comapRight_compProd_id_prod {δ : Type*} {mδ : MeasurableSpace δ} + (κ : kernel α β) [IsSFiniteKernel κ] (η : kernel (α × β) γ) [IsSFiniteKernel η] + {f : δ → γ} (hf : MeasurableEmbedding f) : + comapRight (κ ⊗ₖ η) (MeasurableEmbedding.id.prod_mk hf) = κ ⊗ₖ (comapRight η hf) := by + ext a t ht + rw [comapRight_apply' _ _ _ ht, compProd_apply, compProd_apply _ _ _ ht] + swap; · exact (MeasurableEmbedding.id.prod_mk hf).measurableSet_image.mpr ht + refine lintegral_congr (fun b ↦ ?_) + simp only [id_eq, Set.mem_image, Prod.mk.injEq, Prod.exists] + rw [comapRight_apply'] + swap; · exact measurable_prod_mk_left ht + congr with x + simp only [Set.mem_setOf_eq, Set.mem_image] + constructor + · rintro ⟨b', c, h, rfl, rfl⟩ + exact ⟨c, h, rfl⟩ + · rintro ⟨c, h, rfl⟩ + exact ⟨b, c, h, rfl, rfl⟩ + end CompositionProduct section MapComap diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 1738b9f832f88..f8787a9bb0324 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.Probability.Kernel.MeasureCompProd import Mathlib.Probability.Kernel.Disintegration.CondCdf import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.Disintegration.CdfToKernel +import Mathlib.Probability.ConditionalProbability /-! # Disintegration of kernels and measures @@ -95,115 +96,142 @@ section BorelSnd Since every standard Borel space embeds measurably into `ℝ`, we can generalize a disintegration property on `ℝ` to all these spaces. -/ +open Classical in /-- Auxiliary definition for `ProbabilityTheory.kernel.condKernel`. -For `κ` a kernel from `α` to `β × Ω` where `Ω` is standard borel, and `f` a conditional kernel -CDF of the map of `κ` by an embedding of `Ω` into `ℝ` (a kernel from `α` to `β × ℝ`), we build a -`kernel (α × β) Ω`. - -`f` gives a Markov kernel from `α × β` to `ℝ`: we can't take directly the comap of that kernel by -the embedding `embeddingReal Ω`, because it would not be a Markov kernel (its value is a probability -measure only almost everywhere). We thus take the comap of a slight modification. -/ +A Borel space `Ω` embedds measurably into `ℝ` (with embedding `e`), hence we can get a `kernel α Ω` +from a `kernel α ℝ` by taking the comap by `e`. +Here we take the comap of a modification of `η : kernel α ℝ`, useful when `η a` is a probability +measure with all its mass on `range e` almost everywhere with respect to some measure and we want to +ensure that the comap is a Markov kernel. +We thus take the comap by `e` of a kernel defined piecewise: `η` when +`η a (range (embeddingReal Ω))ᶜ = 0`, and an arbitrary deterministic kernel otherwise. -/ noncomputable -def condKernelBorelSnd (κ : kernel α (β × Ω)) {f : α × β → StieltjesFunction} - (hf : IsCondKernelCDF f - (map κ (Prod.map (id : β → β) (embeddingReal Ω)) - (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) - (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) - (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)))) : - kernel (α × β) Ω := +def borelMarkovFromReal (Ω : Type*) [Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω] + (η : kernel α ℝ) : + kernel α Ω := have he := measurableEmbedding_embeddingReal Ω let x₀ := (range_nonempty (embeddingReal Ω)).choose - comapRight (piecewise (measurableSet_toKernel_eq_one hf he.measurableSet_range) - (hf.toKernel f) (deterministic (fun _ ↦ x₀) measurable_const)) he - -/-- `condKernelBorelSnd` is a Markov kernel. -/ -instance instIsMarkovKernel_condKernelBorelSnd (κ : kernel α (β × Ω)) - {f : α × β → StieltjesFunction} - (hf : IsCondKernelCDF f - (map κ (Prod.map (id : β → β) (embeddingReal Ω)) - (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) - (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) - (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)))) : - IsMarkovKernel (condKernelBorelSnd κ hf) := by - rw [condKernelBorelSnd] - refine IsMarkovKernel.comapRight _ _ fun a ↦ ?_ - rw [piecewise_apply'] - split_ifs with h_mem - · exact h_mem - · classical - rw [deterministic_apply' _ _ (measurableEmbedding_embeddingReal Ω).measurableSet_range, - indicator_apply, if_pos] - exact (range_nonempty (embeddingReal Ω)).choose_spec - -/-- `fst κ ⊗ₖ condKernelBorelSnd κ hf = κ`. -/ -lemma compProd_fst_condKernelBorelSnd (κ : kernel α (β × Ω)) [IsFiniteKernel κ] - {f : α × β → StieltjesFunction} - (hf : IsCondKernelCDF f - (map κ (Prod.map (id : β → β) (embeddingReal Ω)) - (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) - (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) - (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)))) : - fst κ ⊗ₖ condKernelBorelSnd κ hf = κ := by + comapRight + (piecewise ((kernel.measurable_coe η he.measurableSet_range.compl) (measurableSet_singleton 0) : + MeasurableSet {a | η a (range (embeddingReal Ω))ᶜ = 0}) + η (deterministic (fun _ ↦ x₀) measurable_const)) he + +lemma borelMarkovFromReal_apply (Ω : Type*) [Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω] + (η : kernel α ℝ) (a : α) : + borelMarkovFromReal Ω η a + = if η a (range (embeddingReal Ω))ᶜ = 0 then (η a).comap (embeddingReal Ω) + else (Measure.dirac (range_nonempty (embeddingReal Ω)).choose).comap (embeddingReal Ω) := by + classical + rw [borelMarkovFromReal, comapRight_apply, piecewise_apply, deterministic_apply] + simp only [mem_preimage, mem_singleton_iff] + split_ifs <;> rfl + +lemma borelMarkovFromReal_apply' (Ω : Type*) [Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω] + (η : kernel α ℝ) (a : α) {s : Set Ω} (hs : MeasurableSet s) : + borelMarkovFromReal Ω η a s + = if η a (range (embeddingReal Ω))ᶜ = 0 then η a (embeddingReal Ω '' s) + else (embeddingReal Ω '' s).indicator 1 (range_nonempty (embeddingReal Ω)).choose := by + have he := measurableEmbedding_embeddingReal Ω + rw [borelMarkovFromReal_apply] + split_ifs with h + · rw [Measure.comap_apply _ he.injective he.measurableSet_image' _ hs] + · rw [Measure.comap_apply _ he.injective he.measurableSet_image' _ hs, Measure.dirac_apply] + +/-- When `η` is an s-finite kernel, `borelMarkovFromReal Ω η` is an s-finite kernel. -/ +instance instIsSFiniteKernelBorelMarkovFromReal (η : kernel α ℝ) [IsSFiniteKernel η] : + IsSFiniteKernel (borelMarkovFromReal Ω η) := + IsSFiniteKernel.comapRight _ (measurableEmbedding_embeddingReal Ω) + +/-- When `η` is a finite kernel, `borelMarkovFromReal Ω η` is a finite kernel. -/ +instance instIsFiniteKernelBorelMarkovFromReal (η : kernel α ℝ) [IsFiniteKernel η] : + IsFiniteKernel (borelMarkovFromReal Ω η) := + IsFiniteKernel.comapRight _ (measurableEmbedding_embeddingReal Ω) + +/-- When `η` is a Markov kernel, `borelMarkovFromReal Ω η` is a Markov kernel. -/ +instance instIsMarkovKernelBorelMarkovFromReal (η : kernel α ℝ) [IsMarkovKernel η] : + IsMarkovKernel (borelMarkovFromReal Ω η) := by + refine IsMarkovKernel.comapRight _ (measurableEmbedding_embeddingReal Ω) (fun a ↦ ?_) + classical + rw [piecewise_apply] + split_ifs with h + · rwa [← prob_compl_eq_zero_iff (measurableEmbedding_embeddingReal Ω).measurableSet_range] + · rw [deterministic_apply] + simp [(range_nonempty (embeddingReal Ω)).choose_spec] + +/-- For `κ' := map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable)`, the +hypothesis `hη` is `fst κ' ⊗ₖ η = κ'`. The conclusion of the lemma is +`fst κ ⊗ₖ borelMarkovFromReal Ω η = comapRight (fst κ' ⊗ₖ η) _`. -/ +lemma compProd_fst_borelMarkovFromReal_eq_comapRight_compProd + (κ : kernel α (β × Ω)) [IsSFiniteKernel κ] (η : kernel (α × β) ℝ) [IsSFiniteKernel η] + (hη : (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable))) ⊗ₖ η + = map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) : + fst κ ⊗ₖ borelMarkovFromReal Ω η + = comapRight (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) ⊗ₖ η) + (MeasurableEmbedding.id.prod_mk (measurableEmbedding_embeddingReal Ω)) := by let e := embeddingReal Ω let he := measurableEmbedding_embeddingReal Ω let κ' := map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable) + have hη' : fst κ' ⊗ₖ η = κ' := hη have h_prod_embed : MeasurableEmbedding (Prod.map (id : β → β) e) := MeasurableEmbedding.id.prod_mk he + change fst κ ⊗ₖ borelMarkovFromReal Ω η = comapRight (fst κ' ⊗ₖ η) h_prod_embed + rw [comapRight_compProd_id_prod _ _ he] have h_fst : fst κ' = fst κ := by - ext a u hu - unfold_let κ' - rw [fst_apply' _ _ hu, fst_apply' _ _ hu, map_apply' κ h_prod_embed.measurable] - · rfl - · exact measurable_fst hu - have : κ = comapRight κ' h_prod_embed := by - ext c t ht : 2 + ext a u unfold_let κ' - rw [comapRight_apply' _ _ _ ht, map_apply' κ h_prod_embed.measurable - _ (h_prod_embed.measurableSet_image.mpr ht)] - congr with x : 1 - rw [← @Prod.mk.eta _ _ x] - simp only [id.def, mem_preimage, Prod.map_mk, mem_image, Prod.mk.inj_iff, Prod.exists] - refine' ⟨fun h => ⟨x.1, x.2, h, rfl, rfl⟩, _⟩ - rintro ⟨a, b, h_mem, rfl, he_eq⟩ - rwa [he.injective he_eq] at h_mem - conv_rhs => rw [this] - unfold_let κ' - conv_rhs => rw [← compProd_toKernel hf] - change fst κ ⊗ₖ condKernelBorelSnd κ hf = comapRight (fst κ' ⊗ₖ hf.toKernel f) h_prod_embed - ext c t ht : 2 - rw [comapRight_apply' _ _ _ ht, compProd_apply _ _ _ (h_prod_embed.measurableSet_image.mpr ht)] - simp_rw [h_fst, compProd_apply _ _ _ ht] + rw [fst_apply, map_apply, Measure.map_map measurable_fst h_prod_embed.measurable, fst_apply] + congr + rw [h_fst] + ext a t ht : 2 + simp_rw [compProd_apply _ _ _ ht] refine lintegral_congr_ae ?_ - let ρ_set := {p : α × β | hf.toKernel f p (range e) = 1} - have h_ae : ∀ a, ∀ᵐ t ∂(fst κ a), (a, t) ∈ ρ_set := by - intro a + have h_ae : ∀ᵐ t ∂(fst κ a), (a, t) ∈ {p : α × β | η p (range e)ᶜ = 0} := by rw [← h_fst] - refine ae_toKernel_eq_one hf a he.measurableSet_range ?_ - simp only [mem_compl_iff, mem_range, not_exists] - rw [map_apply'] - · have h_empty : {a : β × Ω | ∀ (x : Ω), ¬e x = e a.2} = ∅ := by - ext x - simp only [mem_setOf_eq, mem_empty_iff_false, iff_false, not_forall, not_not, - exists_apply_eq_apply] - simp [h_empty] - · have : {x : β × ℝ | ∀ (y : Ω), ¬ e y = x.2} = univ ×ˢ (range e)ᶜ := by - ext x - simp only [mem_setOf_eq, mem_prod, mem_univ, mem_compl_iff, mem_range, not_exists, true_and] - rw [this] - exact MeasurableSet.univ.prod he.measurableSet_range.compl - filter_upwards [h_ae c] with a ha - rw [condKernelBorelSnd, comapRight_apply'] - swap; · exact measurable_prod_mk_left ht - have h1 : {c : ℝ | (a, c) ∈ Prod.map id e '' t} = e '' {c : Ω | (a, c) ∈ t} := by - ext1 x - simp only [Prod_map, id.def, mem_image, Prod.mk.inj_iff, Prod.exists, mem_setOf_eq] - constructor - · rintro ⟨a', b, h_mem, rfl, hf_eq⟩ - exact ⟨b, h_mem, hf_eq⟩ - · rintro ⟨b, h_mem, hf_eq⟩ - exact ⟨a, b, h_mem, rfl, hf_eq⟩ - rw [h1, piecewise_apply, if_pos ha] - rfl + have h_compProd : κ' a (univ ×ˢ range e)ᶜ = 0 := by + unfold_let κ' + rw [map_apply'] + swap; · exact (MeasurableSet.univ.prod he.measurableSet_range).compl + suffices Prod.map id e ⁻¹' (univ ×ˢ range e)ᶜ = ∅ by rw [this]; simp + ext x + simp + rw [← hη', compProd_null] at h_compProd + swap; · exact (MeasurableSet.univ.prod he.measurableSet_range).compl + simp only [preimage_compl, mem_univ, mk_preimage_prod_right] at h_compProd + exact h_compProd + filter_upwards [h_ae] with a ha + rw [borelMarkovFromReal, comapRight_apply', comapRight_apply'] + rotate_left + · exact measurable_prod_mk_left ht + · exact measurable_prod_mk_left ht + classical + rw [piecewise_apply, if_pos] + exact ha + +/-- For `κ' := map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable)`, the +hypothesis `hη` is `fst κ' ⊗ₖ η = κ'`. With that hypothesis, +`fst κ ⊗ₖ borelMarkovFromReal κ η = κ`.-/ +lemma compProd_fst_borelMarkovFromReal (κ : kernel α (β × Ω)) [IsSFiniteKernel κ] + (η : kernel (α × β) ℝ) [IsSFiniteKernel η] + (hη : (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable))) ⊗ₖ η + = map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) : + fst κ ⊗ₖ borelMarkovFromReal Ω η = κ := by + let e := embeddingReal Ω + let he := measurableEmbedding_embeddingReal Ω + let κ' := map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable) + have hη' : fst κ' ⊗ₖ η = κ' := hη + have h_prod_embed : MeasurableEmbedding (Prod.map (id : β → β) e) := + MeasurableEmbedding.id.prod_mk he + have : κ = comapRight κ' h_prod_embed := by + ext c t : 2 + unfold_let κ' + rw [comapRight_apply, map_apply, h_prod_embed.comap_map] + conv_rhs => rw [this, ← hη'] + exact compProd_fst_borelMarkovFromReal_eq_comapRight_compProd κ η hη end BorelSnd @@ -264,47 +292,81 @@ lemma isCondKernelCDF_condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel IsCondKernelCDF (condKernelCDF κ) κ (fst κ) := isCondKernelCDF_stieltjesOfMeasurableRat (isRatCondKernelCDF_density_Iic κ) +/-- Auxiliary definition for `ProbabilityTheory.kernel.condKernel`. +A conditional kernel for `κ : kernel α (γ × ℝ)` where `γ` is countably generated. -/ +noncomputable +def condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : kernel (α × γ) ℝ := + (isCondKernelCDF_condKernelCDF κ).toKernel + +instance instIsMarkovKernelCondKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsMarkovKernel (condKernelReal κ) := by + rw [condKernelReal] + infer_instance + +lemma compProd_fst_condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + fst κ ⊗ₖ condKernelReal κ = κ := by + rw [condKernelReal, compProd_toKernel] + /-- Auxiliary definition for `ProbabilityTheory.kernel.condKernel`. A conditional kernel for `κ : kernel α (γ × Ω)` where `γ` is countably generated and `Ω` is standard Borel. -/ noncomputable def condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : kernel (α × γ) Ω := - let f := embeddingReal Ω - let hf := measurableEmbedding_embeddingReal Ω - let κ' := map κ (Prod.map (id : γ → γ) f) (measurable_id.prod_map hf.measurable) - condKernelBorelSnd κ (isCondKernelCDF_condKernelCDF κ') + let e := embeddingReal Ω + let he := measurableEmbedding_embeddingReal Ω + let κ' := map κ (Prod.map (id : γ → γ) e) (measurable_id.prod_map he.measurable) + borelMarkovFromReal Ω (condKernelReal κ') -instance instIsMarkovKernel_condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : +instance instIsMarkovKernelCondKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : IsMarkovKernel (condKernelBorel κ) := by rw [condKernelBorel] infer_instance lemma compProd_fst_condKernelBorel (κ : kernel α (γ × Ω)) [IsFiniteKernel κ] : fst κ ⊗ₖ condKernelBorel κ = κ := by - rw [condKernelBorel, compProd_fst_condKernelBorelSnd] + rw [condKernelBorel, compProd_fst_borelMarkovFromReal _ _ (compProd_fst_condKernelReal _)] end CountablyGenerated section Unit +/-- Auxiliary definition for `MeasureTheory.Measure.condKernel` and +`ProbabilityTheory.kernel.condKernel`. +A conditional kernel for `κ : kernel Unit (α × ℝ)`. -/ +noncomputable +def condKernelUnitReal (κ : kernel Unit (α × ℝ)) [IsFiniteKernel κ] : kernel (Unit × α) ℝ := + (isCondKernelCDF_condCDF (κ ())).toKernel + +instance instIsMarkovKernelCondKernelUnitReal (κ : kernel Unit (α × ℝ)) [IsFiniteKernel κ] : + IsMarkovKernel (condKernelUnitReal κ) := by + rw [condKernelUnitReal] + infer_instance + +lemma compProd_fst_condKernelUnitReal (κ : kernel Unit (α × ℝ)) [IsFiniteKernel κ] : + fst κ ⊗ₖ condKernelUnitReal κ = κ := by + rw [condKernelUnitReal, compProd_toKernel] + ext a + simp + /-- Auxiliary definition for `MeasureTheory.Measure.condKernel` and `ProbabilityTheory.kernel.condKernel`. A conditional kernel for `κ : kernel Unit (α × Ω)` where `Ω` is standard Borel. -/ noncomputable def condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : kernel (Unit × α) Ω := - let f := embeddingReal Ω - let hf := measurableEmbedding_embeddingReal Ω - let κ' := map κ (Prod.map (id : α → α) f) (measurable_id.prod_map hf.measurable) - condKernelBorelSnd κ (isCondKernelCDF_condCDF (κ' ())) + let e := embeddingReal Ω + let he := measurableEmbedding_embeddingReal Ω + let κ' := map κ (Prod.map (id : α → α) e) (measurable_id.prod_map he.measurable) + borelMarkovFromReal Ω (condKernelUnitReal κ') -instance instIsMarkovKernel_condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : +instance instIsMarkovKernelCondKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : IsMarkovKernel (condKernelUnitBorel κ) := by rw [condKernelUnitBorel] infer_instance lemma compProd_fst_condKernelUnitBorel (κ : kernel Unit (α × Ω)) [IsFiniteKernel κ] : fst κ ⊗ₖ condKernelUnitBorel κ = κ := by - rw [condKernelUnitBorel, compProd_fst_condKernelBorelSnd] + rw [condKernelUnitBorel, + compProd_fst_borelMarkovFromReal _ _ (compProd_fst_condKernelUnitReal _)] end Unit @@ -324,9 +386,8 @@ lemma _root_.MeasureTheory.Measure.condKernel_apply (ρ : Measure (α × Ω)) [I (a : α) : ρ.condKernel a = condKernelUnitBorel (const Unit ρ) ((), a) := rfl -instance _root_.MeasureTheory.Measure.instIsMarkovKernel_condKernel - (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : - IsMarkovKernel ρ.condKernel := by +instance _root_.MeasureTheory.Measure.instIsMarkovKernelCondKernel + (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : IsMarkovKernel ρ.condKernel := by rw [Measure.condKernel] infer_instance @@ -339,8 +400,7 @@ lemma _root_.MeasureTheory.Measure.compProd_fst_condKernel have h1 : const Unit (Measure.fst ρ) = fst (const Unit ρ) := by ext simp only [fst_apply, Measure.fst, const_apply] - have h2 : prodMkLeft Unit (Measure.condKernel ρ) - = condKernelUnitBorel (const Unit ρ) := by + have h2 : prodMkLeft Unit (Measure.condKernel ρ) = condKernelUnitBorel (const Unit ρ) := by ext simp only [prodMkLeft_apply, Measure.condKernel_apply] rw [Measure.compProd, h1, h2, compProd_fst_condKernelUnitBorel] @@ -407,9 +467,9 @@ def condKernelCountable (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : kernel lemma condKernelCountable_apply (κ : kernel α (β × Ω)) [IsFiniteKernel κ] (p : α × β) : condKernelCountable κ p = (κ p.1).condKernel p.2 := rfl -instance instIsMarkovKernel_condKernelCountable (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : +instance instIsMarkovKernelCondKernelCountable (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : IsMarkovKernel (condKernelCountable κ) := - ⟨fun p ↦ (Measure.instIsMarkovKernel_condKernel (κ p.1)).isProbabilityMeasure p.2⟩ + ⟨fun p ↦ (Measure.instIsMarkovKernelCondKernel (κ p.1)).isProbabilityMeasure p.2⟩ lemma compProd_fst_condKernelCountable (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : fst κ ⊗ₖ condKernelCountable κ = κ := by @@ -437,7 +497,7 @@ irreducible_def condKernel [h : CountableOrCountablyGenerated α β] else letI := h.countableOrCountablyGenerated.resolve_left hα; condKernelBorel κ /-- `condKernel κ` is a Markov kernel. -/ -instance instIsMarkovKernel_condKernel [CountableOrCountablyGenerated α β] +instance instIsMarkovKernelCondKernel [CountableOrCountablyGenerated α β] (κ : kernel α (β × Ω)) [IsFiniteKernel κ] : IsMarkovKernel (condKernel κ) := by rw [condKernel_def] diff --git a/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean index 54ff97d4c54f3..d879838a0e853 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean @@ -668,28 +668,6 @@ lemma compProd_toKernel [IsFiniteKernel κ] [IsSFiniteKernel ν] (hf : IsCondKer ext a s hs rw [kernel.compProd_apply _ _ _ hs, lintegral_toKernel_mem hf a hs] -lemma ae_toKernel_eq_one [IsFiniteKernel κ] [IsSFiniteKernel ν] (hf : IsCondKernelCDF f κ ν) (a : α) - {s : Set ℝ} (hs : MeasurableSet s) (hκs : κ a {x | x.snd ∈ sᶜ} = 0) : - ∀ᵐ b ∂(ν a), hf.toKernel f (a, b) s = 1 := by - have h_eq : ν ⊗ₖ hf.toKernel f = κ := compProd_toKernel hf - have h : κ a {x | x.snd ∈ sᶜ} = (ν ⊗ₖ hf.toKernel f) a {x | x.snd ∈ sᶜ} := by rw [h_eq] - rw [hκs, kernel.compProd_apply] at h - swap; · exact measurable_snd hs.compl - rw [eq_comm, lintegral_eq_zero_iff] at h - swap - · simp only [mem_compl_iff, mem_setOf_eq] - change Measurable ((fun p ↦ hf.toKernel f p {c | c ∉ s}) ∘ (fun b : β ↦ (a, b))) - exact (kernel.measurable_coe _ hs.compl).comp measurable_prod_mk_left - simp only [mem_compl_iff, mem_setOf_eq, kernel.prodMkLeft_apply'] at h - filter_upwards [h] with b hb - change hf.toKernel f (a, b) sᶜ = 0 at hb - rwa [prob_compl_eq_zero_iff hs] at hb - -lemma measurableSet_toKernel_eq_one (hf : IsCondKernelCDF f κ ν) - {s : Set ℝ} (hs : MeasurableSet s) : - MeasurableSet {p | hf.toKernel f p s = 1} := - (kernel.measurable_coe _ hs) (measurableSet_singleton 1) - end end ProbabilityTheory From a28997d19ad23029313abb9e98e1a9aa9674340f Mon Sep 17 00:00:00 2001 From: RemyDegenne Date: Sun, 7 Apr 2024 09:12:16 +0200 Subject: [PATCH 124/129] remove unused import --- Mathlib/Probability/Kernel/Disintegration/Basic.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index f8787a9bb0324..d84663153f987 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -7,7 +7,6 @@ import Mathlib.Probability.Kernel.MeasureCompProd import Mathlib.Probability.Kernel.Disintegration.CondCdf import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.Disintegration.CdfToKernel -import Mathlib.Probability.ConditionalProbability /-! # Disintegration of kernels and measures From 427687ca2318f2a14c6cf1678ff6882b885efcf9 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Tue, 16 Apr 2024 10:44:20 +0200 Subject: [PATCH 125/129] remove some completeness assumptions --- Mathlib/Probability/Kernel/Disintegration/Integral.lean | 6 +++--- Mathlib/Probability/Kernel/IntegralCompProd.lean | 8 +++++--- Mathlib/Probability/Kernel/MeasurableIntegral.lean | 4 +++- Mathlib/Probability/Kernel/MeasureCompProd.lean | 4 ++-- 4 files changed, 13 insertions(+), 9 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index 4644193c0fa57..02aa9b7301509 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -84,7 +84,7 @@ end Lintegral section Integral variable [CountableOrCountablyGenerated α β] {κ : kernel α (β × Ω)} [IsFiniteKernel κ] - {E : Type*} {f : β × Ω → E} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] + {E : Type*} {f : β × Ω → E} [NormedAddCommGroup E] [NormedSpace ℝ E] lemma _root_.MeasureTheory.AEStronglyMeasurable.integral_kernel_condKernel (a : α) (hf : AEStronglyMeasurable f (κ a)) : @@ -188,7 +188,7 @@ end Lintegral section Integral variable {ρ : Measure (β × Ω)} [IsFiniteMeasure ρ] - {E : Type*} {f : β × Ω → E} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] + {E : Type*} {f : β × Ω → E} [NormedAddCommGroup E] [NormedSpace ℝ E] lemma _root_.MeasureTheory.AEStronglyMeasurable.integral_condKernel (hf : AEStronglyMeasurable f ρ) : @@ -238,7 +238,7 @@ open ProbabilityTheory variable {α Ω E F : Type*} {mα : MeasurableSpace α} [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] [NormedAddCommGroup E] [NormedSpace ℝ E] - [CompleteSpace E] [NormedAddCommGroup F] {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] + [NormedAddCommGroup F] {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] theorem AEStronglyMeasurable.ae_integrable_condKernel_iff {f : α × Ω → F} (hf : AEStronglyMeasurable f ρ) : diff --git a/Mathlib/Probability/Kernel/IntegralCompProd.lean b/Mathlib/Probability/Kernel/IntegralCompProd.lean index 02d8491599012..1bcbfc3789852 100644 --- a/Mathlib/Probability/Kernel/IntegralCompProd.lean +++ b/Mathlib/Probability/Kernel/IntegralCompProd.lean @@ -69,7 +69,7 @@ theorem integrable_kernel_prod_mk_left (a : α) {s : Set (β × γ)} (hs : Measu #align probability_theory.integrable_kernel_prod_mk_left ProbabilityTheory.integrable_kernel_prod_mk_left theorem _root_.MeasureTheory.AEStronglyMeasurable.integral_kernel_compProd [NormedSpace ℝ E] - [CompleteSpace E] ⦃f : β × γ → E⦄ (hf : AEStronglyMeasurable f ((κ ⊗ₖ η) a)) : + ⦃f : β × γ → E⦄ (hf : AEStronglyMeasurable f ((κ ⊗ₖ η) a)) : AEStronglyMeasurable (fun x => ∫ y, f (x, y) ∂η (a, x)) (κ a) := ⟨fun x => ∫ y, hf.mk f (x, y) ∂η (a, x), hf.stronglyMeasurable_mk.integral_kernel_prod_right'', by filter_upwards [ae_ae_of_ae_compProd hf.ae_eq_mk] with _ hx using integral_congr_ae hx⟩ @@ -138,7 +138,7 @@ theorem _root_.MeasureTheory.Integrable.integral_norm_compProd ⦃f : β × γ ((integrable_compProd_iff hf.aestronglyMeasurable).mp hf).2 #align measure_theory.integrable.integral_norm_comp_prod MeasureTheory.Integrable.integral_norm_compProd -theorem _root_.MeasureTheory.Integrable.integral_compProd [NormedSpace ℝ E] [CompleteSpace E] +theorem _root_.MeasureTheory.Integrable.integral_compProd [NormedSpace ℝ E] ⦃f : β × γ → E⦄ (hf : Integrable f ((κ ⊗ₖ η) a)) : Integrable (fun x => ∫ y, f (x, y) ∂η (a, x)) (κ a) := Integrable.mono hf.integral_norm_compProd hf.aestronglyMeasurable.integral_kernel_compProd <| @@ -152,7 +152,7 @@ theorem _root_.MeasureTheory.Integrable.integral_compProd [NormedSpace ℝ E] [C /-! ### Bochner integral -/ -variable [NormedSpace ℝ E] [CompleteSpace E] {E' : Type*} [NormedAddCommGroup E'] +variable [NormedSpace ℝ E] {E' : Type*} [NormedAddCommGroup E'] [CompleteSpace E'] [NormedSpace ℝ E'] theorem kernel.integral_fn_integral_add ⦃f g : β × γ → E⦄ (F : E → E') @@ -244,6 +244,8 @@ theorem kernel.continuous_integral_integral : theorem integral_compProd : ∀ {f : β × γ → E} (_ : Integrable f ((κ ⊗ₖ η) a)), ∫ z, f z ∂(κ ⊗ₖ η) a = ∫ x, ∫ y, f (x, y) ∂η (a, x) ∂κ a := by + by_cases hE : CompleteSpace E; swap + · simp [integral, hE] apply Integrable.induction · intro c s hs h2s simp_rw [integral_indicator hs, ← indicator_comp_right, Function.comp, diff --git a/Mathlib/Probability/Kernel/MeasurableIntegral.lean b/Mathlib/Probability/Kernel/MeasurableIntegral.lean index 58eba3e4f0e72..150b967a55c19 100644 --- a/Mathlib/Probability/Kernel/MeasurableIntegral.lean +++ b/Mathlib/Probability/Kernel/MeasurableIntegral.lean @@ -253,12 +253,14 @@ open ProbabilityTheory ProbabilityTheory.kernel namespace MeasureTheory -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] [IsSFiniteKernel κ] +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [IsSFiniteKernel κ] [IsSFiniteKernel η] theorem StronglyMeasurable.integral_kernel_prod_right ⦃f : α → β → E⦄ (hf : StronglyMeasurable (uncurry f)) : StronglyMeasurable fun x => ∫ y, f x y ∂κ x := by classical + by_cases hE : CompleteSpace E; swap + · simp [integral, hE, stronglyMeasurable_const] borelize E haveI : TopologicalSpace.SeparableSpace (range (uncurry f) ∪ {0} : Set E) := hf.separableSpace_range_union_singleton diff --git a/Mathlib/Probability/Kernel/MeasureCompProd.lean b/Mathlib/Probability/Kernel/MeasureCompProd.lean index 9ef775fe51976..fd8d8d52fb4c6 100644 --- a/Mathlib/Probability/Kernel/MeasureCompProd.lean +++ b/Mathlib/Probability/Kernel/MeasureCompProd.lean @@ -113,14 +113,14 @@ lemma integrable_compProd_iff [SFinite μ] [IsSFiniteKernel κ] {E : Type*} [Nor rfl lemma integral_compProd [SFinite μ] [IsSFiniteKernel κ] {E : Type*} - [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] + [NormedAddCommGroup E] [NormedSpace ℝ E] {f : α × β → E} (hf : Integrable f (μ ⊗ₘ κ)) : ∫ x, f x ∂(μ ⊗ₘ κ) = ∫ a, ∫ b, f (a, b) ∂(κ a) ∂μ := by rw [compProd, ProbabilityTheory.integral_compProd hf] simp lemma set_integral_compProd [SFinite μ] [IsSFiniteKernel κ] {E : Type*} - [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] + [NormedAddCommGroup E] [NormedSpace ℝ E] {s : Set α} (hs : MeasurableSet s) {t : Set β} (ht : MeasurableSet t) {f : α × β → E} (hf : IntegrableOn f (s ×ˢ t) (μ ⊗ₘ κ)) : ∫ x in s ×ˢ t, f x ∂(μ ⊗ₘ κ) = ∫ a in s, ∫ b in t, f (a, b) ∂(κ a) ∂μ := by From 94bf67e7742c08ca09dfe049e5ca36b576da144b Mon Sep 17 00:00:00 2001 From: sgouezel Date: Tue, 16 Apr 2024 11:06:14 +0200 Subject: [PATCH 126/129] lint --- Mathlib/Probability/Kernel/Condexp.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Probability/Kernel/Condexp.lean b/Mathlib/Probability/Kernel/Condexp.lean index 1253d438ac867..e3008b03845ca 100644 --- a/Mathlib/Probability/Kernel/Condexp.lean +++ b/Mathlib/Probability/Kernel/Condexp.lean @@ -97,7 +97,7 @@ theorem stronglyMeasurable_condexpKernel {s : Set Ω} (hs : MeasurableSet s) : Measurable.stronglyMeasurable (measurable_condexpKernel hs) theorem _root_.MeasureTheory.AEStronglyMeasurable.integral_condexpKernel [NormedSpace ℝ F] - [CompleteSpace F] (hf : AEStronglyMeasurable f μ) : + (hf : AEStronglyMeasurable f μ) : AEStronglyMeasurable (fun ω => ∫ y, f y ∂condexpKernel μ m ω) μ := by simp_rw [condexpKernel_apply_eq_condDistrib] exact AEStronglyMeasurable.integral_condDistrib @@ -105,7 +105,7 @@ theorem _root_.MeasureTheory.AEStronglyMeasurable.integral_condexpKernel [Normed (hf.comp_snd_map_prod_id inf_le_right) #align measure_theory.ae_strongly_measurable.integral_condexp_kernel MeasureTheory.AEStronglyMeasurable.integral_condexpKernel -theorem aestronglyMeasurable'_integral_condexpKernel [NormedSpace ℝ F] [CompleteSpace F] +theorem aestronglyMeasurable'_integral_condexpKernel [NormedSpace ℝ F] (hf : AEStronglyMeasurable f μ) : AEStronglyMeasurable' m (fun ω => ∫ y, f y ∂condexpKernel μ m ω) μ := by rw [condexpKernel] @@ -139,7 +139,7 @@ theorem _root_.MeasureTheory.Integrable.integral_norm_condexpKernel (hf_int : In #align measure_theory.integrable.integral_norm_condexp_kernel MeasureTheory.Integrable.integral_norm_condexpKernel theorem _root_.MeasureTheory.Integrable.norm_integral_condexpKernel [NormedSpace ℝ F] - [CompleteSpace F] (hf_int : Integrable f μ) : + (hf_int : Integrable f μ) : Integrable (fun ω => ‖∫ y, f y ∂condexpKernel μ m ω‖) μ := by rw [condexpKernel] exact Integrable.norm_integral_condDistrib @@ -147,7 +147,7 @@ theorem _root_.MeasureTheory.Integrable.norm_integral_condexpKernel [NormedSpace (hf_int.comp_snd_map_prod_id (inf_le_right : m ⊓ mΩ ≤ mΩ)) #align measure_theory.integrable.norm_integral_condexp_kernel MeasureTheory.Integrable.norm_integral_condexpKernel -theorem _root_.MeasureTheory.Integrable.integral_condexpKernel [NormedSpace ℝ F] [CompleteSpace F] +theorem _root_.MeasureTheory.Integrable.integral_condexpKernel [NormedSpace ℝ F] (hf_int : Integrable f μ) : Integrable (fun ω => ∫ y, f y ∂condexpKernel μ m ω) μ := by rw [condexpKernel] From 51952fc18af913cb927efbdf814700250cf50fc6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Tue, 16 Apr 2024 12:48:08 +0200 Subject: [PATCH 127/129] Apply suggestions from code review Co-authored-by: sgouezel --- Mathlib/Probability/Kernel/Disintegration/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index d84663153f987..8d317a3971785 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -43,7 +43,7 @@ The first step (building the measurable function on `ℚ`) is done differently d The conditional kernel is defined under the typeclass assumption `CountableOrCountablyGenerated α β`, which encodes the property -`(Countable α ∧ MeasurableSingletonClass α) ∨ CountablyGenerated β`. +`Countable α ∨ CountablyGenerated β`. Properties of integrals involving `condKernel` are collated in the file `Integral.lean`. The conditional kernel is unique (almost everywhere w.r.t. `fst κ`): this is proved in the file @@ -97,7 +97,7 @@ property on `ℝ` to all these spaces. -/ open Classical in /-- Auxiliary definition for `ProbabilityTheory.kernel.condKernel`. -A Borel space `Ω` embedds measurably into `ℝ` (with embedding `e`), hence we can get a `kernel α Ω` +A Borel space `Ω` embeds measurably into `ℝ` (with embedding `e`), hence we can get a `kernel α Ω` from a `kernel α ℝ` by taking the comap by `e`. Here we take the comap of a modification of `η : kernel α ℝ`, useful when `η a` is a probability measure with all its mass on `range e` almost everywhere with respect to some measure and we want to From 4c40319ebbb11ae0e230f4140d3c79885c53d920 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Tue, 16 Apr 2024 13:11:13 +0200 Subject: [PATCH 128/129] create Real section at the top of Disintegration/Basic --- .../Kernel/Disintegration/Basic.lean | 178 +++++++++--------- 1 file changed, 92 insertions(+), 86 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 8d317a3971785..8ff57b2ae1540 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -88,6 +88,98 @@ variable {α β γ Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace {mγ : MeasurableSpace γ} [MeasurableSpace.CountablyGenerated γ] [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] +section Real + +/-! ### Disintegration of kernels from `α` to `γ × ℝ` for countably generated `γ` -/ + +lemma isRatCondKernelCDFAux_density_Iic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsRatCondKernelCDFAux (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) κ (fst κ) + where + measurable := measurable_pi_iff.mpr fun _ ↦ measurable_density κ (fst κ) measurableSet_Iic + mono' a q r hqr := + ae_of_all _ fun c ↦ density_mono_set le_rfl a c (Iic_subset_Iic.mpr (by exact_mod_cast hqr)) + nonneg' a q := ae_of_all _ fun c ↦ density_nonneg le_rfl _ _ _ + le_one' a q := ae_of_all _ fun c ↦ density_le_one le_rfl _ _ _ + tendsto_integral_of_antitone a s hs_anti hs_tendsto := by + let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) + refine tendsto_integral_density_of_antitone le_rfl a s' ?_ ?_ (fun _ ↦ measurableSet_Iic) + · refine fun i j hij ↦ Iic_subset_Iic.mpr ?_ + exact mod_cast hs_anti hij + · ext x + simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, s'] + rw [tendsto_atTop_atBot] at hs_tendsto + have ⟨q, hq⟩ := exists_rat_lt x + obtain ⟨i, hi⟩ := hs_tendsto q + refine ⟨i, lt_of_le_of_lt ?_ hq⟩ + exact mod_cast hi i le_rfl + tendsto_integral_of_monotone a s hs_mono hs_tendsto := by + rw [fst_apply' _ _ MeasurableSet.univ] + let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) + refine tendsto_integral_density_of_monotone (le_rfl : fst κ ≤ fst κ) + a s' ?_ ?_ (fun _ ↦ measurableSet_Iic) + · exact fun i j hij ↦ Iic_subset_Iic.mpr (by exact mod_cast hs_mono hij) + · ext x + simp only [mem_iUnion, mem_Iic, mem_univ, iff_true] + rw [tendsto_atTop_atTop] at hs_tendsto + have ⟨q, hq⟩ := exists_rat_gt x + obtain ⟨i, hi⟩ := hs_tendsto q + refine ⟨i, hq.le.trans ?_⟩ + exact mod_cast hi i le_rfl + integrable a q := integrable_density le_rfl a measurableSet_Iic + set_integral a A hA q := set_integral_density le_rfl a measurableSet_Iic hA + +/-- Taking the kernel density of intervals `Iic q` for `q : ℚ` gives a function with the property +`isRatCondKernelCDF`. -/ +lemma isRatCondKernelCDF_density_Iic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsRatCondKernelCDF (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) κ (fst κ) := + (isRatCondKernelCDFAux_density_Iic κ).isRatCondKernelCDF + +/-- The conditional kernel CDF of a kernel `κ : kernel α (γ × ℝ)`, where `γ` is countably generated. +-/ +noncomputable +def condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : α × γ → StieltjesFunction := + stieltjesOfMeasurableRat (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) + (isRatCondKernelCDF_density_Iic κ).measurable + +lemma isCondKernelCDF_condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsCondKernelCDF (condKernelCDF κ) κ (fst κ) := + isCondKernelCDF_stieltjesOfMeasurableRat (isRatCondKernelCDF_density_Iic κ) + +/-- Auxiliary definition for `ProbabilityTheory.kernel.condKernel`. +A conditional kernel for `κ : kernel α (γ × ℝ)` where `γ` is countably generated. -/ +noncomputable +def condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : kernel (α × γ) ℝ := + (isCondKernelCDF_condKernelCDF κ).toKernel + +instance instIsMarkovKernelCondKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsMarkovKernel (condKernelReal κ) := by + rw [condKernelReal] + infer_instance + +lemma compProd_fst_condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : + fst κ ⊗ₖ condKernelReal κ = κ := by + rw [condKernelReal, compProd_toKernel] + +/-- Auxiliary definition for `MeasureTheory.Measure.condKernel` and +`ProbabilityTheory.kernel.condKernel`. +A conditional kernel for `κ : kernel Unit (α × ℝ)`. -/ +noncomputable +def condKernelUnitReal (κ : kernel Unit (α × ℝ)) [IsFiniteKernel κ] : kernel (Unit × α) ℝ := + (isCondKernelCDF_condCDF (κ ())).toKernel + +instance instIsMarkovKernelCondKernelUnitReal (κ : kernel Unit (α × ℝ)) [IsFiniteKernel κ] : + IsMarkovKernel (condKernelUnitReal κ) := by + rw [condKernelUnitReal] + infer_instance + +lemma compProd_fst_condKernelUnitReal (κ : kernel Unit (α × ℝ)) [IsFiniteKernel κ] : + fst κ ⊗ₖ condKernelUnitReal κ = κ := by + rw [condKernelUnitReal, compProd_toKernel] + ext a + simp + +end Real + section BorelSnd /-! ### Disintegration of kernels on standard Borel spaces @@ -238,74 +330,6 @@ section CountablyGenerated open ProbabilityTheory.kernel -lemma isRatCondKernelCDFAux_density_Iic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsRatCondKernelCDFAux (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) κ (fst κ) - where - measurable := measurable_pi_iff.mpr fun _ ↦ measurable_density κ (fst κ) measurableSet_Iic - mono' a q r hqr := - ae_of_all _ fun c ↦ density_mono_set le_rfl a c (Iic_subset_Iic.mpr (by exact_mod_cast hqr)) - nonneg' a q := ae_of_all _ fun c ↦ density_nonneg le_rfl _ _ _ - le_one' a q := ae_of_all _ fun c ↦ density_le_one le_rfl _ _ _ - tendsto_integral_of_antitone a s hs_anti hs_tendsto := by - let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) - refine tendsto_integral_density_of_antitone le_rfl a s' ?_ ?_ (fun _ ↦ measurableSet_Iic) - · refine fun i j hij ↦ Iic_subset_Iic.mpr ?_ - exact mod_cast hs_anti hij - · ext x - simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, s'] - rw [tendsto_atTop_atBot] at hs_tendsto - have ⟨q, hq⟩ := exists_rat_lt x - obtain ⟨i, hi⟩ := hs_tendsto q - refine ⟨i, lt_of_le_of_lt ?_ hq⟩ - exact mod_cast hi i le_rfl - tendsto_integral_of_monotone a s hs_mono hs_tendsto := by - rw [fst_apply' _ _ MeasurableSet.univ] - let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) - refine tendsto_integral_density_of_monotone (le_rfl : fst κ ≤ fst κ) - a s' ?_ ?_ (fun _ ↦ measurableSet_Iic) - · exact fun i j hij ↦ Iic_subset_Iic.mpr (by exact mod_cast hs_mono hij) - · ext x - simp only [mem_iUnion, mem_Iic, mem_univ, iff_true] - rw [tendsto_atTop_atTop] at hs_tendsto - have ⟨q, hq⟩ := exists_rat_gt x - obtain ⟨i, hi⟩ := hs_tendsto q - refine ⟨i, hq.le.trans ?_⟩ - exact mod_cast hi i le_rfl - integrable a q := integrable_density le_rfl a measurableSet_Iic - set_integral a A hA q := set_integral_density le_rfl a measurableSet_Iic hA - -/-- Taking the kernel density of intervals `Iic q` for `q : ℚ` gives a function with the property -`isRatCondKernelCDF`. -/ -lemma isRatCondKernelCDF_density_Iic (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsRatCondKernelCDF (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) κ (fst κ) := - (isRatCondKernelCDFAux_density_Iic κ).isRatCondKernelCDF - -/-- The conditional kernel CDF of a kernel `κ : kernel α (γ × ℝ)`, where `γ` is countably generated. --/ -noncomputable -def condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : α × γ → StieltjesFunction := - stieltjesOfMeasurableRat (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) - (isRatCondKernelCDF_density_Iic κ).measurable - -lemma isCondKernelCDF_condKernelCDF (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsCondKernelCDF (condKernelCDF κ) κ (fst κ) := - isCondKernelCDF_stieltjesOfMeasurableRat (isRatCondKernelCDF_density_Iic κ) - -/-- Auxiliary definition for `ProbabilityTheory.kernel.condKernel`. -A conditional kernel for `κ : kernel α (γ × ℝ)` where `γ` is countably generated. -/ -noncomputable -def condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : kernel (α × γ) ℝ := - (isCondKernelCDF_condKernelCDF κ).toKernel - -instance instIsMarkovKernelCondKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsMarkovKernel (condKernelReal κ) := by - rw [condKernelReal] - infer_instance - -lemma compProd_fst_condKernelReal (κ : kernel α (γ × ℝ)) [IsFiniteKernel κ] : - fst κ ⊗ₖ condKernelReal κ = κ := by - rw [condKernelReal, compProd_toKernel] - /-- Auxiliary definition for `ProbabilityTheory.kernel.condKernel`. A conditional kernel for `κ : kernel α (γ × Ω)` where `γ` is countably generated and `Ω` is standard Borel. -/ @@ -329,24 +353,6 @@ end CountablyGenerated section Unit -/-- Auxiliary definition for `MeasureTheory.Measure.condKernel` and -`ProbabilityTheory.kernel.condKernel`. -A conditional kernel for `κ : kernel Unit (α × ℝ)`. -/ -noncomputable -def condKernelUnitReal (κ : kernel Unit (α × ℝ)) [IsFiniteKernel κ] : kernel (Unit × α) ℝ := - (isCondKernelCDF_condCDF (κ ())).toKernel - -instance instIsMarkovKernelCondKernelUnitReal (κ : kernel Unit (α × ℝ)) [IsFiniteKernel κ] : - IsMarkovKernel (condKernelUnitReal κ) := by - rw [condKernelUnitReal] - infer_instance - -lemma compProd_fst_condKernelUnitReal (κ : kernel Unit (α × ℝ)) [IsFiniteKernel κ] : - fst κ ⊗ₖ condKernelUnitReal κ = κ := by - rw [condKernelUnitReal, compProd_toKernel] - ext a - simp - /-- Auxiliary definition for `MeasureTheory.Measure.condKernel` and `ProbabilityTheory.kernel.condKernel`. A conditional kernel for `κ : kernel Unit (α × Ω)` where `Ω` is standard Borel. -/ From 0081cac92b3e6da36f63921248b1baa9b154d1ae Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Tue, 16 Apr 2024 13:12:38 +0200 Subject: [PATCH 129/129] irreducible_def --- Mathlib/Probability/Kernel/Disintegration/Basic.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 8ff57b2ae1540..4d56a6bd7318f 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -382,14 +382,15 @@ variable {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] /-- Conditional kernel of a measure on a product space: a Markov kernel such that `ρ = ρ.fst ⊗ₘ ρ.condKernel` (see `MeasureTheory.Measure.compProd_fst_condKernel`). -/ noncomputable -def _root_.MeasureTheory.Measure.condKernel (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : +irreducible_def _root_.MeasureTheory.Measure.condKernel (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : kernel α Ω := comap (condKernelUnitBorel (const Unit ρ)) (fun a ↦ ((), a)) measurable_prod_mk_left #align measure_theory.measure.cond_kernel MeasureTheory.Measure.condKernel lemma _root_.MeasureTheory.Measure.condKernel_apply (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] (a : α) : - ρ.condKernel a = condKernelUnitBorel (const Unit ρ) ((), a) := rfl + ρ.condKernel a = condKernelUnitBorel (const Unit ρ) ((), a) := by + rw [Measure.condKernel]; rfl instance _root_.MeasureTheory.Measure.instIsMarkovKernelCondKernel (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : IsMarkovKernel ρ.condKernel := by