@@ -32,7 +32,7 @@ open_locale classical big_operators ennreal
32
32
33
33
open classical set filter
34
34
35
- variables {α β γ δ ε : Type *}
35
+ variables {α β : Type *}
36
36
37
37
namespace measure_theory
38
38
@@ -42,38 +42,38 @@ variables [measurable_space α] [measurable_space β]
42
42
43
43
/-- Measurability structure on `measure`: Measures are measurable w.r.t. all projections -/
44
44
instance : measurable_space (measure α) :=
45
- ⨆ (s : set α) (hs : measurable_set s), (borel ℝ≥0 ∞).comap (λμ , μ s)
45
+ ⨆ (s : set α) (hs : measurable_set s), (borel ℝ≥0 ∞).comap (λ μ , μ s)
46
46
47
- lemma measurable_coe {s : set α} (hs : measurable_set s) : measurable (λμ : measure α, μ s) :=
47
+ lemma measurable_coe {s : set α} (hs : measurable_set s) : measurable (λ μ : measure α, μ s) :=
48
48
measurable.of_comap_le $ le_supr_of_le s $ le_supr_of_le hs $ le_rfl
49
49
50
50
lemma measurable_of_measurable_coe (f : β → measure α)
51
- (h : ∀(s : set α) (hs : measurable_set s), measurable (λb , f b s)) :
51
+ (h : ∀ (s : set α) (hs : measurable_set s), measurable (λ b , f b s)) :
52
52
measurable f :=
53
53
measurable.of_le_map $ supr₂_le $ assume s hs, measurable_space.comap_le_iff_le_map.2 $
54
54
by rw [measurable_space.map_comp]; exact h s hs
55
55
56
56
lemma measurable_measure {μ : α → measure β} :
57
- measurable μ ↔ ∀(s : set β) (hs : measurable_set s), measurable (λb , μ b s) :=
57
+ measurable μ ↔ ∀ (s : set β) (hs : measurable_set s), measurable (λ b , μ b s) :=
58
58
⟨λ hμ s hs, (measurable_coe hs).comp hμ, measurable_of_measurable_coe μ⟩
59
59
60
60
lemma measurable_map (f : α → β) (hf : measurable f) :
61
- measurable (λμ : measure α, map f μ) :=
62
- measurable_of_measurable_coe _ $ assume s hs,
63
- suffices measurable (λ (μ : measure α), μ (f ⁻¹' s)),
64
- by simpa [map_apply, hs, hf],
65
- measurable_coe (hf hs)
61
+ measurable (λ μ : measure α, map f μ) :=
62
+ begin
63
+ refine measurable_of_measurable_coe _ (λ s hs, _),
64
+ simp_rw map_apply hf hs,
65
+ exact measurable_coe (hf hs),
66
+ end
66
67
67
- lemma measurable_dirac :
68
- measurable (measure.dirac : α → measure α) :=
69
- measurable_of_measurable_coe _ $ assume s hs,
70
- begin
71
- simp only [dirac_apply', hs],
72
- exact measurable_one.indicator hs
73
- end
68
+ lemma measurable_dirac : measurable (measure.dirac : α → measure α) :=
69
+ begin
70
+ refine measurable_of_measurable_coe _ (λ s hs, _),
71
+ simp_rw [dirac_apply' _ hs],
72
+ exact measurable_one.indicator hs
73
+ end
74
74
75
75
lemma measurable_lintegral {f : α → ℝ≥0 ∞} (hf : measurable f) :
76
- measurable (λμ : measure α, ∫⁻ x, f x ∂μ) :=
76
+ measurable (λ μ : measure α, ∫⁻ x, f x ∂μ) :=
77
77
begin
78
78
simp only [lintegral_eq_supr_eapprox_lintegral, hf, simple_func.lintegral],
79
79
refine measurable_supr (λ n, finset.measurable_sum _ (λ i _, _)),
85
85
functions. -/
86
86
def join (m : measure (measure α)) : measure α :=
87
87
measure.of_measurable
88
- (λs hs, ∫⁻ μ, μ s ∂m)
89
- (by simp)
88
+ (λ s hs, ∫⁻ μ, μ s ∂m)
89
+ (by simp only [measure_empty, lintegral_const, zero_mul] )
90
90
begin
91
91
assume f hf h,
92
- simp [measure_Union h hf],
92
+ simp_rw [measure_Union h hf],
93
93
apply lintegral_tsum,
94
94
assume i, exact measurable_coe (hf i)
95
95
end
96
96
97
- @[simp] lemma join_apply {m : measure (measure α)} :
98
- ∀{s : set α}, measurable_set s → join m s = ∫⁻ μ, μ s ∂m :=
99
- measure.of_measurable_apply
97
+ @[simp] lemma join_apply {m : measure (measure α)} {s : set α} (hs : measurable_set s) :
98
+ join m s = ∫⁻ μ, μ s ∂m :=
99
+ measure.of_measurable_apply s hs
100
100
101
101
@[simp] lemma join_zero : (0 : measure (measure α)).join = 0 :=
102
- by { ext1 s hs, simp [hs] }
102
+ by { ext1 s hs, simp only [hs, join_apply, lintegral_zero_measure, coe_zero, pi.zero_apply], }
103
103
104
104
lemma measurable_join : measurable (join : measure (measure α) → measure α) :=
105
105
measurable_of_measurable_coe _ $ assume s hs,
@@ -108,47 +108,23 @@ measurable_of_measurable_coe _ $ assume s hs,
108
108
lemma lintegral_join {m : measure (measure α)} {f : α → ℝ≥0 ∞} (hf : measurable f) :
109
109
∫⁻ x, f x ∂(join m) = ∫⁻ μ, ∫⁻ x, f x ∂μ ∂m :=
110
110
begin
111
- rw [lintegral_eq_supr_eapprox_lintegral hf],
112
- have : ∀n x,
113
- join m (⇑(simple_func.eapprox (λ (a : α), f a) n) ⁻¹' {x}) =
114
- ∫⁻ μ, μ ((⇑(simple_func.eapprox (λ (a : α), f a) n) ⁻¹' {x})) ∂m :=
115
- assume n x, join_apply (simple_func.measurable_set_preimage _ _),
116
- simp only [simple_func.lintegral, this ],
117
- transitivity,
118
- have : ∀(s : ℕ → finset ℝ≥0 ∞) (f : ℕ → ℝ≥0 ∞ → measure α → ℝ≥0 ∞)
119
- (hf : ∀n r, measurable (f n r)) (hm : monotone (λn μ, ∑ r in s n, r * f n r μ)),
120
- (⨆n:ℕ, ∑ r in s n, r * ∫⁻ μ, f n r μ ∂m) =
121
- ∫⁻ μ, ⨆n:ℕ, ∑ r in s n, r * f n r μ ∂m,
122
- { assume s f hf hm,
123
- symmetry,
124
- transitivity,
125
- apply lintegral_supr,
126
- { assume n,
127
- exact finset.measurable_sum _ (assume r _, (hf _ _).const_mul _) },
128
- { exact hm },
129
- congr, funext n,
130
- transitivity,
131
- apply lintegral_finset_sum,
132
- { assume r _, exact (hf _ _).const_mul _ },
133
- congr, funext r,
134
- apply lintegral_const_mul,
135
- exact hf _ _ },
136
- specialize this (λn, simple_func.range (simple_func.eapprox f n)),
137
- specialize this
138
- (λn r μ, μ (⇑(simple_func.eapprox (λ (a : α), f a) n) ⁻¹' {r})),
139
- refine this _ _; clear this ,
140
- { assume n r,
141
- apply measurable_coe,
142
- exact simple_func.measurable_set_preimage _ _ },
143
- { change monotone (λn μ, (simple_func.eapprox f n).lintegral μ),
144
- assume n m h μ,
145
- refine simple_func.lintegral_mono _ le_rfl,
146
- apply simple_func.monotone_eapprox,
147
- assumption },
148
- congr, funext μ,
149
- symmetry,
150
- apply lintegral_eq_supr_eapprox_lintegral,
151
- exact hf
111
+ simp_rw [lintegral_eq_supr_eapprox_lintegral hf,
112
+ simple_func.lintegral, join_apply (simple_func.measurable_set_preimage _ _)],
113
+ suffices : ∀ (s : ℕ → finset ℝ≥0 ∞) (f : ℕ → ℝ≥0 ∞ → measure α → ℝ≥0 ∞)
114
+ (hf : ∀ n r, measurable (f n r)) (hm : monotone (λ n μ, ∑ r in s n, r * f n r μ)),
115
+ (⨆ n, ∑ r in s n, r * ∫⁻ μ, f n r μ ∂m) = ∫⁻ μ, ⨆ n, ∑ r in s n, r * f n r μ ∂m,
116
+ { refine this (λ n, simple_func.range (simple_func.eapprox f n))
117
+ (λ n r μ, μ (simple_func.eapprox f n ⁻¹' {r})) _ _,
118
+ { exact λ n r, measurable_coe (simple_func.measurable_set_preimage _ _), },
119
+ { exact λ n m h μ, simple_func.lintegral_mono (simple_func.monotone_eapprox _ h) le_rfl, }, },
120
+ intros s f hf hm,
121
+ rw lintegral_supr _ hm,
122
+ swap, { exact λ n, finset.measurable_sum _ (λ r _, (hf _ _).const_mul _) },
123
+ congr,
124
+ funext n,
125
+ rw lintegral_finset_sum (s n),
126
+ { simp_rw lintegral_const_mul _ (hf _ _), },
127
+ { exact λ r _, (hf _ _).const_mul _ },
152
128
end
153
129
154
130
/-- Monadic bind on `measure`, only works in the category of measurable spaces and measurable
@@ -164,7 +140,7 @@ begin
164
140
ext1 s hs,
165
141
simp only [bind, hs, join_apply, coe_zero, pi.zero_apply],
166
142
rw [lintegral_map (measurable_coe hs) measurable_zero],
167
- simp
143
+ simp only [pi.zero_apply, coe_zero, lintegral_const, zero_mul],
168
144
end
169
145
170
146
@[simp] lemma bind_zero_right' (m : measure α) :
@@ -176,44 +152,44 @@ bind_zero_right m
176
152
bind m f s = ∫⁻ a, f a s ∂m :=
177
153
by rw [bind, join_apply hs, lintegral_map (measurable_coe hs) hf]
178
154
179
- lemma measurable_bind' {g : α → measure β} (hg : measurable g) : measurable (λm , bind m g) :=
155
+ lemma measurable_bind' {g : α → measure β} (hg : measurable g) : measurable (λ m , bind m g) :=
180
156
measurable_join.comp (measurable_map _ hg)
181
157
182
158
lemma lintegral_bind {m : measure α} {μ : α → measure β} {f : β → ℝ≥0 ∞}
183
159
(hμ : measurable μ) (hf : measurable f) :
184
- ∫⁻ x, f x ∂ (bind m μ) = ∫⁻ a, ∫⁻ x, f x ∂(μ a) ∂m:=
160
+ ∫⁻ x, f x ∂ (bind m μ) = ∫⁻ a, ∫⁻ x, f x ∂(μ a) ∂m :=
185
161
(lintegral_join hf).trans (lintegral_map (measurable_lintegral hf) hμ)
186
162
187
163
lemma bind_bind {γ} [measurable_space γ] {m : measure α} {f : α → measure β} {g : β → measure γ}
188
164
(hf : measurable f) (hg : measurable g) :
189
- bind (bind m f) g = bind m (λa, bind (f a) g) :=
190
- measure.ext $ assume s hs,
165
+ bind (bind m f) g = bind m (λ a, bind (f a) g) :=
191
166
begin
192
- rw [bind_apply hs hg, bind_apply hs ((measurable_bind' hg).comp hf), lintegral_bind hf],
193
- { congr, funext a,
194
- exact (bind_apply hs hg).symm },
195
- exact (measurable_coe hs).comp hg
167
+ ext1 s hs,
168
+ simp_rw [bind_apply hs hg, bind_apply hs ((measurable_bind' hg).comp hf),
169
+ lintegral_bind hf ((measurable_coe hs).comp hg), (bind_apply hs hg)],
196
170
end
197
171
198
172
lemma bind_dirac {f : α → measure β} (hf : measurable f) (a : α) : bind (dirac a) f = f a :=
199
- measure.ext $ λ s hs, by rw [bind_apply hs hf, lintegral_dirac' a ((measurable_coe hs).comp hf)]
173
+ by { ext1 s hs, rw [bind_apply hs hf, lintegral_dirac' a ((measurable_coe hs).comp hf)], }
200
174
201
175
lemma dirac_bind {m : measure α} : bind m dirac = m :=
202
- measure.ext $ assume s hs,
203
- by simp [bind_apply hs measurable_dirac, dirac_apply' _ hs, lintegral_indicator 1 hs]
176
+ begin
177
+ ext1 s hs,
178
+ simp only [bind_apply hs measurable_dirac, dirac_apply' _ hs, lintegral_indicator 1 hs,
179
+ pi.one_apply, lintegral_one, restrict_apply, measurable_set.univ, univ_inter],
180
+ end
204
181
205
182
lemma join_eq_bind (μ : measure (measure α)) : join μ = bind μ id :=
206
183
by rw [bind, map_id]
207
184
208
185
lemma join_map_map {f : α → β} (hf : measurable f) (μ : measure (measure α)) :
209
186
join (map (map f) μ) = map f (join μ) :=
210
- measure.ext $ assume s hs,
211
- begin
212
- rw [join_apply hs, map_apply hf hs, join_apply,
213
- lintegral_map (measurable_coe hs) (measurable_map f hf)],
214
- { congr, funext ν, exact map_apply hf hs },
215
- exact hf hs
216
- end
187
+ begin
188
+ ext1 s hs,
189
+ rw [join_apply hs, map_apply hf hs, join_apply (hf hs),
190
+ lintegral_map (measurable_coe hs) (measurable_map f hf)],
191
+ simp_rw map_apply hf hs,
192
+ end
217
193
218
194
lemma join_map_join (μ : measure (measure (measure α))) :
219
195
join (map join μ) = join (join μ) :=
@@ -229,7 +205,7 @@ lemma join_map_dirac (μ : measure α) : join (map dirac μ) = μ :=
229
205
dirac_bind
230
206
231
207
lemma join_dirac (μ : measure α) : join (dirac μ) = μ :=
232
- eq.trans (join_eq_bind (dirac μ)) (bind_dirac measurable_id _)
208
+ (join_eq_bind (dirac μ)).trans (bind_dirac measurable_id _)
233
209
234
210
end measure
235
211
0 commit comments