Documentation

Mathlib.Analysis.Distribution.TemperateGrowth

Functions and measures of temperate growth #

A function is called of temperate growth if it is smooth and all iterated derivatives are polynomially bounded.

Equations
Instances For
    theorem Function.hasTemperateGrowth_iff_isBigO {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] {f : EF} :
    HasTemperateGrowth f ContDiff (↑) f ∀ (n : ), ∃ (k : ), iteratedFDeriv n f =O[] fun (x : E) => (1 + x) ^ k

    A function has temperate growth if and only if it is smooth and its n-th iterated derivative is O((1 + ‖x‖) ^ k) for some k : ℕ (depending on n).

    Note that the O here is with respect to the filter, meaning that the bound holds everywhere.

    TODO: when E is finite dimensional, this is equivalent to the derivatives being O(‖x‖ ^ k) as ‖x‖ → ∞.

    theorem Function.HasTemperateGrowth.isBigO {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] {f : EF} (hf_temperate : HasTemperateGrowth f) (n : ) :
    ∃ (k : ), iteratedFDeriv n f =O[] fun (x : E) => (1 + x) ^ k

    If f has temperate growth, then its n-th iterated derivative is O((1 + ‖x‖) ^ k) for some k : ℕ (depending on n).

    Note that the O here is with respect to the filter, meaning that the bound holds everywhere.

    theorem Function.HasTemperateGrowth.isBigO_uniform {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] {f : EF} (hf_temperate : HasTemperateGrowth f) (N : ) :
    ∃ (k : ), nN, iteratedFDeriv n f =O[] fun (x : E) => (1 + x) ^ k

    If f has temperate growth, then for any N : ℕ one can find k such that all iterated derivatives of f of order ≤ N are O((1 + ‖x‖) ^ k).

    Note that the O here is with respect to the filter, meaning that the bound holds everywhere.

    theorem Function.HasTemperateGrowth.norm_iteratedFDeriv_le_uniform {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] {f : EF} (hf_temperate : HasTemperateGrowth f) (n : ) :
    ∃ (k : ) (C : ), 0 C Nn, ∀ (x : E), iteratedFDeriv N f x C * (1 + x) ^ k
    @[deprecated Function.HasTemperateGrowth.norm_iteratedFDeriv_le_uniform (since := "2025-10-30")]
    theorem Function.HasTemperateGrowth.norm_iteratedFDeriv_le_uniform_aux {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] {f : EF} (hf_temperate : HasTemperateGrowth f) (n : ) :
    ∃ (k : ) (C : ), 0 C Nn, ∀ (x : E), iteratedFDeriv N f x C * (1 + x) ^ k

    Alias of Function.HasTemperateGrowth.norm_iteratedFDeriv_le_uniform.

    theorem Function.HasTemperateGrowth.of_fderiv {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] {f : EF} (h'f : HasTemperateGrowth (fderiv f)) (hf : Differentiable f) {k : } {C : } (h : ∀ (x : E), f x C * (1 + x) ^ k) :
    @[simp]
    theorem Function.HasTemperateGrowth.comp' {D : Type u_4} {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedAddCommGroup D] [NormedSpace D] {g : EF} {f : DE} {t : Set E} (ht : Set.range f t) (ht' : UniqueDiffOn t) (hg₁ : ContDiffOn (↑) g t) (hg₂ : ∀ (N : ), ∃ (k : ) (C : ) (_ : 0 C), nN, xt, iteratedFDerivWithin n g t x C * (1 + x) ^ k) (hf : HasTemperateGrowth f) :

    Composition of two temperate growth functions is of temperate growth.

    Version where the outer function g is only of temperate growth on the image of inner function f.

    Composition of two temperate growth functions is of temperate growth.

    theorem Function.HasTemperateGrowth.fun_add {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] {f g : EF} (hf : HasTemperateGrowth f) (hg : HasTemperateGrowth g) :
    HasTemperateGrowth fun (i : E) => f i + g i

    Eta-expanded form of Function.HasTemperateGrowth.add

    theorem Function.HasTemperateGrowth.fun_sub {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] {f g : EF} (hf : HasTemperateGrowth f) (hg : HasTemperateGrowth g) :
    HasTemperateGrowth fun (i : E) => f i - g i

    Eta-expanded form of Function.HasTemperateGrowth.sub

    theorem Function.HasTemperateGrowth.sum {ι : Type u_1} {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] {f : ιEF} {s : Finset ι} (hf : is, HasTemperateGrowth (f i)) :
    HasTemperateGrowth fun (x : E) => is, f i x
    theorem ContinuousLinearMap.bilinear_hasTemperateGrowth {𝕜 : Type u_2} {D : Type u_4} {E : Type u_5} {F : Type u_6} {G : Type u_7} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NontriviallyNormedField 𝕜] [NormedAlgebra 𝕜] [NormedAddCommGroup D] [NormedSpace D] [NormedAddCommGroup G] [NormedSpace G] [NormedSpace 𝕜 F] [NormedSpace 𝕜 G] [NormedSpace 𝕜 E] (B : E →L[𝕜] F →L[𝕜] G) {f : DE} {g : DF} (hf : Function.HasTemperateGrowth f) (hg : Function.HasTemperateGrowth g) :
    Function.HasTemperateGrowth fun (x : D) => (B (f x)) (g x)

    The product of two functions of temperate growth is again of temperate growth.

    Version for bilinear maps.

    theorem Function.HasTemperateGrowth.smul {𝕜 : Type u_2} {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NontriviallyNormedField 𝕜] [NormedAlgebra 𝕜] [NormedSpace 𝕜 F] {f : E𝕜} {g : EF} (hf : HasTemperateGrowth f) (hg : HasTemperateGrowth g) :

    The product of two functions of temperate growth is again of temperate growth.

    Version for scalar multiplication.

    theorem Function.HasTemperateGrowth.fun_smul {𝕜 : Type u_2} {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NontriviallyNormedField 𝕜] [NormedAlgebra 𝕜] [NormedSpace 𝕜 F] {f : E𝕜} {g : EF} (hf : HasTemperateGrowth f) (hg : HasTemperateGrowth g) :
    HasTemperateGrowth fun (i : E) => f i g i

    Eta-expanded form of Function.HasTemperateGrowth.smul


    The product of two functions of temperate growth is again of temperate growth.

    Version for scalar multiplication.

    The product of two functions of temperate growth is again of temperate growth.

    theorem Function.HasTemperateGrowth.fun_mul {R : Type u_3} {E : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedRing R] [NormedAlgebra R] {f g : ER} (hf : HasTemperateGrowth f) (hg : HasTemperateGrowth g) :
    HasTemperateGrowth fun (i : E) => f i * g i

    Eta-expanded form of Function.HasTemperateGrowth.mul


    The product of two functions of temperate growth is again of temperate growth.

    theorem Function.HasTemperateGrowth.fun_pow {R : Type u_3} {E : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedRing R] [NormedAlgebra R] {f : ER} (hf : HasTemperateGrowth f) (k : ) :
    HasTemperateGrowth fun (i : E) => f i ^ k

    Eta-expanded form of Function.HasTemperateGrowth.pow

    The Bessel potential x ↦ (1 + ‖x‖ ^ 2) ^ r has temperate growth.

    A measure μ has temperate growth if there is an n : ℕ such that (1 + ‖x‖) ^ (- n) is μ-integrable.

    Instances

      An integer exponent l such that (1 + ‖x‖) ^ (-l) is integrable if μ has temperate growth.

      Equations
      Instances For
        theorem pow_mul_le_of_le_of_pow_mul_le {C₁ C₂ : } {k l : } {x f : } (hx : 0 x) (hf : 0 f) (h₁ : f C₁) (h₂ : x ^ (k + l) * f C₂) :
        x ^ k * f 2 ^ l * (C₁ + C₂) * (1 + x) ^ (-l)

        Pointwise inequality to control x ^ k * f in terms of 1 / (1 + x) ^ l if one controls both f (with a bound C₁) and x ^ (k + l) * f (with a bound C₂). This will be used to check integrability of x ^ k * f x when f is a Schwartz function, and to control explicitly its integral in terms of suitable seminorms of f.

        theorem integrable_of_le_of_pow_mul_le {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [MeasurableSpace E] [NormedAddCommGroup F] [BorelSpace E] [SecondCountableTopology E] {μ : MeasureTheory.Measure E} [μ.HasTemperateGrowth] {f : EF} {C₁ C₂ : } {k : } (hf : ∀ (x : E), f x C₁) (h'f : ∀ (x : E), x ^ (k + μ.integrablePower) * f x C₂) (h''f : MeasureTheory.AEStronglyMeasurable f μ) :
        MeasureTheory.Integrable (fun (x : E) => x ^ k * f x) μ

        Given a function such that f and x ^ (k + l) * f are bounded for a suitable l, then x ^ k * f is integrable. The bounds are not relevant for the integrability conclusion, but they are relevant for bounding the integral in integral_pow_mul_le_of_le_of_pow_mul_le. We formulate the two lemmas with the same set of assumptions for ease of applications.

        theorem integral_pow_mul_le_of_le_of_pow_mul_le {E : Type u_5} {F : Type u_6} [NormedAddCommGroup E] [MeasurableSpace E] [NormedAddCommGroup F] {μ : MeasureTheory.Measure E} [μ.HasTemperateGrowth] {f : EF} {C₁ C₂ : } {k : } (hf : ∀ (x : E), f x C₁) (h'f : ∀ (x : E), x ^ (k + μ.integrablePower) * f x C₂) :
        (x : E), x ^ k * f x μ (2 ^ μ.integrablePower * (x : E), (1 + x) ^ (-μ.integrablePower) μ) * (C₁ + C₂)

        Given a function such that f and x ^ (k + l) * f are bounded for a suitable l, then one can bound explicitly the integral of x ^ k * f.

        theorem MeasureTheory.Measure.HasTemperateGrowth.exists_eLpNorm_lt_top {E : Type u_5} [NormedAddCommGroup E] [MeasurableSpace E] (p : ENNReal) {μ : Measure E} ( : μ.HasTemperateGrowth) :
        ∃ (k : ), eLpNorm (fun (x : E) => (1 + x) ^ (-k)) p μ <

        For any HasTemperateGrowth measure and p, there exists an integer power k such that (1 + ‖x‖) ^ (-k) is in L^p.