Documentation

Mathlib.Analysis.Distribution.SchwartzSpace

Schwartz space #

This file defines the Schwartz space. Usually, the Schwartz space is defined as the set of smooth functions $f : ℝ^n → ℂ$ such that there exists $C_{αβ} > 0$ with $$|x^α ∂^β f(x)| < C_{αβ}$$ for all $x ∈ ℝ^n$ and for all multiindices $α, β$. In mathlib, we use a slightly different approach and define the Schwartz space as all smooth functions f : E → F, where E and F are real normed vector spaces such that for all natural numbers k and n we have uniform bounds ‖x‖^k * ‖iteratedFDeriv ℝ n f x‖ < C. This approach completely avoids using partial derivatives as well as polynomials. We construct the topology on the Schwartz space by a family of seminorms, which are the best constants in the above estimates. The abstract theory of topological vector spaces developed in SeminormFamily.moduleFilterBasis and WithSeminorms.toLocallyConvexSpace turns the Schwartz space into a locally convex topological vector space.

Main definitions #

Main statements #

Implementation details #

The implementation of the seminorms is taken almost literally from ContinuousLinearMap.opNorm.

Notation #

Tags #

Schwartz space, tempered distributions

structure SchwartzMap (E : Type u_4) (F : Type u_5) [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] :
Type (max u_4 u_5)

A function is a Schwartz function if it is smooth and all derivatives decay faster than any power of ‖x‖.

Instances For

    A function is a Schwartz function if it is smooth and all derivatives decay faster than any power of ‖x‖.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      Equations
      theorem SchwartzMap.decay {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] (f : SchwartzMap E F) (k n : ) :
      ∃ (C : ), 0 < C ∀ (x : E), x ^ k * iteratedFDeriv n (⇑f) x C

      All derivatives of a Schwartz function are rapidly decaying.

      theorem SchwartzMap.smooth {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] (f : SchwartzMap E F) (n : ℕ∞) :
      ContDiff n f

      Every Schwartz function is smooth.

      Every Schwartz function is continuous.

      Every Schwartz function is differentiable.

      Every Schwartz function is differentiable at any point.

      theorem SchwartzMap.ext {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] {f g : SchwartzMap E F} (h : ∀ (x : E), f x = g x) :
      f = g
      theorem SchwartzMap.isBigO_cocompact_zpow_neg_nat {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] (f : SchwartzMap E F) (k : ) :
      f =O[Filter.cocompact E] fun (x : E) => x ^ (-k)

      Auxiliary lemma, used in proving the more general result isBigO_cocompact_rpow.

      theorem SchwartzMap.bounds_nonempty {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] (k n : ) (f : SchwartzMap E F) :
      ∃ (c : ), c {c : | 0 c ∀ (x : E), x ^ k * iteratedFDeriv n (⇑f) x c}
      theorem SchwartzMap.bounds_bddBelow {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] (k n : ) (f : SchwartzMap E F) :
      BddBelow {c : | 0 c ∀ (x : E), x ^ k * iteratedFDeriv n (⇑f) x c}
      theorem SchwartzMap.decay_add_le_aux {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] (k n : ) (f g : SchwartzMap E F) (x : E) :
      theorem SchwartzMap.decay_smul_aux {𝕜 : Type u_1} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (k n : ) (f : SchwartzMap E F) (c : 𝕜) (x : E) :

      Helper definition for the seminorms of the Schwartz space.

      Equations
      Instances For
        theorem SchwartzMap.seminormAux_le_bound {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] (k n : ) (f : SchwartzMap E F) {M : } (hMp : 0 M) (hM : ∀ (x : E), x ^ k * iteratedFDeriv n (⇑f) x M) :

        If one controls the norm of every A x, then one controls the norm of A.

        Algebraic properties #

        instance SchwartzMap.instSMul {𝕜 : Type u_1} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] :
        SMul 𝕜 (SchwartzMap E F)
        Equations
        @[simp]
        theorem SchwartzMap.smul_apply {𝕜 : Type u_1} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] {f : SchwartzMap E F} {c : 𝕜} {x : E} :
        (c f) x = c f x
        instance SchwartzMap.instIsScalarTower {𝕜 : Type u_1} {𝕜' : Type u_2} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] [NormedField 𝕜'] [NormedSpace 𝕜' F] [SMulCommClass 𝕜' F] [SMul 𝕜 𝕜'] [IsScalarTower 𝕜 𝕜' F] :
        IsScalarTower 𝕜 𝕜' (SchwartzMap E F)
        instance SchwartzMap.instSMulCommClass {𝕜 : Type u_1} {𝕜' : Type u_2} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] [NormedField 𝕜'] [NormedSpace 𝕜' F] [SMulCommClass 𝕜' F] [SMulCommClass 𝕜 𝕜' F] :
        SMulCommClass 𝕜 𝕜' (SchwartzMap E F)
        theorem SchwartzMap.seminormAux_smul_le {𝕜 : Type u_1} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (k n : ) (c : 𝕜) (f : SchwartzMap E F) :
        Equations
        Equations
        Equations
        @[simp]
        @[simp]
        theorem SchwartzMap.zero_apply {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] {x : E} :
        0 x = 0
        Equations
        Equations
        @[simp]
        theorem SchwartzMap.add_apply {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] {f g : SchwartzMap E F} {x : E} :
        (f + g) x = f x + g x
        Equations
        @[simp]
        theorem SchwartzMap.sub_apply {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] {f g : SchwartzMap E F} {x : E} :
        (f - g) x = f x - g x

        Coercion as an additive homomorphism.

        Equations
        Instances For

          Seminorms on Schwartz space #

          def SchwartzMap.seminorm (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (k n : ) :

          The seminorms of the Schwartz space given by the best constants in the definition of 𝓢(E, F).

          Equations
          Instances For
            theorem SchwartzMap.seminorm_le_bound (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (k n : ) (f : SchwartzMap E F) {M : } (hMp : 0 M) (hM : ∀ (x : E), x ^ k * iteratedFDeriv n (⇑f) x M) :
            (SchwartzMap.seminorm 𝕜 k n) f M

            If one controls the seminorm for every x, then one controls the seminorm.

            theorem SchwartzMap.seminorm_le_bound' (𝕜 : Type u_1) {F : Type u_5} [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (k n : ) (f : SchwartzMap F) {M : } (hMp : 0 M) (hM : ∀ (x : ), |x| ^ k * iteratedDeriv n (⇑f) x M) :
            (SchwartzMap.seminorm 𝕜 k n) f M

            If one controls the seminorm for every x, then one controls the seminorm.

            Variant for functions 𝓢(ℝ, F).

            theorem SchwartzMap.le_seminorm (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (k n : ) (f : SchwartzMap E F) (x : E) :

            The seminorm controls the Schwartz estimate for any fixed x.

            theorem SchwartzMap.le_seminorm' (𝕜 : Type u_1) {F : Type u_5} [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (k n : ) (f : SchwartzMap F) (x : ) :
            |x| ^ k * iteratedDeriv n (⇑f) x (SchwartzMap.seminorm 𝕜 k n) f

            The seminorm controls the Schwartz estimate for any fixed x.

            Variant for functions 𝓢(ℝ, F).

            theorem SchwartzMap.norm_iteratedFDeriv_le_seminorm (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (f : SchwartzMap E F) (n : ) (x₀ : E) :
            theorem SchwartzMap.norm_pow_mul_le_seminorm (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (f : SchwartzMap E F) (k : ) (x₀ : E) :
            x₀ ^ k * f x₀ (SchwartzMap.seminorm 𝕜 k 0) f
            theorem SchwartzMap.norm_le_seminorm (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (f : SchwartzMap E F) (x₀ : E) :
            f x₀ (SchwartzMap.seminorm 𝕜 0 0) f

            The family of Schwartz seminorms.

            Equations
            Instances For
              theorem SchwartzMap.one_add_le_sup_seminorm_apply {𝕜 : Type u_1} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] {m : × } {k n : } (hk : k m.1) (hn : n m.2) (f : SchwartzMap E F) (x : E) :
              (1 + x) ^ k * iteratedFDeriv n (⇑f) x 2 ^ m.1 * ((Finset.Iic m).sup fun (m : × ) => SchwartzMap.seminorm 𝕜 m.1 m.2) f

              A more convenient version of le_sup_seminorm_apply.

              The set Finset.Iic m is the set of all pairs (k', n') with k' ≤ m.1 and n' ≤ m.2. Note that the constant is far from optimal.

              The topology on the Schwartz space #

              Functions 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.norm_iteratedFDeriv_le_uniform_aux {E : Type u_4} {F : Type u_5} [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
                theorem Function.HasTemperateGrowth.of_fderiv {E : Type u_4} {F : Type u_5} [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) :

                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 SchwartzMap.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 SchwartzMap.integrable_of_le_of_pow_mul_le {D : Type u_3} [NormedAddCommGroup D] [MeasurableSpace D] [BorelSpace D] [SecondCountableTopology D] {E : Type u_8} [NormedAddCommGroup E] {μ : MeasureTheory.Measure D} [μ.HasTemperateGrowth] {f : DE} {C₁ C₂ : } {k : } (hf : ∀ (x : D), f x C₁) (h'f : ∀ (x : D), x ^ (k + μ.integrablePower) * f x C₂) (h''f : MeasureTheory.AEStronglyMeasurable f μ) :
                    MeasureTheory.Integrable (fun (x : D) => 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 SchwartzMap.integral_pow_mul_le_of_le_of_pow_mul_le {D : Type u_3} [NormedAddCommGroup D] [MeasurableSpace D] {E : Type u_8} [NormedAddCommGroup E] {μ : MeasureTheory.Measure D} [μ.HasTemperateGrowth] {f : DE} {C₁ C₂ : } {k : } (hf : ∀ (x : D), f x C₁) (h'f : ∀ (x : D), x ^ (k + μ.integrablePower) * f x C₂) :
                    (x : D), x ^ k * f x μ (2 ^ μ.integrablePower * (x : D), (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 {D : Type u_3} [NormedAddCommGroup D] [MeasurableSpace D] (p : ENNReal) {μ : Measure D} (hμ : μ.HasTemperateGrowth) :
                    ∃ (k : ), eLpNorm (fun (x : D) => (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.

                    Construction of continuous linear maps between Schwartz spaces #

                    def SchwartzMap.mkLM {𝕜 : Type u_1} {𝕜' : Type u_2} {D : Type u_3} {E : Type u_4} {F : Type u_5} {G : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedField 𝕜'] [NormedAddCommGroup D] [NormedSpace D] [NormedSpace 𝕜 E] [SMulCommClass 𝕜 E] [NormedAddCommGroup G] [NormedSpace G] [NormedSpace 𝕜' G] [SMulCommClass 𝕜' G] {σ : 𝕜 →+* 𝕜'} (A : (DE)FG) (hadd : ∀ (f g : SchwartzMap D E) (x : F), A (f + g) x = A (⇑f) x + A (⇑g) x) (hsmul : ∀ (a : 𝕜) (f : SchwartzMap D E) (x : F), A (a f) x = σ a A (⇑f) x) (hsmooth : ∀ (f : SchwartzMap D E), ContDiff (↑) (A f)) (hbound : ∀ (n : × ), ∃ (s : Finset ( × )) (C : ), 0 C ∀ (f : SchwartzMap D E) (x : F), x ^ n.1 * iteratedFDeriv n.2 (A f) x C * (s.sup (schwartzSeminormFamily 𝕜 D E)) f) :

                    Create a semilinear map between Schwartz spaces.

                    Note: This is a helper definition for mkCLM.

                    Equations
                    • SchwartzMap.mkLM A hadd hsmul hsmooth hbound = { toFun := fun (f : SchwartzMap D E) => { toFun := A f, smooth' := , decay' := }, map_add' := , map_smul' := }
                    Instances For
                      def SchwartzMap.mkCLM {𝕜 : Type u_1} {𝕜' : Type u_2} {D : Type u_3} {E : Type u_4} {F : Type u_5} {G : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedField 𝕜'] [NormedAddCommGroup D] [NormedSpace D] [NormedSpace 𝕜 E] [SMulCommClass 𝕜 E] [NormedAddCommGroup G] [NormedSpace G] [NormedSpace 𝕜' G] [SMulCommClass 𝕜' G] {σ : 𝕜 →+* 𝕜'} [RingHomIsometric σ] (A : (DE)FG) (hadd : ∀ (f g : SchwartzMap D E) (x : F), A (f + g) x = A (⇑f) x + A (⇑g) x) (hsmul : ∀ (a : 𝕜) (f : SchwartzMap D E) (x : F), A (a f) x = σ a A (⇑f) x) (hsmooth : ∀ (f : SchwartzMap D E), ContDiff (↑) (A f)) (hbound : ∀ (n : × ), ∃ (s : Finset ( × )) (C : ), 0 C ∀ (f : SchwartzMap D E) (x : F), x ^ n.1 * iteratedFDeriv n.2 (A f) x C * (s.sup (schwartzSeminormFamily 𝕜 D E)) f) :

                      Create a continuous semilinear map between Schwartz spaces.

                      For an example of using this definition, see fderivCLM.

                      Equations
                      Instances For
                        def SchwartzMap.mkCLMtoNormedSpace {𝕜 : Type u_1} {𝕜' : Type u_2} {D : Type u_3} {E : Type u_4} {G : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedField 𝕜] [NormedField 𝕜'] [NormedAddCommGroup D] [NormedSpace D] [NormedSpace 𝕜 E] [SMulCommClass 𝕜 E] [NormedAddCommGroup G] [NormedSpace 𝕜' G] {σ : 𝕜 →+* 𝕜'} [RingHomIsometric σ] (A : SchwartzMap D EG) (hadd : ∀ (f g : SchwartzMap D E), A (f + g) = A f + A g) (hsmul : ∀ (a : 𝕜) (f : SchwartzMap D E), A (a f) = σ a A f) (hbound : ∃ (s : Finset ( × )) (C : ), 0 C ∀ (f : SchwartzMap D E), A f C * (s.sup (schwartzSeminormFamily 𝕜 D E)) f) :

                        Define a continuous semilinear map from Schwartz space to a normed space.

                        Equations
                        Instances For
                          def SchwartzMap.evalCLM {𝕜 : Type u_1} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (m : E) :

                          The map applying a vector to Hom-valued Schwartz function as a continuous linear map.

                          Equations
                          Instances For
                            def SchwartzMap.bilinLeftCLM {𝕜 : Type u_1} {D : Type u_3} {E : Type u_4} {F : Type u_5} {G : Type u_6} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [NontriviallyNormedField 𝕜] [NormedAlgebra 𝕜] [NormedAddCommGroup D] [NormedSpace D] [NormedAddCommGroup G] [NormedSpace G] [NormedSpace 𝕜 E] [NormedSpace 𝕜 F] [NormedSpace 𝕜 G] (B : E →L[𝕜] F →L[𝕜] G) {g : DF} (hg : Function.HasTemperateGrowth g) :

                            The map f ↦ (x ↦ B (f x) (g x)) as a continuous 𝕜-linear map on Schwartz space, where B is a continuous 𝕜-linear map and g is a function of temperate growth.

                            Equations
                            Instances For
                              def SchwartzMap.compCLM (𝕜 : Type u_1) {D : Type u_3} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedAddCommGroup D] [NormedSpace D] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] {g : DE} (hg : Function.HasTemperateGrowth g) (hg_upper : ∃ (k : ) (C : ), ∀ (x : D), x C * (1 + g x) ^ k) :

                              Composition with a function on the right is a continuous linear map on Schwartz space provided that the function is temperate and growths polynomially near infinity.

                              Equations
                              Instances For
                                @[simp]
                                theorem SchwartzMap.compCLM_apply (𝕜 : Type u_1) {D : Type u_3} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedAddCommGroup D] [NormedSpace D] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] {g : DE} (hg : Function.HasTemperateGrowth g) (hg_upper : ∃ (k : ) (C : ), ∀ (x : D), x C * (1 + g x) ^ k) (f : SchwartzMap E F) :
                                ((compCLM 𝕜 hg hg_upper) f) = f g
                                def SchwartzMap.compCLMOfAntilipschitz (𝕜 : Type u_1) {D : Type u_3} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedAddCommGroup D] [NormedSpace D] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] {K : NNReal} {g : DE} (hg : Function.HasTemperateGrowth g) (h'g : AntilipschitzWith K g) :

                                Composition with a function on the right is a continuous linear map on Schwartz space provided that the function is temperate and antilipschitz.

                                Equations
                                Instances For
                                  @[simp]
                                  theorem SchwartzMap.compCLMOfAntilipschitz_apply (𝕜 : Type u_1) {D : Type u_3} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedAddCommGroup D] [NormedSpace D] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] {K : NNReal} {g : DE} (hg : Function.HasTemperateGrowth g) (h'g : AntilipschitzWith K g) (f : SchwartzMap E F) :
                                  ((compCLMOfAntilipschitz 𝕜 hg h'g) f) = f g

                                  Composition with a continuous linear equiv on the right is a continuous linear map on Schwartz space.

                                  Equations
                                  Instances For
                                    @[simp]
                                    theorem SchwartzMap.compCLMOfContinuousLinearEquiv_apply (𝕜 : Type u_1) {D : Type u_3} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedAddCommGroup D] [NormedSpace D] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (g : D ≃L[] E) (f : SchwartzMap E F) :
                                    ((compCLMOfContinuousLinearEquiv 𝕜 g) f) = f g

                                    Derivatives of Schwartz functions #

                                    def SchwartzMap.fderivCLM (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] :

                                    The Fréchet derivative on Schwartz space as a continuous 𝕜-linear map.

                                    Equations
                                    Instances For
                                      @[simp]
                                      theorem SchwartzMap.fderivCLM_apply (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (f : SchwartzMap E F) (x : E) :
                                      ((fderivCLM 𝕜) f) x = fderiv (⇑f) x

                                      The 1-dimensional derivative on Schwartz space as a continuous 𝕜-linear map.

                                      Equations
                                      Instances For
                                        @[simp]
                                        theorem SchwartzMap.derivCLM_apply (𝕜 : Type u_1) {F : Type u_5} [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (f : SchwartzMap F) (x : ) :
                                        ((derivCLM 𝕜) f) x = deriv (⇑f) x
                                        def SchwartzMap.pderivCLM (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (m : E) :

                                        The partial derivative (or directional derivative) in the direction m : E as a continuous linear map on Schwartz space.

                                        Equations
                                        Instances For
                                          @[simp]
                                          theorem SchwartzMap.pderivCLM_apply (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (m : E) (f : SchwartzMap E F) (x : E) :
                                          ((pderivCLM 𝕜 m) f) x = (fderiv (⇑f) x) m
                                          theorem SchwartzMap.pderivCLM_eq_lineDeriv (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (m : E) (f : SchwartzMap E F) (x : E) :
                                          ((pderivCLM 𝕜 m) f) x = lineDeriv (⇑f) x m
                                          def SchwartzMap.iteratedPDeriv (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] {n : } :
                                          (Fin nE)SchwartzMap E F →L[𝕜] SchwartzMap E F

                                          The iterated partial derivative (or directional derivative) as a continuous linear map on Schwartz space.

                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For
                                            @[simp]
                                            theorem SchwartzMap.iteratedPDeriv_zero (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (m : Fin 0E) (f : SchwartzMap E F) :
                                            (iteratedPDeriv 𝕜 m) f = f
                                            @[simp]
                                            theorem SchwartzMap.iteratedPDeriv_one (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (m : Fin 1E) (f : SchwartzMap E F) :
                                            (iteratedPDeriv 𝕜 m) f = (pderivCLM 𝕜 (m 0)) f
                                            theorem SchwartzMap.iteratedPDeriv_succ_left (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] {n : } (m : Fin (n + 1)E) (f : SchwartzMap E F) :
                                            (iteratedPDeriv 𝕜 m) f = (pderivCLM 𝕜 (m 0)) ((iteratedPDeriv 𝕜 (Fin.tail m)) f)
                                            theorem SchwartzMap.iteratedPDeriv_succ_right (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] {n : } (m : Fin (n + 1)E) (f : SchwartzMap E F) :
                                            (iteratedPDeriv 𝕜 m) f = (iteratedPDeriv 𝕜 (Fin.init m)) ((pderivCLM 𝕜 (m (Fin.last n))) f)
                                            theorem SchwartzMap.iteratedPDeriv_eq_iteratedFDeriv (𝕜 : Type u_1) {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] {n : } {m : Fin nE} {f : SchwartzMap E F} {x : E} :
                                            ((iteratedPDeriv 𝕜 m) f) x = (iteratedFDeriv n (⇑f) x) m

                                            Integration #

                                            theorem SchwartzMap.integral_pow_mul_iteratedFDeriv_le (𝕜 : Type u_1) {D : Type u_3} {V : Type u_7} [RCLike 𝕜] [NormedAddCommGroup D] [NormedSpace D] [NormedAddCommGroup V] [NormedSpace V] [NormedSpace 𝕜 V] [MeasurableSpace D] (μ : MeasureTheory.Measure D) [hμ : μ.HasTemperateGrowth] (f : SchwartzMap D V) (k n : ) :
                                            (x : D), x ^ k * iteratedFDeriv n (⇑f) x μ (2 ^ μ.integrablePower * (x : D), (1 + x) ^ (-μ.integrablePower) μ) * ((SchwartzMap.seminorm 𝕜 0 n) f + (SchwartzMap.seminorm 𝕜 (k + μ.integrablePower) n) f)

                                            The integral as a continuous linear map from Schwartz space to the codomain.

                                            Equations
                                            Instances For
                                              @[simp]
                                              theorem SchwartzMap.integralCLM_apply (𝕜 : Type u_1) {D : Type u_3} {V : Type u_7} [RCLike 𝕜] [NormedAddCommGroup D] [NormedSpace D] [NormedAddCommGroup V] [NormedSpace V] [NormedSpace 𝕜 V] [MeasurableSpace D] {μ : MeasureTheory.Measure D} [hμ : μ.HasTemperateGrowth] [BorelSpace D] [SecondCountableTopology D] (f : SchwartzMap D V) :
                                              (integralCLM 𝕜 μ) f = (x : D), f x μ

                                              Inclusion into the space of bounded continuous functions #

                                              Schwartz functions as continuous functions

                                              Equations
                                              Instances For

                                                The inclusion map from Schwartz functions to bounded continuous functions as a continuous linear map.

                                                Equations
                                                Instances For
                                                  @[simp]
                                                  theorem SchwartzMap.toBoundedContinuousFunctionCLM_apply (𝕜 : Type u_1) (E : Type u_4) (F : Type u_5) [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (f : SchwartzMap E F) (x : E) :
                                                  ((toBoundedContinuousFunctionCLM 𝕜 E F) f) x = f x
                                                  def SchwartzMap.delta (𝕜 : Type u_1) {E : Type u_4} (F : Type u_5) [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (x : E) :

                                                  The Dirac delta distribution

                                                  Equations
                                                  Instances For
                                                    @[simp]
                                                    theorem SchwartzMap.delta_apply (𝕜 : Type u_1) {E : Type u_4} (F : Type u_5) [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (x₀ : E) (f : SchwartzMap E F) :
                                                    (delta 𝕜 F x₀) f = f x₀
                                                    @[simp]

                                                    Integrating against the Dirac measure is equal to the delta distribution.

                                                    Schwartz functions as continuous functions vanishing at infinity.

                                                    Equations
                                                    • f.toZeroAtInfty = { toFun := f, continuous_toFun := , zero_at_infty' := }
                                                    Instances For

                                                      The inclusion map from Schwartz functions to continuous functions vanishing at infinity as a continuous linear map.

                                                      Equations
                                                      Instances For
                                                        @[simp]
                                                        theorem SchwartzMap.toZeroAtInftyCLM_apply (𝕜 : Type u_1) (E : Type u_4) (F : Type u_5) [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [ProperSpace E] [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (f : SchwartzMap E F) (x : E) :
                                                        ((toZeroAtInftyCLM 𝕜 E F) f) x = f x

                                                        Inclusion into L^p space #

                                                        theorem SchwartzMap.eLpNorm_le_seminorm (𝕜 : Type u_1) {E : Type u_4} (F : Type u_5) [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [MeasurableSpace E] [OpensMeasurableSpace E] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] (p : ENNReal) (μ : MeasureTheory.Measure E := by volume_tac) [hμ : MeasureTheory.Measure.HasTemperateGrowth μ] :
                                                        ∃ (k : ) (C : NNReal), ∀ (f : SchwartzMap E F), MeasureTheory.eLpNorm (⇑f) p μ C * ENNReal.ofReal (((Finset.Iic (k, 0)).sup (schwartzSeminormFamily 𝕜 E F)) f)

                                                        The L^p norm of a Schwartz function is controlled by a finite family of Schwartz seminorms.

                                                        The maximum index k and the constant C depend on p and μ.

                                                        The L^p norm of a Schwartz function is finite.

                                                        Schwartz functions are in L^∞; does not require hμ.HasTemperateGrowth.

                                                        @[deprecated SchwartzMap.memLp_top (since := "2025-02-21")]

                                                        Alias of SchwartzMap.memLp_top.


                                                        Schwartz functions are in L^∞; does not require hμ.HasTemperateGrowth.

                                                        Schwartz functions are in L^p for any p.

                                                        @[deprecated SchwartzMap.memLp (since := "2025-02-21")]

                                                        Alias of SchwartzMap.memLp.


                                                        Schwartz functions are in L^p for any p.

                                                        Map a Schwartz function to an Lp function for any p.

                                                        Equations
                                                        Instances For

                                                          Continuous linear map from Schwartz functions to L^p.

                                                          Equations
                                                          Instances For
                                                            @[simp]
                                                            theorem SchwartzMap.toLpCLM_apply {𝕜 : Type u_1} {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [NormedAddCommGroup F] [NormedSpace F] [MeasurableSpace E] [OpensMeasurableSpace E] [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass 𝕜 F] [SecondCountableTopologyEither E F] {p : ENNReal} [Fact (1 p)] {μ : MeasureTheory.Measure E} [hμ : μ.HasTemperateGrowth] {f : SchwartzMap E F} :
                                                            (toLpCLM 𝕜 F p μ) f = f.toLp p μ