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

Commit 62f94ad

Browse files
committed
feat(measure_theory/measurable_space): define measurable_embedding (#10023)
This way we can generalize our theorems about `measurable_equiv` and `closed_embedding`s.
1 parent c50c2c3 commit 62f94ad

File tree

11 files changed

+398
-175
lines changed

11 files changed

+398
-175
lines changed

src/data/set/function.lean

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -87,6 +87,23 @@ by convert restrict_dite _ _
8787
restrict (extend f g g') (range f)ᶜ = g' ∘ coe :=
8888
by convert restrict_dite_compl _ _
8989

90+
lemma range_extend_subset (f : α → β) (g : α → γ) (g' : β → γ) :
91+
range (extend f g g') ⊆ range g ∪ g' '' (range f)ᶜ :=
92+
begin
93+
classical,
94+
rintro _ ⟨y, rfl⟩,
95+
rw extend_def, split_ifs,
96+
exacts [or.inl (mem_range_self _), or.inr (mem_image_of_mem _ h)]
97+
end
98+
99+
lemma range_extend {f : α → β} (hf : injective f) (g : α → γ) (g' : β → γ) :
100+
range (extend f g g') = range g ∪ g' '' (range f)ᶜ :=
101+
begin
102+
refine (range_extend_subset _ _ _).antisymm _,
103+
rintro z (⟨x, rfl⟩|⟨y, hy, rfl⟩),
104+
exacts [⟨f x, extend_apply hf _ _ _⟩, ⟨y, extend_apply' _ _ _ hy⟩]
105+
end
106+
90107
/-- Restrict codomain of a function `f` to a set `s`. Same as `subtype.coind` but this version
91108
has codomain `↥s` instead of `subtype s`. -/
92109
def cod_restrict (f : α → β) (s : set β) (h : ∀ x, f x ∈ s) : α → s :=

src/measure_theory/constructions/borel_space.lean

Lines changed: 12 additions & 51 deletions
Original file line numberDiff line numberDiff line change
@@ -810,57 +810,18 @@ instance prod.borel_space [second_countable_topology α] [second_countable_topol
810810
borel_space (α × β) :=
811811
⟨le_antisymm prod_le_borel_prod opens_measurable_space.borel_le⟩
812812

813-
lemma closed_embedding.measurable_inv_fun [n : nonempty β] {g : β → γ} (hg : closed_embedding g) :
814-
measurable (function.inv_fun g) :=
815-
begin
816-
refine measurable_of_is_closed (λ s hs, _),
817-
by_cases h : classical.choice n ∈ s,
818-
{ rw preimage_inv_fun_of_mem hg.to_embedding.inj h,
819-
exact (hg.closed_iff_image_closed.mp hs).measurable_set.union
820-
hg.closed_range.measurable_set.compl },
821-
{ rw preimage_inv_fun_of_not_mem hg.to_embedding.inj h,
822-
exact (hg.closed_iff_image_closed.mp hs).measurable_set }
823-
end
824-
825-
lemma measurable_comp_iff_of_closed_embedding {f : δ → β} (g : β → γ) (hg : closed_embedding g) :
826-
measurable (g ∘ f) ↔ measurable f :=
827-
begin
828-
refine ⟨λ hf, _, λ hf, hg.measurable.comp hf⟩,
829-
apply measurable_of_is_closed, intros s hs,
830-
convert hf (hg.is_closed_map s hs).measurable_set,
831-
rw [@preimage_comp _ _ _ f g, preimage_image_eq _ hg.to_embedding.inj]
832-
end
813+
protected lemma embedding.measurable_embedding {f : α → β} (h₁ : embedding f)
814+
(h₂ : measurable_set (range f)) : measurable_embedding f :=
815+
show measurable_embedding (coe ∘ (homeomorph.of_embedding f h₁).to_measurable_equiv),
816+
from (measurable_embedding.subtype_coe h₂).comp (measurable_equiv.measurable_embedding _)
833817

834-
lemma ae_measurable_comp_iff_of_closed_embedding {f : δ → β} {μ : measure δ}
835-
(g : β → γ) (hg : closed_embedding g) : ae_measurable (g ∘ f) μ ↔ ae_measurable f μ :=
836-
begin
837-
casesI is_empty_or_nonempty β,
838-
{ haveI := function.is_empty f,
839-
simp only [(measurable_of_empty (g ∘ f)).ae_measurable,
840-
(measurable_of_empty f).ae_measurable] },
841-
{ refine ⟨λ hf, _, λ hf, hg.measurable.comp_ae_measurable hf⟩,
842-
convert hg.measurable_inv_fun.comp_ae_measurable hf,
843-
ext x,
844-
exact (function.left_inverse_inv_fun hg.to_embedding.inj (f x)).symm },
845-
end
818+
protected lemma closed_embedding.measurable_embedding {f : α → β} (h : closed_embedding f) :
819+
measurable_embedding f :=
820+
h.to_embedding.measurable_embedding h.closed_range.measurable_set
846821

847-
lemma ae_measurable_comp_right_iff_of_closed_embedding {g : α → β} {μ : measure α}
848-
{f : β → δ} (hg : closed_embedding g) :
849-
ae_measurable (f ∘ g) μ ↔ ae_measurable f (measure.map g μ) :=
850-
begin
851-
refine ⟨λ h, _, λ h, h.comp_measurable hg.measurable⟩,
852-
casesI is_empty_or_nonempty α, { simp [μ.eq_zero_of_is_empty] },
853-
refine ⟨(h.mk _) ∘ (function.inv_fun g), h.measurable_mk.comp hg.measurable_inv_fun, _⟩,
854-
have : μ = measure.map (function.inv_fun g) (measure.map g μ),
855-
by rw [measure.map_map hg.measurable_inv_fun hg.measurable,
856-
(function.left_inverse_inv_fun hg.to_embedding.inj).comp_eq_id, measure.map_id],
857-
rw this at h,
858-
filter_upwards [ae_of_ae_map hg.measurable_inv_fun h.ae_eq_mk,
859-
ae_map_mem_range g hg.closed_range.measurable_set μ],
860-
assume x hx₁ hx₂,
861-
convert hx₁,
862-
exact ((function.left_inverse_inv_fun hg.to_embedding.inj).right_inv_on_range hx₂).symm,
863-
end
822+
protected lemma open_embedding.measurable_embedding {f : α → β} (h : open_embedding f) :
823+
measurable_embedding f :=
824+
h.to_embedding.measurable_embedding h.open_range.measurable_set
864825

865826
section linear_order
866827

@@ -1855,10 +1816,10 @@ variables {E : Type*} [normed_group E] [normed_space 𝕜 E] [measurable_space E
18551816

18561817
lemma measurable_smul_const {f : α → 𝕜} {c : E} (hc : c ≠ 0) :
18571818
measurable (λ x, f x • c) ↔ measurable f :=
1858-
measurable_comp_iff_of_closed_embedding (λ y : 𝕜, y • c) (closed_embedding_smul_left hc)
1819+
(closed_embedding_smul_left hc).measurable_embedding.measurable_comp_iff
18591820

18601821
lemma ae_measurable_smul_const {f : α → 𝕜} {μ : measure α} {c : E} (hc : c ≠ 0) :
18611822
ae_measurable (λ x, f x • c) μ ↔ ae_measurable f μ :=
1862-
ae_measurable_comp_iff_of_closed_embedding (λ y : 𝕜, y • c) (closed_embedding_smul_left hc)
1823+
(closed_embedding_smul_left hc).measurable_embedding.ae_measurable_comp_iff
18631824

18641825
end normed_space

src/measure_theory/function/l1_space.lean

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -457,9 +457,14 @@ lemma integrable_map_measure [opens_measurable_space β] {f : α → δ} {g : δ
457457
integrable g (measure.map f μ) ↔ integrable (g ∘ f) μ :=
458458
by simp [integrable, hg, hg.comp_measurable hf, has_finite_integral, lintegral_map' hg.ennnorm hf]
459459

460+
lemma _root_.measurable_embedding.integrable_map_iff {f : α → δ} (hf : measurable_embedding f)
461+
{g : δ → β} :
462+
integrable g (measure.map f μ) ↔ integrable (g ∘ f) μ :=
463+
by simp only [integrable, hf.ae_measurable_map_iff, has_finite_integral, hf.lintegral_map]
464+
460465
lemma integrable_map_equiv (f : α ≃ᵐ δ) (g : δ → β) :
461466
integrable g (measure.map f μ) ↔ integrable (g ∘ f) μ :=
462-
by simp only [integrable, ae_measurable_map_equiv_iff, has_finite_integral, lintegral_map_equiv]
467+
f.measurable_embedding.integrable_map_iff
463468

464469
lemma lintegral_edist_lt_top [second_countable_topology β] [opens_measurable_space β] {f g : α → β}
465470
(hf : integrable f μ) (hg : integrable g μ) :

src/measure_theory/function/lp_space.lean

Lines changed: 5 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -1730,11 +1730,11 @@ lemma measure_theory.mem_ℒp.of_comp_antilipschitz_with {α E F} {K'}
17301730
(hg : uniform_continuous g) (hg' : antilipschitz_with K' g) (g0 : g 0 = 0) : mem_ℒp f p μ :=
17311731
begin
17321732
have : ∀ᵐ x ∂μ, ∥f x∥ ≤ K' * ∥g (f x)∥,
1733-
{ apply filter.eventually_of_forall (λ x, _),
1734-
rw [← dist_zero_right, ← dist_zero_right, ← g0],
1735-
apply hg'.le_mul_dist },
1736-
exact hL.of_le_mul ((ae_measurable_comp_iff_of_closed_embedding g
1737-
(hg'.closed_embedding hg)).1 hL.1) this,
1733+
{ apply filter.eventually_of_forall (λ x, _),
1734+
rw [← dist_zero_right, ← dist_zero_right, ← g0],
1735+
apply hg'.le_mul_dist },
1736+
exact hL.of_le_mul ((hg'.closed_embedding hg).measurable_embedding.ae_measurable_comp_iff.1
1737+
hL.1) this,
17381738
end
17391739

17401740
namespace lipschitz_with

src/measure_theory/integral/bochner.lean

Lines changed: 15 additions & 15 deletions
Original file line numberDiff line numberDiff line change
@@ -1286,25 +1286,25 @@ let g := hfm.mk f in calc
12861286
... = ∫ x, g (φ x) ∂μ : integral_map_of_measurable hφ hfm.measurable_mk
12871287
... = ∫ x, f (φ x) ∂μ : integral_congr_ae $ ae_eq_comp hφ (hfm.ae_eq_mk).symm
12881288

1289-
lemma integral_map_of_closed_embedding {β} [topological_space α] [borel_space α]
1289+
lemma _root_.measurable_embedding.integral_map {β} {_ : measurable_space β} {f : α → β}
1290+
(hf : measurable_embedding f) (g : β → E) :
1291+
∫ y, g y ∂(measure.map f μ) = ∫ x, g (f x) ∂μ :=
1292+
begin
1293+
by_cases hgm : ae_measurable g (measure.map f μ),
1294+
{ exact integral_map hf.measurable hgm },
1295+
{ rw [integral_non_ae_measurable hgm, integral_non_ae_measurable],
1296+
rwa ← hf.ae_measurable_map_iff }
1297+
end
1298+
1299+
lemma _root_.closed_embedding.integral_map {β} [topological_space α] [borel_space α]
12901300
[topological_space β] [measurable_space β] [borel_space β]
12911301
{φ : α → β} (hφ : closed_embedding φ) (f : β → E) :
12921302
∫ y, f y ∂(measure.map φ μ) = ∫ x, f (φ x) ∂μ :=
1293-
begin
1294-
by_cases hfm : ae_measurable f (measure.map φ μ),
1295-
{ exact integral_map hφ.continuous.measurable hfm },
1296-
{ rw [integral_non_ae_measurable hfm, integral_non_ae_measurable],
1297-
rwa ae_measurable_comp_right_iff_of_closed_embedding hφ }
1298-
end
1303+
hφ.measurable_embedding.integral_map _
12991304

13001305
lemma integral_map_equiv {β} [measurable_space β] (e : α ≃ᵐ β) (f : β → E) :
13011306
∫ y, f y ∂(measure.map e μ) = ∫ x, f (e x) ∂μ :=
1302-
begin
1303-
by_cases hfm : ae_measurable f (measure.map e μ),
1304-
{ exact integral_map e.measurable hfm },
1305-
{ rw [integral_non_ae_measurable hfm, integral_non_ae_measurable],
1306-
rwa ← ae_measurable_map_equiv_iff }
1307-
end
1307+
e.measurable_embedding.integral_map f
13081308

13091309
@[simp] lemma integral_dirac' [measurable_space α] (f : α → E) (a : α) (hfm : measurable f) :
13101310
∫ x, f x ∂(measure.dirac a) = f a :=
@@ -1338,7 +1338,7 @@ begin
13381338
{ rw ← map_mul_left_eq_self at hμ,
13391339
exact hμ g },
13401340
have h_mul : closed_embedding (λ x, g * x) := (homeomorph.mul_left g).closed_embedding,
1341-
rw [← integral_map_of_closed_embedding h_mul, hgμ],
1341+
rw [← h_mul.integral_map, hgμ],
13421342
apply_instance,
13431343
end
13441344

@@ -1352,7 +1352,7 @@ begin
13521352
{ rw ← map_mul_right_eq_self at hμ,
13531353
exact hμ g },
13541354
have h_mul : closed_embedding (λ x, x * g) := (homeomorph.mul_right g).closed_embedding,
1355-
rw [← integral_map_of_closed_embedding h_mul, hgμ],
1355+
rw [← h_mul.integral_map, hgμ],
13561356
apply_instance,
13571357
end
13581358

src/measure_theory/integral/integrable_on.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -171,6 +171,11 @@ by { delta integrable_on, rw measure.restrict_add, exact hμ.integrable.add_meas
171171
h.mono_measure (measure.le_add_left (le_refl _))⟩,
172172
λ h, h.1.add_measure h.2
173173

174+
lemma _root_.measurable_embedding.integrable_on_map_iff [measurable_space β] {e : α → β}
175+
(he : measurable_embedding e) {f : β → E} {μ : measure α} {s : set β} :
176+
integrable_on f s (measure.map e μ) ↔ integrable_on (f ∘ e) (e ⁻¹' s) μ :=
177+
by simp only [integrable_on, he.restrict_map, he.integrable_map_iff]
178+
174179
lemma integrable_on_map_equiv [measurable_space β] (e : α ≃ᵐ β) {f : β → E} {μ : measure α}
175180
{s : set β} :
176181
integrable_on f s (measure.map e μ) ↔ integrable_on (f ∘ e) (e ⁻¹' s) μ :=

src/measure_theory/integral/interval_integral.lean

Lines changed: 8 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -576,14 +576,14 @@ variables {a b c d : ℝ} (f : ℝ → E)
576576
@[simp] lemma integral_comp_mul_right (hc : c ≠ 0) :
577577
∫ x in a..b, f (x * c) = c⁻¹ • ∫ x in a*c..b*c, f x :=
578578
begin
579-
have A : closed_embedding (λ x, x * c) := (homeomorph.mul_right₀ c hc).closed_embedding,
579+
have A : measurable_embedding (λ x, x * c) :=
580+
(homeomorph.mul_right₀ c hc).closed_embedding.measurable_embedding,
580581
conv_rhs { rw [← real.smul_map_volume_mul_right hc] },
581-
simp_rw [integral_smul_measure, interval_integral,
582-
set_integral_map_of_closed_embedding measurable_set_Ioc A,
582+
simp_rw [integral_smul_measure, interval_integral, A.set_integral_map,
583583
ennreal.to_real_of_real (abs_nonneg c)],
584-
cases lt_or_gt_of_ne hc,
584+
cases hc.lt_or_lt,
585585
{ simp [h, mul_div_cancel, hc, abs_of_neg, restrict_congr_set Ico_ae_eq_Ioc] },
586-
{ simp [(show 0 < c, from h), mul_div_cancel, hc, abs_of_pos] }
586+
{ simp [h, mul_div_cancel, hc, abs_of_pos] }
587587
end
588588

589589
@[simp] lemma smul_integral_comp_mul_right (c) :
@@ -608,10 +608,11 @@ by by_cases hc : c = 0; simp [hc]
608608

609609
@[simp] lemma integral_comp_add_right (d) :
610610
∫ x in a..b, f (x + d) = ∫ x in a+d..b+d, f x :=
611-
have A : closed_embedding (λ x, x + d) := (homeomorph.add_right d).closed_embedding,
611+
have A : measurable_embedding (λ x, x + d) :=
612+
(homeomorph.add_right d).closed_embedding.measurable_embedding,
612613
calc ∫ x in a..b, f (x + d)
613614
= ∫ x in a+d..b+d, f x ∂(measure.map (λ x, x + d) volume)
614-
: by simp [interval_integral, set_integral_map_of_closed_embedding _ A]
615+
: by simp [interval_integral, A.set_integral_map]
615616
... = ∫ x in a+d..b+d, f x : by rw [real.map_volume_add_right]
616617

617618
@[simp] lemma integral_comp_add_left (d) :

src/measure_theory/integral/lebesgue.lean

Lines changed: 50 additions & 49 deletions
Original file line numberDiff line numberDiff line change
@@ -262,6 +262,31 @@ lemma range_comp_subset_range [measurable_space β] (f : β →ₛ γ) {g : α
262262
(f.comp g hgm).range ⊆ f.range :=
263263
finset.coe_subset.1 $ by simp only [coe_range, coe_comp, set.range_comp_subset_range]
264264

265+
/-- Extend a `simple_func` along a measurable embedding: `f₁.extend g hg f₂` is the function
266+
`F : β →ₛ γ` such that `F ∘ g = f₁` and `F y = f₂ y` whenever `y ∉ range g`. -/
267+
def extend [measurable_space β] (f₁ : α →ₛ γ) (g : α → β)
268+
(hg : measurable_embedding g) (f₂ : β →ₛ γ) : β →ₛ γ :=
269+
{ to_fun := function.extend g f₁ f₂,
270+
finite_range' := (f₁.finite_range.union $ f₂.finite_range.subset
271+
(image_subset_range _ _)).subset (range_extend_subset _ _ _),
272+
measurable_set_fiber' :=
273+
begin
274+
letI : measurable_space γ := ⊤, haveI : measurable_singleton_class γ := ⟨λ _, trivial⟩,
275+
exact λ x, hg.measurable_extend f₁.measurable f₂.measurable (measurable_set_singleton _)
276+
end }
277+
278+
@[simp] lemma extend_apply [measurable_space β] (f₁ : α →ₛ γ) {g : α → β}
279+
(hg : measurable_embedding g) (f₂ : β →ₛ γ) (x : α) : (f₁.extend g hg f₂) (g x) = f₁ x :=
280+
function.extend_apply hg.injective _ _ _
281+
282+
@[simp] lemma extend_comp_eq' [measurable_space β] (f₁ : α →ₛ γ) {g : α → β}
283+
(hg : measurable_embedding g) (f₂ : β →ₛ γ) : (f₁.extend g hg f₂) ∘ g = f₁ :=
284+
funext $ λ x, extend_apply _ _ _ _
285+
286+
@[simp] lemma extend_comp_eq [measurable_space β] (f₁ : α →ₛ γ) {g : α → β}
287+
(hg : measurable_embedding g) (f₂ : β →ₛ γ) : (f₁.extend g hg f₂).comp g hg.measurable = f₁ :=
288+
coe_injective $ extend_comp_eq' _ _ _
289+
265290
/-- If `f` is a simple function taking values in `β → γ` and `g` is another simple function
266291
with the same domain and codomain `β`, then `f.seq g = f a (g a)`. -/
267292
def seq (f : α →ₛ (β → γ)) (g : α →ₛ β) : α →ₛ γ := f.bind (λf, g.map f)
@@ -787,32 +812,15 @@ lemma lintegral_congr {f g : α →ₛ ℝ≥0∞} (h : f =ᵐ[μ] g) :
787812
lintegral_eq_of_measure_preimage $ λ y, measure_congr $
788813
eventually.set_eq $ h.mono $ λ x hx, by simp [hx]
789814

790-
lemma lintegral_map {β} [measurable_space β] {μ' : measure β} (f : α →ₛ ℝ≥0∞) (g : β →ₛ ℝ≥0∞)
815+
lemma lintegral_map' {β} [measurable_space β] {μ' : measure β} (f : α →ₛ ℝ≥0∞) (g : β →ₛ ℝ≥0∞)
791816
(m' : α → β) (eq : ∀ a, f a = g (m' a)) (h : ∀s, measurable_set s → μ' s = μ (m' ⁻¹' s)) :
792817
f.lintegral μ = g.lintegral μ' :=
793818
lintegral_eq_of_measure_preimage $ λ y,
794819
by { simp only [preimage, eq], exact (h (g ⁻¹' {y}) (g.measurable_set_preimage _)).symm }
795820

796-
/-- The `lintegral` of simple functions transforms appropriately under a measurable equivalence.
797-
(Compare `lintegral_map`, which applies to a broader class of transformations of the domain, but
798-
requires measurability of the function being integrated.) -/
799-
lemma lintegral_map_equiv {β} [measurable_space β] (g : β →ₛ ℝ≥0∞) (m' : α ≃ᵐ β) :
800-
(g.comp m' m'.measurable).lintegral μ = g.lintegral (measure.map m' μ) :=
801-
begin
802-
simp [simple_func.lintegral],
803-
have : (g.comp m' m'.measurable).range = g.range,
804-
{ refine le_antisymm _ _,
805-
{ exact g.range_comp_subset_range m'.measurable },
806-
convert (g.comp m' m'.measurable).range_comp_subset_range m'.symm.measurable,
807-
apply simple_func.ext,
808-
intros a,
809-
exact congr_arg g (congr_fun m'.self_comp_symm.symm a) },
810-
rw this,
811-
congr' 1,
812-
funext,
813-
rw [m'.map_apply (g ⁻¹' {x})],
814-
refl,
815-
end
821+
lemma lintegral_map {β} [measurable_space β] (g : β →ₛ ℝ≥0∞) {f : α → β} (hf : measurable f) :
822+
g.lintegral (measure.map f μ) = (g.comp f hf).lintegral μ :=
823+
eq.symm $ lintegral_map' _ _ f (λ a, rfl) (λ s hs, measure.map_apply hf hs)
816824

817825
end measure
818826

@@ -1864,10 +1872,9 @@ lemma lintegral_map [measurable_space β] {f : β → ℝ≥0∞} {g : α → β
18641872
(hf : measurable f) (hg : measurable g) : ∫⁻ a, f a ∂(map g μ) = ∫⁻ a, f (g a) ∂μ :=
18651873
begin
18661874
simp only [lintegral_eq_supr_eapprox_lintegral, hf, hf.comp hg],
1867-
{ congr, funext n, symmetry,
1868-
apply simple_func.lintegral_map,
1869-
{ assume a, exact congr_fun (simple_func.eapprox_comp hf hg) a },
1870-
{ assume s hs, exact map_apply hg hs } },
1875+
congr' with n : 1,
1876+
convert simple_func.lintegral_map _ hg,
1877+
ext1 x, simp only [eapprox_comp hf hg, coe_comp]
18711878
end
18721879

18731880
lemma lintegral_map' [measurable_space β] {f : β → ℝ≥0∞} {g : α → β}
@@ -1887,35 +1894,29 @@ lemma set_lintegral_map [measurable_space β] {f : β → ℝ≥0∞} {g : α
18871894
∫⁻ y in s, f y ∂(map g μ) = ∫⁻ x in g ⁻¹' s, f (g x) ∂μ :=
18881895
by rw [restrict_map hg hs, lintegral_map hf hg]
18891896

1897+
/-- If `g : α → β` is a measurable embedding and `f : β → ℝ≥0∞` is any function (not necessarily
1898+
measurable), then `∫⁻ a, f a ∂(map g μ) = ∫⁻ a, f (g a) ∂μ`. Compare with `lintegral_map` wich
1899+
applies to any measurable `g : α → β` but requires that `f` is measurable as well. -/
1900+
lemma _root_.measurable_embedding.lintegral_map [measurable_space β] {g : α → β}
1901+
(hg : measurable_embedding g) (f : β → ℝ≥0∞) :
1902+
∫⁻ a, f a ∂(map g μ) = ∫⁻ a, f (g a) ∂μ :=
1903+
begin
1904+
refine le_antisymm (bsupr_le $ λ f₀ hf₀, _) (bsupr_le $ λ f₀ hf₀, _),
1905+
{ rw [simple_func.lintegral_map _ hg.measurable, lintegral],
1906+
have : (f₀.comp g hg.measurable : α → ℝ≥0∞) ≤ f ∘ g, from λ x, hf₀ (g x),
1907+
exact le_supr_of_le (comp f₀ g hg.measurable) (le_supr _ this) },
1908+
{ rw [← f₀.extend_comp_eq hg (const _ 0), ← simple_func.lintegral_map,
1909+
← simple_func.lintegral_eq_lintegral],
1910+
refine lintegral_mono_ae (hg.ae_map_iff.2 $ eventually_of_forall $ λ x, _),
1911+
exact (extend_apply _ _ _ _).trans_le (hf₀ _) }
1912+
end
1913+
18901914
/-- The `lintegral` transforms appropriately under a measurable equivalence `g : α ≃ᵐ β`.
18911915
(Compare `lintegral_map`, which applies to a wider class of functions `g : α → β`, but requires
18921916
measurability of the function being integrated.) -/
18931917
lemma lintegral_map_equiv [measurable_space β] (f : β → ℝ≥0∞) (g : α ≃ᵐ β) :
18941918
∫⁻ a, f a ∂(map g μ) = ∫⁻ a, f (g a) ∂μ :=
1895-
begin
1896-
refine le_antisymm _ _,
1897-
{ refine supr_le_supr2 _,
1898-
intros f₀,
1899-
use f₀.comp g g.measurable,
1900-
refine supr_le_supr2 _,
1901-
intros hf₀,
1902-
use λ x, hf₀ (g x),
1903-
exact (lintegral_map_equiv f₀ g).symm.le },
1904-
{ refine supr_le_supr2 _,
1905-
intros f₀,
1906-
use f₀.comp g.symm g.symm.measurable,
1907-
refine supr_le_supr2 _,
1908-
intros hf₀,
1909-
have : (λ a, (f₀.comp (g.symm) g.symm.measurable) a) ≤ λ (a : β), f a,
1910-
{ convert λ x, hf₀ (g.symm x),
1911-
funext,
1912-
simp [congr_arg f (congr_fun g.self_comp_symm a)] },
1913-
use this,
1914-
convert (lintegral_map_equiv (f₀.comp g.symm g.symm.measurable) g).le,
1915-
apply simple_func.ext,
1916-
intros a,
1917-
convert congr_arg f₀ (congr_fun g.symm_comp_self a).symm using 1 }
1918-
end
1919+
g.measurable_embedding.lintegral_map f
19191920

19201921
section dirac_and_count
19211922
variable [measurable_space α]

0 commit comments

Comments
 (0)