Strongly measurable and finitely strongly measurable functions #
A function f
is said to be almost everywhere strongly measurable if f
is almost everywhere
equal to a strongly measurable function, i.e. the sequential limit of simple functions.
It is said to be almost everywhere finitely strongly measurable with respect to a measure μ
if the supports of those simple functions have finite measure.
Almost everywhere strongly measurable functions form the largest class of functions that can be integrated using the Bochner integral.
Main definitions #
AEStronglyMeasurable f μ
:f
is almost everywhere equal to aStronglyMeasurable
function.AEFinStronglyMeasurable f μ
:f
is almost everywhere equal to aFinStronglyMeasurable
function.AEFinStronglyMeasurable.sigmaFiniteSet
: a measurable sett
such thatf =ᵐ[μ.restrict tᶜ] 0
andμ.restrict t
is sigma-finite.
Main statements #
AEFinStronglyMeasurable.exists_set_sigmaFinite
: there exists a measurable sett
such thatf =ᵐ[μ.restrict tᶜ] 0
andμ.restrict t
is sigma-finite.
We provide a solid API for almost everywhere strongly measurable functions, as a basis for the Bochner integral.
References #
- Hytönen, Tuomas, Jan Van Neerven, Mark Veraar, and Lutz Weis. Analysis in Banach spaces. Springer, 2016.
A function is AEStronglyMeasurable
with respect to a measure μ
if it is almost everywhere
equal to the limit of a sequence of simple functions.
One can specify the sigma-algebra according to which simple functions are taken using the
AEStronglyMeasurable[m]
notation in the MeasureTheory
scope.
Equations
- MeasureTheory.AEStronglyMeasurable f μ = ∃ (g : α → β), MeasureTheory.StronglyMeasurable g ∧ f =ᵐ[μ] g
Instances For
A function is m
-AEStronglyMeasurable
with respect to a measure μ
if it is almost
everywhere equal to the limit of a sequence of m
-simple functions.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A function is AEFinStronglyMeasurable
with respect to a measure if it is almost everywhere
equal to the limit of a sequence of simple functions with support with finite measure.
Equations
- MeasureTheory.AEFinStronglyMeasurable f μ = ∃ (g : α → β), MeasureTheory.FinStronglyMeasurable g μ ∧ f =ᵐ[μ] g
Instances For
Almost everywhere strongly measurable functions #
A StronglyMeasurable
function such that f =ᵐ[μ] hf.mk f
. See lemmas
stronglyMeasurable_mk
and ae_eq_mk
.
Equations
Instances For
The composition of a continuous function and an ae strongly measurable function is ae strongly measurable.
A continuous function from α
to β
is ae strongly measurable when one of the two spaces is
second countable.
The composition of a continuous function of two variables and two ae strongly measurable functions is ae strongly measurable.
In a space with second countable topology, measurable implies ae strongly measurable.
If the restriction to a set s
of a σ-algebra m
is included in the restriction to s
of
another σ-algebra m₂
(hypothesis hs
), the set s
is m
measurable and a function f
almost
everywhere supported on s
is m
-ae-strongly-measurable, then f
is also
m₂
-ae-strongly-measurable.
Big operators: ∏
and ∑
#
In a space with second countable topology, measurable implies strongly measurable.
In a space with second countable topology, strongly measurable and measurable are equivalent.
Alias of MeasureTheory.AEStronglyMeasurable.enorm
.
A function is almost everywhere strongly measurable if and only if it is almost everywhere measurable, and up to a zero measure set its range is contained in a separable set.
Alias of Topology.IsEmbedding.aestronglyMeasurable_comp_iff
.
An almost everywhere sequential limit of almost everywhere strongly measurable functions is almost everywhere strongly measurable.
If a sequence of almost everywhere strongly measurable functions converges almost everywhere, one can select a strongly measurable function as the almost everywhere limit.
Almost everywhere finitely strongly measurable functions #
A fin_strongly_measurable
function such that f =ᵐ[μ] hf.mk f
. See lemmas
fin_strongly_measurable_mk
and ae_eq_mk
.
Equations
Instances For
A measurable set t
such that f =ᵐ[μ.restrict tᶜ] 0
and sigma_finite (μ.restrict t)
.
Equations
- hf.sigmaFiniteSet = ⋯.choose
Instances For
In a space with second countable topology and a sigma-finite measure,
AEFinStronglyMeasurable
and AEMeasurable
are equivalent.
In a space with second countable topology and a sigma-finite measure,
an AEMeasurable
function is AEFinStronglyMeasurable
.