|
1 | 1 | /-
|
2 | 2 | Copyright (c) 2023 Rémy Degenne. All rights reserved.
|
3 | 3 | Released under Apache 2.0 license as described in the file LICENSE.
|
4 |
| -Authors: Rémy Degenne |
| 4 | +Authors: Rémy Degenne, Kexing Ying |
5 | 5 | -/
|
6 | 6 | import Mathlib.Probability.Kernel.CondCdf
|
7 | 7 | import Mathlib.MeasureTheory.Constructions.Polish
|
@@ -42,6 +42,7 @@ function `condCdf ρ a` (the conditional cumulative distribution function).
|
42 | 42 | `kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prodMkLeft γ ρ.condKernel)`
|
43 | 43 | * `ProbabilityTheory.measure_eq_compProd`:
|
44 | 44 | `ρ = ((kernel.const Unit ρ.fst) ⊗ₖ (kernel.prodMkLeft Unit ρ.condKernel)) ()`
|
| 45 | +* `ProbabilityTheory.eq_condKernel_of_measure_eq_compProd`: a.e. uniqueness of `condKernel` |
45 | 46 |
|
46 | 47 | -/
|
47 | 48 |
|
@@ -477,6 +478,121 @@ theorem set_integral_condKernel_univ_left {ρ : Measure (α × Ω)} [IsFiniteMea
|
477 | 478 |
|
478 | 479 | end IntegralCondKernel
|
479 | 480 |
|
| 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 | + |
480 | 596 | end Polish
|
481 | 597 |
|
482 | 598 | end ProbabilityTheory
|
|
0 commit comments