Documentation

Mathlib.Analysis.InnerProductSpace.Laplacian

The Laplacian #

This file defines the Laplacian for functions f : E → F on real, finite-dimensional, inner product spaces E. In essence, we define the Laplacian of f as the second derivative, applied to the canonical covariant tensor of E, as defined and discussed in Mathlib.Analysis.InnerProductSpace.CanonicalTensor.

We show that the Laplacian is -linear on continuously differentiable functions, and establish the standard formula for computing the Laplacian in terms of orthonormal bases of E.

Supporting API #

The definition of the Laplacian of a function f : E → F involves the notion of the second derivative, which can be seen as a continous multilinear map ContinuousMultilinearMap 𝕜 (fun (i : Fin 2) ↦ E) F, a bilinear map E →ₗ[𝕜] E →ₗ[𝕜] F, or a linear map on tensors E ⊗[𝕜] E →ₗ[𝕜] F. This section provides convenience API to convert between these notions.

noncomputable def bilinearIteratedFDerivWithinTwo (𝕜 : Type u_1) [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type u_3} [NormedAddCommGroup F] [NormedSpace 𝕜 F] (f : EF) (s : Set E) :
EE →ₗ[𝕜] E →ₗ[𝕜] F

Convenience reformulation of the second iterated derivative, as a map from E to bilinear maps `E →ₗ[ℝ] E →ₗ[ℝ] ℝ

Equations
Instances For
    noncomputable def bilinearIteratedFDerivTwo (𝕜 : Type u_1) [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type u_3} [NormedAddCommGroup F] [NormedSpace 𝕜 F] (f : EF) :
    EE →ₗ[𝕜] E →ₗ[𝕜] F

    Convenience reformulation of the second iterated derivative, as a map from E to bilinear maps `E →ₗ[ℝ] E →ₗ[ℝ] ℝ

    Equations
    Instances For
      theorem bilinearIteratedFDerivWithinTwo_eq_iteratedFDeriv {𝕜 : Type u_1} [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type u_3} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {e : E} {s : Set E} (f : EF) (hs : UniqueDiffOn 𝕜 s) (he : e s) (e₁ e₂ : E) :
      ((bilinearIteratedFDerivWithinTwo 𝕜 f s e) e₁) e₂ = (iteratedFDerivWithin 𝕜 2 f s e) ![e₁, e₂]

      Expression of bilinearIteratedFDerivWithinTwo in terms of iteratedFDerivWithin.

      theorem bilinearIteratedFDerivTwo_eq_iteratedFDeriv {𝕜 : Type u_1} [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type u_3} [NormedAddCommGroup F] [NormedSpace 𝕜 F] (f : EF) (e e₁ e₂ : E) :
      ((bilinearIteratedFDerivTwo 𝕜 f e) e₁) e₂ = (iteratedFDeriv 𝕜 2 f e) ![e₁, e₂]

      Expression of bilinearIteratedFDerivTwo in terms of iteratedFDeriv.

      noncomputable def tensorIteratedFDerivWithinTwo (𝕜 : Type u_1) [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type u_3} [NormedAddCommGroup F] [NormedSpace 𝕜 F] (f : EF) (s : Set E) :
      ETensorProduct 𝕜 E E →ₗ[𝕜] F

      Convenience reformulation of the second iterated derivative, as a map from E to linear maps E ⊗[𝕜] E →ₗ[𝕜] F.

      Equations
      Instances For
        noncomputable def tensorIteratedFDerivTwo (𝕜 : Type u_1) [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type u_3} [NormedAddCommGroup F] [NormedSpace 𝕜 F] (f : EF) :
        ETensorProduct 𝕜 E E →ₗ[𝕜] F

        Convenience reformulation of the second iterated derivative, as a map from E to linear maps E ⊗[𝕜] E →ₗ[𝕜] F.

        Equations
        Instances For
          theorem tensorIteratedFDerivWithinTwo_eq_iteratedFDerivWithin {𝕜 : Type u_1} [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type u_3} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {e : E} {s : Set E} (f : EF) (hs : UniqueDiffOn 𝕜 s) (he : e s) (e₁ e₂ : E) :
          (tensorIteratedFDerivWithinTwo 𝕜 f s e) (e₁ ⊗ₜ[𝕜] e₂) = (iteratedFDerivWithin 𝕜 2 f s e) ![e₁, e₂]

          Expression of tensorIteratedFDerivTwo in terms of iteratedFDerivWithin.

          theorem tensorIteratedFDerivTwo_eq_iteratedFDeriv {𝕜 : Type u_1} [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type u_3} [NormedAddCommGroup F] [NormedSpace 𝕜 F] (f : EF) (e e₁ e₂ : E) :
          (tensorIteratedFDerivTwo 𝕜 f e) (e₁ ⊗ₜ[𝕜] e₂) = (iteratedFDeriv 𝕜 2 f e) ![e₁, e₂]

          Expression of tensorIteratedFDerivTwo in terms of iteratedFDeriv.

          Definition of the Laplacian #

          noncomputable def InnerProductSpace.laplacianWithin {E : Type u_1} [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E] {F : Type u_2} [NormedAddCommGroup F] [NormedSpace F] (f : EF) (s : Set E) :
          EF

          Laplacian for functions on real inner product spaces, with respect to a set s. Use open InnerProductSpace to access the notation Δ[s] for InnerProductSpace.LaplacianWithin.

          Equations
          Instances For

            Laplacian for functions on real inner product spaces, with respect to a set s. Use open InnerProductSpace to access the notation Δ[s] for InnerProductSpace.LaplacianWithin.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              noncomputable def InnerProductSpace.laplacian {E : Type u_1} [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E] {F : Type u_2} [NormedAddCommGroup F] [NormedSpace F] (f : EF) :
              EF

              Laplacian for functions on real inner product spaces. Use open InnerProductSpace to access the notation Δ for InnerProductSpace.Laplacian.

              Equations
              Instances For

                Laplacian for functions on real inner product spaces. Use open InnerProductSpace to access the notation Δ for InnerProductSpace.Laplacian.

                Equations
                Instances For
                  @[simp]

                  The Laplacian equals the Laplacian with respect to Set.univ.

                  Computation of Δ in Terms of Orthonormal Bases #

                  theorem InnerProductSpace.laplacianWithin_eq_iteratedFDerivWithin_orthonormalBasis {E : Type u_1} [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E] {F : Type u_2} [NormedAddCommGroup F] [NormedSpace F] (f : EF) {s : Set E} {ι : Type u_4} [Fintype ι] {e : E} (hs : UniqueDiffOn s) (he : e s) (v : OrthonormalBasis ι E) :
                  laplacianWithin f s e = i : ι, (iteratedFDerivWithin 2 f s e) ![v i, v i]

                  Standard formula, computing the Laplacian from any orthonormal basis.

                  theorem InnerProductSpace.laplacian_eq_iteratedFDeriv_orthonormalBasis {E : Type u_1} [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E] {F : Type u_2} [NormedAddCommGroup F] [NormedSpace F] (f : EF) {ι : Type u_4} [Fintype ι] (v : OrthonormalBasis ι E) :
                  laplacian f = fun (x : E) => i : ι, (iteratedFDeriv 2 f x) ![v i, v i]

                  Standard formula, computing the Laplacian from any orthonormal basis.

                  Standard formula, computing the Laplacian from the standard orthonormal basis of a real inner product space.

                  Standard formula, computing the Laplacian from the standard orthonormal basis of a real inner product space.

                  Special case of the standard formula for functions on , with the standard real inner product structure.

                  Special case of the standard formula for functions on , with the standard real inner product structure.

                  Congruence Lemmata for Δ #

                  If two functions agree in a neighborhood of a point, then so do their Laplacians.

                  theorem InnerProductSpace.laplacian_congr_nhds {E : Type u_1} [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E] {F : Type u_2} [NormedAddCommGroup F] [NormedSpace F] {f₁ f₂ : EF} {x : E} (h : f₁ =ᶠ[nhds x] f₂) :

                  If two functions agree in a neighborhood of a point, then so do their Laplacians.

                  ℝ-Linearity of Δ on Continuously Differentiable Functions #

                  theorem ContDiffWithinAt.laplacianWithin_add {E : Type u_1} [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E] {F : Type u_2} [NormedAddCommGroup F] [NormedSpace F] {f₁ f₂ : EF} {x : E} {s : Set E} (h₁ : ContDiffWithinAt 2 f₁ s x) (h₂ : ContDiffWithinAt 2 f₂ s x) (hs : UniqueDiffOn s) (hx : x s) :

                  The Laplacian commutes with addition.

                  theorem ContDiffAt.laplacian_add {E : Type u_1} [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E] {F : Type u_2} [NormedAddCommGroup F] [NormedSpace F] {f₁ f₂ : EF} {x : E} (h₁ : ContDiffAt 2 f₁ x) (h₂ : ContDiffAt 2 f₂ x) :

                  The Laplacian commutes with addition.

                  The Laplacian commutes with addition.

                  The Laplacian commutes with addition.

                  theorem InnerProductSpace.laplacianWithin_smul {E : Type u_1} [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E] {F : Type u_2} [NormedAddCommGroup F] [NormedSpace F] {f : EF} {x : E} {s : Set E} (v : ) (hf : ContDiffWithinAt 2 f s x) (hs : UniqueDiffOn s) (hx : x s) :

                  The Laplacian commutes with scalar multiplication.

                  theorem InnerProductSpace.laplacian_smul {E : Type u_1} [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E] {F : Type u_2} [NormedAddCommGroup F] [NormedSpace F] {f : EF} {x : E} (v : ) (hf : ContDiffAt 2 f x) :
                  laplacian (v f) x = v laplacian f x

                  The Laplacian commutes with scalar multiplication.

                  Commutativity of Δ with Linear Operators #

                  This section establishes commutativity with linear operators, showing in particular that Δ commutes with taking real and imaginary parts of complex-valued functions.

                  The Laplacian commutes with left composition by continuous linear maps.

                  The Laplacian commutes with left composition by continuous linear maps.

                  theorem InnerProductSpace.laplacianWithin_CLE_comp_left {E : Type u_1} [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E] {F : Type u_2} [NormedAddCommGroup F] [NormedSpace F] {G : Type u_3} [NormedAddCommGroup G] [NormedSpace G] {f : EF} {x : E} {s : Set E} {l : F ≃L[] G} (hs : UniqueDiffOn s) (hx : x s) :
                  laplacianWithin (l f) s x = (l laplacianWithin f s) x

                  The Laplacian commutes with left composition by continuous linear equivalences.

                  The Laplacian commutes with left composition by continuous linear equivalences.