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

Commit da4010c

Browse files
RemyDegennesgouezel
andcommitted
feat(probability/independence): Kolmogorov's 0-1 law (#16648)
We prove that any event in the tail σ-algebra of an independent sequence of sub-σ-algebras has probability 0 or 1. Co-authored-by: Rémy Degenne <remydegenne@gmail.com> Co-authored-by: sgouezel <sebastien.gouezel@univ-rennes1.fr>
1 parent e6f28f5 commit da4010c

File tree

6 files changed

+401
-23
lines changed

6 files changed

+401
-23
lines changed

docs/undergrad.yaml

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -536,7 +536,7 @@ Probability Theory:
536536
independent events: 'probability_theory.Indep_set'
537537
sigma-algebra: 'measurable_space'
538538
independent sigma-algebra: 'probability_theory.Indep'
539-
$0$-$1$ law: ''
539+
$0$-$1$ law: 'probability_theory.measure_zero_or_one_of_measurable_set_limsup_at_top'
540540
Borel-Cantelli lemma (easy direction): 'measure_theory.measure_limsup_eq_zero'
541541
Borel-Cantelli lemma (difficult direction): ''
542542
conditional probability: 'probability_theory.cond'

src/algebra/big_operators/basic.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -821,6 +821,11 @@ lemma prod_extend_by_one [decidable_eq α] (s : finset α) (f : α → β) :
821821
∏ i in s, (if i ∈ s then f i else 1) = ∏ i in s, f i :=
822822
prod_congr rfl $ λ i hi, if_pos hi
823823

824+
@[simp, to_additive]
825+
lemma prod_ite_mem [decidable_eq α] (s t : finset α) (f : α → β) :
826+
∏ i in s, (if i ∈ t then f i else 1) = ∏ i in (s ∩ t), f i :=
827+
by rw [← finset.prod_filter, finset.filter_mem_eq_inter]
828+
824829
@[simp, to_additive]
825830
lemma prod_dite_eq [decidable_eq α] (s : finset α) (a : α) (b : Π x : α, a = x → β) :
826831
(∏ x in s, (if h : a = x then b x h else 1)) = ite (a ∈ s) (b a rfl) 1 :=

src/data/set/basic.lean

Lines changed: 18 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1232,6 +1232,24 @@ ext $ λ s, subset_empty_iff
12321232
@[simp] theorem powerset_univ : 𝒫 (univ : set α) = univ :=
12331233
eq_univ_of_forall subset_univ
12341234

1235+
/-! ### Sets defined as an if-then-else -/
1236+
1237+
lemma mem_dite_univ_right (p : Prop) [decidable p] (t : p → set α) (x : α) :
1238+
(x ∈ if h : p then t h else univ) ↔ (∀ h : p, x ∈ t h) :=
1239+
by split_ifs; simp [h]
1240+
1241+
@[simp] lemma mem_ite_univ_right (p : Prop) [decidable p] (t : set α) (x : α) :
1242+
x ∈ ite p t set.univ ↔ (p → x ∈ t) :=
1243+
mem_dite_univ_right p (λ _, t) x
1244+
1245+
lemma mem_dite_univ_left (p : Prop) [decidable p] (t : ¬ p → set α) (x : α) :
1246+
(x ∈ if h : p then univ else t h) ↔ (∀ h : ¬ p, x ∈ t h) :=
1247+
by split_ifs; simp [h]
1248+
1249+
@[simp] lemma mem_ite_univ_left (p : Prop) [decidable p] (t : set α) (x : α) :
1250+
x ∈ ite p set.univ t ↔ (¬ p → x ∈ t) :=
1251+
mem_dite_univ_left p (λ _, t) x
1252+
12351253
/-! ### If-then-else for sets -/
12361254

12371255
/-- `ite` for sets: `set.ite t s s' ∩ t = s ∩ t`, `set.ite t s s' ∩ tᶜ = s' ∩ tᶜ`.

src/measure_theory/measurable_space_def.lean

Lines changed: 17 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -182,6 +182,11 @@ h₁.inter h₂.compl
182182
measurable_set (t.ite s₁ s₂) :=
183183
(h₁.inter ht).union (h₂.diff ht)
184184

185+
lemma measurable_set.ite' {s t : set α} {p : Prop}
186+
(hs : p → measurable_set s) (ht : ¬ p → measurable_set t) :
187+
measurable_set (ite p s t) :=
188+
by { split_ifs, exacts [hs h, ht h], }
189+
185190
@[simp] lemma measurable_set.cond {s₁ s₂ : set α} (h₁ : measurable_set s₁) (h₂ : measurable_set s₂)
186191
{i : bool} : measurable_set (cond i s₁ s₂) :=
187192
by { cases i, exacts [h₂, h₁] }
@@ -371,6 +376,12 @@ begin
371376
{ exact measurable_set_generate_from ht, },
372377
end
373378

379+
@[simp] lemma generate_from_singleton_empty : generate_from {∅} = (⊥ : measurable_space α) :=
380+
by { rw [eq_bot_iff, generate_from_le_iff], simp, }
381+
382+
@[simp] lemma generate_from_singleton_univ : generate_from {set.univ} = (⊥ : measurable_space α) :=
383+
by { rw [eq_bot_iff, generate_from_le_iff], simp, }
384+
374385
lemma measurable_set_bot_iff {s : set α} : @measurable_set α ⊥ s ↔ (s = ∅ ∨ s = univ) :=
375386
let b : measurable_space α :=
376387
{ measurable_set' := λ s, s = ∅ ∨ s = univ,
@@ -419,12 +430,12 @@ theorem measurable_set_supr {ι} {m : ι → measurable_space α} {s : set α} :
419430
by simp only [supr, measurable_set_Sup, exists_range_iff]
420431

421432
lemma measurable_space_supr_eq (m : ι → measurable_space α) :
422-
(⨆ n, m n) = measurable_space.generate_from {s | ∃ n, measurable_set[m n] s} :=
423-
begin
424-
ext s,
425-
rw measurable_space.measurable_set_supr,
426-
refl,
427-
end
433+
(⨆ n, m n) = generate_from {s | ∃ n, measurable_set[m n] s} :=
434+
by { ext s, rw measurable_set_supr, refl, }
435+
436+
lemma generate_from_Union_measurable_set (m : ι → measurable_space α) :
437+
generate_from (⋃ n, {t | measurable_set[m n] t}) = ⨆ n, m n :=
438+
(@gi_generate_from α).l_supr_u m
428439

429440
end complete_lattice
430441

src/measure_theory/pi_system.lean

Lines changed: 37 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -113,6 +113,23 @@ begin
113113
simpa using hst,
114114
end
115115

116+
lemma is_pi_system_Union_of_directed_le {α ι} (p : ι → set (set α))
117+
(hp_pi : ∀ n, is_pi_system (p n)) (hp_directed : directed (≤) p) :
118+
is_pi_system (⋃ n, p n) :=
119+
begin
120+
intros t1 ht1 t2 ht2 h,
121+
rw set.mem_Union at ht1 ht2 ⊢,
122+
cases ht1 with n ht1,
123+
cases ht2 with m ht2,
124+
obtain ⟨k, hpnk, hpmk⟩ : ∃ k, p n ≤ p k ∧ p m ≤ p k := hp_directed n m,
125+
exact ⟨k, hp_pi k t1 (hpnk ht1) t2 (hpmk ht2) h⟩,
126+
end
127+
128+
lemma is_pi_system_Union_of_monotone {α ι} [semilattice_sup ι] (p : ι → set (set α))
129+
(hp_pi : ∀ n, is_pi_system (p n)) (hp_mono : monotone p) :
130+
is_pi_system (⋃ n, p n) :=
131+
is_pi_system_Union_of_directed_le p hp_pi (monotone.directed_le hp_mono)
132+
116133
section order
117134

118135
variables {α : Type*} {ι ι' : Sort*} [linear_order α]
@@ -435,6 +452,26 @@ begin
435452
exact generate_from_mono (mem_pi_Union_Inter_of_measurable_set m hpS hpi), },
436453
end
437454

455+
lemma generate_from_pi_Union_Inter_subsets {α ι} (m : ι → measurable_space α) (S : set ι) :
456+
generate_from (pi_Union_Inter (λ i, {t | measurable_set[m i] t}) {t : finset ι | ↑t ⊆ S})
457+
= ⨆ i ∈ S, m i :=
458+
begin
459+
rw generate_from_pi_Union_Inter_measurable_space,
460+
simp only [set.mem_set_of_eq, exists_prop, supr_exists],
461+
congr' 1,
462+
ext1 i,
463+
by_cases hiS : i ∈ S,
464+
{ simp only [hiS, csupr_pos],
465+
refine le_antisymm (supr₂_le (λ t ht, le_rfl)) _,
466+
rw le_supr_iff,
467+
intros m' hm',
468+
specialize hm' {i},
469+
simpa only [hiS, finset.coe_singleton, set.singleton_subset_iff, finset.mem_singleton,
470+
eq_self_iff_true, and_self, csupr_pos] using hm', },
471+
{ simp only [hiS, supr_false, supr₂_eq_bot, and_imp],
472+
exact λ t htS hit, absurd (htS hit) hiS, },
473+
end
474+
438475
end Union_Inter
439476

440477
namespace measurable_space

0 commit comments

Comments
 (0)