Documentation

Mathlib.Topology.ContinuousMap.Bounded.Basic

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 C(α, β) :
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.

α →ᵇ β 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.

Equations
  • One or more equations did not get rendered due to their size.
class BoundedContinuousMapClass (F : Type u_2) (α : outParam (Type u_3)) (β : outParam (Type u_4)) [TopologicalSpace α] [PseudoMetricSpace β] [FunLike F α β] extends ContinuousMapClass F α β :

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

You should also extend this typeclass when you extend BoundedContinuousFunction.

Instances
    Equations
    Equations
    @[deprecated BoundedContinuousFunction.coe_toContinuousMap (since := "2024-11-23")]

    Alias of BoundedContinuousFunction.coe_toContinuousMap.

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

    Equations
    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 g : BoundedContinuousFunction α β} (h : ∀ (x : α), f x = g x) :
    f = g
    theorem BoundedContinuousFunction.ext_iff {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f g : BoundedContinuousFunction α β} :
    f = g ∀ (x : α), f x = g x
    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
    @[simp]
    theorem BoundedContinuousFunction.mkOfBound_coe {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f : C(α, β)} {C : } {h : ∀ (x y : α), dist (f x) (f y) C} :
    (mkOfBound f C h) = f

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

    Equations
    @[simp]
    theorem BoundedContinuousFunction.mkOfCompact_apply {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [CompactSpace α] (f : C(α, β)) (a : α) :
    (mkOfCompact f) 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
    @[simp]
    theorem BoundedContinuousFunction.mkOfDiscrete_apply {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [DiscreteTopology α] (f : αβ) (C : ) (h : ∀ (x y : α), dist (f x) (f y) C) (a✝ : α) :
    (mkOfDiscrete f C h) a✝ = f a✝

    The uniform distance between two bounded continuous functions.

    Equations
    theorem BoundedContinuousFunction.dist_eq {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f 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 g : BoundedContinuousFunction α β} :
    ∃ (C : ), 0 C ∀ (x : α), dist (f x) (g x) C
    theorem BoundedContinuousFunction.dist_coe_le_dist {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f g : BoundedContinuousFunction α β} (x : α) :
    dist (f x) (g x) dist f g

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

    theorem BoundedContinuousFunction.dist_le {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f 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 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 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 g : BoundedContinuousFunction α β} {C : } [CompactSpace α] (C0 : 0 < C) :
    dist f g < C ∀ (x : α), dist (f x) (g x) < C
    theorem BoundedContinuousFunction.dist_lt_iff_of_nonempty_compact {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f g : BoundedContinuousFunction α β} {C : } [Nonempty α] [CompactSpace α] :
    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 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 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 g : BoundedContinuousFunction α β} :
    dist f g = ⨆ (x : α), dist (f x) (g x)
    theorem BoundedContinuousFunction.nndist_eq_iSup {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {f 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

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

    @[deprecated BoundedContinuousFunction.isInducing_coeFn (since := "2024-10-28")]

    Alias of BoundedContinuousFunction.isInducing_coeFn.


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

    @[deprecated BoundedContinuousFunction.isEmbedding_coeFn (since := "2024-10-26")]

    Alias of BoundedContinuousFunction.isEmbedding_coeFn.

    Constant as a continuous bounded function.

    Equations
    @[simp]
    theorem BoundedContinuousFunction.const_apply (α : Type u) {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] (b : β) :
    (const α b) = fun (x : α) => b
    theorem BoundedContinuousFunction.const_apply' {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] (a : α) (b : β) :
    (const α b) a = b

    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.

    Bounded continuous functions taking values in a complete space form a complete space.

    Composition of a bounded continuous function and a continuous function.

    Equations
    @[simp]
    theorem BoundedContinuousFunction.coe_compContinuous {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {δ : Type u_2} [TopologicalSpace δ] (f : BoundedContinuousFunction α β) (g : C(δ, α)) :
    (f.compContinuous g) = f g
    @[simp]
    theorem BoundedContinuousFunction.compContinuous_apply {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {δ : Type u_2} [TopologicalSpace δ] (f : BoundedContinuousFunction α β) (g : C(δ, α)) (x : δ) :
    (f.compContinuous g) x = f (g x)

    Restrict a bounded continuous function to a set.

    Equations
    @[simp]
    theorem BoundedContinuousFunction.restrict_apply {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] (f : BoundedContinuousFunction α β) (s : Set α) (x : s) :
    (f.restrict s) x = f x

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

    Equations
    @[simp]
    theorem BoundedContinuousFunction.comp_apply {α : Type u} {β : Type v} {γ : Type w} [TopologicalSpace α] [PseudoMetricSpace β] [PseudoMetricSpace γ] (G : βγ) {C : NNReal} (H : LipschitzWith C G) (f : BoundedContinuousFunction α β) (a : α) :
    (comp G H f) a = G (f a)
    theorem BoundedContinuousFunction.lipschitz_comp {α : Type u} {β : Type v} {γ : Type w} [TopologicalSpace α] [PseudoMetricSpace β] [PseudoMetricSpace γ] {G : βγ} {C : NNReal} (H : LipschitzWith C G) :

    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.

    theorem BoundedContinuousFunction.continuous_comp {α : Type u} {β : Type v} {γ : Type w} [TopologicalSpace α] [PseudoMetricSpace β] [PseudoMetricSpace γ] {G : βγ} {C : NNReal} (H : LipschitzWith C G) :

    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

    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
    @[simp]
    theorem BoundedContinuousFunction.extend_apply {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {δ : Type u_2} [TopologicalSpace δ] [DiscreteTopology δ] (f : α δ) (g : BoundedContinuousFunction α β) (h : BoundedContinuousFunction δ β) (x : α) :
    (extend f g h) (f x) = g x
    @[simp]
    theorem BoundedContinuousFunction.extend_comp {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {δ : Type u_2} [TopologicalSpace δ] [DiscreteTopology δ] (f : α δ) (g : BoundedContinuousFunction α β) (h : BoundedContinuousFunction δ β) :
    (extend f g h) f = g
    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 δ β) :
    (extend f g h) x = h x
    @[simp]
    theorem BoundedContinuousFunction.dist_extend_extend {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] {δ : Type u_2} [TopologicalSpace δ] [DiscreteTopology δ] (f : α δ) (g₁ g₂ : BoundedContinuousFunction α β) (h₁ h₂ : BoundedContinuousFunction δ β) :
    dist (extend f g₁ h₁) (extend f g₂ h₂) = max (dist g₁ g₂) (dist (h₁.restrict (Set.range f)) (h₂.restrict (Set.range f)))

    The indicator function of a clopen set, as a bounded continuous function.

    Equations
    @[simp]
    theorem BoundedContinuousFunction.indicator_apply {α : Type u} [TopologicalSpace α] (s : Set α) (hs : IsClopen s) (x : α) :
    (indicator s hs) x = s.indicator 1 x
    @[simp]
    theorem BoundedContinuousFunction.coe_one {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [One β] :
    1 = 1
    @[simp]
    theorem BoundedContinuousFunction.coe_zero {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [Zero β] :
    0 = 0
    theorem BoundedContinuousFunction.forall_coe_one_iff_one {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [One β] (f : BoundedContinuousFunction α β) :
    (∀ (x : α), f x = 1) f = 1
    theorem BoundedContinuousFunction.forall_coe_zero_iff_zero {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [Zero β] (f : BoundedContinuousFunction α β) :
    (∀ (x : α), f x = 0) f = 0
    @[simp]
    @[simp]
    Equations
    Equations
    @[simp]
    theorem BoundedContinuousFunction.coe_mul {α : Type u} {R : Type u_2} [TopologicalSpace α] [PseudoMetricSpace R] [Mul R] [BoundedMul R] [ContinuousMul R] (f g : BoundedContinuousFunction α R) :
    ⇑(f * g) = f * g
    @[simp]
    theorem BoundedContinuousFunction.coe_add {α : Type u} {R : Type u_2} [TopologicalSpace α] [PseudoMetricSpace R] [Add R] [BoundedAdd R] [ContinuousAdd R] (f g : BoundedContinuousFunction α R) :
    ⇑(f + g) = f + g
    theorem BoundedContinuousFunction.mul_apply {α : Type u} {R : Type u_2} [TopologicalSpace α] [PseudoMetricSpace R] [Mul R] [BoundedMul R] [ContinuousMul R] (f g : BoundedContinuousFunction α R) (x : α) :
    (f * g) x = f x * g x
    theorem BoundedContinuousFunction.add_apply {α : Type u} {R : Type u_2} [TopologicalSpace α] [PseudoMetricSpace R] [Add R] [BoundedAdd R] [ContinuousAdd R] (f g : BoundedContinuousFunction α R) (x : α) :
    (f + g) x = f x + g x
    @[simp]
    Equations
    • One or more equations did not get rendered due to their size.
    Equations
    theorem BoundedContinuousFunction.coe_pow {α : Type u} {R : Type u_2} [TopologicalSpace α] [PseudoMetricSpace R] [Monoid R] [BoundedMul R] [ContinuousMul R] (n : ) (f : BoundedContinuousFunction α R) :
    ⇑(f ^ n) = f ^ n
    @[simp]
    theorem BoundedContinuousFunction.pow_apply {α : Type u} {R : Type u_2} [TopologicalSpace α] [PseudoMetricSpace R] [Monoid R] [BoundedMul R] [ContinuousMul R] (n : ) (f : BoundedContinuousFunction α R) (x : α) :
    (f ^ n) x = f x ^ n
    @[simp]
    theorem BoundedContinuousFunction.nsmul_apply {α : Type u} {R : Type u_2} [TopologicalSpace α] [PseudoMetricSpace R] [AddMonoid R] [BoundedAdd R] [ContinuousAdd R] (n : ) (f : BoundedContinuousFunction α R) (x : α) :
    (n f) x = n f x

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

    Equations

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

    Equations

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

    Equations
    @[simp]

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

    Equations
    @[simp]
    theorem BoundedContinuousFunction.coe_sum {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [AddCommMonoid β] [BoundedAdd β] [ContinuousAdd β] {ι : Type u_2} (s : Finset ι) (f : ιBoundedContinuousFunction α β) :
    (∑ is, f i) = is, (f i)
    theorem BoundedContinuousFunction.sum_apply {α : Type u} {β : Type v} [TopologicalSpace α] [PseudoMetricSpace β] [AddCommMonoid β] [BoundedAdd β] [ContinuousAdd β] {ι : Type u_2} (s : Finset ι) (f : ιBoundedContinuousFunction α β) (a : α) :
    (∑ is, f i) a = is, (f i) a

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

    Equations
    theorem BoundedContinuousFunction.sub_apply {α : Type u} [TopologicalSpace α] {R : Type u_2} [PseudoMetricSpace R] [Sub R] [BoundedSub R] [ContinuousSub R] (f g : BoundedContinuousFunction α R) {x : α} :
    (f - g) x = f x - g x
    @[simp]
    theorem BoundedContinuousFunction.coe_sub {α : Type u} [TopologicalSpace α] {R : Type u_2} [PseudoMetricSpace R] [Sub R] [BoundedSub R] [ContinuousSub R] (f g : BoundedContinuousFunction α R) :
    ⇑(f - g) = f - g
    @[simp]
    theorem BoundedContinuousFunction.natCast_apply {α : Type u} [TopologicalSpace α] {β : Type u_2} [PseudoMetricSpace β] [NatCast β] (n : ) (x : α) :
    n x = n
    @[simp]
    theorem BoundedContinuousFunction.intCast_apply {α : Type u} [TopologicalSpace α] {β : Type u_2} [PseudoMetricSpace β] [IntCast β] (m : ) (x : α) :
    m x = m

    IsBoundedSMul (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 IsBoundedSMul 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.

    instance BoundedContinuousFunction.instSMul {α : Type u} {β : Type v} {𝕜 : Type u_2} [PseudoMetricSpace 𝕜] [TopologicalSpace α] [PseudoMetricSpace β] [Zero 𝕜] [Zero β] [SMul 𝕜 β] [IsBoundedSMul 𝕜 β] :
    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 𝕜 β] [IsBoundedSMul 𝕜 β] (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 𝕜 β] [IsBoundedSMul 𝕜 β] (c : 𝕜) (f : BoundedContinuousFunction α β) (x : α) :
    (c f) x = c f x
    instance BoundedContinuousFunction.instIsScalarTower {α : Type u} {β : Type v} {𝕜 : Type u_2} [PseudoMetricSpace 𝕜] [TopologicalSpace α] [PseudoMetricSpace β] [Zero 𝕜] [Zero β] [SMul 𝕜 β] [IsBoundedSMul 𝕜 β] {𝕜' : Type u_3} [PseudoMetricSpace 𝕜'] [Zero 𝕜'] [SMul 𝕜' β] [IsBoundedSMul 𝕜' β] [SMul 𝕜' 𝕜] [IsScalarTower 𝕜' 𝕜 β] :
    instance BoundedContinuousFunction.instSMulCommClass {α : Type u} {β : Type v} {𝕜 : Type u_2} [PseudoMetricSpace 𝕜] [TopologicalSpace α] [PseudoMetricSpace β] [Zero 𝕜] [Zero β] [SMul 𝕜 β] [IsBoundedSMul 𝕜 β] {𝕜' : Type u_3} [PseudoMetricSpace 𝕜'] [Zero 𝕜'] [SMul 𝕜' β] [IsBoundedSMul 𝕜' β] [SMulCommClass 𝕜' 𝕜 β] :
    Equations
    def BoundedContinuousFunction.evalCLM {α : Type u} {β : Type v} (𝕜 : Type u_2) [PseudoMetricSpace 𝕜] [TopologicalSpace α] [PseudoMetricSpace β] [Semiring 𝕜] [AddCommMonoid β] [Module 𝕜 β] [IsBoundedSMul 𝕜 β] [BoundedAdd β] [ContinuousAdd β] (x : α) :

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

    Equations
    @[simp]
    theorem BoundedContinuousFunction.evalCLM_apply {α : Type u} {β : Type v} (𝕜 : Type u_2) [PseudoMetricSpace 𝕜] [TopologicalSpace α] [PseudoMetricSpace β] [Semiring 𝕜] [AddCommMonoid β] [Module 𝕜 β] [IsBoundedSMul 𝕜 β] [BoundedAdd β] [ContinuousAdd β] (x : α) (f : BoundedContinuousFunction α β) :
    (evalCLM 𝕜 x) f = f x

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

    Equations