@@ -262,6 +262,31 @@ lemma range_comp_subset_range [measurable_space β] (f : β →ₛ γ) {g : α
262
262
(f.comp g hgm).range ⊆ f.range :=
263
263
finset.coe_subset.1 $ by simp only [coe_range, coe_comp, set.range_comp_subset_range]
264
264
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
+
265
290
/-- If `f` is a simple function taking values in `β → γ` and `g` is another simple function
266
291
with the same domain and codomain `β`, then `f.seq g = f a (g a)`. -/
267
292
def seq (f : α →ₛ (β → γ)) (g : α →ₛ β) : α →ₛ γ := f.bind (λf, g.map f)
@@ -787,32 +812,15 @@ lemma lintegral_congr {f g : α →ₛ ℝ≥0∞} (h : f =ᵐ[μ] g) :
787
812
lintegral_eq_of_measure_preimage $ λ y, measure_congr $
788
813
eventually.set_eq $ h.mono $ λ x hx, by simp [hx]
789
814
790
- lemma lintegral_map {β} [measurable_space β] {μ' : measure β} (f : α →ₛ ℝ≥0 ∞) (g : β →ₛ ℝ≥0 ∞)
815
+ lemma lintegral_map' {β} [measurable_space β] {μ' : measure β} (f : α →ₛ ℝ≥0 ∞) (g : β →ₛ ℝ≥0 ∞)
791
816
(m' : α → β) (eq : ∀ a, f a = g (m' a)) (h : ∀s, measurable_set s → μ' s = μ (m' ⁻¹' s)) :
792
817
f.lintegral μ = g.lintegral μ' :=
793
818
lintegral_eq_of_measure_preimage $ λ y,
794
819
by { simp only [preimage, eq], exact (h (g ⁻¹' {y}) (g.measurable_set_preimage _)).symm }
795
820
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)
816
824
817
825
end measure
818
826
@@ -1864,10 +1872,9 @@ lemma lintegral_map [measurable_space β] {f : β → ℝ≥0∞} {g : α → β
1864
1872
(hf : measurable f) (hg : measurable g) : ∫⁻ a, f a ∂(map g μ) = ∫⁻ a, f (g a) ∂μ :=
1865
1873
begin
1866
1874
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]
1871
1878
end
1872
1879
1873
1880
lemma lintegral_map' [measurable_space β] {f : β → ℝ≥0 ∞} {g : α → β}
@@ -1887,35 +1894,29 @@ lemma set_lintegral_map [measurable_space β] {f : β → ℝ≥0∞} {g : α
1887
1894
∫⁻ y in s, f y ∂(map g μ) = ∫⁻ x in g ⁻¹' s, f (g x) ∂μ :=
1888
1895
by rw [restrict_map hg hs, lintegral_map hf hg]
1889
1896
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
+
1890
1914
/-- The `lintegral` transforms appropriately under a measurable equivalence `g : α ≃ᵐ β`.
1891
1915
(Compare `lintegral_map`, which applies to a wider class of functions `g : α → β`, but requires
1892
1916
measurability of the function being integrated.) -/
1893
1917
lemma lintegral_map_equiv [measurable_space β] (f : β → ℝ≥0 ∞) (g : α ≃ᵐ β) :
1894
1918
∫⁻ 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
1919
1920
1920
1921
section dirac_and_count
1921
1922
variable [measurable_space α]
0 commit comments