@@ -25,147 +25,188 @@ indeed generates this sigma-algebra.
25
25
-/
26
26
27
27
28
- universe u
28
+ universe u v
29
29
30
30
variable {α : Type u}
31
31
32
- open Cardinal Set
33
-
34
- -- Porting note: fix universe below, not here
35
- local notation "ω₁" => (Ordinal.toType <| Cardinal.ord (aleph 1 ))
32
+ open Cardinal Ordinal Set MeasureTheory
36
33
37
34
namespace MeasurableSpace
38
35
39
36
/-- Transfinite induction construction of the sigma-algebra generated by a set of sets `s`. At each
40
37
step, we add all elements of `s`, the empty set, the complements of already constructed sets, and
41
- countable unions of already constructed sets. We index this construction by an ordinal `< ω₁`, as
42
- this will be enough to generate all sets in the sigma-algebra.
38
+ countable unions of already constructed sets.
39
+
40
+ We index this construction by an arbitrary ordinal for simplicity, but by `ω₁` we will have
41
+ generated all the sets in the sigma-algebra.
43
42
44
43
This construction is very similar to that of the Borel hierarchy. -/
45
- def generateMeasurableRec (s : Set (Set α)) (i : (ω₁ : Type u) ) : Set (Set α) :=
46
- let S := ⋃ j : Iio i, generateMeasurableRec s (j. 1 )
44
+ def generateMeasurableRec (s : Set (Set α)) (i : Ordinal ) : Set (Set α) :=
45
+ let S := ⋃ j < i, generateMeasurableRec s j
47
46
s ∪ {∅} ∪ compl '' S ∪ Set.range fun f : ℕ → S => ⋃ n, (f n).1
48
47
termination_by i
49
- decreasing_by exact j.2
50
48
51
- theorem self_subset_generateMeasurableRec (s : Set (Set α)) (i : ω₁ ) :
49
+ theorem self_subset_generateMeasurableRec (s : Set (Set α)) (i : Ordinal ) :
52
50
s ⊆ generateMeasurableRec s i := by
53
51
unfold generateMeasurableRec
54
52
apply_rules [subset_union_of_subset_left]
55
53
exact subset_rfl
56
54
57
- theorem empty_mem_generateMeasurableRec (s : Set (Set α)) (i : ω₁ ) :
55
+ theorem empty_mem_generateMeasurableRec (s : Set (Set α)) (i : Ordinal ) :
58
56
∅ ∈ generateMeasurableRec s i := by
59
57
unfold generateMeasurableRec
60
58
exact mem_union_left _ (mem_union_left _ (mem_union_right _ (mem_singleton ∅)))
61
59
62
- theorem compl_mem_generateMeasurableRec {s : Set (Set α)} {i j : ω₁ } (h : j < i) {t : Set α}
60
+ theorem compl_mem_generateMeasurableRec {s : Set (Set α)} {i j : Ordinal } (h : j < i) {t : Set α}
63
61
(ht : t ∈ generateMeasurableRec s j) : tᶜ ∈ generateMeasurableRec s i := by
64
62
unfold generateMeasurableRec
65
- exact mem_union_left _ (mem_union_right _ ⟨t, mem_iUnion.2 ⟨⟨ j, h⟩ , ht⟩, rfl⟩)
63
+ exact mem_union_left _ (mem_union_right _ ⟨t, mem_iUnion₂ .2 ⟨j, h, ht⟩, rfl⟩)
66
64
67
- theorem iUnion_mem_generateMeasurableRec {s : Set (Set α)} {i : ω₁ } {f : ℕ → Set α}
65
+ theorem iUnion_mem_generateMeasurableRec {s : Set (Set α)} {i : Ordinal } {f : ℕ → Set α}
68
66
(hf : ∀ n, ∃ j < i, f n ∈ generateMeasurableRec s j) :
69
- ( ⋃ n, f n) ∈ generateMeasurableRec s i := by
67
+ ⋃ n, f n ∈ generateMeasurableRec s i := by
70
68
unfold generateMeasurableRec
71
- exact mem_union_right _ ⟨fun n => ⟨f n, let ⟨j, hj, hf⟩ := hf n; mem_iUnion.2 ⟨⟨ j, hj⟩ , hf⟩⟩, rfl⟩
69
+ exact mem_union_right _ ⟨fun n => ⟨f n, let ⟨j, hj, hf⟩ := hf n; mem_iUnion₂ .2 ⟨j, hj, hf⟩⟩, rfl⟩
72
70
73
- theorem generateMeasurableRec_subset (s : Set (Set α)) {i j : ω₁} (h : i ≤ j ) :
74
- generateMeasurableRec s i ⊆ generateMeasurableRec s j := fun x hx => by
75
- rcases eq_or_lt_of_le h with (rfl | h)
71
+ theorem generateMeasurableRec_mono (s : Set (Set α)) : Monotone (generateMeasurableRec s ) := by
72
+ intro i j h x hx
73
+ rcases h.eq_or_lt with (rfl | h)
76
74
· exact hx
77
75
· convert iUnion_mem_generateMeasurableRec fun _ => ⟨i, h, hx⟩
78
76
exact (iUnion_const x).symm
79
77
80
- /-- At each step of the inductive construction, the cardinality bound `≤ (max #s 2) ^ ℵ₀` holds. -/
81
- theorem cardinal_generateMeasurableRec_le (s : Set (Set α)) (i : ω₁) :
78
+ /-- An inductive principle for the elements of `generateMeasurableRec`. -/
79
+ @[elab_as_elim]
80
+ theorem generateMeasurableRec_induction {s : Set (Set α)} {i : Ordinal} {t : Set α}
81
+ {p : Set α → Prop } (hs : ∀ t ∈ s, p t) (h0 : p ∅)
82
+ (hc : ∀ u, p u → (∃ j < i, u ∈ generateMeasurableRec s j) → p uᶜ)
83
+ (hn : ∀ f : ℕ → Set α,
84
+ (∀ n, p (f n) ∧ ∃ j < i, f n ∈ generateMeasurableRec s j) → p (⋃ n, f n)) :
85
+ t ∈ generateMeasurableRec s i → p t := by
86
+ suffices H : ∀ k ≤ i, ∀ t ∈ generateMeasurableRec s k, p t from H i le_rfl t
87
+ intro k
88
+ apply WellFoundedLT.induction k
89
+ intro k IH hk t
90
+ replace IH := fun j hj => IH j hj (hj.le.trans hk)
91
+ unfold generateMeasurableRec
92
+ rintro (((ht | rfl) | ht) | ⟨f, rfl⟩)
93
+ · exact hs t ht
94
+ · exact h0
95
+ · simp_rw [mem_image, mem_iUnion₂] at ht
96
+ obtain ⟨u, ⟨⟨j, hj, hj'⟩, rfl⟩⟩ := ht
97
+ exact hc u (IH j hj u hj') ⟨j, hj.trans_le hk, hj'⟩
98
+ · apply hn
99
+ intro n
100
+ obtain ⟨j, hj, hj'⟩ := mem_iUnion₂.1 (f n).2
101
+ use IH j hj _ hj', j, hj.trans_le hk
102
+
103
+ theorem generateMeasurableRec_omega1 (s : Set (Set α)) :
104
+ generateMeasurableRec s (ω₁ : Ordinal.{v}) =
105
+ ⋃ i < (ω₁ : Ordinal.{v}), generateMeasurableRec s i := by
106
+ apply (iUnion₂_subset fun i h => generateMeasurableRec_mono s h.le).antisymm'
107
+ intro t ht
108
+ rw [mem_iUnion₂]
109
+ refine generateMeasurableRec_induction ?_ ?_ ?_ ?_ ht
110
+ · intro t ht
111
+ exact ⟨0 , omega_pos 1 , self_subset_generateMeasurableRec s 0 ht⟩
112
+ · exact ⟨0 , omega_pos 1 , empty_mem_generateMeasurableRec s 0 ⟩
113
+ · rintro u - ⟨j, hj, hj'⟩
114
+ exact ⟨_, (isLimit_omega 1 ).succ_lt hj,
115
+ compl_mem_generateMeasurableRec (Order.lt_succ j) hj'⟩
116
+ · intro f H
117
+ choose I hI using fun n => (H n).1
118
+ simp_rw [exists_prop] at hI
119
+ refine ⟨_, Ordinal.lsub_lt_ord_lift ?_ fun n => (hI n).1 ,
120
+ iUnion_mem_generateMeasurableRec fun n => ⟨_, Ordinal.lt_lsub I n, (hI n).2 ⟩⟩
121
+ rw [mk_nat, lift_aleph0, isRegular_aleph_one.cof_omega_eq]
122
+ exact aleph0_lt_aleph_one
123
+
124
+ theorem generateMeasurableRec_subset (s : Set (Set α)) (i : Ordinal) :
125
+ generateMeasurableRec s i ⊆ { t | GenerateMeasurable s t } := by
126
+ apply WellFoundedLT.induction i
127
+ exact fun i IH t ht => generateMeasurableRec_induction .basic .empty
128
+ (fun u _ ⟨j, hj, hj'⟩ => .compl _ (IH j hj hj')) (fun f H => .iUnion _ fun n => (H n).1 ) ht
129
+
130
+ /-- `generateMeasurableRec s ω₁` generates precisely the smallest sigma-algebra containing `s`. -/
131
+ theorem generateMeasurable_eq_rec (s : Set (Set α)) :
132
+ { t | GenerateMeasurable s t } = generateMeasurableRec s ω₁ := by
133
+ apply (generateMeasurableRec_subset s _).antisymm'
134
+ intro t ht
135
+ induction' ht with u hu u _ IH f _ IH
136
+ · exact self_subset_generateMeasurableRec s _ hu
137
+ · exact empty_mem_generateMeasurableRec s _
138
+ · rw [generateMeasurableRec_omega1, mem_iUnion₂] at IH
139
+ obtain ⟨i, hi, hi'⟩ := IH
140
+ exact generateMeasurableRec_mono _ ((isLimit_omega 1 ).succ_lt hi).le
141
+ (compl_mem_generateMeasurableRec (Order.lt_succ i) hi')
142
+ · simp_rw [generateMeasurableRec_omega1, mem_iUnion₂, exists_prop] at IH
143
+ exact iUnion_mem_generateMeasurableRec IH
144
+
145
+ /-- `generateMeasurableRec` is constant for ordinals `≥ ω₁`. -/
146
+ theorem generateMeasurableRec_of_omega1_le (s : Set (Set α)) {i : Ordinal.{v}} (hi : ω₁ ≤ i) :
147
+ generateMeasurableRec s i = generateMeasurableRec s (ω₁ : Ordinal.{v}) := by
148
+ apply (generateMeasurableRec_mono s hi).antisymm'
149
+ rw [← generateMeasurable_eq_rec]
150
+ exact generateMeasurableRec_subset s i
151
+
152
+ /-- At each step of the inductive construction, the cardinality bound `≤ #s ^ ℵ₀` holds. -/
153
+ theorem cardinal_generateMeasurableRec_le (s : Set (Set α)) (i : Ordinal.{v}) :
82
154
#(generateMeasurableRec s i) ≤ max #s 2 ^ ℵ₀ := by
155
+ suffices ∀ i ≤ ω₁, #(generateMeasurableRec s i) ≤ max #s 2 ^ ℵ₀ by
156
+ obtain hi | hi := le_or_lt i ω₁
157
+ · exact this i hi
158
+ · rw [generateMeasurableRec_of_omega1_le s hi.le]
159
+ exact this _ le_rfl
160
+ intro i
83
161
apply WellFoundedLT.induction i
84
- intro i IH
85
- have A := aleph0_le_aleph 1
86
- have B : aleph 1 ≤ max #s 2 ^ ℵ₀ :=
87
- aleph_one_le_continuum.trans (power_le_power_right (le_max_right _ _))
88
- have C : ℵ₀ ≤ max #s 2 ^ ℵ₀ := A.trans B
89
- have J : #(⋃ j : Iio i, generateMeasurableRec s j.1 ) ≤ max #s 2 ^ ℵ₀ := by
90
- refine (mk_iUnion_le _).trans ?_
91
- have D : ⨆ j : Iio i, #(generateMeasurableRec s j) ≤ _ := ciSup_le' fun ⟨j, hj⟩ => IH j hj
92
- apply (mul_le_mul' ((mk_subtype_le _).trans (aleph 1 ).mk_ord_toType.le) D).trans
93
- rw [mul_eq_max A C]
94
- exact max_le B le_rfl
162
+ intro i IH hi
163
+ have A : 𝔠 ≤ max #s 2 ^ ℵ₀ := power_le_power_right (le_max_right _ _)
164
+ have B := aleph0_le_continuum.trans A
165
+ have C : #(⋃ j < i, generateMeasurableRec s j) ≤ max #s 2 ^ ℵ₀ := by
166
+ apply mk_iUnion_Ordinal_lift_le_of_le _ B _
167
+ · intro j hj
168
+ exact IH j hj (hj.trans_le hi).le
169
+ · rw [lift_power, lift_aleph0]
170
+ rw [← Ordinal.lift_le.{u}, lift_omega, Ordinal.lift_one, ← ord_aleph] at hi
171
+ have H := card_le_of_le_ord hi
172
+ rw [← Ordinal.lift_card] at H
173
+ apply H.trans <| aleph_one_le_continuum.trans <| power_le_power_right _
174
+ rw [lift_max, Cardinal.lift_ofNat]
175
+ exact le_max_right _ _
95
176
rw [generateMeasurableRec]
96
- apply_rules [(mk_union_le _ _).trans, add_le_of_le C, mk_image_le.trans]
97
- · exact (le_max_left _ _).trans (self_le_power _ one_lt_aleph0.le)
177
+ apply_rules [(mk_union_le _ _).trans, add_le_of_le (aleph_one_le_continuum.trans A),
178
+ mk_image_le.trans]
179
+ · exact (self_le_power _ one_le_aleph0).trans (power_le_power_right (le_max_left _ _))
98
180
· rw [mk_singleton]
99
- exact one_lt_aleph0.le.trans C
181
+ exact one_lt_aleph0.le.trans B
100
182
· apply mk_range_le.trans
101
- simp only [mk_pi, prod_const, lift_uzero, mk_denumerable, lift_aleph0]
102
- have := @power_le_power_right _ _ ℵ₀ J
183
+ simp only [mk_pi, prod_const, Cardinal. lift_uzero, mk_denumerable, lift_aleph0]
184
+ have := @power_le_power_right _ _ ℵ₀ C
103
185
rwa [← power_mul, aleph0_mul_aleph0] at this
104
186
105
- /-- `generateMeasurableRec s` generates precisely the smallest sigma-algebra containing `s`. -/
106
- theorem generateMeasurable_eq_rec (s : Set (Set α)) :
107
- { t | GenerateMeasurable s t } =
108
- ⋃ (i : (aleph 1 ).ord.toType), generateMeasurableRec s i := by
109
- ext t; refine ⟨fun ht => ?_, fun ht => ?_⟩
110
- · inhabit ω₁
111
- induction ht with
112
- | basic u hu => exact mem_iUnion.2 ⟨default, self_subset_generateMeasurableRec s _ hu⟩
113
- | empty => exact mem_iUnion.2 ⟨default, empty_mem_generateMeasurableRec s _⟩
114
- | compl _ _ IH =>
115
- rcases mem_iUnion.1 IH with ⟨i, hi⟩
116
- obtain ⟨j, hj⟩ := exists_gt i
117
- exact mem_iUnion.2 ⟨j, compl_mem_generateMeasurableRec hj hi⟩
118
- | iUnion f _ IH =>
119
- have : ∀ n, ∃ i, f n ∈ generateMeasurableRec s i := fun n => by simpa using IH n
120
- choose I hI using this
121
- refine mem_iUnion.2 ⟨Ordinal.enum (α := ω₁) (· < ·)
122
- ⟨Ordinal.lsub fun n => Ordinal.typein.{u} (α := ω₁) (· < ·) (I n), ?_⟩,
123
- iUnion_mem_generateMeasurableRec fun n => ⟨I n, ?_, hI n⟩⟩
124
- · rw [Ordinal.type_toType]
125
- refine Ordinal.lsub_lt_ord_lift ?_ fun i => Ordinal.typein_lt_self _
126
- rw [mk_denumerable, lift_aleph0, isRegular_aleph_one.cof_eq]
127
- exact aleph0_lt_aleph_one
128
- · rw [← Ordinal.typein_lt_typein (· < ·), Ordinal.typein_enum]
129
- apply Ordinal.lt_lsub fun n : ℕ => _
130
- · rcases ht with ⟨t, ⟨i, rfl⟩, hx⟩
131
- revert t
132
- apply WellFoundedLT.induction i
133
- intro j H t ht
134
- unfold generateMeasurableRec at ht
135
- rcases ht with (((h | (rfl : t = ∅)) | ⟨u, ⟨-, ⟨⟨k, hk⟩, rfl⟩, hu⟩, rfl⟩) | ⟨f, rfl⟩)
136
- · exact .basic t h
137
- · exact .empty
138
- · exact .compl u (H k hk u hu)
139
- · refine .iUnion _ @fun n => ?_
140
- obtain ⟨-, ⟨⟨k, hk⟩, rfl⟩, hf⟩ := (f n).prop
141
- exact H k hk _ hf
142
-
143
187
/-- If a sigma-algebra is generated by a set of sets `s`, then the sigma-algebra has cardinality at
144
- most `( max #s 2) ^ ℵ₀`. -/
188
+ most `max #s 2 ^ ℵ₀`. -/
145
189
theorem cardinal_generateMeasurable_le (s : Set (Set α)) :
146
190
#{ t | GenerateMeasurable s t } ≤ max #s 2 ^ ℵ₀ := by
147
- rw [generateMeasurable_eq_rec]
148
- apply (mk_iUnion_le _).trans
149
- rw [(aleph 1 ).mk_ord_toType]
150
- refine le_trans (mul_le_mul' aleph_one_le_continuum
151
- (ciSup_le' fun i => cardinal_generateMeasurableRec_le s i)) ?_
152
- refine (mul_le_max_of_aleph0_le_left aleph0_le_continuum).trans (max_le ?_ le_rfl)
153
- exact power_le_power_right (le_max_right _ _)
191
+ rw [generateMeasurable_eq_rec.{u, 0 }]
192
+ exact cardinal_generateMeasurableRec_le s _
154
193
155
194
/-- If a sigma-algebra is generated by a set of sets `s`, then the sigma
156
- algebra has cardinality at most `( max #s 2) ^ ℵ₀`. -/
157
- theorem cardinalMeasurableSet_le (s : Set (Set α)) :
195
+ algebra has cardinality at most `max #s 2 ^ ℵ₀`. -/
196
+ theorem cardinal_measurableSet_le (s : Set (Set α)) :
158
197
#{ t | @MeasurableSet α (generateFrom s) t } ≤ max #s 2 ^ ℵ₀ :=
159
198
cardinal_generateMeasurable_le s
160
199
200
+ @[deprecated cardinal_measurableSet_le (since := "2024-08-30")]
201
+ alias cardinalMeasurableSet_le := cardinal_measurableSet_le
202
+
161
203
/-- If a sigma-algebra is generated by a set of sets `s` with cardinality at most the continuum,
162
204
then the sigma algebra has the same cardinality bound. -/
163
205
theorem cardinal_generateMeasurable_le_continuum {s : Set (Set α)} (hs : #s ≤ 𝔠) :
164
- #{ t | GenerateMeasurable s t } ≤ 𝔠 :=
165
- (cardinal_generateMeasurable_le s).trans
166
- (by
167
- rw [← continuum_power_aleph0]
168
- exact mod_cast power_le_power_right (max_le hs (nat_lt_continuum 2 ).le))
206
+ #{ t | GenerateMeasurable s t } ≤ 𝔠 := by
207
+ apply (cardinal_generateMeasurable_le s).trans
208
+ rw [← continuum_power_aleph0]
209
+ exact_mod_cast power_le_power_right (max_le hs (nat_lt_continuum 2 ).le)
169
210
170
211
/-- If a sigma-algebra is generated by a set of sets `s` with cardinality at most the continuum,
171
212
then the sigma algebra has the same cardinality bound. -/
0 commit comments