Documentation

Mathlib.Topology.ContinuousFunction.Bounded

Bounded continuous functions #

The type of bounded continuous functions taking values in a metric space, with the uniform distance.

structure BoundedContinuousFunction (α : Type u) (β : Type v) [TopologicalSpace α] [PseudoMetricSpace β] extends ContinuousMap :
Type (max u v)

α →ᵇ β is the type of bounded continuous functions α → β from a topological space to a metric space.

When possible, instead of parametrizing results over (f : α →ᵇ β), you should parametrize over (F : Type*) [BoundedContinuousMapClass F α β] (f : F).

When you extend this structure, make sure to extend BoundedContinuousMapClass.

  • toFun : αβ
  • continuous_toFun : Continuous self.toFun
  • map_bounded' : ∃ (C : ), ∀ (x y : α), dist (self.toContinuousMap.toFun x) (self.toContinuousMap.toFun y) C
Instances For
    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      class BoundedContinuousMapClass (F : Type u_2) (α : outParam (Type u_3)) (β : outParam (Type u_4)) [TopologicalSpace α] [PseudoMetricSpace β] [FunLike F α β] extends ContinuousMapClass :

      BoundedContinuousMapClass F α β states that F is a type of bounded continuous maps.

      You should also extend this typeclass when you extend BoundedContinuousFunction.

      • map_continuous : ∀ (f : F), Continuous f
      • map_bounded : ∀ (f : F), ∃ (C : ), ∀ (x y : α), dist (f x) (f y) C
      Instances
        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_to_continuous_fun {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] (f : BoundedContinuousFunction α β) :
        f.toContinuousMap = f

        See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections.

        Equations
        Instances For
          theorem BoundedContinuousFunction.bounded {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] (f : BoundedContinuousFunction α β) :
          ∃ (C : ), ∀ (x y : α), dist (f x) (f y) C
          theorem BoundedContinuousFunction.ext {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f : BoundedContinuousFunction α β} {g : BoundedContinuousFunction α β} (h : ∀ (x : α), f x = g x) :
          f = g
          def BoundedContinuousFunction.mkOfBound {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] (f : C(α, β)) (C : ) (h : ∀ (x y : α), dist (f x) (f y) C) :

          A continuous function with an explicit bound is a bounded continuous function.

          Equations
          Instances For
            @[simp]
            theorem BoundedContinuousFunction.mkOfBound_coe {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f : C(α, β)} {C : } {h : ∀ (x y : α), dist (f x) (f y) C} :

            A continuous function on a compact space is automatically a bounded continuous function.

            Equations
            Instances For
              @[simp]
              theorem BoundedContinuousFunction.mkOfDiscrete_apply {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [DiscreteTopology α] (f : αβ) (C : ) (h : ∀ (x y : α), dist (f x) (f y) C) :
              ∀ (a : α), (BoundedContinuousFunction.mkOfDiscrete f C h) a = f a
              @[simp]
              theorem BoundedContinuousFunction.mkOfDiscrete_toFun {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [DiscreteTopology α] (f : αβ) (C : ) (h : ∀ (x y : α), dist (f x) (f y) C) :
              ∀ (a : α), (BoundedContinuousFunction.mkOfDiscrete f C h) a = f a
              def BoundedContinuousFunction.mkOfDiscrete {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [DiscreteTopology α] (f : αβ) (C : ) (h : ∀ (x y : α), dist (f x) (f y) C) :

              If a function is bounded on a discrete space, it is automatically continuous, and therefore gives rise to an element of the type of bounded continuous functions.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For

                The uniform distance between two bounded continuous functions.

                Equations
                theorem BoundedContinuousFunction.dist_eq {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f : BoundedContinuousFunction α β} {g : BoundedContinuousFunction α β} :
                dist f g = sInf {C : | 0 C ∀ (x : α), dist (f x) (g x) C}
                theorem BoundedContinuousFunction.dist_set_exists {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f : BoundedContinuousFunction α β} {g : BoundedContinuousFunction α β} :
                ∃ (C : ), 0 C ∀ (x : α), dist (f x) (g x) C

                The pointwise distance is controlled by the distance between functions, by definition.

                theorem BoundedContinuousFunction.dist_le {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f : BoundedContinuousFunction α β} {g : BoundedContinuousFunction α β} {C : } (C0 : 0 C) :
                dist f g C ∀ (x : α), dist (f x) (g x) C

                The distance between two functions is controlled by the supremum of the pointwise distances.

                theorem BoundedContinuousFunction.dist_le_iff_of_nonempty {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f : BoundedContinuousFunction α β} {g : BoundedContinuousFunction α β} {C : } [Nonempty α] :
                dist f g C ∀ (x : α), dist (f x) (g x) C
                theorem BoundedContinuousFunction.dist_lt_of_nonempty_compact {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f : BoundedContinuousFunction α β} {g : BoundedContinuousFunction α β} {C : } [Nonempty α] [CompactSpace α] (w : ∀ (x : α), dist (f x) (g x) < C) :
                dist f g < C
                theorem BoundedContinuousFunction.dist_lt_iff_of_compact {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f : BoundedContinuousFunction α β} {g : BoundedContinuousFunction α β} {C : } [CompactSpace α] (C0 : 0 < C) :
                dist f g < C ∀ (x : α), dist (f x) (g x) < C

                The type of bounded continuous functions, with the uniform distance, is a pseudometric space.

                Equations
                • One or more equations did not get rendered due to their size.

                The type of bounded continuous functions, with the uniform distance, is a metric space.

                Equations
                theorem BoundedContinuousFunction.nndist_eq {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f : BoundedContinuousFunction α β} {g : BoundedContinuousFunction α β} :
                nndist f g = sInf {C : NNReal | ∀ (x : α), nndist (f x) (g x) C}
                theorem BoundedContinuousFunction.nndist_set_exists {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f : BoundedContinuousFunction α β} {g : BoundedContinuousFunction α β} :
                ∃ (C : NNReal), ∀ (x : α), nndist (f x) (g x) C

                On an empty space, bounded continuous functions are at distance 0.

                theorem BoundedContinuousFunction.dist_eq_iSup {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f : BoundedContinuousFunction α β} {g : BoundedContinuousFunction α β} :
                dist f g = ⨆ (x : α), dist (f x) (g x)
                theorem BoundedContinuousFunction.nndist_eq_iSup {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f : BoundedContinuousFunction α β} {g : BoundedContinuousFunction α β} :
                nndist f g = ⨆ (x : α), nndist (f x) (g x)
                theorem BoundedContinuousFunction.tendsto_iff_tendstoUniformly {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {ι : Type u_2} {F : ιBoundedContinuousFunction α β} {f : BoundedContinuousFunction α β} {l : Filter ι} :
                Filter.Tendsto F l (nhds f) TendstoUniformly (fun (i : ι) => (F i)) (f) l
                theorem BoundedContinuousFunction.inducing_coeFn {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] :
                Inducing (UniformFun.ofFun DFunLike.coe)

                The topology on α →ᵇ β is exactly the topology induced by the natural map to α →ᵤ β.

                theorem BoundedContinuousFunction.embedding_coeFn {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] :
                Embedding (UniformFun.ofFun DFunLike.coe)
                @[simp]
                theorem BoundedContinuousFunction.const_apply (α : Type u) {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] (b : β) :
                (BoundedContinuousFunction.const α b) = fun (x : α) => b
                @[simp]
                theorem BoundedContinuousFunction.const_toFun (α : Type u) {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] (b : β) :
                (BoundedContinuousFunction.const α b) = fun (x : α) => b

                Constant as a continuous bounded function.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For

                  If the target space is inhabited, so is the space of bounded continuous functions.

                  Equations

                  When x is fixed, (f : α →ᵇ β) ↦ f x is continuous.

                  The evaluation map is continuous, as a joint function of u and x.

                  Composition of a bounded continuous function and a continuous function.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    @[simp]

                    Composition (in the target) of a bounded continuous function with a Lipschitz map again gives a bounded continuous function.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For

                      The composition operator (in the target) with a Lipschitz map is Lipschitz.

                      The composition operator (in the target) with a Lipschitz map is uniformly continuous.

                      The composition operator (in the target) with a Lipschitz map is continuous.

                      def BoundedContinuousFunction.codRestrict {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] (s : Set β) (f : BoundedContinuousFunction α β) (H : ∀ (x : α), f x s) :

                      Restriction (in the target) of a bounded continuous function taking values in a subset.

                      Equations
                      Instances For

                        A version of Function.extend for bounded continuous maps. We assume that the domain has discrete topology, so we only need to verify boundedness.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          @[simp]
                          theorem BoundedContinuousFunction.extend_apply' {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {δ : Type u_2} [TopologicalSpace δ] [DiscreteTopology δ] {f : α δ} {x : δ} (hx : xSet.range f) (g : BoundedContinuousFunction α β) (h : BoundedContinuousFunction δ β) :
                          theorem BoundedContinuousFunction.arzela_ascoli₁ {α : Type u} {β : Type v} [TopologicalSpace α] [CompactSpace α] [PseudoMetricSpace β] [CompactSpace β] (A : Set (BoundedContinuousFunction α β)) (closed : IsClosed A) (H : Equicontinuous fun (x : A) => x) :

                          First version, with pointwise equicontinuity and range in a compact space.

                          theorem BoundedContinuousFunction.arzela_ascoli₂ {α : Type u} {β : Type v} [TopologicalSpace α] [CompactSpace α] [PseudoMetricSpace β] (s : Set β) (hs : IsCompact s) (A : Set (BoundedContinuousFunction α β)) (closed : IsClosed A) (in_s : ∀ (f : BoundedContinuousFunction α β) (x : α), f Af x s) (H : Equicontinuous fun (x : A) => x) :

                          Second version, with pointwise equicontinuity and range in a compact subset.

                          theorem BoundedContinuousFunction.arzela_ascoli {α : Type u} {β : Type v} [TopologicalSpace α] [CompactSpace α] [PseudoMetricSpace β] [T2Space β] (s : Set β) (hs : IsCompact s) (A : Set (BoundedContinuousFunction α β)) (in_s : ∀ (f : BoundedContinuousFunction α β) (x : α), f Af x s) (H : Equicontinuous fun (x : A) => x) :

                          Third (main) version, with pointwise equicontinuity and range in a compact subset, but without closedness. The closure is then compact.

                          Equations
                          Equations
                          @[simp]
                          theorem BoundedContinuousFunction.coe_zero {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [Zero β] :
                          0 = 0
                          @[simp]
                          theorem BoundedContinuousFunction.coe_one {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [One β] :
                          1 = 1
                          theorem BoundedContinuousFunction.forall_coe_zero_iff_zero {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [Zero β] (f : BoundedContinuousFunction α β) :
                          (∀ (x : α), f x = 0) f = 0
                          theorem BoundedContinuousFunction.forall_coe_one_iff_one {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [One β] (f : BoundedContinuousFunction α β) :
                          (∀ (x : α), f x = 1) f = 1

                          The pointwise sum of two bounded continuous functions is again bounded continuous.

                          Equations
                          • One or more equations did not get rendered due to their size.
                          @[simp]
                          theorem BoundedContinuousFunction.coe_add {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [AddMonoid β] [LipschitzAdd β] (f : BoundedContinuousFunction α β) (g : BoundedContinuousFunction α β) :
                          (f + g) = f + g
                          theorem BoundedContinuousFunction.add_apply {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [AddMonoid β] [LipschitzAdd β] (f : BoundedContinuousFunction α β) (g : BoundedContinuousFunction α β) {x : α} :
                          (f + g) x = f x + g x
                          @[simp]
                          Equations
                          • One or more equations did not get rendered due to their size.
                          @[simp]
                          theorem BoundedContinuousFunction.coe_nsmul {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [AddMonoid β] [LipschitzAdd β] (r : ) (f : BoundedContinuousFunction α β) :
                          (r f) = r f
                          @[simp]
                          theorem BoundedContinuousFunction.nsmul_apply {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [AddMonoid β] [LipschitzAdd β] (r : ) (f : BoundedContinuousFunction α β) (v : α) :
                          (r f) v = r f v
                          Equations
                          • One or more equations did not get rendered due to their size.
                          @[simp]
                          theorem BoundedContinuousFunction.coeFnAddHom_apply {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [AddMonoid β] [LipschitzAdd β] :
                          ∀ (a : BoundedContinuousFunction α β) (a_1 : α), BoundedContinuousFunction.coeFnAddHom a a_1 = a a_1

                          Coercion of a NormedAddGroupHom is an AddMonoidHom. Similar to AddMonoidHom.coeFn.

                          Equations
                          • BoundedContinuousFunction.coeFnAddHom = { toZeroHom := { toFun := DFunLike.coe, map_zero' := (_ : 0 = 0) }, map_add' := (_ : ∀ (f g : BoundedContinuousFunction α β), (f + g) = f + g) }
                          Instances For

                            The additive map forgetting that a bounded continuous function is bounded.

                            Equations
                            • One or more equations did not get rendered due to their size.
                            Instances For
                              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_sum {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [AddCommMonoid β] [LipschitzAdd β] {ι : Type u_2} (s : Finset ι) (f : ιBoundedContinuousFunction α β) :
                              (Finset.sum s fun (i : ι) => f i) = Finset.sum s fun (i : ι) => (f i)
                              theorem BoundedContinuousFunction.sum_apply {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [AddCommMonoid β] [LipschitzAdd β] {ι : Type u_2} (s : Finset ι) (f : ιBoundedContinuousFunction α β) (a : α) :
                              (Finset.sum s fun (i : ι) => f i) a = Finset.sum s fun (i : ι) => (f i) a
                              Equations
                              theorem BoundedContinuousFunction.norm_eq {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] (f : BoundedContinuousFunction α β) :
                              f = sInf {C : | 0 C ∀ (x : α), f x C}

                              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.

                              theorem BoundedContinuousFunction.norm_eq_of_nonempty {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] (f : BoundedContinuousFunction α β) [h : Nonempty α] :
                              f = sInf {C : | ∀ (x : α), f x C}

                              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
                              theorem BoundedContinuousFunction.dist_le_two_norm {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] (f : BoundedContinuousFunction α β) (x : α) (y : α) :
                              dist (f x) (f y) 2 * f

                              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
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                @[simp]
                                theorem BoundedContinuousFunction.coe_ofNormedAddCommGroup {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] (f : αβ) (Hf : Continuous f) (C : ) (H : ∀ (x : α), f x C) :
                                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

                                  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

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

                                    Equations
                                    • One or more equations did not get rendered due to their size.

                                    The pointwise difference of two bounded continuous functions is again bounded continuous.

                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    @[simp]
                                    theorem BoundedContinuousFunction.coe_sub {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] (f : BoundedContinuousFunction α β) (g : BoundedContinuousFunction α β) :
                                    (f - g) = f - g
                                    theorem BoundedContinuousFunction.sub_apply {α : Type u} {β : Type v} [TopologicalSpace α] [SeminormedAddCommGroup β] (f : BoundedContinuousFunction α β) (g : BoundedContinuousFunction α β) {x : α} :
                                    (f - g) x = f x - g x
                                    @[simp]
                                    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
                                    • BoundedContinuousFunction.seminormedAddCommGroup = SeminormedAddCommGroup.mk
                                    Equations
                                    • BoundedContinuousFunction.normedAddCommGroup = let src := BoundedContinuousFunction.seminormedAddCommGroup; NormedAddCommGroup.mk

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

                                    BoundedSMul (in particular, topological module) structure #

                                    In this section, if β is a metric space and a 𝕜-module whose addition and scalar multiplication are compatible with the metric structure, then we show that the space of bounded continuous functions from α to β inherits a so-called BoundedSMul structure (in particular, a ContinuousMul structure, which is the mathlib formulation of being a topological module), by using pointwise operations and checking that they are compatible with the uniform distance.

                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    @[simp]
                                    theorem BoundedContinuousFunction.coe_smul {α : Type u} {β : Type v} {𝕜 : Type u_2} [PseudoMetricSpace 𝕜] [TopologicalSpace α] [PseudoMetricSpace β] [Zero 𝕜] [Zero β] [SMul 𝕜 β] [BoundedSMul 𝕜 β] (c : 𝕜) (f : BoundedContinuousFunction α β) :
                                    (c f) = fun (x : α) => c f x
                                    theorem BoundedContinuousFunction.smul_apply {α : Type u} {β : Type v} {𝕜 : Type u_2} [PseudoMetricSpace 𝕜] [TopologicalSpace α] [PseudoMetricSpace β] [Zero 𝕜] [Zero β] [SMul 𝕜 β] [BoundedSMul 𝕜 β] (c : 𝕜) (f : BoundedContinuousFunction α β) (x : α) :
                                    (c f) x = c 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.
                                    instance BoundedContinuousFunction.module {α : Type u} {β : Type v} {𝕜 : Type u_2} [PseudoMetricSpace 𝕜] [TopologicalSpace α] [PseudoMetricSpace β] [Semiring 𝕜] [AddCommMonoid β] [Module 𝕜 β] [BoundedSMul 𝕜 β] [LipschitzAdd β] :
                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    def BoundedContinuousFunction.evalCLM {α : Type u} {β : Type v} (𝕜 : Type u_2) [PseudoMetricSpace 𝕜] [TopologicalSpace α] [PseudoMetricSpace β] [Semiring 𝕜] [AddCommMonoid β] [Module 𝕜 β] [BoundedSMul 𝕜 β] [LipschitzAdd β] (x : α) :

                                    The evaluation at a point, as a continuous linear map from α →ᵇ β to β.

                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    Instances For
                                      @[simp]
                                      theorem BoundedContinuousFunction.evalCLM_apply {α : Type u} {β : Type v} (𝕜 : Type u_2) [PseudoMetricSpace 𝕜] [TopologicalSpace α] [PseudoMetricSpace β] [Semiring 𝕜] [AddCommMonoid β] [Module 𝕜 β] [BoundedSMul 𝕜 β] [LipschitzAdd β] (x : α) (f : BoundedContinuousFunction α β) :
                                      @[simp]

                                      The linear map forgetting that a bounded continuous function is bounded.

                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For

                                        Normed space structure #

                                        In this section, if β is a normed space, then we show that the space of bounded continuous functions from α to β inherits a normed space structure, by using pointwise operations and checking that they are compatible with the uniform distance.

                                        Equations
                                        • One or more equations did not get rendered due to their size.

                                        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]

                                          Normed ring structure #

                                          In this section, if R is a normed ring, then we show that the space of bounded continuous functions from α to R inherits a normed ring structure, by using pointwise operations and checking that they are compatible with the uniform distance.

                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          @[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.
                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          @[simp]
                                          theorem BoundedContinuousFunction.coe_npowRec {α : Type u} [TopologicalSpace α] {R : Type u_2} [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_pow {α : Type u} [TopologicalSpace α] {R : Type u_2} [SeminormedRing R] (n : ) (f : BoundedContinuousFunction α R) :
                                          (f ^ n) = f ^ n
                                          @[simp]
                                          theorem BoundedContinuousFunction.pow_apply {α : Type u} [TopologicalSpace α] {R : Type u_2} [SeminormedRing R] (n : ) (f : BoundedContinuousFunction α R) (v : α) :
                                          (f ^ n) v = f v ^ n
                                          Equations
                                          @[simp]
                                          theorem BoundedContinuousFunction.coe_natCast {α : Type u} [TopologicalSpace α] {R : Type u_2} [SeminormedRing R] (n : ) :
                                          n = n
                                          Equations
                                          @[simp]
                                          theorem BoundedContinuousFunction.coe_intCast {α : Type u} [TopologicalSpace α] {R : Type u_2} [SeminormedRing R] (n : ) :
                                          n = n
                                          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.

                                          Normed commutative ring structure #

                                          In this section, if R is a normed commutative ring, then we show that the space of bounded continuous functions from α to R inherits a normed commutative ring structure, by using pointwise operations and checking that they are compatible with the uniform distance.

                                          Equations

                                          Normed algebra structure #

                                          In this section, if γ is a normed algebra, then we show that the space of bounded continuous functions from α to γ inherits a normed algebra structure, by using pointwise operations and checking that they are compatible with the uniform distance.

                                          def BoundedContinuousFunction.C {α : Type u} {γ : Type w} {𝕜 : Type u_2} [NormedField 𝕜] [TopologicalSpace α] [NormedRing γ] [NormedAlgebra 𝕜 γ] :

                                          BoundedContinuousFunction.const as a RingHom.

                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For
                                            instance BoundedContinuousFunction.algebra {α : Type u} {γ : Type w} {𝕜 : Type u_2} [NormedField 𝕜] [TopologicalSpace α] [NormedRing γ] [NormedAlgebra 𝕜 γ] :
                                            Equations
                                            • One or more equations did not get rendered due to their size.
                                            @[simp]
                                            theorem BoundedContinuousFunction.algebraMap_apply {α : Type u} {γ : Type w} {𝕜 : Type u_2} [NormedField 𝕜] [TopologicalSpace α] [NormedRing γ] [NormedAlgebra 𝕜 γ] (k : 𝕜) (a : α) :
                                            ((algebraMap 𝕜 (BoundedContinuousFunction α γ)) k) a = k 1

                                            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
                                            • One or more equations did not get rendered due to their size.

                                            Star structures #

                                            In this section, if β is a normed ⋆-group, then so is the space of bounded continuous functions from α to β, by using the star operation pointwise.

                                            If 𝕜 is normed field and a ⋆-ring over which β is a normed algebra and a star module, then the space of bounded continuous functions from α to β is a star module.

                                            If β is a ⋆-ring in addition to being a normed ⋆-group, then α →ᵇ β inherits a ⋆-ring structure.

                                            In summary, if β is a C⋆-algebra over 𝕜, then so is α →ᵇ β; note that completeness is guaranteed when β is complete (see BoundedContinuousFunction.complete).

                                            Equations
                                            @[simp]

                                            The right-hand side of this equality can be parsed star ∘ ⇑f because of the instance Pi.instStarForAll. Upon inspecting the goal, one sees ⊢ ↑(star f) = star ↑f.

                                            Equations

                                            Continuous normed lattice group valued functions form a meet-semilattice.

                                            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.

                                            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

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