Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 00abe06

Browse files
committed
feat(probability/kernel/cond_distrib): regular conditional probability distributions (#19090)
We define the regular conditional probability distribution `cond_distrib Y X μ` of `Y : α → Ω` given `X : α → β`, where `Ω` is a standard Borel space. This is a `kernel β Ω` such that for almost all `a`, for all measurable set `s`, `cond_distrib Y X μ (X a) s` is equal to the conditional expectation `μ⟦Y ⁻¹' s | mβ.comap X⟧` evaluated at `a`. Also define the above notation for the conditional expectation of the indicator of a set. Co-authored-by: Rémy Degenne <remydegenne@gmail.com>
1 parent a99f852 commit 00abe06

File tree

5 files changed

+534
-0
lines changed

5 files changed

+534
-0
lines changed

src/measure_theory/constructions/prod/basic.lean

Lines changed: 30 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -749,6 +749,21 @@ instance [is_finite_measure ρ] : is_finite_measure ρ.fst := by { rw fst, apply
749749
instance [is_probability_measure ρ] : is_probability_measure ρ.fst :=
750750
{ measure_univ := by { rw fst_univ, exact measure_univ, } }
751751

752+
lemma fst_map_prod_mk₀ {X : α → β} {Y : α → γ} {μ : measure α}
753+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ) :
754+
(μ.map (λ a, (X a, Y a))).fst = μ.map X :=
755+
begin
756+
ext1 s hs,
757+
rw [measure.fst_apply hs, measure.map_apply_of_ae_measurable (hX.prod_mk hY) (measurable_fst hs),
758+
measure.map_apply_of_ae_measurable hX hs, ← prod_univ, mk_preimage_prod, preimage_univ,
759+
inter_univ],
760+
end
761+
762+
lemma fst_map_prod_mk {X : α → β} {Y : α → γ} {μ : measure α}
763+
(hX : measurable X) (hY : measurable Y) :
764+
(μ.map (λ a, (X a, Y a))).fst = μ.map X :=
765+
fst_map_prod_mk₀ hX.ae_measurable hY.ae_measurable
766+
752767
/-- Marginal measure on `β` obtained from a measure on `ρ` `α × β`, defined by `ρ.map prod.snd`. -/
753768
noncomputable
754769
def snd (ρ : measure (α × β)) : measure β := ρ.map prod.snd
@@ -764,6 +779,21 @@ instance [is_finite_measure ρ] : is_finite_measure ρ.snd := by { rw snd, apply
764779
instance [is_probability_measure ρ] : is_probability_measure ρ.snd :=
765780
{ measure_univ := by { rw snd_univ, exact measure_univ, } }
766781

782+
lemma snd_map_prod_mk₀ {X : α → β} {Y : α → γ} {μ : measure α}
783+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ) :
784+
(μ.map (λ a, (X a, Y a))).snd = μ.map Y :=
785+
begin
786+
ext1 s hs,
787+
rw [measure.snd_apply hs, measure.map_apply_of_ae_measurable (hX.prod_mk hY) (measurable_snd hs),
788+
measure.map_apply_of_ae_measurable hY hs, ← univ_prod, mk_preimage_prod, preimage_univ,
789+
univ_inter],
790+
end
791+
792+
lemma snd_map_prod_mk {X : α → β} {Y : α → γ} {μ : measure α}
793+
(hX : measurable X) (hY : measurable Y) :
794+
(μ.map (λ a, (X a, Y a))).snd = μ.map Y :=
795+
snd_map_prod_mk₀ hX.ae_measurable hY.ae_measurable
796+
767797
end measure
768798

769799

src/measure_theory/function/conditional_expectation/basic.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -187,6 +187,14 @@ lemma ae_eq_trim_iff_of_ae_strongly_measurable' {α β} [topological_space β] [
187187
⟨λ h, hfm.ae_eq_mk.trans (h.trans hgm.ae_eq_mk.symm),
188188
λ h, hfm.ae_eq_mk.symm.trans (h.trans hgm.ae_eq_mk)⟩
189189

190+
lemma ae_strongly_measurable.comp_ae_measurable'
191+
{α β γ : Type*} [topological_space β] {mα : measurable_space α} {mγ : measurable_space γ}
192+
{f : α → β} {μ : measure γ} {g : γ → α}
193+
(hf : ae_strongly_measurable f (μ.map g)) (hg : ae_measurable g μ) :
194+
ae_strongly_measurable' (mα.comap g) (f ∘ g) μ :=
195+
⟨(hf.mk f) ∘ g, hf.strongly_measurable_mk.comp_measurable (measurable_iff_comap_le.mpr le_rfl),
196+
ae_eq_comp hg hf.ae_eq_mk⟩
197+
190198
/-- If the restriction to a set `s` of a σ-algebra `m` is included in the restriction to `s` of
191199
another σ-algebra `m₂` (hypothesis `hs`), the set `s` is `m` measurable and a function `f` almost
192200
everywhere supported on `s` is `m`-ae-strongly-measurable, then `f` is also
Lines changed: 315 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,315 @@
1+
/-
2+
Copyright (c) 2023 Rémy Degenne. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Rémy Degenne
5+
-/
6+
import probability.kernel.disintegration
7+
import probability.notation
8+
9+
/-!
10+
# Regular conditional probability distribution
11+
12+
We define the regular conditional probability distribution of `Y : α → Ω` given `X : α → β`, where
13+
`Ω` is a standard Borel space. This is a `kernel β Ω` such that for almost all `a`, `cond_distrib`
14+
evaluated at `X a` and a measurable set `s` is equal to the conditional expectation
15+
`μ⟦Y ⁻¹' s | mβ.comap X⟧` evaluated at `a`.
16+
17+
`μ⟦Y ⁻¹' s | mβ.comap X⟧` maps a measurable set `s` to a function `α → ℝ≥0∞`, and for all `s` that
18+
map is unique up to a `μ`-null set. For all `a`, the map from sets to `ℝ≥0∞` that we obtain that way
19+
verifies some of the properties of a measure, but in general the fact that the `μ`-null set depends
20+
on `s` can prevent us from finding versions of the conditional expectation that combine into a true
21+
measure. The standard Borel space assumption on `Ω` allows us to do so.
22+
23+
The case `Y = X = id` is developed in more detail in `probability/kernel/condexp.lean`: here `X` is
24+
understood as a map from `Ω` with a sub-σ-algebra to `Ω` with its default σ-algebra and the
25+
conditional distribution defines a kernel associated with the conditional expectation with respect
26+
to `m`.
27+
28+
## Main definitions
29+
30+
* `cond_distrib Y X μ`: regular conditional probability distribution of `Y : α → Ω` given
31+
`X : α → β`, where `Ω` is a standard Borel space.
32+
33+
## Main statements
34+
35+
* `cond_distrib_ae_eq_condexp`: for almost all `a`, `cond_distrib` evaluated at `X a` and a
36+
measurable set `s` is equal to the conditional expectation `μ⟦Y ⁻¹' s | mβ.comap X⟧ a`.
37+
* `condexp_prod_ae_eq_integral_cond_distrib`: the conditional expectation
38+
`μ[(λ a, f (X a, Y a)) | X ; mβ]` is almost everywhere equal to the integral
39+
`∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a))`.
40+
41+
-/
42+
43+
open measure_theory set filter topological_space
44+
45+
open_locale ennreal measure_theory probability_theory
46+
47+
namespace probability_theory
48+
49+
variables {α β Ω F : Type*}
50+
[topological_space Ω] [measurable_space Ω] [polish_space Ω] [borel_space Ω] [nonempty Ω]
51+
[normed_add_comm_group F]
52+
{mα : measurable_space α} {μ : measure α} [is_finite_measure μ] {X : α → β} {Y : α → Ω}
53+
54+
/-- **Regular conditional probability distribution**: kernel associated with the conditional
55+
expectation of `Y` given `X`.
56+
For almost all `a`, `cond_distrib Y X μ` evaluated at `X a` and a measurable set `s` is equal to
57+
the conditional expectation `μ⟦Y ⁻¹' s | mβ.comap X⟧ a`. It also satisfies the equality
58+
`μ[(λ a, f (X a, Y a)) | mβ.comap X] =ᵐ[μ] λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a))` for
59+
all integrable functions `f`. -/
60+
@[irreducible] noncomputable
61+
def cond_distrib {mα : measurable_space α} [measurable_space β]
62+
(Y : α → Ω) (X : α → β) (μ : measure α) [is_finite_measure μ] :
63+
kernel β Ω :=
64+
(μ.map (λ a, (X a, Y a))).cond_kernel
65+
66+
instance [measurable_space β] : is_markov_kernel (cond_distrib Y X μ) :=
67+
by { rw cond_distrib, apply_instance, }
68+
69+
variables {mβ : measurable_space β} {s : set Ω} {t : set β} {f : β × Ω → F}
70+
include
71+
72+
section measurability
73+
74+
lemma measurable_cond_distrib (hs : measurable_set s) :
75+
measurable[mβ.comap X] (λ a, cond_distrib Y X μ (X a) s) :=
76+
(kernel.measurable_coe _ hs).comp (measurable.of_comap_le le_rfl)
77+
78+
lemma _root_.measure_theory.ae_strongly_measurable.ae_integrable_cond_distrib_map_iff
79+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ)
80+
(hf : ae_strongly_measurable f (μ.map (λ a, (X a, Y a)))) :
81+
(∀ᵐ a ∂(μ.map X), integrable (λ ω, f (a, ω)) (cond_distrib Y X μ a))
82+
∧ integrable (λ a, ∫ ω, ‖f (a, ω)‖ ∂(cond_distrib Y X μ a)) (μ.map X)
83+
↔ integrable f (μ.map (λ a, (X a, Y a))) :=
84+
by rw [cond_distrib, ← hf.ae_integrable_cond_kernel_iff, measure.fst_map_prod_mk₀ hX hY]
85+
86+
variables [normed_space ℝ F] [complete_space F]
87+
88+
lemma _root_.measure_theory.ae_strongly_measurable.integral_cond_distrib_map
89+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ)
90+
(hf : ae_strongly_measurable f (μ.map (λ a, (X a, Y a)))) :
91+
ae_strongly_measurable (λ x, ∫ y, f (x, y) ∂(cond_distrib Y X μ x)) (μ.map X) :=
92+
by { rw [← measure.fst_map_prod_mk₀ hX hY, cond_distrib], exact hf.integral_cond_kernel, }
93+
94+
lemma _root_.measure_theory.ae_strongly_measurable.integral_cond_distrib
95+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ)
96+
(hf : ae_strongly_measurable f (μ.map (λ a, (X a, Y a)))) :
97+
ae_strongly_measurable (λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a))) μ :=
98+
(hf.integral_cond_distrib_map hX hY).comp_ae_measurable hX
99+
100+
lemma ae_strongly_measurable'_integral_cond_distrib
101+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ)
102+
(hf : ae_strongly_measurable f (μ.map (λ a, (X a, Y a)))) :
103+
ae_strongly_measurable' (mβ.comap X) (λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a))) μ :=
104+
(hf.integral_cond_distrib_map hX hY).comp_ae_measurable' hX
105+
106+
end measurability
107+
108+
section integrability
109+
110+
lemma integrable_to_real_cond_distrib (hX : ae_measurable X μ) (hs : measurable_set s) :
111+
integrable (λ a, (cond_distrib Y X μ (X a) s).to_real) μ :=
112+
begin
113+
refine integrable_to_real_of_lintegral_ne_top _ _,
114+
{ exact measurable.comp_ae_measurable (kernel.measurable_coe _ hs) hX, },
115+
{ refine ne_of_lt _,
116+
calc ∫⁻ a, cond_distrib Y X μ (X a) s ∂μ
117+
≤ ∫⁻ a, 1 ∂μ : lintegral_mono (λ a, prob_le_one)
118+
... = μ univ : lintegral_one
119+
... < ∞ : measure_lt_top _ _, },
120+
end
121+
122+
lemma _root_.measure_theory.integrable.cond_distrib_ae_map
123+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ)
124+
(hf_int : integrable f (μ.map (λ a, (X a, Y a)))) :
125+
∀ᵐ b ∂(μ.map X), integrable (λ ω, f (b, ω)) (cond_distrib Y X μ b) :=
126+
by { rw [cond_distrib, ← measure.fst_map_prod_mk₀ hX hY], exact hf_int.cond_kernel_ae, }
127+
128+
lemma _root_.measure_theory.integrable.cond_distrib_ae
129+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ)
130+
(hf_int : integrable f (μ.map (λ a, (X a, Y a)))) :
131+
∀ᵐ a ∂μ, integrable (λ ω, f (X a, ω)) (cond_distrib Y X μ (X a)) :=
132+
ae_of_ae_map hX (hf_int.cond_distrib_ae_map hX hY)
133+
134+
lemma _root_.measure_theory.integrable.integral_norm_cond_distrib_map
135+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ)
136+
(hf_int : integrable f (μ.map (λ a, (X a, Y a)))) :
137+
integrable (λ x, ∫ y, ‖f (x, y)‖ ∂(cond_distrib Y X μ x)) (μ.map X) :=
138+
by { rw [cond_distrib, ← measure.fst_map_prod_mk₀ hX hY], exact hf_int.integral_norm_cond_kernel, }
139+
140+
lemma _root_.measure_theory.integrable.integral_norm_cond_distrib
141+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ)
142+
(hf_int : integrable f (μ.map (λ a, (X a, Y a)))) :
143+
integrable (λ a, ∫ y, ‖f (X a, y)‖ ∂(cond_distrib Y X μ (X a))) μ :=
144+
(hf_int.integral_norm_cond_distrib_map hX hY).comp_ae_measurable hX
145+
146+
variables [normed_space ℝ F] [complete_space F]
147+
148+
lemma _root_.measure_theory.integrable.norm_integral_cond_distrib_map
149+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ)
150+
(hf_int : integrable f (μ.map (λ a, (X a, Y a)))) :
151+
integrable (λ x, ‖∫ y, f (x, y) ∂(cond_distrib Y X μ x)‖) (μ.map X) :=
152+
by { rw [cond_distrib, ← measure.fst_map_prod_mk₀ hX hY], exact hf_int.norm_integral_cond_kernel, }
153+
154+
lemma _root_.measure_theory.integrable.norm_integral_cond_distrib
155+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ)
156+
(hf_int : integrable f (μ.map (λ a, (X a, Y a)))) :
157+
integrable (λ a, ‖∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a))‖) μ :=
158+
(hf_int.norm_integral_cond_distrib_map hX hY).comp_ae_measurable hX
159+
160+
lemma _root_.measure_theory.integrable.integral_cond_distrib_map
161+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ)
162+
(hf_int : integrable f (μ.map (λ a, (X a, Y a)))) :
163+
integrable (λ x, ∫ y, f (x, y) ∂(cond_distrib Y X μ x)) (μ.map X) :=
164+
(integrable_norm_iff (hf_int.1.integral_cond_distrib_map hX hY)).mp
165+
(hf_int.norm_integral_cond_distrib_map hX hY)
166+
167+
lemma _root_.measure_theory.integrable.integral_cond_distrib
168+
(hX : ae_measurable X μ) (hY : ae_measurable Y μ)
169+
(hf_int : integrable f (μ.map (λ a, (X a, Y a)))) :
170+
integrable (λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a))) μ :=
171+
(hf_int.integral_cond_distrib_map hX hY).comp_ae_measurable hX
172+
173+
end integrability
174+
175+
lemma set_lintegral_preimage_cond_distrib (hX : measurable X) (hY : ae_measurable Y μ)
176+
(hs : measurable_set s) (ht : measurable_set t) :
177+
∫⁻ a in X ⁻¹' t, cond_distrib Y X μ (X a) s ∂μ = μ (X ⁻¹' t ∩ Y ⁻¹' s) :=
178+
by rw [lintegral_comp (kernel.measurable_coe _ hs) hX, cond_distrib,
179+
← measure.restrict_map hX ht, ← measure.fst_map_prod_mk₀ hX.ae_measurable hY,
180+
set_lintegral_cond_kernel_eq_measure_prod _ ht hs,
181+
measure.map_apply_of_ae_measurable (hX.ae_measurable.prod_mk hY) (ht.prod hs),
182+
mk_preimage_prod]
183+
184+
lemma set_lintegral_cond_distrib_of_measurable_set (hX : measurable X) (hY : ae_measurable Y μ)
185+
(hs : measurable_set s) {t : set α} (ht : measurable_set[mβ.comap X] t) :
186+
∫⁻ a in t, cond_distrib Y X μ (X a) s ∂μ = μ (t ∩ Y ⁻¹' s) :=
187+
by { obtain ⟨t', ht', rfl⟩ := ht, rw set_lintegral_preimage_cond_distrib hX hY hs ht', }
188+
189+
/-- For almost every `a : α`, the `cond_distrib Y X μ` kernel applied to `X a` and a measurable set
190+
`s` is equal to the conditional expectation of the indicator of `Y ⁻¹' s`. -/
191+
lemma cond_distrib_ae_eq_condexp (hX : measurable X) (hY : measurable Y) (hs : measurable_set s) :
192+
(λ a, (cond_distrib Y X μ (X a) s).to_real) =ᵐ[μ] μ⟦Y ⁻¹' s | mβ.comap X⟧ :=
193+
begin
194+
refine ae_eq_condexp_of_forall_set_integral_eq hX.comap_le _ _ _ _,
195+
{ exact (integrable_const _).indicator (hY hs), },
196+
{ exact λ t ht _, (integrable_to_real_cond_distrib hX.ae_measurable hs).integrable_on, },
197+
{ intros t ht _,
198+
rw [integral_to_real ((measurable_cond_distrib hs).mono hX.comap_le le_rfl).ae_measurable
199+
(eventually_of_forall (λ ω, measure_lt_top (cond_distrib Y X μ (X ω)) _)),
200+
integral_indicator_const _ (hY hs), measure.restrict_apply (hY hs), smul_eq_mul, mul_one,
201+
inter_comm, set_lintegral_cond_distrib_of_measurable_set hX hY.ae_measurable hs ht], },
202+
{ refine (measurable.strongly_measurable _).ae_strongly_measurable',
203+
exact @measurable.ennreal_to_real _ (mβ.comap X) _ (measurable_cond_distrib hs), },
204+
end
205+
206+
/-- The conditional expectation of a function `f` of the product `(X, Y)` is almost everywhere equal
207+
to the integral of `y ↦ f(X, y)` against the `cond_distrib` kernel. -/
208+
lemma condexp_prod_ae_eq_integral_cond_distrib' [normed_space ℝ F] [complete_space F]
209+
(hX : measurable X) (hY : ae_measurable Y μ)
210+
(hf_int : integrable f (μ.map (λ a, (X a, Y a)))) :
211+
μ[(λ a, f (X a, Y a)) | mβ.comap X] =ᵐ[μ] λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a)) :=
212+
begin
213+
have hf_int' : integrable (λ a, f (X a, Y a)) μ,
214+
{ exact (integrable_map_measure hf_int.1 (hX.ae_measurable.prod_mk hY)).mp hf_int, },
215+
refine (ae_eq_condexp_of_forall_set_integral_eq hX.comap_le hf_int' (λ s hs hμs, _) _ _).symm,
216+
{ exact (hf_int.integral_cond_distrib hX.ae_measurable hY).integrable_on, },
217+
{ rintros s ⟨t, ht, rfl⟩ _,
218+
change ∫ a in X ⁻¹' t, ((λ x', ∫ y, f (x', y) ∂(cond_distrib Y X μ) x') ∘ X) a ∂μ
219+
= ∫ a in X ⁻¹' t, f (X a, Y a) ∂μ,
220+
rw ← integral_map hX.ae_measurable,
221+
swap,
222+
{ rw ← measure.restrict_map hX ht,
223+
exact (hf_int.1.integral_cond_distrib_map hX.ae_measurable hY).restrict, },
224+
rw [← measure.restrict_map hX ht, ← measure.fst_map_prod_mk₀ hX.ae_measurable hY, cond_distrib,
225+
set_integral_cond_kernel_univ_right ht hf_int.integrable_on,
226+
set_integral_map (ht.prod measurable_set.univ) hf_int.1 (hX.ae_measurable.prod_mk hY),
227+
mk_preimage_prod, preimage_univ, inter_univ], },
228+
{ exact ae_strongly_measurable'_integral_cond_distrib hX.ae_measurable hY hf_int.1, },
229+
end
230+
231+
/-- The conditional expectation of a function `f` of the product `(X, Y)` is almost everywhere equal
232+
to the integral of `y ↦ f(X, y)` against the `cond_distrib` kernel. -/
233+
lemma condexp_prod_ae_eq_integral_cond_distrib₀ [normed_space ℝ F] [complete_space F]
234+
(hX : measurable X) (hY : ae_measurable Y μ)
235+
(hf : ae_strongly_measurable f (μ.map (λ a, (X a, Y a))))
236+
(hf_int : integrable (λ a, f (X a, Y a)) μ) :
237+
μ[(λ a, f (X a, Y a)) | mβ.comap X] =ᵐ[μ] λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a)) :=
238+
begin
239+
have hf_int' : integrable f (μ.map (λ a, (X a, Y a))),
240+
{ rwa integrable_map_measure hf (hX.ae_measurable.prod_mk hY), },
241+
exact condexp_prod_ae_eq_integral_cond_distrib' hX hY hf_int',
242+
end
243+
244+
/-- The conditional expectation of a function `f` of the product `(X, Y)` is almost everywhere equal
245+
to the integral of `y ↦ f(X, y)` against the `cond_distrib` kernel. -/
246+
lemma condexp_prod_ae_eq_integral_cond_distrib [normed_space ℝ F] [complete_space F]
247+
(hX : measurable X) (hY : ae_measurable Y μ)
248+
(hf : strongly_measurable f) (hf_int : integrable (λ a, f (X a, Y a)) μ) :
249+
μ[(λ a, f (X a, Y a)) | mβ.comap X] =ᵐ[μ] λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a)) :=
250+
begin
251+
have hf_int' : integrable f (μ.map (λ a, (X a, Y a))),
252+
{ rwa integrable_map_measure hf.ae_strongly_measurable (hX.ae_measurable.prod_mk hY), },
253+
exact condexp_prod_ae_eq_integral_cond_distrib' hX hY hf_int',
254+
end
255+
256+
lemma condexp_ae_eq_integral_cond_distrib [normed_space ℝ F] [complete_space F]
257+
(hX : measurable X) (hY : ae_measurable Y μ)
258+
{f : Ω → F} (hf : strongly_measurable f) (hf_int : integrable (λ a, f (Y a)) μ) :
259+
μ[(λ a, f (Y a)) | mβ.comap X] =ᵐ[μ] λ a, ∫ y, f y ∂(cond_distrib Y X μ (X a)) :=
260+
condexp_prod_ae_eq_integral_cond_distrib hX hY (hf.comp_measurable measurable_snd) hf_int
261+
262+
/-- The conditional expectation of `Y` given `X` is almost everywhere equal to the integral
263+
`∫ y, y ∂(cond_distrib Y X μ (X a))`. -/
264+
lemma condexp_ae_eq_integral_cond_distrib' {Ω} [normed_add_comm_group Ω] [normed_space ℝ Ω]
265+
[complete_space Ω] [measurable_space Ω] [borel_space Ω] [second_countable_topology Ω] {Y : α → Ω}
266+
(hX : measurable X) (hY_int : integrable Y μ) :
267+
μ[Y | mβ.comap X] =ᵐ[μ] λ a, ∫ y, y ∂(cond_distrib Y X μ (X a)) :=
268+
condexp_ae_eq_integral_cond_distrib hX hY_int.1.ae_measurable strongly_measurable_id hY_int
269+
270+
lemma _root_.measure_theory.ae_strongly_measurable.comp_snd_map_prod_mk
271+
{Ω F} {mΩ : measurable_space Ω} {X : Ω → β} {μ : measure Ω}
272+
[topological_space F] (hX : measurable X) {f : Ω → F} (hf : ae_strongly_measurable f μ) :
273+
ae_strongly_measurable (λ x : β × Ω, f x.2) (μ.map (λ ω, (X ω, ω))) :=
274+
begin
275+
refine ⟨λ x, hf.mk f x.2, hf.strongly_measurable_mk.comp_measurable measurable_snd, _⟩,
276+
suffices h : measure.quasi_measure_preserving prod.snd (μ.map (λ ω, (X ω, ω))) μ,
277+
{ exact measure.quasi_measure_preserving.ae_eq h hf.ae_eq_mk, },
278+
refine ⟨measurable_snd, measure.absolutely_continuous.mk (λ s hs hμs, _)⟩,
279+
rw measure.map_apply _ hs,
280+
swap, { exact measurable_snd, },
281+
rw measure.map_apply,
282+
{ rw [← univ_prod, mk_preimage_prod, preimage_univ, univ_inter, preimage_id'],
283+
exact hμs, },
284+
{ exact hX.prod_mk measurable_id, },
285+
{ exact measurable_snd hs, },
286+
end
287+
288+
lemma _root_.measure_theory.integrable.comp_snd_map_prod_mk {Ω} {mΩ : measurable_space Ω}
289+
{X : Ω → β} {μ : measure Ω} (hX : measurable X) {f : Ω → F} (hf_int : integrable f μ) :
290+
integrable (λ x : β × Ω, f x.2) (μ.map (λ ω, (X ω, ω))) :=
291+
begin
292+
have hf := hf_int.1.comp_snd_map_prod_mk hX,
293+
refine ⟨hf, _⟩,
294+
rw [has_finite_integral, lintegral_map' hf.ennnorm (hX.prod_mk measurable_id).ae_measurable],
295+
exact hf_int.2,
296+
end
297+
298+
lemma ae_strongly_measurable_comp_snd_map_prod_mk_iff {Ω F} {mΩ : measurable_space Ω}
299+
[topological_space F] {X : Ω → β} {μ : measure Ω} (hX : measurable X) {f : Ω → F} :
300+
ae_strongly_measurable (λ x : β × Ω, f x.2) (μ.map (λ ω, (X ω, ω)))
301+
↔ ae_strongly_measurable f μ :=
302+
⟨λ h, h.comp_measurable (hX.prod_mk measurable_id), λ h, h.comp_snd_map_prod_mk hX⟩
303+
304+
lemma integrable_comp_snd_map_prod_mk_iff {Ω} {mΩ : measurable_space Ω} {X : Ω → β} {μ : measure Ω}
305+
(hX : measurable X) {f : Ω → F} :
306+
integrable (λ x : β × Ω, f x.2) (μ.map (λ ω, (X ω, ω))) ↔ integrable f μ :=
307+
⟨λ h, h.comp_measurable (hX.prod_mk measurable_id), λ h, h.comp_snd_map_prod_mk hX⟩
308+
309+
lemma condexp_ae_eq_integral_cond_distrib_id [normed_space ℝ F] [complete_space F]
310+
{X : Ω → β} {μ : measure Ω} [is_finite_measure μ]
311+
(hX : measurable X) {f : Ω → F} (hf_int : integrable f μ) :
312+
μ[f | mβ.comap X] =ᵐ[μ] λ a, ∫ y, f y ∂(cond_distrib id X μ (X a)) :=
313+
condexp_prod_ae_eq_integral_cond_distrib' hX ae_measurable_id (hf_int.comp_snd_map_prod_mk hX)
314+
315+
end probability_theory

0 commit comments

Comments
 (0)