Skip to content

Commit 042888f

Browse files
kex-yRemyDegenne
andcommitted
feat(Probability/Kernel/Disintegration): Uniqueness of the disintegration kernel (#6110)
Co-authored-by: JasonKYi <kexing.ying19@imperial.ac.uk> Co-authored-by: RemyDegenne <Remydegenne@gmail.com>
1 parent d8905e7 commit 042888f

File tree

3 files changed

+157
-1
lines changed

3 files changed

+157
-1
lines changed

Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean

Lines changed: 24 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -423,6 +423,10 @@ theorem ae_all_iff {ι : Sort*} [Countable ι] {p : α → ι → Prop} :
423423
eventually_countable_forall
424424
#align measure_theory.ae_all_iff MeasureTheory.ae_all_iff
425425

426+
theorem all_ae_of {ι : Sort _} {p : α → ι → Prop} (hp : ∀ᵐ a ∂μ, ∀ i, p a i) (i : ι) :
427+
∀ᵐ a ∂μ, p a i := by
428+
filter_upwards [hp] with a ha using ha i
429+
426430
theorem ae_ball_iff {S : Set ι} (hS : S.Countable) {p : ∀ (_x : α), ∀ i ∈ S, Prop} :
427431
(∀ᵐ x ∂μ, ∀ i (hi : i ∈ S), p x i hi) ↔ ∀ i (hi : i ∈ S), ∀ᵐ x ∂μ, p x i hi :=
428432
eventually_countable_ball hS
@@ -559,6 +563,26 @@ theorem inter_ae_eq_empty_of_ae_eq_empty_right (h : t =ᵐ[μ] (∅ : Set α)) :
559563
rw [inter_empty]
560564
#align measure_theory.inter_ae_eq_empty_of_ae_eq_empty_right MeasureTheory.inter_ae_eq_empty_of_ae_eq_empty_right
561565

566+
/-- Given a predicate on `β` and `set α` where both `α` and `β` are measurable spaces, if the
567+
predicate holds for almost every `x : β` and
568+
- `∅ : set α`
569+
- a family of sets generating the σ-algebra of `α`
570+
Moreover, if for almost every `x : β`, the predicate is closed under complements and countable
571+
disjoint unions, then the predicate holds for almost every `x : β` and all measurable sets of `α`.
572+
573+
This is an AE version of `MeasurableSpace.induction_on_inter` where the condition is dependent
574+
on a measurable space `β`. -/
575+
theorem _root_.MeasurableSpace.ae_induction_on_inter {β} [MeasurableSpace β] {μ : Measure β}
576+
{C : β → Set α → Prop} {s : Set (Set α)} [m : MeasurableSpace α]
577+
(h_eq : m = MeasurableSpace.generateFrom s)
578+
(h_inter : IsPiSystem s) (h_empty : ∀ᵐ x ∂μ, C x ∅) (h_basic : ∀ᵐ x ∂μ, ∀ t ∈ s, C x t)
579+
(h_compl : ∀ᵐ x ∂μ, ∀ t, MeasurableSet t → C x t → C x tᶜ)
580+
(h_union : ∀ᵐ x ∂μ, ∀ f : ℕ → Set α,
581+
Pairwise (Disjoint on f) → (∀ i, MeasurableSet (f i)) → (∀ i, C x (f i)) → C x (⋃ i, f i)) :
582+
∀ᵐ x ∂μ, ∀ ⦃t⦄, MeasurableSet t → C x t := by
583+
filter_upwards [h_empty, h_basic, h_compl, h_union] with x hx_empty hx_basic hx_compl hx_union
584+
using MeasurableSpace.induction_on_inter h_eq h_inter hx_empty hx_basic hx_compl hx_union
585+
562586
@[to_additive]
563587
theorem _root_.Set.mulIndicator_ae_eq_one {M : Type*} [One M] {f : α → M} {s : Set α} :
564588
s.mulIndicator f =ᵐ[μ] 1 ↔ μ (s ∩ f.mulSupport) = 0 := by

Mathlib/Probability/Kernel/CondDistrib.lean

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -106,6 +106,22 @@ theorem aestronglyMeasurable'_integral_condDistrib (hX : AEMeasurable X μ) (hY
106106

107107
end Measurability
108108

109+
/-- `condDistrib` is a.e. uniquely defined as the kernel satisfying the defining property of
110+
`condKernel`. -/
111+
theorem condDistrib_ae_eq_of_measure_eq_compProd (hX : Measurable X) (hY : Measurable Y)
112+
(κ : kernel β Ω) [IsFiniteKernel κ]
113+
(hκ : μ.map (fun x => (X x, Y x)) =
114+
(kernel.const Unit (μ.map X) ⊗ₖ kernel.prodMkLeft Unit κ) ()) :
115+
∀ᵐ x ∂μ.map X, κ x = condDistrib Y X μ x := by
116+
have heq : μ.map X = (μ.map (fun x => (X x, Y x))).fst
117+
· ext s hs
118+
rw [Measure.map_apply hX hs, Measure.fst_apply hs, Measure.map_apply]
119+
exacts [rfl, Measurable.prod hX hY, measurable_fst hs]
120+
rw [heq, condDistrib]
121+
refine' eq_condKernel_of_measure_eq_compProd _ _ _
122+
convert hκ
123+
exact heq.symm
124+
109125
section Integrability
110126

111127
theorem integrable_toReal_condDistrib (hX : AEMeasurable X μ) (hs : MeasurableSet s) :

Mathlib/Probability/Kernel/Disintegration.lean

Lines changed: 117 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,7 @@
11
/-
22
Copyright (c) 2023 Rémy Degenne. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
4-
Authors: Rémy Degenne
4+
Authors: Rémy Degenne, Kexing Ying
55
-/
66
import Mathlib.Probability.Kernel.CondCdf
77
import Mathlib.MeasureTheory.Constructions.Polish
@@ -42,6 +42,7 @@ function `condCdf ρ a` (the conditional cumulative distribution function).
4242
`kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ ρ.condKernel)`
4343
* `ProbabilityTheory.measure_eq_compProd`:
4444
`ρ = ((kernel.const Unit ρ.fst) ⊗ₖ (kernel.prodMkLeft Unit ρ.condKernel)) ()`
45+
* `ProbabilityTheory.eq_condKernel_of_measure_eq_compProd`: a.e. uniqueness of `condKernel`
4546
4647
-/
4748

@@ -477,6 +478,121 @@ theorem set_integral_condKernel_univ_left {ρ : Measure (α × Ω)} [IsFiniteMea
477478

478479
end IntegralCondKernel
479480

481+
section Unique
482+
483+
/-! ### Uniqueness
484+
485+
This section will prove that the conditional kernel is unique almost everywhere. -/
486+
487+
/-- A s-finite kernel which satisfy the disintegration property of the given measure `ρ` is almost
488+
everywhere equal to the disintegration kernel of `ρ` when evaluated on a measurable set.
489+
490+
This theorem in the case of finite kernels is weaker than `eq_condKernel_of_measure_eq_compProd`
491+
which asserts that the kernels are equal almost everywhere and not just on a given measurable
492+
set. -/
493+
theorem eq_condKernel_of_measure_eq_compProd' (κ : kernel α Ω) [IsSFiniteKernel κ]
494+
(hκ : ρ = (kernel.const Unit ρ.fst ⊗ₖ kernel.prodMkLeft Unit κ) ())
495+
{s : Set Ω} (hs : MeasurableSet s) :
496+
∀ᵐ x ∂ρ.fst, κ x s = ρ.condKernel x s := by
497+
refine' ae_eq_of_forall_set_lintegral_eq_of_sigmaFinite
498+
(kernel.measurable_coe κ hs) (kernel.measurable_coe ρ.condKernel hs) _
499+
intros t ht _
500+
conv_rhs => rw [set_lintegral_condKernel_eq_measure_prod _ ht hs, hκ]
501+
simp only [kernel.compProd_apply _ _ _ (ht.prod hs), kernel.prodMkLeft_apply, Set.mem_prod,
502+
kernel.lintegral_const, ← lintegral_indicator _ ht]
503+
congr; ext x
504+
by_cases hx : x ∈ t
505+
all_goals simp [hx]
506+
507+
-- This lemma establishes uniqueness of the disintegration kernel on ℝ
508+
lemma eq_condKernel_of_measure_eq_compProd_real (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ]
509+
(κ : kernel α ℝ) [IsFiniteKernel κ]
510+
(hκ : ρ = (kernel.const Unit ρ.fst ⊗ₖ kernel.prodMkLeft Unit κ) ()) :
511+
∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by
512+
have huniv : ∀ᵐ x ∂ρ.fst, κ x Set.univ = ρ.condKernel x Set.univ :=
513+
eq_condKernel_of_measure_eq_compProd' ρ κ hκ MeasurableSet.univ
514+
suffices : ∀ᵐ x ∂ρ.fst, ∀ ⦃t⦄, MeasurableSet t → κ x t = ρ.condKernel x t
515+
· filter_upwards [this] with x hx
516+
ext t ht; exact hx ht
517+
apply MeasurableSpace.ae_induction_on_inter Real.borel_eq_generateFrom_Iic_rat
518+
Real.isPiSystem_Iic_rat
519+
· simp only [OuterMeasure.empty', Filter.eventually_true]
520+
· simp only [iUnion_singleton_eq_range, mem_range, forall_exists_index, forall_apply_eq_imp_iff']
521+
exact ae_all_iff.2 <| fun q => eq_condKernel_of_measure_eq_compProd' ρ κ hκ measurableSet_Iic
522+
· filter_upwards [huniv] with x hxuniv t ht heq
523+
rw [measure_compl ht <| measure_ne_top _ _, heq, hxuniv, measure_compl ht <| measure_ne_top _ _]
524+
· refine' ae_of_all _ (fun x f hdisj hf heq => _)
525+
rw [measure_iUnion hdisj hf, measure_iUnion hdisj hf]
526+
exact tsum_congr heq
527+
528+
/-- A finite kernel which satisfies the disintegration property is almost everywhere equal to the
529+
disintegration kernel. -/
530+
theorem eq_condKernel_of_measure_eq_compProd (κ : kernel α Ω) [IsFiniteKernel κ]
531+
(hκ : ρ = (kernel.const Unit ρ.fst ⊗ₖ kernel.prodMkLeft Unit κ) ()) :
532+
∀ᵐ x ∂ρ.fst, κ x = ρ.condKernel x := by
533+
-- The idea is to transporting the question to `ℝ` from `Ω` using `exists_measurableEmbedding_real`
534+
-- and then constructing a measure on `α × ℝ` using the obtained measurable embedding
535+
obtain ⟨f, hf⟩ := exists_measurableEmbedding_real Ω
536+
set ρ' : Measure (α × ℝ) := ρ.map (Prod.map id f) with hρ'def
537+
have hρ' : ρ'.fst = ρ.fst
538+
· ext s hs
539+
rw [hρ'def, Measure.fst_apply, Measure.fst_apply, Measure.map_apply]
540+
exacts [rfl, Measurable.prod measurable_fst <| hf.measurable.comp measurable_snd,
541+
measurable_fst hs, hs, hs]
542+
have hρ'' : ∀ᵐ x ∂ρ'.fst, kernel.map κ f hf.measurable x = ρ'.condKernel x
543+
· refine' eq_condKernel_of_measure_eq_compProd_real ρ' (kernel.map κ f hf.measurable) _
544+
ext s hs
545+
simp only [Measure.map_apply (measurable_id.prod_map hf.measurable) hs]
546+
conv_lhs => congr; rw [hκ]
547+
rw [kernel.compProd_apply _ _ _ hs, kernel.compProd_apply _ _ _
548+
(measurable_id.prod_map hf.measurable hs), (_ : (ρ.map (Prod.map id f)).fst = ρ.fst)]
549+
· congr
550+
ext x
551+
simp only [Set.mem_preimage, Prod_map, id_eq, kernel.prodMkLeft_apply, kernel.map_apply]
552+
rw [Measure.map_apply hf.measurable]
553+
· rfl
554+
· exact measurable_prod_mk_left hs
555+
· rw [Measure.fst_map_prod_mk]
556+
· simp only [Prod_map, id_eq]; rfl
557+
· exact (hf.measurable.comp measurable_snd)
558+
rw [hρ'] at hρ''
559+
suffices : ∀ᵐ x ∂ρ.fst, ∀ s, MeasurableSet s →
560+
((ρ.map (Prod.map id f)).condKernel x) s = (ρ.condKernel x) (f ⁻¹' s)
561+
· filter_upwards [hρ'', this] with x hx h
562+
rw [kernel.map_apply] at hx
563+
ext s hs
564+
rw [← Set.preimage_image_eq s hf.injective,
565+
← Measure.map_apply hf.measurable <| hf.measurableSet_image.2 hs, hx,
566+
h _ <| hf.measurableSet_image.2 hs]
567+
have hprod : (ρ.map (Prod.map id f)).fst = ρ.fst
568+
· ext s hs
569+
rw [Measure.fst_apply hs,
570+
Measure.map_apply (measurable_id.prod_map hf.measurable) (measurable_fst hs),
571+
← Set.preimage_comp, Measure.fst_apply hs]
572+
rfl
573+
suffices : ρ.map (Prod.map id f) =
574+
(kernel.const Unit (ρ.map (Prod.map id f)).fst ⊗ₖ
575+
kernel.prodMkLeft Unit (kernel.map (Measure.condKernel ρ) f hf.measurable)) ()
576+
· have heq := eq_condKernel_of_measure_eq_compProd_real _ _ this
577+
rw [hprod] at heq
578+
filter_upwards [heq] with x hx s hs
579+
rw [← hx, kernel.map_apply, Measure.map_apply hf.measurable hs]
580+
ext s hs
581+
have hinteq : ∀ x, (ρ.condKernel x).map f {c | (x, c) ∈ s} =
582+
ρ.condKernel x {c | (x, c) ∈ Prod.map id f ⁻¹' s}
583+
· intro x
584+
rw [Measure.map_apply hf.measurable]
585+
· rfl
586+
· exact measurable_prod_mk_left hs
587+
simp only [hprod, kernel.compProd_apply _ _ _ hs, kernel.prodMkLeft_apply,
588+
kernel.map_apply _ hf.measurable, hinteq, Set.mem_preimage, Prod_map, id_eq,
589+
kernel.lintegral_const]
590+
rw [Measure.map_apply (measurable_id.prod_map hf.measurable) hs, ← lintegral_condKernel_mem]
591+
· rfl
592+
· exact measurable_id.prod_map hf.measurable hs
593+
594+
end Unique
595+
480596
end Polish
481597

482598
end ProbabilityTheory

0 commit comments

Comments
 (0)