Density of simple functions #
Show that each Lᵖ Borel measurable function can be approximated in Lᵖ norm
by a sequence of simple functions.
Main definitions #
MeasureTheory.Lp.simpleFunc, the type ofLpsimple functionscoeToLp, the embedding ofLp.simpleFunc E p μintoLp E p μ
Main results #
tendsto_approxOn_Lp_eLpNorm(Lᵖ convergence): IfEis aNormedAddCommGroupandfis measurable andMemℒp(forp < ∞), then the simple functionsSimpleFunc.approxOn f hf s 0 h₀ nmay be considered as elements ofLp E p μ, and they tend in Lᵖ tof.Lp.simpleFunc.isDenseEmbedding: the embeddingcoeToLpof theLpsimple functions intoLpis dense.Lp.simpleFunc.induction,Lp.induction,Memℒp.induction,Integrable.induction: to prove a predicate for all elements of one of these classes of functions, it suffices to check that it behaves correctly on simple functions.
TODO #
For E finite-dimensional, simple functions α →ₛ E are dense in L^∞ -- prove this.
Notations #
α →ₛ β(local notation): the type of simple functionsα → β.α →₁ₛ[μ] E: the type ofL1simple functionsα → β.
Lp approximation by simple functions #
Alias of MeasureTheory.SimpleFunc.tendsto_approxOn_Lp_eLpNorm.
Alias of MeasureTheory.SimpleFunc.tendsto_approxOn_range_Lp_eLpNorm.
Any function in ℒp can be approximated by a simple function if p < ∞.
Alias of MeasureTheory.Memℒp.exists_simpleFunc_eLpNorm_sub_lt.
Any function in ℒp can be approximated by a simple function if p < ∞.
L1 approximation by simple functions #
Properties of simple functions in Lp spaces #
A simple function f : α →ₛ E into a normed group E verifies, for a measure μ:
Memℒp f 0 μandMemℒp f ∞ μ, sincefis a.e.-measurable and bounded,- for
0 < p < ∞,Memℒp f p μ ↔ Integrable f μ ↔ f.FinMeasSupp μ ↔ ∀ y, y ≠ 0 → μ (f ⁻¹' {y}) < ∞.
Alias of MeasureTheory.SimpleFunc.eLpNorm'_eq.
Construction of the space of Lp simple functions, and its dense embedding into Lp.
Lp.simpleFunc is a subspace of Lp consisting of equivalence classes of an integrable simple
function.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Simple functions in Lp space form a NormedSpace.
Implementation note: If Lp.simpleFunc E p μ were defined as a 𝕜-submodule of Lp E p μ,
then the next few lemmas, putting a normed 𝕜-group structure on Lp.simpleFunc E p μ, would be
unnecessary. But instead, Lp.simpleFunc E p μ is defined as an AddSubgroup of Lp E p μ,
which does not permit this (but has the advantage of working when E itself is a normed group,
i.e. has no scalar action).
If E is a normed space, Lp.simpleFunc E p μ is a SMul. Not declared as an
instance as it is (as of writing) used only in the construction of the Bochner integral.
Equations
- MeasureTheory.Lp.simpleFunc.smul = { smul := fun (k : 𝕜) (f : ↥(MeasureTheory.Lp.simpleFunc E p μ)) => ⟨k • ↑f, ⋯⟩ }
Instances For
If E is a normed space, Lp.simpleFunc E p μ is a module. Not declared as an
instance as it is (as of writing) used only in the construction of the Bochner integral.
Instances For
If E is a normed space, Lp.simpleFunc E p μ is a normed space. Not declared as an
instance as it is (as of writing) used only in the construction of the Bochner integral.
If E is a normed space, Lp.simpleFunc E p μ is a normed space. Not declared as an
instance as it is (as of writing) used only in the construction of the Bochner integral.
Equations
- MeasureTheory.Lp.simpleFunc.normedSpace = NormedSpace.mk ⋯
Instances For
Construct the equivalence class [f] of a simple function f satisfying Memℒp.
Equations
- MeasureTheory.Lp.simpleFunc.toLp f hf = ⟨MeasureTheory.Memℒp.toLp (⇑f) hf, ⋯⟩
Instances For
Find a representative of a Lp.simpleFunc.
Equations
Instances For
(toSimpleFunc f) is measurable.
toSimpleFunc f satisfies the predicate Memℒp.
The characteristic function of a finite-measure measurable set s, as an Lp simple function.
Equations
- One or more equations did not get rendered due to their size.
Instances For
To prove something for an arbitrary Lp simple function, with 0 < p < ∞, it suffices to show
that the property holds for (multiples of) characteristic functions of finite-measure measurable
sets and is closed under addition (of functions with disjoint support).
The embedding of Lp simple functions into Lp functions, as a continuous linear map.
Equations
- MeasureTheory.Lp.simpleFunc.coeToLp α E 𝕜 = { toFun := (↑(MeasureTheory.Lp.simpleFunc E p μ).subtype).toFun, map_add' := ⋯, map_smul' := ⋯, cont := ⋯ }
Instances For
Equations
- ⋯ = ⋯
Coercion from nonnegative simple functions of Lp to nonnegative functions of Lp.
Equations
- MeasureTheory.Lp.simpleFunc.coeSimpleFuncNonnegToLpNonneg p μ G g = ⟨↑↑g, ⋯⟩
Instances For
To prove something for an arbitrary Lp function in a second countable Borel normed group, it
suffices to show that
- the property holds for (multiples of) characteristic functions;
- is closed under addition;
- the set of functions in
Lpfor which the property holds is closed.
To prove something for an arbitrary Memℒp function in a second countable
Borel normed group, it suffices to show that
- the property holds for (multiples of) characteristic functions;
- is closed under addition;
- the set of functions in the
Lᵖspace for which the property holds is closed. - the property is closed under the almost-everywhere equal relation.
It is possible to make the hypotheses in the induction steps a bit stronger, and such conditions
can be added once we need them (for example in h_add it is only necessary to consider the sum of
a simple function with a multiple of a characteristic function and that the intersection
of their images is a subset of {0}).
If a set of ae strongly measurable functions is stable under addition and approximates
characteristic functions in ℒp, then it is dense in ℒp.
To prove something for an arbitrary integrable function in a normed group, it suffices to show that
- the property holds for (multiples of) characteristic functions;
- is closed under addition;
- the set of functions in the
L¹space for which the property holds is closed. - the property is closed under the almost-everywhere equal relation.
It is possible to make the hypotheses in the induction steps a bit stronger, and such conditions
can be added once we need them (for example in h_add it is only necessary to consider the sum of
a simple function with a multiple of a characteristic function and that the intersection
of their images is a subset of {0}).