@@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Rémy Degenne
5
5
-/
6
6
import Mathlib.MeasureTheory.Constructions.Pi
7
+ import Mathlib.Probability.ConditionalProbability
7
8
import Mathlib.Probability.Kernel.Basic
8
9
import Mathlib.Tactic.Peel
9
10
@@ -1212,4 +1213,56 @@ theorem iIndepSet.iIndepFun_indicator [Zero β] [One β] {m : MeasurableSpace β
1212
1213
1213
1214
end IndepFun
1214
1215
1216
+ variable {ι Ω α β : Type *} {mΩ : MeasurableSpace Ω} {mα : MeasurableSpace α}
1217
+ {mβ : MeasurableSpace β} {κ : Kernel α Ω} {μ : Measure α} {X : ι → Ω → α} {Y : ι → Ω → β}
1218
+ {f : _ → Set Ω} {t : ι → Set β} {s : Finset ι}
1219
+
1220
+ /-- The probability of an intersection of preimages conditioning on another intersection factors
1221
+ into a product. -/
1222
+ lemma iIndepFun.cond_iInter [Finite ι] (hY : ∀ i, Measurable (Y i))
1223
+ (hindep : iIndepFun (fun _ ↦ mα.prod mβ) (fun i ω ↦ (X i ω, Y i ω)) κ μ)
1224
+ (hf : ∀ i ∈ s, MeasurableSet[mα.comap (X i)] (f i))
1225
+ (hy : ∀ᵐ a ∂μ, ∀ i ∉ s, κ a (Y i ⁻¹' t i) ≠ 0 ) (ht : ∀ i, MeasurableSet (t i)) :
1226
+ ∀ᵐ a ∂μ, (κ a)[⋂ i ∈ s, f i | ⋂ i, Y i ⁻¹' t i] = ∏ i ∈ s, (κ a)[f i | Y i in t i] := by
1227
+ classical
1228
+ cases nonempty_fintype ι
1229
+ let g (i' : ι) := if i' ∈ s then Y i' ⁻¹' t i' ∩ f i' else Y i' ⁻¹' t i'
1230
+ have hYt i : MeasurableSet[(mα.prod mβ).comap fun ω ↦ (X i ω, Y i ω)] (Y i ⁻¹' t i) :=
1231
+ ⟨.univ ×ˢ t i, .prod .univ (ht _), by ext; simp [eq_comm]⟩
1232
+ have hg i : MeasurableSet[(mα.prod mβ).comap fun ω ↦ (X i ω, Y i ω)] (g i) := by
1233
+ by_cases hi : i ∈ s <;> simp only [hi, ↓reduceIte, g]
1234
+ · obtain ⟨A, hA, hA'⟩ := hf i hi
1235
+ exact (hYt _).inter ⟨A ×ˢ .univ, hA.prod .univ, by ext; simp [← hA']⟩
1236
+ · exact hYt _
1237
+ filter_upwards [hy, hindep.ae_isProbabilityMeasure, hindep.meas_iInter hYt, hindep.meas_iInter hg]
1238
+ with a hy _ hYt hg
1239
+ calc
1240
+ _ = (κ a (⋂ i, Y i ⁻¹' t i))⁻¹ * κ a ((⋂ i, Y i ⁻¹' t i) ∩ ⋂ i ∈ s, f i) := by
1241
+ rw [cond_apply]; exact .iInter fun i ↦ hY i (ht i)
1242
+ _ = (κ a (⋂ i, Y i ⁻¹' t i))⁻¹ * κ a (⋂ i, g i) := by
1243
+ congr
1244
+ calc
1245
+ _ = (⋂ i, Y i ⁻¹' t i) ∩ ⋂ i, if i ∈ s then f i else .univ := by
1246
+ congr
1247
+ simp only [Set.iInter_ite, Set.iInter_univ, Set.inter_univ]
1248
+ _ = ⋂ i, Y i ⁻¹' t i ∩ (if i ∈ s then f i else .univ) := by rw [Set.iInter_inter_distrib]
1249
+ _ = _ := Set.iInter_congr fun i ↦ by by_cases hi : i ∈ s <;> simp [hi, g]
1250
+ _ = (∏ i, κ a (Y i ⁻¹' t i))⁻¹ * κ a (⋂ i, g i) := by
1251
+ rw [hYt]
1252
+ _ = (∏ i, κ a (Y i ⁻¹' t i))⁻¹ * ∏ i, κ a (g i) := by
1253
+ rw [hg]
1254
+ _ = ∏ i, (κ a (Y i ⁻¹' t i))⁻¹ * κ a (g i) := by
1255
+ rw [Finset.prod_mul_distrib, ENNReal.prod_inv_distrib]
1256
+ exact fun _ _ i _ _ ↦ .inr <| measure_ne_top _ _
1257
+ _ = ∏ i, if i ∈ s then (κ a)[f i | Y i ⁻¹' t i] else 1 := by
1258
+ refine Finset.prod_congr rfl fun i _ ↦ ?_
1259
+ by_cases hi : i ∈ s
1260
+ · simp only [hi, ↓reduceIte, g, cond_apply (hY i (ht i))]
1261
+ · simp only [hi, ↓reduceIte, g, ENNReal.inv_mul_cancel (hy i hi) (measure_ne_top _ _)]
1262
+ _ = _ := by simp
1263
+
1264
+ -- TODO: We can't state `Kernel.iIndepFun.cond` (the `Kernel` analogue of
1265
+ -- `ProbabilityTheory.iIndepFun.cond`) because we don't have a version of `ProbabilityTheory.cond`
1266
+ -- for kernels
1267
+
1215
1268
end ProbabilityTheory.Kernel
0 commit comments