Documentation

Mathlib.Topology.ContinuousMap.Bounded.Normed

Inheritance of normed algebraic structures by bounded continuous functions #

For various types of normed algebraic structures β, we show in this file that the space of bounded continuous functions from α to β inherits the same normed structure, by using pointwise operations and checking that they are compatible with the uniform distance.

The norm of a bounded continuous function is the supremum of ‖f x‖. We use sInf to ensure that the definition works if α has no elements.

When the domain is non-empty, we do not need the 0 ≤ C condition in the formula for ‖f‖ as a sInf.

theorem BoundedContinuousFunction.dist_le_two_norm' {β : Type v} {γ : Type w} [SeminormedAddCommGroup β] {f : γβ} {C : } (hC : ∀ (x : γ), f x C) (x y : γ) :
dist (f x) (f y) 2 * C

Distance between the images of any two points is at most twice the norm of the function.

theorem BoundedContinuousFunction.norm_le {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] {f : BoundedContinuousFunction α β} {C : } (C0 : 0 C) :
f C ∀ (x : α), f x C

The norm of a function is controlled by the supremum of the pointwise norms.

theorem BoundedContinuousFunction.norm_lt_iff_of_compact {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] [CompactSpace α] {f : BoundedContinuousFunction α β} {M : } (M0 : 0 < M) :
f < M ∀ (x : α), f x < M

Norm of const α b is less than or equal to ‖b‖. If α is nonempty, then it is equal to ‖b‖.

def BoundedContinuousFunction.ofNormedAddCommGroup {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] (f : αβ) (Hf : Continuous f) (C : ) (H : ∀ (x : α), f x C) :

Constructing a bounded continuous function from a uniformly bounded continuous function taking values in a normed group.

Equations
Instances For
    @[simp]
    theorem BoundedContinuousFunction.coe_ofNormedAddCommGroup {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] (f : αβ) (Hf : Continuous f) (C : ) (H : ∀ (x : α), f x C) :
    (ofNormedAddCommGroup f Hf C H) = f
    theorem BoundedContinuousFunction.norm_ofNormedAddCommGroup_le {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] {f : αβ} (hfc : Continuous f) {C : } (hC : 0 C) (hfC : ∀ (x : α), f x C) :

    Constructing a bounded continuous function from a uniformly bounded function on a discrete space, taking values in a normed group.

    Equations
    Instances For
      @[simp]
      theorem BoundedContinuousFunction.coe_ofNormedAddCommGroupDiscrete {α : Type u} {β : Type v} [TopologicalSpace α] [DiscreteTopology α] [SeminormedAddCommGroup β] (f : αβ) (C : ) (H : ∀ (x : α), f x C) :

      Taking the pointwise norm of a bounded continuous function with values in a SeminormedAddCommGroup yields a bounded continuous function with values in ℝ.

      Equations
      Instances For

        If ‖(1 : β)‖ = 1, then ‖(1 : α →ᵇ β)‖ = 1 if α is nonempty.

        The pointwise opposite of a bounded continuous function is again bounded continuous.

        Equations
        @[simp]
        theorem BoundedContinuousFunction.coe_zsmulRec {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] (f : BoundedContinuousFunction α β) (z : ) :
        (zsmulRec (fun (x1 : ) (x2 : BoundedContinuousFunction α β) => x1 x2) z f) = z f
        Equations
        • One or more equations did not get rendered due to their size.
        @[simp]
        theorem BoundedContinuousFunction.coe_zsmul {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] (r : ) (f : BoundedContinuousFunction α β) :
        ⇑(r f) = r f
        @[simp]
        theorem BoundedContinuousFunction.zsmul_apply {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] (r : ) (f : BoundedContinuousFunction α β) (v : α) :
        (r f) v = r f v
        Equations
        • One or more equations did not get rendered due to their size.
        Equations
        • One or more equations did not get rendered due to their size.

        The nnnorm of a function is controlled by the supremum of the pointwise nnnorms.

        Postcomposition of bounded continuous functions into a normed module by a continuous linear map is a continuous linear map. Upgraded version of ContinuousLinearMap.compLeftContinuous, similar to LinearMap.compLeft.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          @[simp]
          Equations
          • One or more equations did not get rendered due to their size.
          Equations
          • One or more equations did not get rendered due to their size.
          @[simp]
          theorem BoundedContinuousFunction.coe_npowRec {α : Type u} [TopologicalSpace α] {R : Type u_1} [SeminormedRing R] (f : BoundedContinuousFunction α R) (n : ) :
          (npowRec n f) = f ^ n
          Equations
          • One or more equations did not get rendered due to their size.
          @[simp]
          theorem BoundedContinuousFunction.coe_natCast {α : Type u} [TopologicalSpace α] {R : Type u_1} [SeminormedRing R] (n : ) :
          n = n
          @[simp]
          theorem BoundedContinuousFunction.coe_intCast {α : Type u} [TopologicalSpace α] {R : Type u_1} [SeminormedRing R] (n : ) :
          n = n
          Equations
          Equations
          • One or more equations did not get rendered due to their size.

          Composition on the left by a (lipschitz-continuous) homomorphism of topological semirings, as a RingHom. Similar to RingHom.compLeftContinuous.

          Equations
          Instances For
            @[simp]
            theorem RingHom.compLeftContinuousBounded_apply_apply {β : Type v} {γ : Type w} (α : Type u_2) [TopologicalSpace α] [SeminormedRing β] [SeminormedRing γ] (g : β →+* γ) {C : NNReal} (hg : LipschitzWith C g) (f : BoundedContinuousFunction α β) (x : α) :
            ((RingHom.compLeftContinuousBounded α g hg) f) x = g (f x)
            Equations
            • One or more equations did not get rendered due to their size.
            Equations
            • One or more equations did not get rendered due to their size.
            Equations
            • One or more equations did not get rendered due to their size.
            def BoundedContinuousFunction.C {α : Type u} {γ : Type w} {𝕜 : Type u_1} [NormedField 𝕜] [TopologicalSpace α] [NormedRing γ] [NormedAlgebra 𝕜 γ] :

            BoundedContinuousFunction.const as a RingHom.

            Equations
            Instances For
              instance BoundedContinuousFunction.instAlgebra {α : Type u} {γ : Type w} {𝕜 : Type u_1} [NormedField 𝕜] [TopologicalSpace α] [NormedRing γ] [NormedAlgebra 𝕜 γ] :
              Equations
              @[simp]
              theorem BoundedContinuousFunction.algebraMap_apply {α : Type u} {γ : Type w} {𝕜 : Type u_1} [NormedField 𝕜] [TopologicalSpace α] [NormedRing γ] [NormedAlgebra 𝕜 γ] (k : 𝕜) (a : α) :
              ((algebraMap 𝕜 (BoundedContinuousFunction α γ)) k) a = k 1
              def BoundedContinuousFunction.AlgHom.compLeftContinuousBounded {α : Type u} {β : Type v} {γ : Type w} (𝕜 : Type u_1) [NormedField 𝕜] [TopologicalSpace α] [NormedRing β] [NormedAlgebra 𝕜 β] [NormedRing γ] [NormedAlgebra 𝕜 γ] (g : β →ₐ[𝕜] γ) {C : NNReal} (hg : LipschitzWith C g) :

              Composition on the left by a (lipschitz-continuous) homomorphism of topological R-algebras, as an AlgHom. Similar to AlgHom.compLeftContinuous.

              Equations
              Instances For
                @[simp]
                theorem BoundedContinuousFunction.AlgHom.compLeftContinuousBounded_apply_apply {α : Type u} {β : Type v} {γ : Type w} (𝕜 : Type u_1) [NormedField 𝕜] [TopologicalSpace α] [NormedRing β] [NormedAlgebra 𝕜 β] [NormedRing γ] [NormedAlgebra 𝕜 γ] (g : β →ₐ[𝕜] γ) {C : NNReal} (hg : LipschitzWith C g) (f : BoundedContinuousFunction α β) (x : α) :
                ((AlgHom.compLeftContinuousBounded 𝕜 g hg) f) x = g (f x)

                The algebra-homomorphism forgetting that a bounded continuous function is bounded.

                Equations
                Instances For
                  @[simp]
                  @[simp]
                  theorem BoundedContinuousFunction.coe_toContinuousMapₐ {α : Type u} {γ : Type w} (𝕜 : Type u_1) [NormedField 𝕜] [TopologicalSpace α] [NormedRing γ] [NormedAlgebra 𝕜 γ] (f : BoundedContinuousFunction α γ) :
                  ((toContinuousMapₐ 𝕜) f) = f

                  Structure as normed module over scalar functions #

                  If β is a normed 𝕜-space, then we show that the space of bounded continuous functions from α to β is naturally a module over the algebra of bounded continuous functions from α to 𝕜.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Equations
                  Equations
                  @[simp]
                  theorem BoundedContinuousFunction.coe_sup {α : Type u} {β : Type v} [TopologicalSpace α] [NormedAddCommGroup β] [Lattice β] [HasSolidNorm β] [IsOrderedAddMonoid β] (f g : BoundedContinuousFunction α β) :
                  ⇑(f g) = f g
                  @[simp]
                  theorem BoundedContinuousFunction.coe_inf {α : Type u} {β : Type v} [TopologicalSpace α] [NormedAddCommGroup β] [Lattice β] [HasSolidNorm β] [IsOrderedAddMonoid β] (f g : BoundedContinuousFunction α β) :
                  ⇑(f g) = f g

                  The nonnegative part of a bounded continuous -valued function as a bounded continuous ℝ≥0-valued function.

                  Equations
                  Instances For

                    The absolute value of a bounded continuous -valued function as a bounded continuous ℝ≥0-valued function.

                    Equations
                    Instances For

                      Decompose a bounded continuous function to its positive and negative parts.

                      Express the absolute value of a bounded continuous function in terms of its positive and negative parts.