Indicator of a set as an element of Lp
#
For a set s
with (hs : MeasurableSet s)
and (hμs : μ s < ∞)
, we build
indicatorConstLp p hs hμs c
, the element of Lp
corresponding to s.indicator (fun _ => c)
.
Main definitions #
MeasureTheory.indicatorConstLp
: Indicator of a set as an element ofLp
.MeasureTheory.Lp.const
: Constant function as an element ofLp
for a finite measure.
The eLpNorm
of the indicator of a set is uniformly small if the set itself has small measure,
for any p < ∞
. Given here as an existential ∀ ε > 0, ∃ η > 0, ...
to avoid later
management of ℝ≥0∞
-arithmetic.
A bounded measurable function with compact support is in L^p.
Alias of HasCompactSupport.memLp_of_bound
.
A bounded measurable function with compact support is in L^p.
A continuous function with compact support is in L^p.
Alias of Continuous.memLp_of_hasCompactSupport
.
A continuous function with compact support is in L^p.
Indicator of a set as an element of Lp
.
Equations
- MeasureTheory.indicatorConstLp p hs hμs c = MeasureTheory.MemLp.toLp (s.indicator fun (x : α) => c) ⋯
Instances For
A version of Set.indicator_add
for MeasureTheory.indicatorConstLp
A version of Set.indicator_sub
for MeasureTheory.indicatorConstLp
Alias of MeasureTheory.enorm_indicatorConstLp_le
.
A family of indicatorConstLp
functions tends to an indicatorConstLp
,
if the underlying sets tend to the set in the sense of the measure of the symmetric difference.
A family of indicatorConstLp
functions is continuous in the parameter,
if μ (s y ∆ s x)
tends to zero as y
tends to x
for all x
.
Alias of MeasureTheory.memLp_add_of_disjoint
.
The indicator of a disjoint union of two sets is the sum of the indicators of the sets.
Constant function as an element of MeasureTheory.Lp
for a finite measure.
Equations
- MeasureTheory.Lp.const p μ = { toFun := fun (c : E) => ⟨MeasureTheory.AEEqFun.const α c, ⋯⟩, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Alias of MeasureTheory.MemLp.toLp_const
.
MeasureTheory.Lp.const
as a LinearMap
.
Equations
- MeasureTheory.Lp.constₗ p μ 𝕜 = { toFun := ⇑(MeasureTheory.Lp.const p μ), map_add' := ⋯, map_smul' := ⋯ }
Instances For
Equations
- MeasureTheory.Lp.constL p μ 𝕜 = (MeasureTheory.Lp.constₗ p μ 𝕜).mkContinuous ((μ Set.univ).toReal ^ (1 / p.toReal)) ⋯