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

Commit bda091d

Browse files
sgouezelvihdzp
andcommitted
feat(measure_theory/card_measurable_space): cardinality of generated sigma-algebras (#12422)
If a sigma-algebra is generated by a set of sets `s` whose cardinality is at most the continuum, then the sigma-algebra satisfies the same cardinality bound. Co-authored-by: Violeta Hernández <vi.hdz.p@gmail.com>
1 parent 93451af commit bda091d

File tree

6 files changed

+184
-1
lines changed

6 files changed

+184
-1
lines changed
Lines changed: 161 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,161 @@
1+
/-
2+
Copyright (c) 2022 Sébastien Gouëzel. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Sébastien Gouëzel, Violeta Hernández Palacios
5+
-/
6+
import measure_theory.measurable_space_def
7+
import set_theory.continuum
8+
import set_theory.cofinality
9+
10+
/-!
11+
# Cardinal of sigma-algebras
12+
13+
If a sigma-algebra is generated by a set of sets `s`, then the cardinality of the sigma-algebra is
14+
bounded by `(max (#s) 2) ^ ω`. This is stated in `measurable_space.cardinal_generate_measurable_le`
15+
and `measurable_space.cardinal_measurable_set_le`.
16+
17+
In particular, if `#s ≤ 𝔠`, then the generated sigma-algebra has cardinality at most `𝔠`, see
18+
`measurable_space.cardinal_measurable_set_le_continuum`.
19+
20+
For the proof, we rely on an explicit inductive construction of the sigma-algebra generated by
21+
`s` (instead of the inductive predicate `generate_measurable`). This transfinite inductive
22+
construction is parameterized by an ordinal `< ω₁`, and the cardinality bound is preserved along
23+
each step of the construction.
24+
-/
25+
26+
universe u
27+
variables {α : Type u}
28+
29+
open_locale cardinal
30+
open cardinal set
31+
32+
local notation a `<₁` b := (aleph 1 : cardinal.{u}).ord.out.r a b
33+
local notation `ω₁`:= (aleph 1 : cardinal.{u}).ord.out.α
34+
35+
namespace measurable_space
36+
37+
/-- Transfinite induction construction of the sigma-algebra generated by a set of sets `s`. At each
38+
step, we add all elements of `s`, the empty set, the complements of already constructed sets, and
39+
countable unions of already constructed sets. We index this construction by an ordinal `< ω₁`, as
40+
this will be enough to generate all sets in the sigma-algebra. -/
41+
def generate_measurable_rec (s : set (set α)) : ω₁ → set (set α)
42+
| i := let S := ⋃ j : {j // j <₁ i}, generate_measurable_rec j.1 in
43+
s ∪ {∅} ∪ compl '' S ∪ (set.range (λ (f : ℕ → S), ⋃ n, (f n).1))
44+
using_well_founded {dec_tac := `[exact j.2]}
45+
46+
/-- At each step of the inductive construction, the cardinality bound `≤ (max (#s) 2) ^ ω`
47+
holds. -/
48+
lemma cardinal_generate_measurable_rec_le (s : set (set α)) (i : ω₁) :
49+
#(generate_measurable_rec s i) ≤ (max (#s) 2) ^ omega.{u} :=
50+
begin
51+
apply (aleph 1).ord.out.wo.wf.induction i,
52+
assume i IH,
53+
have A := omega_le_aleph 1,
54+
have B : aleph 1 ≤ (max (#s) 2) ^ omega.{u} :=
55+
aleph_one_le_continuum.trans (power_le_power_right (le_max_right _ _)),
56+
have C : omega.{u} ≤ (max (#s) 2) ^ omega.{u} := A.trans B,
57+
have J : #(⋃ (j : {j // j <₁ i}), generate_measurable_rec s j.1) ≤ (max (#s) 2) ^ omega.{u},
58+
{ apply (mk_Union_le _).trans,
59+
have D : # {j // j <₁ i} ≤ aleph 1 := (mk_subtype_le _).trans (le_of_eq (aleph 1).mk_ord_out),
60+
have E : cardinal.sup.{u u}
61+
(λ (j : {j // j <₁ i}), #(generate_measurable_rec s j.1)) ≤ (max (#s) 2) ^ omega.{u} :=
62+
cardinal.sup_le.2 (λ ⟨j, hj⟩, IH j hj),
63+
apply (mul_le_mul' D E).trans,
64+
rw mul_eq_max A C,
65+
exact max_le B le_rfl },
66+
rw [generate_measurable_rec],
67+
apply_rules [(mk_union_le _ _).trans, add_le_of_le C, mk_image_le.trans],
68+
{ exact (le_max_left _ _).trans (self_le_power _ one_lt_omega.le) },
69+
{ rw [mk_singleton],
70+
exact one_lt_omega.le.trans C },
71+
{ apply mk_range_le.trans,
72+
simp only [mk_pi, subtype.val_eq_coe, prod_const, lift_uzero, mk_denumerable, lift_omega],
73+
have := @power_le_power_right _ _ omega.{u} J,
74+
rwa [← power_mul, omega_mul_omega] at this }
75+
end
76+
77+
lemma cardinal_Union_generate_measurable_rec_le (s : set (set α)) :
78+
#(⋃ i, generate_measurable_rec s i) ≤ (max (#s) 2) ^ omega.{u} :=
79+
begin
80+
apply (mk_Union_le _).trans,
81+
rw [(aleph 1).mk_ord_out],
82+
refine le_trans (mul_le_mul' aleph_one_le_continuum
83+
(cardinal.sup_le.2 (λ i, cardinal_generate_measurable_rec_le s i))) _,
84+
have := power_le_power_right (le_max_right (#s) 2),
85+
rw mul_eq_max omega_le_continuum (omega_le_continuum.trans this),
86+
exact max_le this le_rfl
87+
end
88+
89+
/-- A set in the sigma-algebra generated by a set of sets `s` is constructed at some step of the
90+
transfinite induction defining `generate_measurable_rec`.
91+
The other inclusion is also true, but not proved here as it is not needed for the cardinality
92+
bounds.-/
93+
theorem generate_measurable_subset_rec (s : set (set α)) ⦃t : set α⦄
94+
(ht : generate_measurable s t) :
95+
t ∈ ⋃ i, generate_measurable_rec s i :=
96+
begin
97+
haveI : nonempty ω₁, by simp [← mk_ne_zero_iff, ne_of_gt, (aleph 1).mk_ord_out, aleph_pos 1],
98+
inhabit ω₁,
99+
induction ht with u hu u hu IH f hf IH,
100+
{ refine mem_Union.2 ⟨default, _⟩,
101+
rw generate_measurable_rec,
102+
simp only [hu, mem_union_eq, true_or] },
103+
{ refine mem_Union.2 ⟨default, _⟩,
104+
rw generate_measurable_rec,
105+
simp only [union_singleton, mem_union_eq, mem_insert_iff, eq_self_iff_true, true_or] },
106+
{ rcases mem_Union.1 IH with ⟨i, hi⟩,
107+
obtain ⟨j, hj⟩ : ∃ j, i <₁ j := ordinal.has_succ_of_is_limit
108+
(by { rw ordinal.type_out, exact ord_aleph_is_limit 1 }) _,
109+
apply mem_Union.2 ⟨j, _⟩,
110+
rw generate_measurable_rec,
111+
have : ∃ a, (a <₁ j) ∧ u ∈ generate_measurable_rec s a := ⟨i, hj, hi⟩,
112+
simp [this] },
113+
{ have : ∀ n, ∃ i, f n ∈ generate_measurable_rec s i := λ n, by simpa using IH n,
114+
choose I hI using this,
115+
obtain ⟨j, hj⟩ : ∃ j, ∀ k, k ∈ range I → (k <₁ j),
116+
{ apply ordinal.lt_cof_type,
117+
simp only [is_regular_aleph_one.2, mk_singleton, ordinal.type_out],
118+
have : #(range I) = lift.{0} (#(range I)), by simp only [lift_uzero],
119+
rw this,
120+
apply mk_range_le_lift.trans_lt _,
121+
simp [omega_lt_aleph_one] },
122+
apply mem_Union.2 ⟨j, _⟩,
123+
rw generate_measurable_rec,
124+
have : ∃ (g : ℕ → (↥⋃ (i : {i // i <₁ j}), generate_measurable_rec s i.1)),
125+
(⋃ (n : ℕ), ↑(g n)) = (⋃ n, f n),
126+
{ refine ⟨λ n, ⟨f n, _⟩, rfl⟩,
127+
exact mem_Union.2 ⟨⟨I n, hj (I n) (mem_range_self _)⟩, hI n⟩ },
128+
simp [this] }
129+
end
130+
131+
/-- If a sigma-algebra is generated by a set of sets `s`, then the sigma
132+
algebra has cardinality at most `(max (#s) 2) ^ ω`. -/
133+
theorem cardinal_generate_measurable_le (s : set (set α)) :
134+
#{t | generate_measurable s t} ≤ (max (#s) 2) ^ omega.{u} :=
135+
(mk_subtype_le_of_subset (generate_measurable_subset_rec s)).trans
136+
(cardinal_Union_generate_measurable_rec_le s)
137+
138+
/-- If a sigma-algebra is generated by a set of sets `s`, then the sigma
139+
algebra has cardinality at most `(max (#s) 2) ^ ω`. -/
140+
theorem cardinal_measurable_set_le (s : set (set α)) :
141+
#{t | @measurable_set α (generate_from s) t} ≤ (max (#s) 2) ^ omega.{u} :=
142+
cardinal_generate_measurable_le s
143+
144+
/-- If a sigma-algebra is generated by a set of sets `s` with cardinality at most the continuum,
145+
then the sigma algebra has the same cardinality bound. -/
146+
theorem cardinal_generate_measurable_le_continuum {s : set (set α)} (hs : #s ≤ 𝔠) :
147+
#{t | generate_measurable s t} ≤ 𝔠 :=
148+
calc
149+
#{t | generate_measurable s t}
150+
≤ (max (#s) 2) ^ omega.{u} : cardinal_generate_measurable_le s
151+
... ≤ 𝔠 ^ omega.{u} :
152+
by exact_mod_cast power_le_power_right (max_le hs (nat_lt_continuum 2).le)
153+
... = 𝔠 : continuum_power_omega
154+
155+
/-- If a sigma-algebra is generated by a set of sets `s` with cardinality at most the continuum,
156+
then the sigma algebra has the same cardinality bound. -/
157+
theorem cardinal_measurable_set_le_continuum {s : set (set α)} (hs : #s ≤ 𝔠) :
158+
#{t | @measurable_set α (generate_from s) t} ≤ 𝔠 :=
159+
cardinal_generate_measurable_le_continuum hs
160+
161+
end measurable_space

src/set_theory/cardinal.lean

Lines changed: 8 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -460,11 +460,18 @@ lt_of_le_of_ne (zero_le _) zero_ne_one
460460
lemma zero_power_le (c : cardinal.{u}) : (0 : cardinal.{u}) ^ c ≤ 1 :=
461461
by { by_cases h : c = 0, rw [h, power_zero], rw [zero_power h], apply zero_le }
462462

463-
theorem power_le_power_left : ∀{a b c : cardinal}, a ≠ 0 → b ≤ c → a ^ b ≤ a ^ c :=
463+
theorem power_le_power_left : ∀ {a b c : cardinal}, a ≠ 0 → b ≤ c → a ^ b ≤ a ^ c :=
464464
by rintros ⟨α⟩ ⟨β⟩ ⟨γ⟩ hα ⟨e⟩; exact
465465
let ⟨a⟩ := mk_ne_zero_iff.1in
466466
⟨@embedding.arrow_congr_left _ _ _ ⟨a⟩ e⟩
467467

468+
theorem self_le_power (a : cardinal) {b : cardinal} (hb : 1 ≤ b) : a ≤ a ^ b :=
469+
begin
470+
rcases eq_or_ne a 0 with rfl|ha,
471+
{ exact zero_le _ },
472+
{ convert power_le_power_left ha hb, exact power_one.symm }
473+
end
474+
468475
/-- **Cantor's theorem** -/
469476
theorem cantor (a : cardinal.{u}) : a < 2 ^ a :=
470477
begin

src/set_theory/cardinal_ordinal.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -232,6 +232,9 @@ aleph'_is_normal.trans $ add_is_normal ordinal.omega
232232
theorem succ_omega : succ ω = aleph 1 :=
233233
by rw [← aleph_zero, ← aleph_succ, ordinal.succ_zero]
234234

235+
lemma omega_lt_aleph_one : ω < aleph 1 :=
236+
by { rw ← succ_omega, exact lt_succ_self _ }
237+
235238
lemma countable_iff_lt_aleph_one {α : Type*} (s : set α) : countable s ↔ #s < aleph 1 :=
236239
by rw [← succ_omega, lt_succ, mk_set_le_omega]
237240

@@ -456,6 +459,10 @@ begin
456459
{ exact le_max_of_le_right (le_of_lt (add_lt_omega (lt_of_not_ge ha) (lt_of_not_ge hb))) } }
457460
end
458461

462+
theorem add_le_of_le {a b c : cardinal} (hc : ω ≤ c)
463+
(h1 : a ≤ c) (h2 : b ≤ c) : a + b ≤ c :=
464+
(add_le_add h1 h2).trans $ le_of_eq $ add_eq_self hc
465+
459466
theorem add_lt_of_lt {a b c : cardinal} (hc : ω ≤ c)
460467
(h1 : a < c) (h2 : b < c) : a + b < c :=
461468
lt_of_le_of_lt (add_le_add (le_max_left a b) (le_max_right a b)) $

src/set_theory/cofinality.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -542,6 +542,9 @@ theorem succ_is_regular {c : cardinal.{u}} (h : ω ≤ c) : is_regular (succ c)
542542
apply typein_lt_type }
543543
end
544544

545+
theorem is_regular_aleph_one : is_regular (aleph 1) :=
546+
by { rw ← succ_omega, exact succ_is_regular le_rfl }
547+
545548
theorem aleph'_succ_is_regular {o : ordinal} (h : ordinal.omega ≤ o) : is_regular (aleph' o.succ) :=
546549
by { rw aleph'_succ, exact succ_is_regular (omega_le_aleph'.2 h) }
547550

src/set_theory/continuum.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -48,6 +48,9 @@ lemma continuum_pos : 0 < 𝔠 := nat_lt_continuum 0
4848

4949
lemma continuum_ne_zero : 𝔠 ≠ 0 := continuum_pos.ne'
5050

51+
lemma aleph_one_le_continuum : aleph 1 ≤ 𝔠 :=
52+
by { rw ← succ_omega, exact succ_le.2 omega_lt_continuum }
53+
5154
/-!
5255
### Addition
5356
-/

src/set_theory/ordinal.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -474,6 +474,8 @@ instance ordinal.is_equivalent : setoid Well_order :=
474474
/-- `ordinal.{u}` is the type of well orders in `Type u`, up to order isomorphism. -/
475475
def ordinal : Type (u + 1) := quotient ordinal.is_equivalent
476476

477+
instance (o : ordinal) : has_well_founded o.out.α := ⟨o.out.r, o.out.wo.wf⟩
478+
477479
namespace ordinal
478480

479481
/-- The order type of a well order is an ordinal. -/

0 commit comments

Comments
 (0)