mathlib documentation

topology.bounded_continuous_function

Bounded continuous functions

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

def bounded_continuous_function (α : Type u) (β : Type v) [topological_space α] [metric_space β] :
Type (max u v)

The type of bounded continuous functions from a topological space to a metric space

Equations
@[instance]

Equations
def bounded_continuous_function.mk_of_compact {α : Type u} {β : Type v} [topological_space α] [metric_space β] [compact_space α] (f : α → β) (hf : continuous f) :

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

Equations
def bounded_continuous_function.mk_of_discrete {α : Type u} {β : Type v} [topological_space α] [metric_space β] [discrete_topology α] (f : α → β) (hf : ∃ (C : ), ∀ (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
@[instance]

The uniform distance between two bounded continuous functions

Equations
theorem bounded_continuous_function.dist_eq {α : Type u} {β : Type v} [topological_space α] [metric_space β] {f g : bounded_continuous_function α β} :
dist f g = Inf {C : | 0 C ∀ (x : α), dist (f x) (g x) C}

theorem bounded_continuous_function.dist_set_exists {α : Type u} {β : Type v} [topological_space α] [metric_space β] {f g : bounded_continuous_function α β} :
∃ (C : ), 0 C ∀ (x : α), dist (f x) (g x) C

theorem bounded_continuous_function.dist_coe_le_dist {α : Type u} {β : Type v} [topological_space α] [metric_space β] {f g : bounded_continuous_function α β} (x : α) :
dist (f x) (g x) dist f g

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

@[ext]
theorem bounded_continuous_function.ext {α : Type u} {β : Type v} [topological_space α] [metric_space β] {f g : bounded_continuous_function α β} (H : ∀ (x : α), f x = g x) :
f = g

theorem bounded_continuous_function.ext_iff {α : Type u} {β : Type v} [topological_space α] [metric_space β] {f g : bounded_continuous_function α β} :
f = g ∀ (x : α), f x = g x

theorem bounded_continuous_function.dist_le {α : Type u} {β : Type v} [topological_space α] [metric_space β] {f g : bounded_continuous_function α β} {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 bounded_continuous_function.dist_zero_of_empty {α : Type u} {β : Type v} [topological_space α] [metric_space β] {f g : bounded_continuous_function α β} (e : ¬nonempty α) :
dist f g = 0

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

@[instance]

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

Equations
def bounded_continuous_function.const (α : Type u) {β : Type v} [topological_space α] [metric_space β] (b : β) :

Constant as a continuous bounded function.

Equations
@[simp]

theorem bounded_continuous_function.const_apply {α : Type u} {β : Type v} [topological_space α] [metric_space β] (a : α) (b : β) :

@[instance]

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

Equations
theorem bounded_continuous_function.continuous_eval {α : Type u} {β : Type v} [topological_space α] [metric_space β] :

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

theorem bounded_continuous_function.continuous_evalx {α : Type u} {β : Type v} [topological_space α] [metric_space β] {x : α} :

In particular, when x is fixed, f → f x is continuous

When f is fixed, x → f x is also continuous, by definition

@[instance]

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

def bounded_continuous_function.comp {α : Type u} {β : Type v} {γ : Type w} [topological_space α] [metric_space β] [metric_space γ] (G : β → γ) {C : ℝ≥0} (H : lipschitz_with C G) (f : bounded_continuous_function α β) :

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

Equations
theorem bounded_continuous_function.lipschitz_comp {α : Type u} {β : Type v} {γ : Type w} [topological_space α] [metric_space β] [metric_space γ] {G : β → γ} {C : ℝ≥0} (H : lipschitz_with C G) :

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

theorem bounded_continuous_function.uniform_continuous_comp {α : Type u} {β : Type v} {γ : Type w} [topological_space α] [metric_space β] [metric_space γ] {G : β → γ} {C : ℝ≥0} (H : lipschitz_with C G) :

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

theorem bounded_continuous_function.continuous_comp {α : Type u} {β : Type v} {γ : Type w} [topological_space α] [metric_space β] [metric_space γ] {G : β → γ} {C : ℝ≥0} (H : lipschitz_with C G) :

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

def bounded_continuous_function.cod_restrict {α : Type u} {β : Type v} [topological_space α] [metric_space β] (s : set β) (f : bounded_continuous_function α β) (H : ∀ (x : α), f x s) :

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

Equations
theorem bounded_continuous_function.arzela_ascoli₁ {α : Type u} {β : Type v} [topological_space α] [compact_space α] [metric_space β] [compact_space β] (A : set (bounded_continuous_function α β)) (closed : is_closed A) (H : ∀ (x : α) (ε : ), ε > 0(∃ (U : set α) (H : U 𝓝 x), ∀ (y z : α), y Uz U∀ (f : bounded_continuous_function α β), f Adist (f y) (f z) < ε)) :

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

theorem bounded_continuous_function.arzela_ascoli₂ {α : Type u} {β : Type v} [topological_space α] [compact_space α] [metric_space β] (s : set β) (hs : is_compact s) (A : set (bounded_continuous_function α β)) (closed : is_closed A) (in_s : ∀ (f : bounded_continuous_function α β) (x : α), f Af x s) (H : ∀ (x : α) (ε : ), ε > 0(∃ (U : set α) (H : U 𝓝 x), ∀ (y z : α), y Uz U∀ (f : bounded_continuous_function α β), f Adist (f y) (f z) < ε)) :

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

theorem bounded_continuous_function.arzela_ascoli {α : Type u} {β : Type v} [topological_space α] [compact_space α] [metric_space β] (s : set β) (hs : is_compact s) (A : set (bounded_continuous_function α β)) (in_s : ∀ (f : bounded_continuous_function α β) (x : α), f Af x s) (H : ∀ (x : α) (ε : ), ε > 0(∃ (U : set α) (H : U 𝓝 x), ∀ (y z : α), y Uz U∀ (f : bounded_continuous_function α β), f Adist (f y) (f z) < ε)) :

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

theorem bounded_continuous_function.equicontinuous_of_continuity_modulus {β : Type v} [metric_space β] {α : Type u} [metric_space α] (b : ) (b_lim : b→_{0}0) (A : set (bounded_continuous_function α β)) (H : ∀ (x y : α) (f : bounded_continuous_function α β), f Adist (f x) (f y) b (dist x y)) (x : α) (ε : ) (ε0 : 0 < ε) :
∃ (U : set α) (H : U 𝓝 x), ∀ (y z : α), y Uz U∀ (f : bounded_continuous_function α β), f Adist (f y) (f z) < ε

@[simp]
theorem bounded_continuous_function.coe_zero {α : Type u} {β : Type v} [topological_space α] [normed_group β] {x : α} :
0 x = 0

theorem bounded_continuous_function.norm_def {α : Type u} {β : Type v} [topological_space α] [normed_group β] (f : bounded_continuous_function α β) :

theorem bounded_continuous_function.norm_eq {α : Type u} {β : Type v} [topological_space α] [normed_group β] (f : bounded_continuous_function α β) :
f = Inf {C : | 0 C ∀ (x : α), f x C}

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

theorem bounded_continuous_function.norm_coe_le_norm {α : Type u} {β : Type v} [topological_space α] [normed_group β] (f : bounded_continuous_function α β) (x : α) :

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

theorem bounded_continuous_function.dist_le_two_norm {α : Type u} {β : Type v} [topological_space α] [normed_group β] (f : bounded_continuous_function α β) (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 bounded_continuous_function.norm_le {α : Type u} {β : Type v} [topological_space α] [normed_group β] {f : bounded_continuous_function α β} {C : } (C0 : 0 C) :
f C ∀ (x : α), f x C

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

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

@[simp]

def bounded_continuous_function.of_normed_group {α : Type u} {β : Type v} [topological_space α] [normed_group β] (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
theorem bounded_continuous_function.norm_of_normed_group_le {α : Type u} {β : Type v} [topological_space α] [normed_group β] {f : α → β} (hfc : continuous f) {C : } (hC : 0 C) (hfC : ∀ (x : α), f x C) :

def bounded_continuous_function.of_normed_group_discrete {α : Type u} {β : Type v} [topological_space α] [discrete_topology α] [normed_group β] (f : α → β) (C : ) (H : ∀ (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
@[instance]

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

Equations
@[instance]

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

Equations
@[simp]
theorem bounded_continuous_function.coe_add {α : Type u} {β : Type v} [topological_space α] [normed_group β] (f g : bounded_continuous_function α β) :
(f + g) = λ (x : α), f x + g x

theorem bounded_continuous_function.add_apply {α : Type u} {β : Type v} [topological_space α] [normed_group β] (f g : bounded_continuous_function α β) {x : α} :
(f + g) x = f x + g x

@[simp]
theorem bounded_continuous_function.coe_neg {α : Type u} {β : Type v} [topological_space α] [normed_group β] (f : bounded_continuous_function α β) :
-f = λ (x : α), -f x

theorem bounded_continuous_function.neg_apply {α : Type u} {β : Type v} [topological_space α] [normed_group β] (f : bounded_continuous_function α β) {x : α} :
(-f) x = -f x

theorem bounded_continuous_function.forall_coe_zero_iff_zero {α : Type u} {β : Type v} [topological_space α] [normed_group β] (f : bounded_continuous_function α β) :
(∀ (x : α), f x = 0) f = 0

@[simp]
theorem bounded_continuous_function.coe_sub {α : Type u} {β : Type v} [topological_space α] [normed_group β] (f g : bounded_continuous_function α β) :
(f - g) = λ (x : α), f x - g x

theorem bounded_continuous_function.sub_apply {α : Type u} {β : Type v} [topological_space α] [normed_group β] (f g : bounded_continuous_function α β) {x : α} :
(f - g) x = f x - g x

theorem bounded_continuous_function.abs_diff_coe_le_dist {α : Type u} {β : Type v} [topological_space α] [normed_group β] (f g : bounded_continuous_function α β) {x : α} :
f x - g x dist f g

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.

@[instance]
def bounded_continuous_function.has_scalar {α : Type u} {β : Type v} {𝕜 : Type u_1} [normed_field 𝕜] [topological_space α] [normed_group β] [normed_space 𝕜 β] :

Equations
@[simp]
theorem bounded_continuous_function.coe_smul {α : Type u} {β : Type v} {𝕜 : Type u_1} [normed_field 𝕜] [topological_space α] [normed_group β] [normed_space 𝕜 β] (c : 𝕜) (f : bounded_continuous_function α β) :
(c f) = λ (x : α), c f x

theorem bounded_continuous_function.smul_apply {α : Type u} {β : Type v} {𝕜 : Type u_1} [normed_field 𝕜] [topological_space α] [normed_group β] [normed_space 𝕜 β] (c : 𝕜) (f : bounded_continuous_function α β) (x : α) :
(c f) x = c f x

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.

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 bounded_continuous_function.C {α : Type u} {γ : Type w} {𝕜 : Type u_1} [normed_field 𝕜] [topological_space α] [normed_ring γ] [normed_algebra 𝕜 γ] :

bounded_continuous_function.const as a ring_hom.

Equations

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 𝕜.

theorem bounded_continuous_function.norm_smul_le {α : Type u} {β : Type v} {𝕜 : Type u_1} [normed_field 𝕜] [topological_space α] [normed_group β] [normed_space 𝕜 β] (f : bounded_continuous_function α 𝕜) (g : bounded_continuous_function α β) :