Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - chore(MeasureTheory/Function/ConditionalExpectation): remove unneeded lpMeasSubgroup_coe and lpMeas_coe #22214

Closed
wants to merge 1 commit into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
Original file line number Diff line number Diff line change
Expand Up @@ -215,14 +215,6 @@ theorem mem_lpMeas_self {m0 : MeasurableSpace α} (μ : Measure α) (f : Lp F p
f ∈ lpMeas F 𝕜 m0 p μ :=
mem_lpMeas_iff_aeStronglyMeasurable.mpr (Lp.aestronglyMeasurable f)

theorem lpMeasSubgroup_coe {m _ : MeasurableSpace α} {μ : Measure α} {f : lpMeasSubgroup F m p μ} :
(f : _ → _) = (f : Lp F p μ) :=
rfl

theorem lpMeas_coe {m _ : MeasurableSpace α} {μ : Measure α} {f : lpMeas F 𝕜 m p μ} :
(f : _ → _) = (f : Lp F p μ) :=
rfl

theorem mem_lpMeas_indicatorConstLp {m m0 : MeasurableSpace α} (hm : m ≤ m0) {μ : Measure α}
{s : Set α} (hs : MeasurableSet[m] s) (hμs : μ s ≠ ∞) {c : F} :
indicatorConstLp p (hm s hs) hμs c ∈ lpMeas F 𝕜 m p μ :=
Expand Down Expand Up @@ -334,7 +326,6 @@ theorem lpMeasSubgroupToLpTrim_left_inv (hm : m ≤ m0) :
intro f
ext1
ext1
rw [← lpMeasSubgroup_coe]
exact (lpTrimToLpMeasSubgroup_ae_eq hm _).trans (lpMeasSubgroupToLpTrim_ae_eq hm _)

theorem lpMeasSubgroupToLpTrim_add (hm : m ≤ m0) (f g : lpMeasSubgroup F m p μ) :
Expand All @@ -350,7 +341,6 @@ theorem lpMeasSubgroupToLpTrim_add (hm : m ≤ m0) (f g : lpMeasSubgroup F m p
(EventuallyEq.add (lpMeasSubgroupToLpTrim_ae_eq hm f).symm
(lpMeasSubgroupToLpTrim_ae_eq hm g).symm)
refine (Lp.coeFn_add _ _).trans ?_
simp_rw [lpMeasSubgroup_coe]
filter_upwards with x using rfl

theorem lpMeasSubgroupToLpTrim_neg (hm : m ≤ m0) (f : lpMeasSubgroup F m p μ) :
Expand All @@ -360,7 +350,6 @@ theorem lpMeasSubgroupToLpTrim_neg (hm : m ≤ m0) (f : lpMeasSubgroup F m p μ)
refine (Lp.stronglyMeasurable _).ae_eq_trim_of_stronglyMeasurable hm (Lp.stronglyMeasurable _).neg
<| (lpMeasSubgroupToLpTrim_ae_eq hm _).trans <|
((Lp.coeFn_neg _).trans ?_).trans (lpMeasSubgroupToLpTrim_ae_eq hm f).symm.neg
simp_rw [lpMeasSubgroup_coe]
exact Eventually.of_forall fun x => by rfl

theorem lpMeasSubgroupToLpTrim_sub (hm : m ≤ m0) (f g : lpMeasSubgroup F m p μ) :
Expand All @@ -384,7 +373,7 @@ theorem lpMeasToLpTrim_smul (hm : m ≤ m0) (c : 𝕜) (f : lpMeas F 𝕜 m p μ
theorem lpMeasSubgroupToLpTrim_norm_map [hp : Fact (1 ≤ p)] (hm : m ≤ m0)
(f : lpMeasSubgroup F m p μ) : ‖lpMeasSubgroupToLpTrim F p μ hm f‖ = ‖f‖ := by
rw [Lp.norm_def, eLpNorm_trim hm (Lp.stronglyMeasurable _),
eLpNorm_congr_ae (lpMeasSubgroupToLpTrim_ae_eq hm _), lpMeasSubgroup_coe, ← Lp.norm_def]
eLpNorm_congr_ae (lpMeasSubgroupToLpTrim_ae_eq hm _), ← Lp.norm_def]
congr

theorem isometry_lpMeasSubgroupToLpTrim [hp : Fact (1 ≤ p)] (hm : m ≤ m0) :
Expand Down Expand Up @@ -470,7 +459,6 @@ theorem lpMeasToLpTrimLie_symm_indicator [one_le_p : Fact (1 ≤ p)] [NormedSpac
((lpMeasToLpTrimLie F ℝ p μ hm).symm (indicatorConstLp p hs hμs c) : Lp F p μ) =
indicatorConstLp p (hm s hs) ((le_trim hm).trans_lt hμs.lt_top).ne c := by
ext1
rw [← lpMeas_coe]
change
lpTrimToLpMeas F ℝ p μ hm (indicatorConstLp p hs hμs c) =ᵐ[μ]
(indicatorConstLp p _ _ c : α → F)
Expand All @@ -482,7 +470,6 @@ theorem lpMeasToLpTrimLie_symm_toLp [one_le_p : Fact (1 ≤ p)] [NormedSpace ℝ
((lpMeasToLpTrimLie F ℝ p μ hm).symm (hf.toLp f) : Lp F p μ) =
(memLp_of_memLp_trim hm hf).toLp f := by
ext1
rw [← lpMeas_coe]
refine (lpTrimToLpMeas_ae_eq hm _).trans ?_
exact (ae_eq_of_ae_eq_trim (MemLp.coeFn_toLp hf)).trans (MemLp.coeFn_toLp _).symm

Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -379,7 +379,7 @@ theorem condExpInd_of_measurable (hs : MeasurableSet[m] s) (hμs : μ s ≠ ∞)
refine EventuallyEq.trans ?_ indicatorConstLp_coeFn.symm
refine (condExpInd_ae_eq_condExpIndSMul hm (hm s hs) hμs c).trans ?_
refine (condExpIndSMul_ae_eq_smul hm (hm s hs) hμs c).trans ?_
rw [lpMeas_coe, condExpL2_indicator_of_measurable hm hs hμs (1 : ℝ)]
rw [condExpL2_indicator_of_measurable hm hs hμs (1 : ℝ)]
refine (@indicatorConstLp_coeFn α _ _ 2 μ _ s (hm s hs) hμs (1 : ℝ)).mono fun x hx => ?_
dsimp only
rw [hx]
Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -112,15 +112,15 @@ theorem norm_condExpL2_le (hm : m ≤ m0) (f : α →₂[μ] E) : ‖condExpL2 E

theorem eLpNorm_condExpL2_le (hm : m ≤ m0) (f : α →₂[μ] E) :
eLpNorm (ε := E) (condExpL2 E 𝕜 hm f) 2 μ ≤ eLpNorm f 2 μ := by
rw [lpMeas_coe, ← ENNReal.toReal_le_toReal (Lp.eLpNorm_ne_top _) (Lp.eLpNorm_ne_top _), ←
rw [← ENNReal.toReal_le_toReal (Lp.eLpNorm_ne_top _) (Lp.eLpNorm_ne_top _), ←
Lp.norm_def, ← Lp.norm_def, Submodule.norm_coe]
exact norm_condExpL2_le hm f

@[deprecated (since := "2025-01-21")] alias eLpNorm_condexpL2_le := eLpNorm_condExpL2_le

theorem norm_condExpL2_coe_le (hm : m ≤ m0) (f : α →₂[μ] E) :
‖(condExpL2 E 𝕜 hm f : α →₂[μ] E)‖ ≤ ‖f‖ := by
rw [Lp.norm_def, Lp.norm_def, ← lpMeas_coe]
rw [Lp.norm_def, Lp.norm_def]
exact ENNReal.toReal_mono (Lp.eLpNorm_ne_top _) (eLpNorm_condExpL2_le hm f)

@[deprecated (since := "2025-01-21")] alias norm_condexpL2_coe_le := norm_condExpL2_coe_le
Expand Down Expand Up @@ -209,7 +209,6 @@ theorem condExpL2_ae_eq_zero_of_ae_eq_zero (hs : MeasurableSet[m] s) (hμs : μ
rw [Pi.zero_apply] at hx ⊢
· rwa [ENNReal.coe_eq_zero, nnnorm_eq_zero] at hx
· refine Measurable.coe_nnreal_ennreal (Measurable.nnnorm ?_)
rw [lpMeas_coe]
exact (Lp.stronglyMeasurable _).measurable
refine le_antisymm ?_ (zero_le _)
refine (lintegral_nnnorm_condExpL2_le hs hμs f).trans (le_of_eq ?_)
Expand Down Expand Up @@ -251,9 +250,8 @@ linear maps. -/
theorem condExpL2_const_inner (hm : m ≤ m0) (f : Lp E 2 μ) (c : E) :
condExpL2 𝕜 𝕜 hm (((Lp.memLp f).const_inner c).toLp fun a => ⟪c, f a⟫) =ᵐ[μ]
fun a => ⟪c, (condExpL2 E 𝕜 hm f : α → E) a⟫ := by
rw [lpMeas_coe]
have h_mem_Lp : MemLp (fun a => ⟪c, (condExpL2 E 𝕜 hm f : α → E) a⟫) 2 μ := by
refine MemLp.const_inner _ ?_; rw [lpMeas_coe]; exact Lp.memLp _
refine MemLp.const_inner _ ?_; exact Lp.memLp _
have h_eq : h_mem_Lp.toLp _ =ᵐ[μ] fun a => ⟪c, (condExpL2 E 𝕜 hm f : α → E) a⟫ :=
h_mem_Lp.coeFn_toLp
refine EventuallyEq.trans ?_ h_eq
Expand All @@ -263,14 +261,14 @@ theorem condExpL2_const_inner (hm : m ≤ m0) (f : Lp E 2 μ) (c : E) :
rw [IntegrableOn, integrable_congr (ae_restrict_of_ae h_eq)]
exact (integrableOn_condExpL2_of_measure_ne_top hm hμs.ne _).const_inner _
· intro s hs hμs
rw [← lpMeas_coe, integral_condExpL2_eq_of_fin_meas_real _ hs hμs.ne,
integral_congr_ae (ae_restrict_of_ae h_eq), lpMeas_coe,
rw [integral_condExpL2_eq_of_fin_meas_real _ hs hμs.ne,
integral_congr_ae (ae_restrict_of_ae h_eq), ←
L2.inner_indicatorConstLp_eq_setIntegral_inner 𝕜 (↑(condExpL2 E 𝕜 hm f)) (hm s hs) c hμs.ne,
← inner_condExpL2_left_eq_right, condExpL2_indicator_of_measurable _ hs,
L2.inner_indicatorConstLp_eq_setIntegral_inner 𝕜 f (hm s hs) c hμs.ne,
setIntegral_congr_ae (hm s hs)
((MemLp.coeFn_toLp ((Lp.memLp f).const_inner c)).mono fun x hx _ => hx)]
· rw [← lpMeas_coe]; exact lpMeas.aeStronglyMeasurable _
· exact lpMeas.aeStronglyMeasurable _
· refine AEStronglyMeasurable.congr ?_ h_eq.symm
exact (lpMeas.aeStronglyMeasurable _).const_inner

Expand All @@ -279,7 +277,7 @@ theorem condExpL2_const_inner (hm : m ≤ m0) (f : Lp E 2 μ) (c : E) :
/-- `condExpL2` verifies the equality of integrals defining the conditional expectation. -/
theorem integral_condExpL2_eq (hm : m ≤ m0) (f : Lp E' 2 μ) (hs : MeasurableSet[m] s)
(hμs : μ s ≠ ∞) : ∫ x in s, (condExpL2 E' 𝕜 hm f : α → E') x ∂μ = ∫ x in s, f x ∂μ := by
rw [← sub_eq_zero, lpMeas_coe,
rw [← sub_eq_zero, ←
integral_sub' (integrableOn_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs)
(integrableOn_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs)]
refine integral_eq_zero_of_forall_integral_inner_eq_zero 𝕜 _ ?_ ?_
Expand All @@ -291,7 +289,7 @@ theorem integral_condExpL2_eq (hm : m ≤ m0) (f : Lp E' 2 μ) (hs : MeasurableS
((integrableOn_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs).const_inner c)
((integrableOn_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs).const_inner c)]
have h_ae_eq_f := MemLp.coeFn_toLp (E := 𝕜) ((Lp.memLp f).const_inner c)
rw [← lpMeas_coe, sub_eq_zero, ←
rw [sub_eq_zero, ←
setIntegral_congr_ae (hm s hs) ((condExpL2_const_inner hm f c).mono fun x hx _ => hx), ←
setIntegral_congr_ae (hm s hs) (h_ae_eq_f.mono fun x hx _ => hx)]
exact integral_condExpL2_eq_of_fin_meas_real _ hs hμs
Expand All @@ -313,11 +311,11 @@ theorem condExpL2_comp_continuousLinearMap (hm : m ≤ m0) (T : E' →L[ℝ] E''
rw [T.setIntegral_compLp _ (hm s hs),
T.integral_comp_comm
(integrableOn_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs.ne),
← lpMeas_coe, ← lpMeas_coe, integral_condExpL2_eq hm f hs hμs.ne,
integral_condExpL2_eq hm f hs hμs.ne,
integral_condExpL2_eq hm (T.compLp f) hs hμs.ne, T.setIntegral_compLp _ (hm s hs),
T.integral_comp_comm
(integrableOn_Lp_of_measure_ne_top f fact_one_le_two_ennreal.elim hμs.ne)]
· rw [← lpMeas_coe]; exact lpMeas.aeStronglyMeasurable _
· exact lpMeas.aeStronglyMeasurable _
· have h_coe := T.coeFn_compLp (condExpL2 E' 𝕜 hm f : α →₂[μ] E')
rw [← EventuallyEq] at h_coe
refine AEStronglyMeasurable.congr ?_ h_coe.symm
Expand All @@ -340,7 +338,6 @@ theorem condExpL2_indicator_ae_eq_smul (hm : m ≤ m0) (hs : MeasurableSet s) (h
have h_comp :=
condExpL2_comp_continuousLinearMap ℝ 𝕜 hm (toSpanSingleton ℝ x)
(indicatorConstLp 2 hs hμs (1 : ℝ))
rw [← lpMeas_coe] at h_comp
refine h_comp.trans ?_
exact (toSpanSingleton ℝ x).coeFn_compLp _

Expand All @@ -351,7 +348,6 @@ theorem condExpL2_indicator_eq_toSpanSingleton_comp (hm : m ≤ m0) (hs : Measur
(hμs : μ s ≠ ∞) (x : E') : (condExpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x) : α →₂[μ] E') =
(toSpanSingleton ℝ x).compLp (condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1)) := by
ext1
rw [← lpMeas_coe]
refine (condExpL2_indicator_ae_eq_smul 𝕜 hm hs hμs x).trans ?_
have h_comp := (toSpanSingleton ℝ x).coeFn_compLp
(condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α →₂[μ] ℝ)
Expand All @@ -375,7 +371,7 @@ theorem setLIntegral_nnnorm_condExpL2_indicator_le (hm : m ≤ m0) (hs : Measura
((condExpL2_indicator_ae_eq_smul 𝕜 hm hs hμs x).mono fun a ha _ => by rw [ha])
_ = (∫⁻ a in t, ‖(condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a‖₊ ∂μ) * ‖x‖₊ := by
simp_rw [nnnorm_smul, ENNReal.coe_mul]
rw [lintegral_mul_const, lpMeas_coe]
rw [lintegral_mul_const]
exact (Lp.stronglyMeasurable _).enorm
_ ≤ μ (s ∩ t) * ‖x‖₊ :=
mul_le_mul_right' (lintegral_nnnorm_condExpL2_indicator_le_real hs hμs ht hμt) _
Expand All @@ -401,7 +397,7 @@ theorem integrable_condExpL2_indicator (hm : m ≤ m0) [SigmaFinite (μ.trim hm)
Integrable (ε := E') (condExpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x)) μ := by
refine integrable_of_forall_fin_meas_le' hm (μ s * ‖x‖₊)
(ENNReal.mul_lt_top hμs.lt_top ENNReal.coe_lt_top) ?_ ?_
· rw [lpMeas_coe]; exact Lp.aestronglyMeasurable _
· exact Lp.aestronglyMeasurable _
· refine fun t ht hμt =>
(setLIntegral_nnnorm_condExpL2_indicator_le hm hs hμs x ht hμt).trans ?_
gcongr
Expand Down Expand Up @@ -474,7 +470,7 @@ theorem setLIntegral_nnnorm_condExpIndSMul_le (hm : m ≤ m0) (hs : MeasurableSe
((condExpIndSMul_ae_eq_smul hm hs hμs x).mono fun a ha _ => by rw [ha])
_ = (∫⁻ a in t, ‖(condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a‖₊ ∂μ) * ‖x‖₊ := by
simp_rw [nnnorm_smul, ENNReal.coe_mul]
rw [lintegral_mul_const, lpMeas_coe]
rw [lintegral_mul_const]
exact (Lp.stronglyMeasurable _).enorm
_ ≤ μ (s ∩ t) * ‖x‖₊ :=
mul_le_mul_right' (lintegral_nnnorm_condExpL2_indicator_le_real hs hμs ht hμt) _
Expand Down
Loading