Documentation

Mathlib.Analysis.Normed.Module.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_6) (E : Type u_7) [NormedField π•œ] [SeminormedAddCommGroup E] extends Module π•œ E :
Type (max u_6 u_7)

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 100]
    instance NormedSpace.isBoundedSMul {π•œ : Type u_1} {E : Type u_3} [NormedField π•œ] [SeminormedAddCommGroup E] [NormedSpace π•œ E] :
    IsBoundedSMul π•œ E
    instance NormedField.toNormedSpace {π•œ : Type u_1} [NormedField π•œ] :
    NormedSpace π•œ π•œ
    Equations
    instance NormedField.to_isBoundedSMul {π•œ : Type u_1} [NormedField π•œ] :
    IsBoundedSMul π•œ π•œ
    theorem norm_zsmul (π•œ : Type u_1) {E : Type u_3} [NormedField π•œ] [SeminormedAddCommGroup E] [NormedSpace π•œ E] (n : β„€) (x : E) :
    theorem eventually_nhds_norm_smul_sub_lt {π•œ : Type u_1} {E : Type u_3} [NormedField π•œ] [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} {E : Type u_3} {Ξ± : Type u_5} [NormedField π•œ] [SeminormedAddCommGroup E] [NormedSpace π•œ E] {f : Ξ± β†’ π•œ} {g : Ξ± β†’ E} {l : Filter Ξ±} (hf : Tendsto f l (nhds 0)) (hg : IsBoundedUnder (fun (x1 x2 : ℝ) => x1 ≀ x2) l (Norm.norm ∘ g)) :
    Tendsto (fun (x : Ξ±) => f x β€’ g x) l (nhds 0)
    theorem Filter.IsBoundedUnder.smul_tendsto_zero {π•œ : Type u_1} {E : Type u_3} {Ξ± : Type u_5} [NormedField π•œ] [SeminormedAddCommGroup E] [NormedSpace π•œ E] {f : Ξ± β†’ π•œ} {g : Ξ± β†’ E} {l : Filter Ξ±} (hf : IsBoundedUnder (fun (x1 x2 : ℝ) => x1 ≀ x2) l (norm ∘ f)) (hg : Tendsto g l (nhds 0)) :
    Tendsto (fun (x : Ξ±) => f x β€’ g x) l (nhds 0)
    instance ULift.normedSpace {π•œ : Type u_1} {E : Type u_3} [NormedField π•œ] [SeminormedAddCommGroup E] [NormedSpace π•œ E] :
    Equations
    instance Prod.normedSpace {π•œ : Type u_1} {E : Type u_3} {F : Type u_4} [NormedField π•œ] [SeminormedAddCommGroup E] [SeminormedAddCommGroup F] [NormedSpace π•œ E] [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} [NormedField π•œ] {ΞΉ : Type u_6} {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
    instance Submodule.normedSpace {π•œ : Type u_6} {R : Type u_7} [SMul π•œ R] [NormedField π•œ] [Ring R] {E : Type u_8} [SeminormedAddCommGroup E] [NormedSpace π•œ E] [Module R E] [IsScalarTower π•œ R E] (s : Submodule R E) :
    NormedSpace π•œ β†₯s

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

    Equations
    @[instance 75]
    instance SubmoduleClass.toNormedSpace {S : Type u_6} {π•œ : Type u_7} {R : Type u_8} {E : Type u_9} [SMul π•œ R] [NormedField π•œ] [Ring R] [SeminormedAddCommGroup E] [NormedSpace π•œ E] [Module R E] [IsScalarTower π•œ R E] [SetLike S E] [AddSubgroupClass S E] [SMulMemClass S R E] (s : S) :
    NormedSpace π•œ β†₯s
    Equations
    @[reducible, inline]
    abbrev NormedSpace.induced {F : Type u_6} (π•œ : Type u_7) (E : Type u_8) (G : Type u_9) [NormedField π•œ] [AddCommGroup E] [Module π•œ E] [SeminormedAddCommGroup G] [NormedSpace π•œ G] [FunLike F E G] [LinearMapClass F π•œ E G] (f : F) :
    NormedSpace π•œ E

    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
      theorem NormedSpace.exists_lt_norm (π•œ : Type u_1) (E : Type u_3) [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [Nontrivial E] (c : ℝ) :
      βˆƒ (x : E), 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.

      @[instance 100]
      @[instance 80]
      instance NontriviallyNormedField.infinite (π•œ : Type u_1) [NontriviallyNormedField π•œ] :
      Infinite π•œ
      theorem NormedSpace.noncompactSpace (π•œ : Type u_1) (E : Type u_3) [NormedField π•œ] [Infinite π•œ] [NormedAddCommGroup E] [Nontrivial E] [NormedSpace π•œ E] :

      A normed vector space over an infinite 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 π•œ = ℝ.

      @[instance 100]
      instance NormedField.noncompactSpace (π•œ : Type u_1) [NormedField π•œ] [Infinite π•œ] :
      class NormedAlgebra (π•œ : Type u_6) (π•œ' : Type u_7) [NormedField π•œ] [SeminormedRing π•œ'] extends Algebra π•œ π•œ' :
      Type (max u_6 u_7)

      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:

      variable [NormedField π•œ] [NonUnitalSeminormedRing π•œ']
      variable [NormedSpace π•œ π•œ'] [SMulCommClass π•œ π•œ' π•œ'] [IsScalarTower π•œ π•œ' π•œ']
      
      Instances
        @[instance 100]
        instance NormedAlgebra.toNormedSpace {π•œ : Type u_1} (π•œ' : Type u_2) [NormedField π•œ] [SeminormedRing π•œ'] [NormedAlgebra π•œ π•œ'] :
        NormedSpace π•œ π•œ'
        Equations
        theorem norm_algebraMap {π•œ : Type u_1} (π•œ' : Type u_2) [NormedField π•œ] [SeminormedRing π•œ'] [NormedAlgebra π•œ π•œ'] (x : π•œ) :
        theorem nnnorm_algebraMap {π•œ : Type u_1} (π•œ' : Type u_2) [NormedField π•œ] [SeminormedRing π•œ'] [NormedAlgebra π•œ π•œ'] (x : π•œ) :
        theorem dist_algebraMap {π•œ : Type u_1} (π•œ' : Type u_2) [NormedField π•œ] [SeminormedRing π•œ'] [NormedAlgebra π•œ π•œ'] (x y : π•œ) :
        dist ((algebraMap π•œ π•œ') x) ((algebraMap π•œ π•œ') y) = dist x y * β€–1β€–
        @[simp]
        theorem norm_algebraMap' {π•œ : Type u_1} (π•œ' : Type u_2) [NormedField π•œ] [SeminormedRing π•œ'] [NormedAlgebra π•œ π•œ'] [NormOneClass π•œ'] (x : π•œ) :
        β€–(algebraMap π•œ π•œ') xβ€– = β€–xβ€–

        This is a simpler version of norm_algebraMap when β€–1β€– = 1 in π•œ'.

        @[simp]
        theorem nnnorm_algebraMap' {π•œ : Type u_1} (π•œ' : Type u_2) [NormedField π•œ] [SeminormedRing π•œ'] [NormedAlgebra π•œ π•œ'] [NormOneClass π•œ'] (x : π•œ) :

        This is a simpler version of nnnorm_algebraMap when β€–1β€– = 1 in π•œ'.

        @[simp]
        theorem dist_algebraMap' {π•œ : Type u_1} (π•œ' : Type u_2) [NormedField π•œ] [SeminormedRing π•œ'] [NormedAlgebra π•œ π•œ'] [NormOneClass π•œ'] (x y : π•œ) :
        dist ((algebraMap π•œ π•œ') x) ((algebraMap π•œ π•œ') y) = dist x y

        This is a simpler version of dist_algebraMap when β€–1β€– = 1 in π•œ'.

        @[simp]
        theorem norm_algebraMap_nnreal (π•œ' : Type u_2) [SeminormedRing π•œ'] [NormOneClass π•œ'] [NormedAlgebra ℝ π•œ'] (x : NNReal) :
        β€–(algebraMap NNReal π•œ') xβ€– = ↑x
        @[simp]
        theorem nnnorm_algebraMap_nnreal (π•œ' : Type u_2) [SeminormedRing π•œ'] [NormOneClass π•œ'] [NormedAlgebra ℝ π•œ'] (x : NNReal) :
        theorem algebraMap_isometry (π•œ : Type u_1) (π•œ' : Type u_2) [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_1) [NormedField π•œ] :
        NormedAlgebra π•œ π•œ
        Equations
        instance normedAlgebraRat {π•œ : Type u_6} [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_1) [NormedField π•œ] :
        Equations
        instance instNormedAlgebraULift (π•œ : Type u_1) (π•œ' : Type u_2) [NormedField π•œ] [SeminormedRing π•œ'] [NormedAlgebra π•œ π•œ'] :
        NormedAlgebra π•œ (ULift.{u_6, u_2} π•œ')
        Equations
        instance Prod.normedAlgebra (π•œ : Type u_1) [NormedField π•œ] {E : Type u_6} {F : Type u_7} [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_1) [NormedField π•œ] {ΞΉ : Type u_6} {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.instNormedAlgebra (π•œ : Type u_1) [NormedField π•œ] {E : Type u_6} [SeminormedRing E] [NormedAlgebra π•œ E] :
        Equations
        @[reducible, inline]
        abbrev NormedAlgebra.induced {F : Type u_6} (π•œ : Type u_7) (R : Type u_8) (S : Type u_9) [NormedField π•œ] [Ring R] [Algebra π•œ R] [SeminormedRing S] [NormedAlgebra π•œ S] [FunLike F R S] [NonUnitalAlgHomClass F π•œ R S] (f : F) :
        NormedAlgebra π•œ R

        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_6} {A : Type u_7} [SeminormedRing A] [NormedField π•œ] [NormedAlgebra π•œ A] (S : Subalgebra π•œ A) :
          NormedAlgebra π•œ β†₯S
          Equations
          @[instance 75]
          instance SubalgebraClass.toNormedAlgebra {S : Type u_6} {π•œ : Type u_7} {E : Type u_8} [NormedField π•œ] [SeminormedRing E] [NormedAlgebra π•œ E] [SetLike S E] [SubringClass S E] [SMulMemClass S π•œ E] (s : S) :
          NormedAlgebra π•œ β†₯s
          Equations
          instance instNormedAddCommGroupRestrictScalars {π•œ : Type u_1} {π•œ' : Type u_2} {E : Type u_3} [I : NormedAddCommGroup E] :
          NormedAddCommGroup (RestrictScalars π•œ π•œ' E)
          Equations
          instance instNonUnitalNormedRingRestrictScalars {π•œ : Type u_1} {π•œ' : Type u_2} {E : Type u_3} [I : NonUnitalNormedRing E] :
          NonUnitalNormedRing (RestrictScalars π•œ π•œ' E)
          Equations
          instance instSeminormedRingRestrictScalars {π•œ : Type u_1} {π•œ' : Type u_2} {E : Type u_3} [I : SeminormedRing E] :
          SeminormedRing (RestrictScalars π•œ π•œ' E)
          Equations
          instance instNormedRingRestrictScalars {π•œ : Type u_1} {π•œ' : Type u_2} {E : Type u_3} [I : NormedRing E] :
          NormedRing (RestrictScalars π•œ π•œ' E)
          Equations
          instance instSeminormedCommRingRestrictScalars {π•œ : Type u_1} {π•œ' : Type u_2} {E : Type u_3} [I : SeminormedCommRing E] :
          SeminormedCommRing (RestrictScalars π•œ π•œ' E)
          Equations
          instance instNormedCommRingRestrictScalars {π•œ : Type u_1} {π•œ' : Type u_2} {E : Type u_3} [I : NormedCommRing E] :
          NormedCommRing (RestrictScalars π•œ π•œ' E)
          Equations
          instance RestrictScalars.normedSpace (π•œ : Type u_1) (π•œ' : Type u_2) (E : Type u_3) [NormedField π•œ] [NormedField π•œ'] [NormedAlgebra π•œ π•œ'] [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_6} {π•œ' : Type u_7} {E : Type u_8} [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
          Instances For
            @[reducible, inline]
            abbrev NormedSpace.restrictScalars (π•œ : Type u_1) (π•œ' : Type u_2) (E : Type u_3) [NormedField π•œ] [NormedField π•œ'] [NormedAlgebra π•œ π•œ'] [SeminormedAddCommGroup E] [NormedSpace π•œ' 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.

            See Note [reducible non-instances].

            Equations
            Instances For
              instance RestrictScalars.normedAlgebra (π•œ : Type u_1) (π•œ' : Type u_2) (E : Type u_3) [NormedField π•œ] [NormedField π•œ'] [NormedAlgebra π•œ π•œ'] [SeminormedRing E] [NormedAlgebra π•œ' E] :
              NormedAlgebra π•œ (RestrictScalars π•œ π•œ' E)

              If E is a normed algebra over π•œ' and π•œ is a normed algebra over π•œ', then RestrictScalars.module is additionally a NormedAlgebra.

              Equations
              def Module.RestrictScalars.normedAlgebraOrig {π•œ : Type u_6} {π•œ' : Type u_7} {E : Type u_8} [NormedField π•œ'] [SeminormedRing E] [I : NormedAlgebra π•œ' E] :
              NormedAlgebra π•œ' (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
              Instances For
                @[reducible, inline]
                abbrev NormedAlgebra.restrictScalars (π•œ : Type u_1) (π•œ' : Type u_2) (E : Type u_3) [NormedField π•œ] [NormedField π•œ'] [NormedAlgebra π•œ π•œ'] [SeminormedRing E] [NormedAlgebra π•œ' E] :
                NormedAlgebra π•œ E

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

                This definition allows the RestrictScalars.normedAlgebra 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.

                See Note [reducible non-instances].

                Equations
                Instances For

                  Structures for constructing new normed spaces #

                  This section contains tools meant for constructing new normed spaces. These allow one to easily construct all the relevant instances (distances measures, etc) while proving only a minimal set of axioms. Furthermore, tools are provided to add a norm structure to a type that already has a preexisting uniformity or bornology: in such cases, it is necessary to keep the preexisting instances, while ensuring that the norm induces the same uniformity/bornology.

                  structure SeminormedAddCommGroup.Core (π•œ : Type u_6) (E : Type u_7) [NormedField π•œ] [AddCommGroup E] [Norm E] [Module π•œ E] :

                  A structure encapsulating minimal axioms needed to defined a seminormed vector space, as found in textbooks. This is meant to be used to easily define SeminormedAddCommGroup E instances from scratch on a type with no preexisting distance or topology.

                  Instances For
                    @[reducible, inline]
                    abbrev PseudoMetricSpace.ofSeminormedAddCommGroupCore {π•œ : Type u_6} {E : Type u_7} [NormedField π•œ] [AddCommGroup E] [Norm E] [Module π•œ E] (core : SeminormedAddCommGroup.Core π•œ E) :

                    Produces a PseudoMetricSpace E instance from a SeminormedAddCommGroup.Core. Note that if this is used to define an instance on a type, it also provides a new uniformity and topology on the type. See note [reducible non-instances].

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      @[reducible, inline]
                      abbrev PseudoEMetricSpace.ofSeminormedAddCommGroupCore {π•œ : Type u_6} {E : Type u_7} [NormedField π•œ] [AddCommGroup E] [Norm E] [Module π•œ E] (core : SeminormedAddCommGroup.Core π•œ E) :

                      Produces a PseudoEMetricSpace E instance from a SeminormedAddCommGroup.Core. Note that if this is used to define an instance on a type, it also provides a new uniformity and topology on the type. See note [reducible non-instances].

                      Equations
                      Instances For
                        @[reducible, inline]

                        Produces a PseudoEMetricSpace E instance from a SeminormedAddCommGroup.Core on a type that already has an existing uniform space structure. This requires a proof that the uniformity induced by the norm is equal to the preexisting uniformity. See note [reducible non-instances].

                        Equations
                        Instances For
                          @[reducible, inline]
                          abbrev PseudoMetricSpace.ofSeminormedAddCommGroupCoreReplaceAll {π•œ : Type u_6} {E : Type u_7} [NormedField π•œ] [AddCommGroup E] [Norm E] [Module π•œ E] [U : UniformSpace E] [B : Bornology E] (core : SeminormedAddCommGroup.Core π•œ E) (HU : uniformity E = uniformity E) (HB : βˆ€ (s : Set E), Bornology.IsBounded s ↔ Bornology.IsBounded s) :

                          Produces a PseudoEMetricSpace E instance from a SeminormedAddCommGroup.Core on a type that already has a preexisting uniform space structure and a preexisting bornology. This requires proofs that the uniformity induced by the norm is equal to the preexisting uniformity, and likewise for the bornology. See note [reducible non-instances].

                          Equations
                          Instances For
                            @[reducible, inline]
                            abbrev SeminormedAddCommGroup.ofCore {π•œ : Type u_6} {E : Type u_7} [NormedField π•œ] [AddCommGroup E] [Norm E] [Module π•œ E] (core : Core π•œ E) :

                            Produces a SeminormedAddCommGroup E instance from a SeminormedAddCommGroup.Core. Note that if this is used to define an instance on a type, it also provides a new distance measure from the norm. it must therefore not be used on a type with a preexisting distance measure or topology. See note [reducible non-instances].

                            Equations
                            Instances For
                              @[reducible, inline]
                              abbrev SeminormedAddCommGroup.ofCoreReplaceUniformity {π•œ : Type u_6} {E : Type u_7} [NormedField π•œ] [AddCommGroup E] [Norm E] [Module π•œ E] [U : UniformSpace E] (core : Core π•œ E) (H : uniformity E = uniformity E) :

                              Produces a SeminormedAddCommGroup E instance from a SeminormedAddCommGroup.Core on a type that already has an existing uniform space structure. This requires a proof that the uniformity induced by the norm is equal to the preexisting uniformity. See note [reducible non-instances].

                              Equations
                              Instances For
                                @[reducible, inline]
                                abbrev SeminormedAddCommGroup.ofCoreReplaceAll {π•œ : Type u_6} {E : Type u_7} [NormedField π•œ] [AddCommGroup E] [Norm E] [Module π•œ E] [U : UniformSpace E] [B : Bornology E] (core : Core π•œ E) (HU : uniformity E = uniformity E) (HB : βˆ€ (s : Set E), Bornology.IsBounded s ↔ Bornology.IsBounded s) :

                                Produces a SeminormedAddCommGroup E instance from a SeminormedAddCommGroup.Core on a type that already has a preexisting uniform space structure and a preexisting bornology. This requires proofs that the uniformity induced by the norm is equal to the preexisting uniformity, and likewise for the bornology. See note [reducible non-instances].

                                Equations
                                Instances For
                                  structure NormedSpace.Core (π•œ : Type u_6) (E : Type u_7) [NormedField π•œ] [AddCommGroup E] [Module π•œ E] [Norm E] extends SeminormedAddCommGroup.Core π•œ E :

                                  A structure encapsulating minimal axioms needed to defined a normed vector space, as found in textbooks. This is meant to be used to easily define NormedAddCommGroup E and NormedSpace E instances from scratch on a type with no preexisting distance or topology.

                                  Instances For
                                    @[reducible, inline]
                                    abbrev NormedAddCommGroup.ofCore {π•œ : Type u_6} {E : Type u_7} [NormedField π•œ] [AddCommGroup E] [Module π•œ E] [Norm E] (core : NormedSpace.Core π•œ E) :

                                    Produces a NormedAddCommGroup E instance from a NormedSpace.Core. Note that if this is used to define an instance on a type, it also provides a new distance measure from the norm. it must therefore not be used on a type with a preexisting distance measure. See note [reducible non-instances].

                                    Equations
                                    Instances For
                                      @[reducible, inline]
                                      abbrev NormedAddCommGroup.ofCoreReplaceUniformity {π•œ : Type u_6} {E : Type u_7} [NormedField π•œ] [AddCommGroup E] [Module π•œ E] [Norm E] [U : UniformSpace E] (core : NormedSpace.Core π•œ E) (H : uniformity E = uniformity E) :

                                      Produces a NormedAddCommGroup E instance from a NormedAddCommGroup.Core on a type that already has an existing uniform space structure. This requires a proof that the uniformity induced by the norm is equal to the preexisting uniformity. See note [reducible non-instances].

                                      Equations
                                      Instances For
                                        @[reducible, inline]
                                        abbrev NormedAddCommGroup.ofCoreReplaceAll {π•œ : Type u_6} {E : Type u_7} [NormedField π•œ] [AddCommGroup E] [Module π•œ E] [Norm E] [U : UniformSpace E] [B : Bornology E] (core : NormedSpace.Core π•œ E) (HU : uniformity E = uniformity E) (HB : βˆ€ (s : Set E), Bornology.IsBounded s ↔ Bornology.IsBounded s) :

                                        Produces a NormedAddCommGroup E instance from a NormedAddCommGroup.Core on a type that already has a preexisting uniform space structure and a preexisting bornology. This requires proofs that the uniformity induced by the norm is equal to the preexisting uniformity, and likewise for the bornology. See note [reducible non-instances].

                                        Equations
                                        Instances For
                                          @[reducible, inline]
                                          abbrev NormedSpace.ofCore {π•œ : Type u_8} {E : Type u_9} [NormedField π•œ] [SeminormedAddCommGroup E] [Module π•œ E] (core : Core π•œ E) :
                                          NormedSpace π•œ E

                                          Produces a NormedSpace π•œ E instance from a NormedSpace.Core. This is meant to be used on types where the NormedAddCommGroup E instance has also been defined using core. See note [reducible non-instances].

                                          Equations
                                          Instances For
                                            theorem AddMonoidHom.continuous_of_isBounded_nhds_zero {G : Type u_6} {H : Type u_7} [SeminormedAddCommGroup G] [SeminormedAddCommGroup H] [NormedSpace ℝ H] {s : Set G} (f : G β†’+ H) (hs : s ∈ nhds 0) (hbounded : Bornology.IsBounded (⇑f '' s)) :
                                            Continuous ⇑f

                                            A group homomorphism from a normed group to a real normed space, bounded on a neighborhood of 0, must be continuous.