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

Commit c7d6e64

Browse files
committed
chore(measure_theory/measure/giry_monad): add spaces, golf (#17155)
Minor changes: add spaces after lambdas, squeeze some simps (one of which was non-terminal), golf a proof...
1 parent 44c394b commit c7d6e64

File tree

1 file changed

+62
-86
lines changed

1 file changed

+62
-86
lines changed

src/measure_theory/measure/giry_monad.lean

Lines changed: 62 additions & 86 deletions
Original file line numberDiff line numberDiff line change
@@ -32,7 +32,7 @@ open_locale classical big_operators ennreal
3232

3333
open classical set filter
3434

35-
variables {α β γ δ ε : Type*}
35+
variables {α β : Type*}
3636

3737
namespace measure_theory
3838

@@ -42,38 +42,38 @@ variables [measurable_space α] [measurable_space β]
4242

4343
/-- Measurability structure on `measure`: Measures are measurable w.r.t. all projections -/
4444
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)
4646

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) :=
4848
measurable.of_comap_le $ le_supr_of_le s $ le_supr_of_le hs $ le_rfl
4949

5050
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)) :
5252
measurable f :=
5353
measurable.of_le_map $ supr₂_le $ assume s hs, measurable_space.comap_le_iff_le_map.2 $
5454
by rw [measurable_space.map_comp]; exact h s hs
5555

5656
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) :=
5858
⟨λ hμ s hs, (measurable_coe hs).comp hμ, measurable_of_measurable_coe μ⟩
5959

6060
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
6667

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
7474

7575
lemma measurable_lintegral {f : α → ℝ≥0∞} (hf : measurable f) :
76-
measurable (λμ : measure α, ∫⁻ x, f x ∂μ) :=
76+
measurable (λ μ : measure α, ∫⁻ x, f x ∂μ) :=
7777
begin
7878
simp only [lintegral_eq_supr_eapprox_lintegral, hf, simple_func.lintegral],
7979
refine measurable_supr (λ n, finset.measurable_sum _ (λ i _, _)),
@@ -85,21 +85,21 @@ end
8585
functions. -/
8686
def join (m : measure (measure α)) : measure α :=
8787
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])
9090
begin
9191
assume f hf h,
92-
simp [measure_Union h hf],
92+
simp_rw [measure_Union h hf],
9393
apply lintegral_tsum,
9494
assume i, exact measurable_coe (hf i)
9595
end
9696

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
100100

101101
@[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], }
103103

104104
lemma measurable_join : measurable (join : measure (measure α) → measure α) :=
105105
measurable_of_measurable_coe _ $ assume s hs,
@@ -108,47 +108,23 @@ measurable_of_measurable_coe _ $ assume s hs,
108108
lemma lintegral_join {m : measure (measure α)} {f : α → ℝ≥0∞} (hf : measurable f) :
109109
∫⁻ x, f x ∂(join m) = ∫⁻ μ, ∫⁻ x, f x ∂μ ∂m :=
110110
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 _ },
152128
end
153129

154130
/-- Monadic bind on `measure`, only works in the category of measurable spaces and measurable
@@ -164,7 +140,7 @@ begin
164140
ext1 s hs,
165141
simp only [bind, hs, join_apply, coe_zero, pi.zero_apply],
166142
rw [lintegral_map (measurable_coe hs) measurable_zero],
167-
simp
143+
simp only [pi.zero_apply, coe_zero, lintegral_const, zero_mul],
168144
end
169145

170146
@[simp] lemma bind_zero_right' (m : measure α) :
@@ -176,44 +152,44 @@ bind_zero_right m
176152
bind m f s = ∫⁻ a, f a s ∂m :=
177153
by rw [bind, join_apply hs, lintegral_map (measurable_coe hs) hf]
178154

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) :=
180156
measurable_join.comp (measurable_map _ hg)
181157

182158
lemma lintegral_bind {m : measure α} {μ : α → measure β} {f : β → ℝ≥0∞}
183159
(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 :=
185161
(lintegral_join hf).trans (lintegral_map (measurable_lintegral hf) hμ)
186162

187163
lemma bind_bind {γ} [measurable_space γ] {m : measure α} {f : α → measure β} {g : β → measure γ}
188164
(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) :=
191166
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)],
196170
end
197171

198172
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)], }
200174

201175
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
204181

205182
lemma join_eq_bind (μ : measure (measure α)) : join μ = bind μ id :=
206183
by rw [bind, map_id]
207184

208185
lemma join_map_map {f : α → β} (hf : measurable f) (μ : measure (measure α)) :
209186
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
217193

218194
lemma join_map_join (μ : measure (measure (measure α))) :
219195
join (map join μ) = join (join μ) :=
@@ -229,7 +205,7 @@ lemma join_map_dirac (μ : measure α) : join (map dirac μ) = μ :=
229205
dirac_bind
230206

231207
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 _)
233209

234210
end measure
235211

0 commit comments

Comments
 (0)