Skip to content

Commit 0b15a0d

Browse files
committed
feat: independence of functions under conditioning (#17915)
From PFR
1 parent d1d52a6 commit 0b15a0d

File tree

7 files changed

+125
-5
lines changed

7 files changed

+125
-5
lines changed

Mathlib/Analysis/SpecificLimits/Basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -623,7 +623,7 @@ theorem tendsto_factorial_div_pow_self_atTop :
623623
refine (eventually_gt_atTop 0).mono fun n hn ↦ ?_
624624
rcases Nat.exists_eq_succ_of_ne_zero hn.ne.symm with ⟨k, rfl⟩
625625
rw [← prod_range_add_one_eq_factorial, pow_eq_prod_const, div_eq_mul_inv, ← inv_eq_one_div,
626-
prod_natCast, Nat.cast_succ, ← prod_inv_distrib, ← prod_mul_distrib,
626+
prod_natCast, Nat.cast_succ, ← Finset.prod_inv_distrib, ← prod_mul_distrib,
627627
Finset.prod_range_succ']
628628
simp only [prod_range_succ', one_mul, Nat.cast_add, zero_add, Nat.cast_one]
629629
refine

Mathlib/Data/ENNReal/Inv.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -173,6 +173,16 @@ protected theorem inv_div {a b : ℝ≥0∞} (htop : b ≠ ∞ ∨ a ≠ ∞) (h
173173
rw [← ENNReal.inv_ne_top] at hzero
174174
rw [ENNReal.div_eq_inv_mul, ENNReal.div_eq_inv_mul, ENNReal.mul_inv htop hzero, mul_comm, inv_inv]
175175

176+
lemma prod_inv_distrib {ι : Type*} {f : ι → ℝ≥0∞} {s : Finset ι}
177+
(hf : s.toSet.Pairwise fun i j ↦ f i ≠ 0 ∨ f j ≠ ∞) : (∏ i ∈ s, f i)⁻¹ = ∏ i ∈ s, (f i)⁻¹ := by
178+
induction' s using Finset.cons_induction with i s hi ih
179+
· simp
180+
simp [← ih (hf.mono <| by simp)]
181+
refine ENNReal.mul_inv (not_or_of_imp fun hi₀ ↦ prod_ne_top fun j hj ↦ ?_)
182+
(not_or_of_imp fun hi₀ ↦ Finset.prod_ne_zero_iff.2 fun j hj ↦ ?_)
183+
· exact imp_iff_not_or.2 (hf (by simp) (by simp [hj]) <| .symm <| ne_of_mem_of_not_mem hj hi) hi₀
184+
· exact imp_iff_not_or.2 (hf (by simp [hj]) (by simp) <| ne_of_mem_of_not_mem hj hi).symm hi₀
185+
176186
protected theorem mul_div_mul_left (a b : ℝ≥0∞) (hc : c ≠ 0) (hc' : c ≠ ⊤) :
177187
c * a / (c * b) = a / b := by
178188
rw [div_eq_mul_inv, div_eq_mul_inv, ENNReal.mul_inv (Or.inl hc) (Or.inl hc'), mul_mul_mul_comm,

Mathlib/MeasureTheory/Function/LpSpace.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1192,7 +1192,7 @@ def compLpₗ (L : E →L[𝕜] F) : Lp E p μ →ₗ[𝕜] Lp F p μ where
11921192
ext1
11931193
filter_upwards [Lp.coeFn_smul c f, coeFn_compLp L (c • f), Lp.coeFn_smul c (L.compLp f),
11941194
coeFn_compLp L f] with _ ha1 ha2 ha3 ha4
1195-
simp only [ha1, ha2, ha3, ha4, map_smul, Pi.smul_apply]
1195+
simp only [ha1, ha2, ha3, ha4, _root_.map_smul, Pi.smul_apply]
11961196

11971197
/-- Composing `f : Lp E p μ` with `L : E →L[𝕜] F`, seen as a continuous `𝕜`-linear map on
11981198
`Lp E p μ`. See also the similar

Mathlib/MeasureTheory/Integral/Bochner.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -638,7 +638,7 @@ theorem integral_sub (f g : α →₁[μ] E) : integral (f - g) = integral f - i
638638
theorem integral_smul (c : 𝕜) (f : α →₁[μ] E) : integral (c • f) = c • integral f := by
639639
simp only [integral]
640640
show (integralCLM' (E := E) 𝕜) (c • f) = c • (integralCLM' (E := E) 𝕜) f
641-
exact map_smul (integralCLM' (E := E) 𝕜) c f
641+
exact _root_.map_smul (integralCLM' (E := E) 𝕜) c f
642642

643643
local notation "Integral" => @integralCLM α E _ _ μ _ _
644644

Mathlib/Probability/ConditionalProbability.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -70,8 +70,8 @@ and scaled by the inverse of `μ s` (to make it a probability measure):
7070
def cond (s : Set Ω) : Measure Ω :=
7171
(μ s)⁻¹ • μ.restrict s
7272

73-
@[inherit_doc] scoped notation:max μ "[" t " | " s "]" => ProbabilityTheory.cond μ s t
7473
@[inherit_doc] scoped notation:max μ "[|" s "]" => ProbabilityTheory.cond μ s
74+
@[inherit_doc cond] scoped notation3:max μ "[" t " | " s "]" => ProbabilityTheory.cond μ s t
7575

7676
/-- The conditional probability measure of measure `μ` on `{ω | X ω ∈ s}`.
7777
@@ -172,7 +172,7 @@ theorem inter_pos_of_cond_ne_zero (hms : MeasurableSet s) (hcst : μ[t|s] ≠ 0)
172172
simp [hms, Set.inter_comm, cond]
173173

174174
lemma cond_pos_of_inter_ne_zero [IsFiniteMeasure μ] (hms : MeasurableSet s) (hci : μ (s ∩ t) ≠ 0) :
175-
0 < μ[|s] t := by
175+
0 < μ[t | s] := by
176176
rw [cond_apply hms]
177177
refine ENNReal.mul_pos ?_ hci
178178
exact ENNReal.inv_ne_zero.mpr (measure_ne_top _ _)

Mathlib/Probability/Independence/Basic.lean

Lines changed: 57 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -721,4 +721,61 @@ theorem iIndepSet.iIndepFun_indicator [Zero β] [One β] {m : MeasurableSpace β
721721

722722
end IndepFun
723723

724+
variable {ι Ω α β : Type*} {mΩ : MeasurableSpace Ω} {mα : MeasurableSpace α}
725+
{mβ : MeasurableSpace β} {μ : Measure Ω} {X : ι → Ω → α} {Y : ι → Ω → β} {f : _ → Set Ω}
726+
{t : ι → Set β} {s : Finset ι}
727+
728+
/-- The probability of an intersection of preimages conditioning on another intersection factors
729+
into a product. -/
730+
lemma cond_iInter [Finite ι] (hY : ∀ i, Measurable (Y i))
731+
(hindep : iIndepFun (fun _ ↦ mα.prod mβ) (fun i ω ↦ (X i ω, Y i ω)) μ)
732+
(hf : ∀ i ∈ s, MeasurableSet[mα.comap (X i)] (f i))
733+
(hy : ∀ i ∉ s, μ (Y i ⁻¹' t i) ≠ 0) (ht : ∀ i, MeasurableSet (t i)) :
734+
μ[⋂ i ∈ s, f i | ⋂ i, Y i ⁻¹' t i] = ∏ i ∈ s, μ[f i | Y i in t i] := by
735+
have : IsProbabilityMeasure (μ : Measure Ω) := hindep.isProbabilityMeasure
736+
classical
737+
cases nonempty_fintype ι
738+
let g (i' : ι) := if i' ∈ s then Y i' ⁻¹' t i' ∩ f i' else Y i' ⁻¹' t i'
739+
calc
740+
_ = (μ (⋂ i, Y i ⁻¹' t i))⁻¹ * μ ((⋂ i, Y i ⁻¹' t i) ∩ ⋂ i ∈ s, f i) := by
741+
rw [cond_apply]; exact .iInter fun i ↦ hY i (ht i)
742+
_ = (μ (⋂ i, Y i ⁻¹' t i))⁻¹ * μ (⋂ i, g i) := by
743+
congr
744+
calc
745+
_ = (⋂ i, Y i ⁻¹' t i) ∩ ⋂ i, if i ∈ s then f i else .univ := by
746+
congr
747+
simp only [Set.iInter_ite, Set.iInter_univ, Set.inter_univ]
748+
_ = ⋂ i, Y i ⁻¹' t i ∩ (if i ∈ s then f i else .univ) := by rw [Set.iInter_inter_distrib]
749+
_ = _ := Set.iInter_congr fun i ↦ by by_cases hi : i ∈ s <;> simp [hi, g]
750+
_ = (∏ i, μ (Y i ⁻¹' t i))⁻¹ * μ (⋂ i, g i) := by
751+
rw [hindep.meas_iInter]
752+
exact fun i ↦ ⟨.univ ×ˢ t i, MeasurableSet.univ.prod (ht _), by ext; simp [eq_comm]⟩
753+
_ = (∏ i, μ (Y i ⁻¹' t i))⁻¹ * ∏ i, μ (g i) := by
754+
rw [hindep.meas_iInter]
755+
intro i
756+
by_cases hi : i ∈ s <;> simp only [hi, ↓reduceIte, g]
757+
· obtain ⟨A, hA, hA'⟩ := hf i hi
758+
exact .inter ⟨.univ ×ˢ t i, MeasurableSet.univ.prod (ht _), by ext; simp [eq_comm]⟩
759+
⟨A ×ˢ Set.univ, hA.prod .univ, by ext; simp [← hA']⟩
760+
· exact ⟨.univ ×ˢ t i, MeasurableSet.univ.prod (ht _), by ext; simp [eq_comm]⟩
761+
_ = ∏ i, (μ (Y i ⁻¹' t i))⁻¹ * μ (g i) := by
762+
rw [Finset.prod_mul_distrib, ENNReal.prod_inv_distrib]
763+
exact fun _ _ i _ _ ↦ .inr <| measure_ne_top _ _
764+
_ = ∏ i, if i ∈ s then μ[f i | Y i ⁻¹' t i] else 1 := by
765+
refine Finset.prod_congr rfl fun i _ ↦ ?_
766+
by_cases hi : i ∈ s
767+
· simp only [hi, ↓reduceIte, g, cond_apply (hY i (ht i))]
768+
· simp only [hi, ↓reduceIte, g, ENNReal.inv_mul_cancel (hy i hi) (measure_ne_top μ _)]
769+
_ = _ := by simp
770+
771+
lemma iIndepFun.cond [Finite ι] (hY : ∀ i, Measurable (Y i))
772+
(hindep : iIndepFun (fun _ ↦ mα.prod mβ) (fun i ω ↦ (X i ω, Y i ω)) μ)
773+
(hy : ∀ i, μ (Y i ⁻¹' t i) ≠ 0) (ht : ∀ i, MeasurableSet (t i)) :
774+
iIndepFun (fun _ ↦ mα) X μ[|⋂ i, Y i ⁻¹' t i] := by
775+
rw [iIndepFun_iff]
776+
intro s f hf
777+
convert cond_iInter hY hindep hf (fun i _ ↦ hy _) ht using 2 with i hi
778+
simpa using cond_iInter hY hindep (fun j hj ↦ hf _ <| Finset.mem_singleton.1 hj ▸ hi)
779+
(fun i _ ↦ hy _) ht
780+
724781
end ProbabilityTheory

Mathlib/Probability/Independence/Kernel.lean

Lines changed: 53 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Rémy Degenne
55
-/
66
import Mathlib.MeasureTheory.Constructions.Pi
7+
import Mathlib.Probability.ConditionalProbability
78
import Mathlib.Probability.Kernel.Basic
89
import Mathlib.Tactic.Peel
910

@@ -1212,4 +1213,56 @@ theorem iIndepSet.iIndepFun_indicator [Zero β] [One β] {m : MeasurableSpace β
12121213

12131214
end IndepFun
12141215

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+
12151268
end ProbabilityTheory.Kernel

0 commit comments

Comments
 (0)