Extension of a linear function from indicators to L1 #
Let T : Set α → E →L[ℝ] F
be additive for measurable sets with finite measure, in the sense that
for s, t
two such sets, Disjoint s t → T (s ∪ t) = T s + T t
. T
is akin to a bilinear map on
Set α × E
, or a linear map on indicator functions.
This file constructs an extension of T
to integrable simple functions, which are finite sums of
indicators of measurable sets with finite measure, then to integrable functions, which are limits of
integrable simple functions.
The main result is a continuous linear map (α →₁[μ] E) →L[ℝ] F
. This extension process is used to
define the Bochner integral in the MeasureTheory.Integral.Bochner
file and the conditional
expectation of an integrable function in MeasureTheory.Function.ConditionalExpectation
.
Main Definitions #
FinMeasAdditive μ T
: the property thatT
is additive on measurable sets with finite measure. For two such sets,Disjoint s t → T (s ∪ t) = T s + T t
.DominatedFinMeasAdditive μ T C
:FinMeasAdditive μ T ∧ ∀ s, ‖T s‖ ≤ C * (μ s).toReal
. This is the property needed to perform the extension from indicators to L1.setToL1 (hT : DominatedFinMeasAdditive μ T C) : (α →₁[μ] E) →L[ℝ] F
: the extension ofT
from indicators to L1.setToFun μ T (hT : DominatedFinMeasAdditive μ T C) (f : α → E) : F
: a version of the extension which applies to functions (with value 0 if the function is not integrable).
Properties #
For most properties of setToFun
, we provide two lemmas. One version uses hypotheses valid on
all sets, like T = T'
, and a second version which uses a primed name uses hypotheses on
measurable sets with finite measure, like ∀ s, MeasurableSet s → μ s < ∞ → T s = T' s
.
The lemmas listed here don't show all hypotheses. Refer to the actual lemmas for details.
Linearity:
setToFun_zero_left : setToFun μ 0 hT f = 0
setToFun_add_left : setToFun μ (T + T') _ f = setToFun μ T hT f + setToFun μ T' hT' f
setToFun_smul_left : setToFun μ (fun s ↦ c • (T s)) (hT.smul c) f = c • setToFun μ T hT f
setToFun_zero : setToFun μ T hT (0 : α → E) = 0
setToFun_neg : setToFun μ T hT (-f) = - setToFun μ T hT f
Iff
andg
are integrable:setToFun_add : setToFun μ T hT (f + g) = setToFun μ T hT f + setToFun μ T hT g
setToFun_sub : setToFun μ T hT (f - g) = setToFun μ T hT f - setToFun μ T hT g
IfT
is verifies∀ c : 𝕜, ∀ s x, T s (c • x) = c • T s x
:setToFun_smul : setToFun μ T hT (c • f) = c • setToFun μ T hT f
Other:
setToFun_congr_ae (h : f =ᵐ[μ] g) : setToFun μ T hT f = setToFun μ T hT g
setToFun_measure_zero (h : μ = 0) : setToFun μ T hT f = 0
If the space is a NormedLatticeAddCommGroup
and T
is such that 0 ≤ T s x
for 0 ≤ x
, we
also prove order-related properties:
setToFun_mono_left (h : ∀ s x, T s x ≤ T' s x) : setToFun μ T hT f ≤ setToFun μ T' hT' f
setToFun_nonneg (hf : 0 ≤ᵐ[μ] f) : 0 ≤ setToFun μ T hT f
setToFun_mono (hfg : f ≤ᵐ[μ] g) : setToFun μ T hT f ≤ setToFun μ T hT g
Implementation notes #
The starting object T : Set α → E →L[ℝ] F
matters only through its restriction on measurable sets
with finite measure. Its value on other sets is ignored.
A set function is FinMeasAdditive
if its value on the union of two disjoint measurable
sets with finite measure is the sum of its values on each set.
Equations
- MeasureTheory.FinMeasAdditive μ T = ∀ (s t : Set α), MeasurableSet s → MeasurableSet t → μ s ≠ ⊤ → μ t ≠ ⊤ → Disjoint s t → T (s ∪ t) = T s + T t
Instances For
A FinMeasAdditive
set function whose norm on every set is less than the measure of the
set (up to a multiplicative constant).
Equations
- MeasureTheory.DominatedFinMeasAdditive μ T C = (MeasureTheory.FinMeasAdditive μ T ∧ ∀ (s : Set α), MeasurableSet s → μ s < ⊤ → ‖T s‖ ≤ C * (μ s).toReal)
Instances For
Extend Set α → (F →L[ℝ] F')
to (α →ₛ F) → F'
.
Equations
- MeasureTheory.SimpleFunc.setToSimpleFunc T f = ∑ x_1 ∈ f.range, (T (⇑f ⁻¹' {x_1})) x_1
Instances For
Alias of MeasureTheory.SimpleFunc.norm_setToSimpleFunc_le_sum_opNorm
.
Extend Set α → (E →L[ℝ] F')
to (α →₁ₛ[μ] E) → F'
.
Equations
Instances For
setToL1S
does not change if we replace the measure μ
by μ'
with μ ≪ μ'
. The statement
uses two functions f
and f'
because they have to belong to different types, but morally these
are the same function (we have f =ᵐ[μ] f'
).
Extend Set α → E →L[ℝ] F
to (α →₁ₛ[μ] E) →L[𝕜] F
.
Equations
- MeasureTheory.L1.SimpleFunc.setToL1SCLM' α E 𝕜 μ hT h_smul = { toFun := MeasureTheory.L1.SimpleFunc.setToL1S T, map_add' := ⋯, map_smul' := ⋯ }.mkContinuous C ⋯
Instances For
Extend Set α → E →L[ℝ] F
to (α →₁ₛ[μ] E) →L[ℝ] F
.
Equations
- MeasureTheory.L1.SimpleFunc.setToL1SCLM α E μ hT = { toFun := MeasureTheory.L1.SimpleFunc.setToL1S T, map_add' := ⋯, map_smul' := ⋯ }.mkContinuous C ⋯
Instances For
Extend set α → (E →L[ℝ] F)
to (α →₁[μ] E) →L[𝕜] F
.
Equations
- MeasureTheory.L1.setToL1' 𝕜 hT h_smul = (MeasureTheory.L1.SimpleFunc.setToL1SCLM' α E 𝕜 μ hT h_smul).extend (MeasureTheory.Lp.simpleFunc.coeToLp α E 𝕜) ⋯ ⋯
Instances For
Extend Set α → E →L[ℝ] F
to (α →₁[μ] E) →L[ℝ] F
.
Equations
- MeasureTheory.L1.setToL1 hT = (MeasureTheory.L1.SimpleFunc.setToL1SCLM α E μ hT).extend (MeasureTheory.Lp.simpleFunc.coeToLp α E ℝ) ⋯ ⋯
Instances For
Extend T : Set α → E →L[ℝ] F
to (α → E) → F
(for integrable functions α → E
). We set it to
0 if the function is not integrable.
Equations
- MeasureTheory.setToFun μ T hT f = if hf : MeasureTheory.Integrable f μ then (MeasureTheory.L1.setToL1 hT) (MeasureTheory.Integrable.toL1 f hf) else 0
Instances For
If F i → f
in L1
, then setToFun μ T hT (F i) → setToFun μ T hT f
.
Auxiliary lemma for setToFun_congr_measure
: the function sending f : α →₁[μ] G
to
f : α →₁[μ'] G
is continuous when μ' ≤ c' • μ
for c' ≠ ∞
.
Lebesgue dominated convergence theorem provides sufficient conditions under which almost
everywhere convergence of a sequence of functions implies the convergence of their image by
setToFun
.
We could weaken the condition bound_integrable
to require HasFiniteIntegral bound μ
instead
(i.e. not requiring that bound
is measurable), but in all applications proving integrability
is easier.
Lebesgue dominated convergence theorem for filters with a countable basis