Bochner integral
The Bochner integral extends the definition of the Lebesgue integral to functions that map from a measure space into a Banach space (complete normed vector space). It is constructed here by extending the integral on simple functions.
Main definitions
The Bochner integral is defined following these steps:
Define the integral on simple functions of the type
simple_func α E
(notation :α →ₛ E
) whereE
is a real normed space.(See
simple_func.bintegral
and sectionbintegral
for details. Also seesimple_func.integral
for the integral on simple functions of the typesimple_func α ennreal
.)Use
α →ₛ E
to cut out the simple functions from L1 functions, and define integral on these. The type of simple functions in L1 space is written asα →₁ₛ[μ] E
.Show that the embedding of
α →₁ₛ[μ] E
into L1 is a dense and uniform one.Show that the integral defined on
α →₁ₛ[μ] E
is a continuous linear map.Define the Bochner integral on L1 functions by extending the integral on integrable simple functions
α →₁ₛ[μ] E
usingcontinuous_linear_map.extend
. Define the Bochner integral on functions as the Bochner integral of its equivalence class in L1 space.
Main statements
Basic properties of the Bochner integral on functions of type
α → E
, whereα
is a measure space andE
is a real normed space.integral_zero
:∫ 0 = 0
integral_add
:∫ f + g = ∫ f + ∫ g
integral_neg
:∫ -f = - ∫ f
integral_sub
:∫ f - g = ∫ f - ∫ g
integral_smul
:∫ r • f = r • ∫ f
integral_congr_ae
:∀ᵐ a, f a = g a → ∫ f = ∫ g
norm_integral_le_integral_norm
:∥∫ f∥ ≤ ∫ ∥f∥
Basic properties of the Bochner integral on functions of type
α → ℝ
, whereα
is a measure space.integral_nonneg_of_ae
:∀ᵐ a, 0 ≤ f a → 0 ≤ ∫ f
integral_nonpos_of_nonpos_ae
:∀ᵐ a, f a ≤ 0 → ∫ f ≤ 0
integral_le_integral_of_le_ae
:∀ᵐ a, f a ≤ g a → ∫ f ≤ ∫ g
Propositions connecting the Bochner integral with the integral on
ennreal
-valued functions, which is calledlintegral
and has the notation∫⁻
.integral_eq_lintegral_max_sub_lintegral_min
:∫ f = ∫⁻ f⁺ - ∫⁻ f⁻
, wheref⁺
is the positive part off
andf⁻
is the negative part off
.integral_eq_lintegral_of_nonneg_ae
:∀ᵐ a, 0 ≤ f a → ∫ f = ∫⁻ f
tendsto_integral_of_dominated_convergence
: the Lebesgue dominated convergence theorem
Notes
Some tips on how to prove a proposition if the API for the Bochner integral is not enough so that you need to unfold the definition of the Bochner integral and go back to simple functions.
See integral_eq_lintegral_max_sub_lintegral_min
for a complicated example, which proves that
∫ f = ∫⁻ f⁺ - ∫⁻ f⁻
, with the first integral sign being the Bochner integral of a real-valued
function f : α → ℝ, and second and third integral sign being the integral on ennreal-valued
functions (called lintegral
). The proof of integral_eq_lintegral_max_sub_lintegral_min
is
scattered in sections with the name pos_part
.
Here are the usual steps of proving that a property p
, say ∫ f = ∫⁻ f⁺ - ∫⁻ f⁻
, holds for all
functions :
First go to the
L¹
space.For example, if you see
ennreal.to_real (∫⁻ a, ennreal.of_real $ ∥f a∥)
, that is the norm off
inL¹
space. Rewrite usingl1.norm_of_fun_eq_lintegral_norm
.Show that the set
{f ∈ L¹ | ∫ f = ∫⁻ f⁺ - ∫⁻ f⁻}
is closed inL¹
usingis_closed_eq
.Show that the property holds for all simple functions
s
inL¹
space.Typically, you need to convert various notions to their
simple_func
counterpart, using lemmas likel1.integral_coe_eq_integral
.Since simple functions are dense in
L¹
,
univ = closure {s simple}
= closure {s simple | ∫ s = ∫⁻ s⁺ - ∫⁻ s⁻} : the property holds for all simple functions
⊆ closure {f | ∫ f = ∫⁻ f⁺ - ∫⁻ f⁻}
= {f | ∫ f = ∫⁻ f⁺ - ∫⁻ f⁻} : closure of a closed set is itself
Use is_closed_property
or dense_range.induction_on
for this argument.
Notations
α →ₛ E
: simple functions (defined inmeasure_theory/integration
)α →₁[μ] E
: functions in L1 space, i.e., equivalence classes of integrable functions (defined inmeasure_theory/l1_space
)α →₁ₛ[μ] E
: simple functions in L1 space, i.e., equivalence classes of integrable simple functions
Note : ₛ
is typed using \_s
. Sometimes it shows as a box if font is missing.
Tags
Bochner integral, simple function, function space, Lebesgue dominated convergence theorem
Positive part of a simple function.
Equations
- f.pos_part = measure_theory.simple_func.map (λ (b : E), max b 0) f
Negative part of a simple function.
The Bochner integral of simple functions
Define the Bochner integral of simple functions of the type α →ₛ β
where β
is a normed group,
and prove basic property of this integral.
For simple functions with a normed_group
as codomain, being integrable is the same as having
finite volume support.
Bochner integral of simple functions whose codomain is a real normed_space
.
Calculate the integral of g ∘ f : α →ₛ F
, where f
is an integrable function from α
to E
and g
is a function from E
to F
. We require g 0 = 0
so that g ∘ f
is integrable.
simple_func.integral
and simple_func.lintegral
agree when the integrand has type
α →ₛ ennreal
. But since ennreal
is not a normed_space
, we need some form of coercion.
See integral_eq_lintegral
for a simpler version.
simple_func.bintegral
and simple_func.integral
agree when the integrand has type
α →ₛ ennreal
. But since ennreal
is not a normed_space
, we need some form of coercion.
l1.simple_func
is a subspace of L1 consisting of equivalence classes of an integrable simple
function.
Simple functions in L1 space form a normed_space
.
L1 simple functions forms a emetric_space
, with the emetric being inherited from L1 space,
i.e., edist f g = ∫⁻ a, edist (f a) (g a)
.
Not declared as an instance as α →₁ₛ[μ] β
will only be useful in the construction of the bochner
integral.
L1 simple functions forms a metric_space
, with the metric being inherited from L1 space,
i.e., dist f g = ennreal.to_real (∫⁻ a, edist (f a) (g a)
).
Not declared as an instance as α →₁ₛ[μ] β
will only be useful in the construction of the bochner
integral.
Functions α →₁ₛ[μ] E
form an additive commutative group.
Equations
- measure_theory.l1.simple_func.add_comm_group = {carrier := {f : α →₁[μ] E | ∃ (s : measure_theory.simple_func α E), measure_theory.ae_eq_fun.mk ⇑s _ = ↑f}, zero_mem' := _, add_mem' := _, neg_mem' := _}.to_add_comm_group
Equations
The norm on α →₁ₛ[μ] E
is inherited from L1 space. That is, ∥f∥ = ∫⁻ a, edist (f a) 0
.
Not declared as an instance as α →₁ₛ[μ] E
will only be useful in the construction of the bochner
integral.
Not declared as an instance as α →₁ₛ[μ] E
will only be useful in the construction of the bochner
integral.
Equations
- measure_theory.l1.simple_func.normed_group = normed_group.of_add_dist measure_theory.l1.simple_func.normed_group._proof_1 measure_theory.l1.simple_func.normed_group._proof_2
Not declared as an instance as α →₁ₛ[μ] E
will only be useful in the construction of the bochner
integral.
Not declared as an instance as α →₁ₛ[μ] E
will only be useful in the construction of the bochner
integral.
Equations
- measure_theory.l1.simple_func.semimodule = {to_distrib_mul_action := {to_mul_action := {to_has_scalar := measure_theory.l1.simple_func.has_scalar _inst_11, one_smul := _, mul_smul := _}, smul_add := _, smul_zero := _}, add_smul := _, zero_smul := _}
Not declared as an instance as α →₁ₛ[μ] E
will only be useful in the construction of the bochner
integral.
Equations
Construct the equivalence class [f]
of an integrable simple function f
.
Equations
- measure_theory.l1.simple_func.of_simple_func f hf = ⟨measure_theory.l1.of_fun ⇑f _ hf, _⟩
Find a representative of a l1.simple_func
.
Equations
f.to_simple_func
is measurable.
f.to_simple_func
is integrable.
The uniform and dense embedding of L1 simple functions into L1 functions.
Equations
- measure_theory.l1.simple_func.coe_to_l1 α E 𝕜 = {to_linear_map := {to_fun := coe coe_to_lift, map_add' := _, map_smul' := _}, cont := _}
Positive part of a simple function in L1 space.
Negative part of a simple function in L1 space.
Define the Bochner integral on α →₁ₛ[μ] E
and prove basic properties of this integral.
The Bochner integral over simple functions in l1 space.
Equations
The Bochner integral over simple functions in l1 space as a continuous linear map.
Equations
- measure_theory.l1.simple_func.integral_clm = {to_fun := measure_theory.l1.simple_func.integral _inst_10, map_add' := _, map_smul' := _}.mk_continuous 1 measure_theory.l1.simple_func.integral_clm._proof_1
The Bochner integral in l1 space as a continuous linear map.
The Bochner integral in l1 space
Equations
The Bochner integral
Equations
- measure_theory.integral μ f = dite (measurable f ∧ measure_theory.integrable f μ) (λ (hf : measurable f ∧ measure_theory.integrable f μ), (measure_theory.l1.of_fun f _ _).integral) (λ (hf : ¬(measurable f ∧ measure_theory.integrable f μ)), 0)
If F i → f
in L1
, then ∫ x, F i x ∂μ → ∫ x, f x∂μ
.
Lebesgue dominated convergence theorem provides sufficient conditions under which almost everywhere convergence of a sequence of functions implies the convergence of their integrals.
Lebesgue dominated convergence theorem for filters with a countable basis
The Bochner integral of a real-valued function f : α → ℝ
is the difference between the
integral of the positive part of f
and the integral of the negative part of f
.