Skip to content

Commit

Permalink
Browse files Browse the repository at this point in the history
  • Loading branch information
Parcly-Taxel committed Jun 20, 2023
1 parent 0f82593 commit 76ed1a8
Showing 1 changed file with 27 additions and 1 deletion.
28 changes: 27 additions & 1 deletion Mathlib/MeasureTheory/Constructions/Prod/Basic.lean
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn
! This file was ported from Lean 3 source module measure_theory.constructions.prod.basic
! leanprover-community/mathlib commit 3b88f4005dc2e28d42f974cc1ce838f0dafb39b8
! leanprover-community/mathlib commit 00abe0695d8767201e6d008afa22393978bb324d
! Please do not edit these lines, except to modify the commit id
! if you have ported upstream changes.
-/
Expand Down Expand Up @@ -834,6 +834,19 @@ instance fst.instIsProbabilityMeasure [IsProbabilityMeasure ρ] : IsProbabilityM
exact measure_univ
#align measure_theory.measure.fst.measure_theory.is_probability_measure MeasureTheory.Measure.fst.instIsProbabilityMeasure

theorem fst_map_prod_mk₀ {X : α → β} {Y : α → γ} {μ : Measure α} (hX : AEMeasurable X μ)
(hY : AEMeasurable Y μ) : (μ.map fun a => (X a, Y a)).fst = μ.map X := by
ext1 s hs
rw [Measure.fst_apply hs, Measure.map_apply_of_aemeasurable (hX.prod_mk hY) (measurable_fst hs),
Measure.map_apply_of_aemeasurable hX hs, ← prod_univ, mk_preimage_prod, preimage_univ,
inter_univ]
#align measure_theory.measure.fst_map_prod_mk₀ MeasureTheory.Measure.fst_map_prod_mk₀

theorem fst_map_prod_mk {X : α → β} {Y : α → γ} {μ : Measure α} (hX : Measurable X)
(hY : Measurable Y) : (μ.map fun a => (X a, Y a)).fst = μ.map X :=
fst_map_prod_mk₀ hX.aemeasurable hY.aemeasurable
#align measure_theory.measure.fst_map_prod_mk MeasureTheory.Measure.fst_map_prod_mk

/-- Marginal measure on `β` obtained from a measure on `ρ` `α × β`, defined by `ρ.map Prod.snd`. -/
noncomputable def snd (ρ : Measure (α × β)) : Measure β :=
ρ.map Prod.snd
Expand All @@ -857,6 +870,19 @@ instance snd.instIsProbabilityMeasure [IsProbabilityMeasure ρ] : IsProbabilityM
exact measure_univ
#align measure_theory.measure.snd.measure_theory.is_probability_measure MeasureTheory.Measure.snd.instIsProbabilityMeasure

theorem snd_map_prod_mk₀ {X : α → β} {Y : α → γ} {μ : Measure α} (hX : AEMeasurable X μ)
(hY : AEMeasurable Y μ) : (μ.map fun a => (X a, Y a)).snd = μ.map Y := by
ext1 s hs
rw [Measure.snd_apply hs, Measure.map_apply_of_aemeasurable (hX.prod_mk hY) (measurable_snd hs),
Measure.map_apply_of_aemeasurable hY hs, ← univ_prod, mk_preimage_prod, preimage_univ,
univ_inter]
#align measure_theory.measure.snd_map_prod_mk₀ MeasureTheory.Measure.snd_map_prod_mk₀

theorem snd_map_prod_mk {X : α → β} {Y : α → γ} {μ : Measure α} (hX : Measurable X)
(hY : Measurable Y) : (μ.map fun a => (X a, Y a)).snd = μ.map Y :=
snd_map_prod_mk₀ hX.aemeasurable hY.aemeasurable
#align measure_theory.measure.snd_map_prod_mk MeasureTheory.Measure.snd_map_prod_mk

end Measure

end MeasureTheory

0 comments on commit 76ed1a8

Please sign in to comment.