Documentation

Mathlib.MeasureTheory.Function.LpSpace.Indicator

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 #

theorem MeasureTheory.exists_eLpNorm_indicator_le {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] (hp : p ) (c : E) {ε : ENNReal} ( : ε 0) :
∃ (η : NNReal), 0 < η ∀ (s : Set α), μ s ηeLpNorm (s.indicator fun (x : α) => c) p μ ε

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.

@[deprecated HasCompactSupport.memLp_of_bound (since := "2025-02-21")]

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.

@[deprecated Continuous.memLp_of_hasCompactSupport (since := "2025-02-21")]

Alias of Continuous.memLp_of_hasCompactSupport.


A continuous function with compact support is in L^p.

def MeasureTheory.indicatorConstLp {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} (p : ENNReal) (hs : MeasurableSet s) (hμs : μ s ) (c : E) :
(Lp E p μ)

Indicator of a set as an element of Lp.

Equations
Instances For
    theorem MeasureTheory.indicatorConstLp_add {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c c' : E} :
    indicatorConstLp p hs hμs c + indicatorConstLp p hs hμs c' = indicatorConstLp p hs hμs (c + c')

    A version of Set.indicator_add for MeasureTheory.indicatorConstLp

    theorem MeasureTheory.indicatorConstLp_sub {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c c' : E} :
    indicatorConstLp p hs hμs c - indicatorConstLp p hs hμs c' = indicatorConstLp p hs hμs (c - c')

    A version of Set.indicator_sub for MeasureTheory.indicatorConstLp

    theorem MeasureTheory.indicatorConstLp_coeFn {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} :
    (indicatorConstLp p hs hμs c) =ᶠ[ae μ] s.indicator fun (x : α) => c
    theorem MeasureTheory.indicatorConstLp_coeFn_mem {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} :
    ∀ᵐ (x : α) μ, x s(indicatorConstLp p hs hμs c) x = c
    theorem MeasureTheory.indicatorConstLp_coeFn_nmem {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} :
    ∀ᵐ (x : α) μ, xs(indicatorConstLp p hs hμs c) x = 0
    theorem MeasureTheory.norm_indicatorConstLp {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} (hp_ne_zero : p 0) (hp_ne_top : p ) :
    indicatorConstLp p hs hμs c = c * (μ s).toReal ^ (1 / p.toReal)
    theorem MeasureTheory.norm_indicatorConstLp_top {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} (hμs_ne_zero : μ s 0) :
    theorem MeasureTheory.norm_indicatorConstLp' {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} (hp_pos : p 0) (hμs_pos : μ s 0) :
    indicatorConstLp p hs hμs c = c * (μ s).toReal ^ (1 / p.toReal)
    theorem MeasureTheory.norm_indicatorConstLp_le {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} :
    indicatorConstLp p hs hμs c c * (μ s).toReal ^ (1 / p.toReal)
    theorem MeasureTheory.nnnorm_indicatorConstLp_le {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} :
    theorem MeasureTheory.enorm_indicatorConstLp_le {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} :
    @[deprecated MeasureTheory.enorm_indicatorConstLp_le (since := "2025-01-20")]
    theorem MeasureTheory.ennnorm_indicatorConstLp_le {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} :

    Alias of MeasureTheory.enorm_indicatorConstLp_le.

    theorem MeasureTheory.edist_indicatorConstLp_eq_enorm {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} {t : Set α} {ht : MeasurableSet t} {hμt : μ t } :
    edist (indicatorConstLp p hs hμs c) (indicatorConstLp p ht hμt c) = indicatorConstLp p c‖ₑ
    @[deprecated MeasureTheory.edist_indicatorConstLp_eq_enorm (since := "2025-01-20")]
    theorem MeasureTheory.edist_indicatorConstLp_eq_nnnorm {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} {t : Set α} {ht : MeasurableSet t} {hμt : μ t } :
    edist (indicatorConstLp p hs hμs c) (indicatorConstLp p ht hμt c) = indicatorConstLp p c‖ₑ

    Alias of MeasureTheory.edist_indicatorConstLp_eq_enorm.

    theorem MeasureTheory.dist_indicatorConstLp_eq_norm {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} {t : Set α} {ht : MeasurableSet t} {hμt : μ t } :
    dist (indicatorConstLp p hs hμs c) (indicatorConstLp p ht hμt c) = indicatorConstLp p c
    theorem MeasureTheory.tendsto_indicatorConstLp_set {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s : Set α} {hs : MeasurableSet s} {hμs : μ s } {c : E} [hp₁ : Fact (1 p)] {β : Type u_3} {l : Filter β} {t : βSet α} {ht : ∀ (b : β), MeasurableSet (t b)} {hμt : ∀ (b : β), μ (t b) } (hp : p ) (h : Filter.Tendsto (fun (b : β) => μ (symmDiff (t b) s)) l (nhds 0)) :
    Filter.Tendsto (fun (b : β) => indicatorConstLp p c) l (nhds (indicatorConstLp p hs hμs c))

    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.

    theorem MeasureTheory.continuous_indicatorConstLp_set {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {c : E} [Fact (1 p)] {X : Type u_3} [TopologicalSpace X] {s : XSet α} {hs : ∀ (x : X), MeasurableSet (s x)} {hμs : ∀ (x : X), μ (s x) } (hp : p ) (h : ∀ (x : X), Filter.Tendsto (fun (y : X) => μ (symmDiff (s y) (s x))) (nhds x) (nhds 0)) :
    Continuous fun (x : X) => indicatorConstLp p c

    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.

    @[simp]
    theorem MeasureTheory.indicatorConstLp_empty {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {c : E} :
    indicatorConstLp p c = 0
    theorem MeasureTheory.indicatorConstLp_inj {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s t : Set α} (hs : MeasurableSet s) (hsμ : μ s ) (ht : MeasurableSet t) (htμ : μ t ) {c : E} (hc : c 0) :
    indicatorConstLp p hs hsμ c = indicatorConstLp p ht htμ c s =ᶠ[ae μ] t
    theorem MeasureTheory.memLp_add_of_disjoint {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {f g : αE} (h : Disjoint (Function.support f) (Function.support g)) (hf : StronglyMeasurable f) (hg : StronglyMeasurable g) :
    MemLp (f + g) p μ MemLp f p μ MemLp g p μ
    @[deprecated MeasureTheory.memLp_add_of_disjoint (since := "2025-02-21")]
    theorem MeasureTheory.memℒp_add_of_disjoint {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {f g : αE} (h : Disjoint (Function.support f) (Function.support g)) (hf : StronglyMeasurable f) (hg : StronglyMeasurable g) :
    MemLp (f + g) p μ MemLp f p μ MemLp g p μ

    Alias of MeasureTheory.memLp_add_of_disjoint.

    theorem MeasureTheory.indicatorConstLp_disjoint_union {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {s t : Set α} (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ) (hμt : μ t ) (hst : Disjoint s t) (c : E) :
    indicatorConstLp p c = indicatorConstLp p hs hμs c + indicatorConstLp p ht hμt c

    The indicator of a disjoint union of two sets is the sum of the indicators of the sets.

    def MeasureTheory.Lp.const {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] :
    E →+ (Lp E p μ)

    Constant function as an element of MeasureTheory.Lp for a finite measure.

    Equations
    Instances For
      theorem MeasureTheory.Lp.coeFn_const {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] (c : E) :
      ((Lp.const p μ) c) =ᶠ[ae μ] Function.const α c
      @[simp]
      theorem MeasureTheory.Lp.const_val {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] (c : E) :
      ((Lp.const p μ) c) = AEEqFun.const α c
      @[simp]
      theorem MeasureTheory.MemLp.toLp_const {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] (c : E) :
      toLp (fun (x : α) => c) = (Lp.const p μ) c
      @[deprecated MeasureTheory.MemLp.toLp_const (since := "2025-02-21")]
      theorem MeasureTheory.Memℒp.toLp_const {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] (c : E) :
      MemLp.toLp (fun (x : α) => c) = (Lp.const p μ) c

      Alias of MeasureTheory.MemLp.toLp_const.

      @[simp]
      theorem MeasureTheory.indicatorConstLp_univ {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] (c : E) :
      indicatorConstLp p c = (Lp.const p μ) c
      theorem MeasureTheory.Lp.norm_const {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] (c : E) [NeZero μ] (hp_zero : p 0) :
      theorem MeasureTheory.Lp.norm_const' {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] (c : E) (hp_zero : p 0) (hp_top : p ) :
      theorem MeasureTheory.Lp.norm_const_le {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] (c : E) :
      def MeasureTheory.Lp.constₗ {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] (𝕜 : Type u_3) [NormedRing 𝕜] [Module 𝕜 E] [IsBoundedSMul 𝕜 E] :
      E →ₗ[𝕜] (Lp E p μ)

      MeasureTheory.Lp.const as a LinearMap.

      Equations
      Instances For
        @[simp]
        theorem MeasureTheory.Lp.constₗ_apply {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] (𝕜 : Type u_3) [NormedRing 𝕜] [Module 𝕜 E] [IsBoundedSMul 𝕜 E] (a : E) :
        (Lp.constₗ p μ 𝕜) a = (Lp.const p μ) a
        def MeasureTheory.Lp.constL {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] (𝕜 : Type u_3) [NormedField 𝕜] [NormedSpace 𝕜 E] [Fact (1 p)] :
        E →L[𝕜] (Lp E p μ)
        Equations
        Instances For
          @[simp]
          theorem MeasureTheory.Lp.constL_apply {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] (𝕜 : Type u_3) [NormedField 𝕜] [NormedSpace 𝕜 E] [Fact (1 p)] (a : E) :
          (Lp.constL p μ 𝕜) a = (Lp.const p μ) a
          theorem MeasureTheory.Lp.norm_constL_le {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} (p : ENNReal) (μ : Measure α) [NormedAddCommGroup E] [IsFiniteMeasure μ] (𝕜 : Type u_3) [NontriviallyNormedField 𝕜] [NormedSpace 𝕜 E] [Fact (1 p)] :
          Lp.constL p μ 𝕜 (μ Set.univ).toReal ^ (1 / p.toReal)
          theorem MeasureTheory.Lp.indicatorConstLp_compMeasurePreserving {α : Type u_1} {E : Type u_2} {m : MeasurableSpace α} {p : ENNReal} {μ : Measure α} [NormedAddCommGroup E] {β : Type u_3} [MeasurableSpace β] {μb : Measure β} {f : αβ} {s : Set β} (hs : MeasurableSet s) (hμs : μb s ) (c : E) (hf : MeasurePreserving f μ μb) :
          (compMeasurePreserving f hf) (indicatorConstLp p hs hμs c) = indicatorConstLp p c