mathlib documentation

topology.​uniform_space.​uniform_convergence

topology.​uniform_space.​uniform_convergence

Uniform convergence

A sequence of functions Fₙ (with values in a metric space) converges uniformly on a set s to a function f if, for all ε > 0, for all large enough n, one has for all y ∈ s the inequality dist (f y, Fₙ y) < ε. Under uniform convergence, many properties of the Fₙ pass to the limit, most notably continuity. We prove this in the file, defining the notion of uniform convergence in the more general setting of uniform spaces, and with respect to an arbitrary indexing set endowed with a filter (instead of just with at_top).

Main results

Let α be a topological space, β a uniform space, Fₙ and f be functions from αto β (where the index n belongs to an indexing type ι endowed with a filter p).

We also define notions where the convergence is locally uniform, called tendsto_locally_uniformly_on F f p s and tendsto_locally_uniformly F f p. The previous theorems all have corresponding versions under locally uniform convergence.

Implementation notes

Most results hold under weaker assumptions of locally uniform approximation. In a first section, we prove the results under these weaker assumptions. Then, we derive the results on uniform convergence from them.

Tags

Uniform limit, uniform convergence, tends uniformly to

Different notions of uniform convergence

We define uniform convergence and locally uniform convergence, on a set or in the whole space.

def tendsto_uniformly_on {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] :
(ι → α → β)(α → β)filter ιset α → Prop

A sequence of functions Fₙ converges uniformly on a set s to a limiting function f with respect to the filter p if, for any entourage of the diagonal u, one has p-eventually (f x, Fₙ x) ∈ u for all x ∈ s.

Equations
def tendsto_uniformly {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] :
(ι → α → β)(α → β)filter ι → Prop

A sequence of functions Fₙ converges uniformly to a limiting function f with respect to a filter p if, for any entourage of the diagonal u, one has p-eventually (f x, Fₙ x) ∈ u for all x.

Equations
theorem tendsto_uniformly_on_univ {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {p : filter ι} :

theorem tendsto_uniformly_on.​mono {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s : set α} {p : filter ι} {s' : set α} :
tendsto_uniformly_on F f p ss' stendsto_uniformly_on F f p s'

theorem tendsto_uniformly.​tendsto_uniformly_on {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s : set α} {p : filter ι} :

theorem tendsto_uniformly_on.​comp {α : Type u} {β : Type v} {γ : Type w} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s : set α} {p : filter ι} (h : tendsto_uniformly_on F f p s) (g : γ → α) :
tendsto_uniformly_on (λ (n : ι), F n g) (f g) p (g ⁻¹' s)

Composing on the right by a function preserves uniform convergence on a set

theorem tendsto_uniformly.​comp {α : Type u} {β : Type v} {γ : Type w} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {p : filter ι} (h : tendsto_uniformly F f p) (g : γ → α) :
tendsto_uniformly (λ (n : ι), F n g) (f g) p

Composing on the right by a function preserves uniform convergence

def tendsto_locally_uniformly_on {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] [topological_space α] :
(ι → α → β)(α → β)filter ιset α → Prop

A sequence of functions Fₙ converges locally uniformly on a set s to a limiting function f with respect to a filter p if, for any entourage of the diagonal u, for any x ∈ s, one has p-eventually (f x, Fₙ x) ∈ u for all y in a neighborhood of x in s.

Equations
def tendsto_locally_uniformly {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] [topological_space α] :
(ι → α → β)(α → β)filter ι → Prop

A sequence of functions Fₙ converges locally uniformly to a limiting function f with respect to a filter p if, for any entourage of the diagonal u, for any x, one has p-eventually (f x, Fₙ x) ∈ u for all y in a neighborhood of x.

Equations
theorem tendsto_uniformly_on.​tendsto_locally_uniformly_on {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s : set α} {p : filter ι} [topological_space α] :

theorem tendsto_uniformly.​tendsto_locally_uniformly {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {p : filter ι} [topological_space α] :

theorem tendsto_locally_uniformly_on.​mono {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s s' : set α} {p : filter ι} [topological_space α] :

theorem tendsto_locally_uniformly_on_univ {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {p : filter ι} [topological_space α] :

theorem tendsto_locally_uniformly_on.​comp {α : Type u} {β : Type v} {γ : Type w} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s : set α} {p : filter ι} [topological_space α] [topological_space γ] {t : set γ} (h : tendsto_locally_uniformly_on F f p s) (g : γ → α) :
set.maps_to g t scontinuous_on g ttendsto_locally_uniformly_on (λ (n : ι), F n g) (f g) p t

theorem tendsto_locally_uniformly.​comp {α : Type u} {β : Type v} {γ : Type w} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {p : filter ι} [topological_space α] [topological_space γ] (h : tendsto_locally_uniformly F f p) (g : γ → α) :
continuous gtendsto_locally_uniformly (λ (n : ι), F n g) (f g) p

Uniform approximation

In this section, we give lemmas ensuring that a function is continuous if it can be approximated uniformly by continuous functions. We give various versions, within a set or the whole space, at a single point or at all points, with locally uniform approximation or uniform approximation. All the statements are derived from a statement about locally uniform approximation within a set at a point, called continuous_within_at_of_locally_uniform_approx_of_continuous_within_at.

theorem continuous_within_at_of_locally_uniform_approx_of_continuous_within_at {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s : set α} {x : α} [topological_space α] :
x s(∀ (u : set × β)), u uniformity β(∃ (t : set α) (H : t nhds_within x s) (n : ι), ∀ (y : α), y t(f y, F n y) u))(∀ (n : ι), continuous_within_at (F n) s x)continuous_within_at f s x

A function which can be locally uniformly approximated by functions which are continuous within a set at a point is continuous within this set at this point.

theorem continuous_at_of_locally_uniform_approx_of_continuous_at {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {x : α} [topological_space α] :
(∀ (u : set × β)), u uniformity β(∃ (t : set α) (H : t nhds x) (n : ι), ∀ (y : α), y t(f y, F n y) u))(∀ (n : ι), continuous_at (F n) x)continuous_at f x

A function which can be locally uniformly approximated by functions which are continuous at a point is continuous at this point.

theorem continuous_on_of_locally_uniform_approx_of_continuous_on {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s : set α} [topological_space α] :
(∀ (x : α), x s∀ (u : set × β)), u uniformity β(∃ (t : set α) (H : t nhds_within x s) (n : ι), ∀ (y : α), y t(f y, F n y) u))(∀ (n : ι), continuous_on (F n) s)continuous_on f s

A function which can be locally uniformly approximated by functions which are continuous on a set is continuous on this set.

theorem continuous_on_of_uniform_approx_of_continuous_on {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s : set α} [topological_space α] :
(∀ (u : set × β)), u uniformity β(∃ (n : ι), ∀ (y : α), y s(f y, F n y) u))(∀ (n : ι), continuous_on (F n) s)continuous_on f s

A function which can be uniformly approximated by functions which are continuous on a set is continuous on this set.

theorem continuous_of_locally_uniform_approx_of_continuous {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} [topological_space α] :
(∀ (x : α) (u : set × β)), u uniformity β(∃ (t : set α) (H : t nhds x) (n : ι), ∀ (y : α), y t(f y, F n y) u))(∀ (n : ι), continuous (F n))continuous f

A function which can be locally uniformly approximated by continuous functions is continuous.

theorem continuous_of_uniform_approx_of_continuous {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} [topological_space α] :
(∀ (u : set × β)), u uniformity β(∃ (N : ι), ∀ (y : α), (f y, F N y) u))(∀ (n : ι), continuous (F n))continuous f

A function which can be uniformly approximated by continuous functions is continuous.

Uniform limits

From the previous statements on uniform approximation, we deduce continuity results for uniform limits.

theorem tendsto_locally_uniformly_on.​continuous_on {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s : set α} {p : filter ι} [topological_space α] (h : tendsto_locally_uniformly_on F f p s) (hc : ∀ (n : ι), continuous_on (F n) s) [p.ne_bot] :

A locally uniform limit on a set of functions which are continuous on this set is itself continuous on this set.

theorem tendsto_uniformly_on.​continuous_on {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s : set α} {p : filter ι} [topological_space α] (h : tendsto_uniformly_on F f p s) (hc : ∀ (n : ι), continuous_on (F n) s) [p.ne_bot] :

A uniform limit on a set of functions which are continuous on this set is itself continuous on this set.

theorem tendsto_locally_uniformly.​continuous {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {p : filter ι} [topological_space α] (h : tendsto_locally_uniformly F f p) (hc : ∀ (n : ι), continuous (F n)) [p.ne_bot] :

A locally uniform limit of continuous functions is continuous.

theorem tendsto_uniformly.​continuous {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {p : filter ι} [topological_space α] (h : tendsto_uniformly F f p) (hc : ∀ (n : ι), continuous (F n)) [p.ne_bot] :

A uniform limit of continuous functions is continuous.

Composing limits under uniform convergence

In general, if Fₙ converges pointwise to a function f, and gₙ tends to x, it is not true that Fₙ gₙ tends to f x. It is true however if the convergence of Fₙ to f is uniform. In this paragraph, we prove variations around this statement.

theorem tendsto_comp_of_locally_uniform_limit_within {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s : set α} {x : α} {p : filter ι} {g : ι → α} [topological_space α] :
continuous_within_at f s xfilter.tendsto g p (nhds_within x s)(∀ (u : set × β)), u uniformity β(∃ (t : set α) (H : t nhds_within x s), ∀ᶠ (n : ι) in p, ∀ (y : α), y t(f y, F n y) u))filter.tendsto (λ (n : ι), F n (g n)) p (nhds (f x))

If Fₙ converges locally uniformly on a neighborhood of x within a set s to a function f which is continuous at x within s, and gₙ tends to x within s, then Fₙ (gₙ) tends to f x.

theorem tendsto_comp_of_locally_uniform_limit {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {x : α} {p : filter ι} {g : ι → α} [topological_space α] :
continuous_at f xfilter.tendsto g p (nhds x)(∀ (u : set × β)), u uniformity β(∃ (t : set α) (H : t nhds x), ∀ᶠ (n : ι) in p, ∀ (y : α), y t(f y, F n y) u))filter.tendsto (λ (n : ι), F n (g n)) p (nhds (f x))

If Fₙ converges locally uniformly on a neighborhood of x to a function f which is continuous at x, and gₙ tends to x, then Fₙ (gₙ) tends to f x.

theorem tendsto_locally_uniformly_on.​tendsto_comp {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s : set α} {x : α} {p : filter ι} {g : ι → α} [topological_space α] :
tendsto_locally_uniformly_on F f p scontinuous_within_at f s xx sfilter.tendsto g p (nhds_within x s)filter.tendsto (λ (n : ι), F n (g n)) p (nhds (f x))

If Fₙ tends locally uniformly to f on a set s, and gₙ tends to x within s, then Fₙ gₙ tends to f x if f is continuous at x within s and x ∈ s.

theorem tendsto_uniformly_on.​tendsto_comp {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {s : set α} {x : α} {p : filter ι} {g : ι → α} [topological_space α] :
tendsto_uniformly_on F f p scontinuous_within_at f s xfilter.tendsto g p (nhds_within x s)filter.tendsto (λ (n : ι), F n (g n)) p (nhds (f x))

If Fₙ tends uniformly to f on a set s, and gₙ tends to x within s, then Fₙ gₙ tends to f x if f is continuous at x within s.

theorem tendsto_locally_uniformly.​tendsto_comp {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {x : α} {p : filter ι} {g : ι → α} [topological_space α] :
tendsto_locally_uniformly F f pcontinuous_at f xfilter.tendsto g p (nhds x)filter.tendsto (λ (n : ι), F n (g n)) p (nhds (f x))

If Fₙ tends locally uniformly to f, and gₙ tends to x, then Fₙ gₙ tends to f x.

theorem tendsto_uniformly.​tendsto_comp {α : Type u} {β : Type v} {ι : Type u_1} [uniform_space β] {F : ι → α → β} {f : α → β} {x : α} {p : filter ι} {g : ι → α} [topological_space α] :
tendsto_uniformly F f pcontinuous_at f xfilter.tendsto g p (nhds x)filter.tendsto (λ (n : ι), F n (g n)) p (nhds (f x))

If Fₙ tends uniformly to f, and gₙ tends to x, then Fₙ gₙ tends to f x.