Documentation

Mathlib.Analysis.NormedSpace.Basic

Normed spaces #

In this file we define (semi)normed spaces and algebras. We also prove some theorems about these definitions.

class NormedSpace (α : Type u_5) (β : Type u_6) [NormedField α] [SeminormedAddCommGroup β] extends Module :
Type (max u_5 u_6)
  • smul : αββ
  • one_smul : ∀ (b : β), 1 b = b
  • mul_smul : ∀ (x y : α) (b : β), (x * y) b = x y b
  • smul_zero : ∀ (a : α), a 0 = 0
  • smul_add : ∀ (a : α) (x y : β), a (x + y) = a x + a y
  • add_smul : ∀ (r s_1 : α) (x : β), (r + s_1) x = r x + s_1 x
  • zero_smul : ∀ (x : β), 0 x = 0
  • norm_smul_le : ∀ (a : α) (b : β), a b a * b

    A normed space over a normed field is a vector space endowed with a norm which satisfies the equality ‖c • x‖ = ‖c‖ ‖x‖. We require only ‖c • x‖ ≤ ‖c‖ ‖x‖ in the definition, then prove ‖c • x‖ = ‖c‖ ‖x‖ in norm_smul.

    Note that since this requires SeminormedAddCommGroup and not NormedAddCommGroup, this typeclass can be used for "semi normed spaces" too, just as Module can be used for "semi modules".

A normed space over a normed field is a vector space endowed with a norm which satisfies the equality ‖c • x‖ = ‖c‖ ‖x‖. We require only ‖c • x‖ ≤ ‖c‖ ‖x‖ in the definition, then prove ‖c • x‖ = ‖c‖ ‖x‖ in norm_smul.

Note that since this requires SeminormedAddCommGroup and not NormedAddCommGroup, this typeclass can be used for "semi normed spaces" too, just as Module can be used for "semi modules".

Instances
    instance NormedField.toNormedSpace {α : Type u_1} [NormedField α] :
    Equations
    theorem norm_zsmul {β : Type u_2} [SeminormedAddCommGroup β] (α : Type u_5) [NormedField α] [NormedSpace α β] (n : ) (x : β) :
    @[simp]
    theorem abs_norm {β : Type u_2} [SeminormedAddCommGroup β] (z : β) :
    theorem norm_smul_of_nonneg {β : Type u_2} [SeminormedAddCommGroup β] [NormedSpace β] {t : } (ht : 0 t) (x : β) :
    theorem eventually_nhds_norm_smul_sub_lt {α : Type u_1} [NormedField α] {E : Type u_5} [SeminormedAddCommGroup E] [NormedSpace α E] (c : α) (x : E) {ε : } (h : 0 < ε) :
    ∀ᶠ (y : E) in nhds x, c (y - x) < ε
    theorem Filter.Tendsto.zero_smul_isBoundedUnder_le {α : Type u_1} {ι : Type u_4} [NormedField α] {E : Type u_5} [SeminormedAddCommGroup E] [NormedSpace α E] {f : ια} {g : ιE} {l : Filter ι} (hf : Filter.Tendsto f l (nhds 0)) (hg : Filter.IsBoundedUnder (fun x x_1 => x x_1) l (norm g)) :
    Filter.Tendsto (fun x => f x g x) l (nhds 0)
    theorem Filter.IsBoundedUnder.smul_tendsto_zero {α : Type u_1} {ι : Type u_4} [NormedField α] {E : Type u_5} [SeminormedAddCommGroup E] [NormedSpace α E] {f : ια} {g : ιE} {l : Filter ι} (hf : Filter.IsBoundedUnder (fun x x_1 => x x_1) l (norm f)) (hg : Filter.Tendsto g l (nhds 0)) :
    Filter.Tendsto (fun x => f x g x) l (nhds 0)
    theorem closure_ball {E : Type u_5} [SeminormedAddCommGroup E] [NormedSpace E] (x : E) {r : } (hr : r 0) :
    theorem frontier_ball {E : Type u_5} [SeminormedAddCommGroup E] [NormedSpace E] (x : E) {r : } (hr : r 0) :
    theorem interior_closedBall {E : Type u_5} [SeminormedAddCommGroup E] [NormedSpace E] (x : E) {r : } (hr : r 0) :
    theorem interior_sphere {E : Type u_5} [SeminormedAddCommGroup E] [NormedSpace E] (x : E) {r : } (hr : r 0) :
    theorem frontier_sphere {E : Type u_5} [SeminormedAddCommGroup E] [NormedSpace E] (x : E) {r : } (hr : r 0) :
    Equations
    instance Prod.normedSpace {α : Type u_1} [NormedField α] {E : Type u_5} [SeminormedAddCommGroup E] [NormedSpace α E] {F : Type u_6} [SeminormedAddCommGroup F] [NormedSpace α F] :
    NormedSpace α (E × F)

    The product of two normed spaces is a normed space, with the sup norm.

    Equations
    instance Pi.normedSpace {α : Type u_1} {ι : Type u_4} [NormedField α] {E : ιType u_7} [Fintype ι] [(i : ι) → SeminormedAddCommGroup (E i)] [(i : ι) → NormedSpace α (E i)] :
    NormedSpace α ((i : ι) → E i)

    The product of finitely many normed spaces is a normed space, with the sup norm.

    Equations
    Equations
    • One or more equations did not get rendered due to their size.
    instance Submodule.normedSpace {𝕜 : Type u_7} {R : Type u_8} [SMul 𝕜 R] [NormedField 𝕜] [Ring R] {E : Type u_9} [SeminormedAddCommGroup E] [NormedSpace 𝕜 E] [Module R E] [IsScalarTower 𝕜 R E] (s : Submodule R E) :
    NormedSpace 𝕜 { x // x s }

    A subspace of a normed space is also a normed space, with the restriction of the norm.

    Equations
    @[reducible]
    def NormedSpace.induced {F : Type u_5} (α : Type u_6) (β : Type u_7) (γ : Type u_8) [NormedField α] [AddCommGroup β] [Module α β] [SeminormedAddCommGroup γ] [NormedSpace α γ] [LinearMapClass F α β γ] (f : F) :

    A linear map from a Module to a NormedSpace induces a NormedSpace structure on the domain, using the SeminormedAddCommGroup.induced norm.

    See note [reducible non-instances]

    Equations
    Instances For
      instance NormedSpace.toModule' {α : Type u_1} [NormedField α] {F : Type u_6} [NormedAddCommGroup F] [NormedSpace α F] :
      Module α F

      While this may appear identical to NormedSpace.toModule, it contains an implicit argument involving NormedAddCommGroup.toSeminormedAddCommGroup that typeclass inference has trouble inferring.

      Specifically, the following instance cannot be found without this NormedSpace.toModule':

      example
        (𝕜 ι : Type*) (E : ι → Type*)
        [NormedField 𝕜] [Π i, NormedAddCommGroup (E i)] [Π i, NormedSpace 𝕜 (E i)] :
        Π i, Module 𝕜 (E i) := by infer_instance
      

      This Zulip thread gives some more context.

      Equations
      • NormedSpace.toModule' = NormedSpace.toModule
      theorem exists_norm_eq (E : Type u_5) [NormedAddCommGroup E] [NormedSpace E] [Nontrivial E] {c : } (hc : 0 c) :
      x, x = c
      @[simp]
      @[simp]
      theorem range_nnnorm (E : Type u_5) [NormedAddCommGroup E] [NormedSpace E] [Nontrivial E] :
      Set.range nnnorm = Set.univ

      If E is a nontrivial topological module over , then E has no isolated points. This is a particular case of Module.punctured_nhds_neBot.

      Equations
      @[simp]
      theorem interior_sphere' {E : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [Nontrivial E] (x : E) (r : ) :
      @[simp]
      theorem NormedSpace.exists_lt_norm (𝕜 : Type u_5) (E : Type u_6) [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [Nontrivial E] (c : ) :
      x, c < x

      If E is a nontrivial normed space over a nontrivially normed field 𝕜, then E is unbounded: for any c : ℝ, there exists a vector x : E with norm strictly greater than c.

      A normed vector space over a nontrivially normed field is a noncompact space. This cannot be an instance because in order to apply it, Lean would have to search for NormedSpace 𝕜 E with unknown 𝕜. We register this as an instance in two cases: 𝕜 = E and 𝕜 = ℝ.

      class NormedAlgebra (𝕜 : Type u_5) (𝕜' : Type u_6) [NormedField 𝕜] [SeminormedRing 𝕜'] extends Algebra :
      Type (max u_5 u_6)
      • smul : 𝕜𝕜'𝕜'
      • toFun : 𝕜𝕜'
      • map_one' : OneHom.toFun (Algebra.toRingHom) 1 = 1
      • map_mul' : ∀ (x y : 𝕜), OneHom.toFun (Algebra.toRingHom) (x * y) = OneHom.toFun (Algebra.toRingHom) x * OneHom.toFun (Algebra.toRingHom) y
      • map_zero' : OneHom.toFun (Algebra.toRingHom) 0 = 0
      • map_add' : ∀ (x y : 𝕜), OneHom.toFun (Algebra.toRingHom) (x + y) = OneHom.toFun (Algebra.toRingHom) x + OneHom.toFun (Algebra.toRingHom) y
      • commutes' : ∀ (r : 𝕜) (x : (fun x => 𝕜') r), Algebra.toRingHom r * x = x * Algebra.toRingHom r
      • smul_def' : ∀ (r : 𝕜) (x : (fun x => 𝕜') r), r x = Algebra.toRingHom r * x
      • norm_smul_le : ∀ (r : 𝕜) (x : 𝕜'), r x r * x

        A normed algebra 𝕜' over 𝕜 is normed module that is also an algebra.

        See the implementation notes for Algebra for a discussion about non-unital algebras. Following the strategy there, a non-unital normed algebra can be written as:

        variables [NormedField 𝕜] [NonunitalSeminormedRing 𝕜']
        variables [NormedModule 𝕜 𝕜'] [SMulCommClass 𝕜 𝕜' 𝕜'] [IsScalarTower 𝕜 𝕜' 𝕜']
        

      A normed algebra 𝕜' over 𝕜 is normed module that is also an algebra.

      See the implementation notes for Algebra for a discussion about non-unital algebras. Following the strategy there, a non-unital normed algebra can be written as:

      variables [NormedField 𝕜] [NonunitalSeminormedRing 𝕜']
      variables [NormedModule 𝕜 𝕜'] [SMulCommClass 𝕜 𝕜' 𝕜'] [IsScalarTower 𝕜 𝕜' 𝕜']
      
      Instances
        instance NormedAlgebra.toNormedSpace {𝕜 : Type u_5} (𝕜' : Type u_6) [NormedField 𝕜] [SeminormedRing 𝕜'] [NormedAlgebra 𝕜 𝕜'] :
        NormedSpace 𝕜 𝕜'
        Equations
        instance NormedAlgebra.toNormedSpace' {𝕜 : Type u_5} [NormedField 𝕜] {𝕜' : Type u_7} [NormedRing 𝕜'] [NormedAlgebra 𝕜 𝕜'] :
        NormedSpace 𝕜 𝕜'

        While this may appear identical to NormedAlgebra.toNormedSpace, it contains an implicit argument involving NormedRing.toSeminormedRing that typeclass inference has trouble inferring.

        Specifically, the following instance cannot be found without this NormedSpace.toModule':

        example
          (𝕜 ι : Type*) (E : ι → Type*)
          [NormedField 𝕜] [Π i, NormedRing (E i)] [Π i, NormedAlgebra 𝕜 (E i)] :
          Π i, Module 𝕜 (E i) := by infer_instance
        

        See NormedSpace.toModule' for a similar situation.

        Equations
        • NormedAlgebra.toNormedSpace' = inferInstance
        theorem norm_algebraMap {𝕜 : Type u_5} (𝕜' : Type u_6) [NormedField 𝕜] [SeminormedRing 𝕜'] [NormedAlgebra 𝕜 𝕜'] (x : 𝕜) :
        ↑(algebraMap 𝕜 𝕜') x = x * 1
        theorem nnnorm_algebraMap {𝕜 : Type u_5} (𝕜' : Type u_6) [NormedField 𝕜] [SeminormedRing 𝕜'] [NormedAlgebra 𝕜 𝕜'] (x : 𝕜) :
        @[simp]
        theorem norm_algebraMap' {𝕜 : Type u_5} (𝕜' : Type u_6) [NormedField 𝕜] [SeminormedRing 𝕜'] [NormedAlgebra 𝕜 𝕜'] [NormOneClass 𝕜'] (x : 𝕜) :
        ↑(algebraMap 𝕜 𝕜') x = x
        @[simp]
        theorem nnnorm_algebraMap' {𝕜 : Type u_5} (𝕜' : Type u_6) [NormedField 𝕜] [SeminormedRing 𝕜'] [NormedAlgebra 𝕜 𝕜'] [NormOneClass 𝕜'] (x : 𝕜) :
        ↑(algebraMap 𝕜 𝕜') x‖₊ = x‖₊
        @[simp]
        theorem norm_algebraMap_nNReal (𝕜' : Type u_6) [SeminormedRing 𝕜'] [NormOneClass 𝕜'] [NormedAlgebra 𝕜'] (x : NNReal) :
        ↑(algebraMap NNReal 𝕜') x = x
        @[simp]
        theorem nnnorm_algebraMap_nNReal (𝕜' : Type u_6) [SeminormedRing 𝕜'] [NormOneClass 𝕜'] [NormedAlgebra 𝕜'] (x : NNReal) :
        theorem algebraMap_isometry (𝕜 : Type u_5) (𝕜' : Type u_6) [NormedField 𝕜] [SeminormedRing 𝕜'] [NormedAlgebra 𝕜 𝕜'] [NormOneClass 𝕜'] :
        Isometry ↑(algebraMap 𝕜 𝕜')

        In a normed algebra, the inclusion of the base field in the extended field is an isometry.

        instance NormedAlgebra.id (𝕜 : Type u_5) [NormedField 𝕜] :
        NormedAlgebra 𝕜 𝕜
        Equations
        instance normedAlgebraRat {𝕜 : Type u_7} [NormedDivisionRing 𝕜] [CharZero 𝕜] [NormedAlgebra 𝕜] :

        Any normed characteristic-zero division ring that is a normed algebra over the reals is also a normed algebra over the rationals.

        Phrased another way, if 𝕜 is a normed algebra over the reals, then AlgebraRat respects that norm.

        Equations
        instance PUnit.normedAlgebra (𝕜 : Type u_5) [NormedField 𝕜] :
        Equations
        instance instNormedAlgebraULiftSeminormedRing (𝕜 : Type u_5) (𝕜' : Type u_6) [NormedField 𝕜] [SeminormedRing 𝕜'] [NormedAlgebra 𝕜 𝕜'] :
        Equations
        instance Prod.normedAlgebra (𝕜 : Type u_5) [NormedField 𝕜] {E : Type u_7} {F : Type u_8} [SeminormedRing E] [SeminormedRing F] [NormedAlgebra 𝕜 E] [NormedAlgebra 𝕜 F] :
        NormedAlgebra 𝕜 (E × F)

        The product of two normed algebras is a normed algebra, with the sup norm.

        Equations
        instance Pi.normedAlgebra {ι : Type u_4} (𝕜 : Type u_5) [NormedField 𝕜] {E : ιType u_7} [Fintype ι] [(i : ι) → SeminormedRing (E i)] [(i : ι) → NormedAlgebra 𝕜 (E i)] :
        NormedAlgebra 𝕜 ((i : ι) → E i)

        The product of finitely many normed algebras is a normed algebra, with the sup norm.

        Equations
        instance MulOpposite.normedAlgebra (𝕜 : Type u_5) [NormedField 𝕜] {E : Type u_8} [SeminormedRing E] [NormedAlgebra 𝕜 E] :
        Equations
        @[reducible]
        def NormedAlgebra.induced {F : Type u_5} (α : Type u_6) (β : Type u_7) (γ : Type u_8) [NormedField α] [Ring β] [Algebra α β] [SeminormedRing γ] [NormedAlgebra α γ] [NonUnitalAlgHomClass F α β γ] (f : F) :

        A non-unital algebra homomorphism from an Algebra to a NormedAlgebra induces a NormedAlgebra structure on the domain, using the SeminormedRing.induced norm.

        See note [reducible non-instances]

        Equations
        Instances For
          instance Subalgebra.toNormedAlgebra {𝕜 : Type u_5} {A : Type u_6} [SeminormedRing A] [NormedField 𝕜] [NormedAlgebra 𝕜 A] (S : Subalgebra 𝕜 A) :
          NormedAlgebra 𝕜 { x // x S }
          Equations
          instance instSeminormedAddCommGroupRestrictScalars {𝕜 : Type u_8} {𝕜' : Type u_9} {E : Type u_10} [I : SeminormedAddCommGroup E] :
          Equations
          • instSeminormedAddCommGroupRestrictScalars = I
          instance instNormedAddCommGroupRestrictScalars {𝕜 : Type u_8} {𝕜' : Type u_9} {E : Type u_10} [I : NormedAddCommGroup E] :
          Equations
          • instNormedAddCommGroupRestrictScalars = I
          instance RestrictScalars.normedSpace (𝕜 : Type u_5) (𝕜' : Type u_6) [NormedField 𝕜] [NormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] (E : Type u_7) [SeminormedAddCommGroup E] [NormedSpace 𝕜' E] :
          NormedSpace 𝕜 (RestrictScalars 𝕜 𝕜' E)

          If E is a normed space over 𝕜' and 𝕜 is a normed algebra over 𝕜', then RestrictScalars.module is additionally a NormedSpace.

          Equations
          def Module.RestrictScalars.normedSpaceOrig {𝕜 : Type u_8} {𝕜' : Type u_9} {E : Type u_10} [NormedField 𝕜'] [SeminormedAddCommGroup E] [I : NormedSpace 𝕜' E] :
          NormedSpace 𝕜' (RestrictScalars 𝕜 𝕜' E)

          The action of the original normed_field on RestrictScalars 𝕜 𝕜' E. This is not an instance as it would be contrary to the purpose of RestrictScalars.

          Equations
          • Module.RestrictScalars.normedSpaceOrig = I
          Instances For
            def NormedSpace.restrictScalars (𝕜 : Type u_5) (𝕜' : Type u_6) [NormedField 𝕜] [NormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] (E : Type u_7) [SeminormedAddCommGroup E] [NormedSpace 𝕜' E] :

            Warning: This declaration should be used judiciously. Please consider using IsScalarTower and/or RestrictScalars 𝕜 𝕜' E instead.

            This definition allows the RestrictScalars.normedSpace instance to be put directly on E rather on RestrictScalars 𝕜 𝕜' E. This would be a very bad instance; both because 𝕜' cannot be inferred, and because it is likely to create instance diamonds.

            Equations
            Instances For