mathlib documentation

topology.metric_space.basic

Metric spaces #

This file defines metric spaces. Many definitions and theorems expected on metric spaces are already introduced on uniform spaces and topological spaces. For example: open and closed sets, compactness, completeness, continuity and uniform continuity

Main definitions #

Additional useful definitions:

TODO (anyone): Add "Main results" section.

Implementation notes #

Since a lot of elementary properties don't require eq_of_dist_eq_zero we start setting up the theory of pseudo_metric_space, where we don't require dist x y = 0 → x = y and we specialize to metric_space at the end.

Tags #

metric, pseudo_metric, dist

def uniform_space.core_of_dist {α : Type u_1} (dist : α → α → ) (dist_self : ∀ (x : α), dist x x = 0) (dist_comm : ∀ (x y : α), dist x y = dist y x) (dist_triangle : ∀ (x y z : α), dist x z dist x y + dist y z) :

Construct a uniform structure core from a distance function and metric space axioms. This is a technical construction that can be immediately used to construct a uniform structure from a distance function and metric space axioms but is also useful when discussing metrizable topologies, see pseudo_metric_space.of_metrizable.

Equations
def uniform_space_of_dist {α : Type u} (dist : α → α → ) (dist_self : ∀ (x : α), dist x x = 0) (dist_comm : ∀ (x y : α), dist x y = dist y x) (dist_triangle : ∀ (x y z : α), dist x z dist x y + dist y z) :

Construct a uniform structure from a distance function and metric space axioms

Equations
def bornology.of_dist {α : Type u_1} (dist : α → α → ) (dist_self : ∀ (x : α), dist x x = 0) (dist_comm : ∀ (x y : α), dist x y = dist y x) (dist_triangle : ∀ (x y z : α), dist x z dist x y + dist y z) :

Construct a bornology from a distance function and metric space axioms.

Equations
@[ext, class]
structure has_dist (α : Type u_2) :
Type u_2
  • dist : α → α →

The distance function (given an ambient metric space on α), which returns a nonnegative real number dist x y given x y : α.

Instances of this typeclass
Instances of other typeclasses for has_dist
  • has_dist.has_sizeof_inst
theorem has_dist.ext_iff {α : Type u_2} (x y : has_dist α) :
theorem has_dist.ext {α : Type u_2} (x y : has_dist α) (h : has_dist.dist = has_dist.dist) :
x = y
@[protected]

This tactic is used to populate pseudo_metric_space.edist_dist when the default edist is used.

@[class]
structure pseudo_metric_space (α : Type u) :
Type u

Metric space

Each metric space induces a canonical uniform_space and hence a canonical topological_space. This is enforced in the type class definition, by extending the uniform_space structure. When instantiating a metric_space structure, the uniformity fields are not necessary, they will be filled in by default. In the same way, each metric space induces an emetric space structure. It is included in the structure, but filled in by default.

Instances of this typeclass
Instances of other typeclasses for pseudo_metric_space
  • pseudo_metric_space.has_sizeof_inst
@[ext]

Two pseudo metric space structures with the same distance function coincide.

def pseudo_metric_space.of_metrizable {α : Type u_1} [topological_space α] (dist : α → α → ) (dist_self : ∀ (x : α), dist x x = 0) (dist_comm : ∀ (x y : α), dist x y = dist y x) (dist_triangle : ∀ (x y z : α), dist x z dist x y + dist y z) (H : ∀ (s : set α), is_open s ∀ (x : α), x s(∃ (ε : ) (H : ε > 0), ∀ (y : α), dist x y < εy s)) :

Construct a pseudo-metric space structure whose underlying topological space structure (definitionally) agrees which a pre-existing topology which is compatible with a given distance function.

Equations
@[simp]
theorem dist_self {α : Type u} [pseudo_metric_space α] (x : α) :
theorem dist_comm {α : Type u} [pseudo_metric_space α] (x y : α) :
theorem edist_dist {α : Type u} [pseudo_metric_space α] (x y : α) :
theorem dist_triangle {α : Type u} [pseudo_metric_space α] (x y z : α) :
theorem dist_triangle_left {α : Type u} [pseudo_metric_space α] (x y z : α) :
theorem dist_triangle_right {α : Type u} [pseudo_metric_space α] (x y z : α) :
theorem dist_triangle4 {α : Type u} [pseudo_metric_space α] (x y z w : α) :
theorem dist_triangle4_left {α : Type u} [pseudo_metric_space α] (x₁ y₁ x₂ y₂ : α) :
has_dist.dist x₂ y₂ has_dist.dist x₁ y₁ + (has_dist.dist x₁ x₂ + has_dist.dist y₁ y₂)
theorem dist_triangle4_right {α : Type u} [pseudo_metric_space α] (x₁ y₁ x₂ y₂ : α) :
has_dist.dist x₁ y₁ has_dist.dist x₁ x₂ + has_dist.dist y₁ y₂ + has_dist.dist x₂ y₂
theorem dist_le_Ico_sum_dist {α : Type u} [pseudo_metric_space α] (f : → α) {m n : } (h : m n) :
has_dist.dist (f m) (f n) (finset.Ico m n).sum (λ (i : ), has_dist.dist (f i) (f (i + 1)))

The triangle (polygon) inequality for sequences of points; finset.Ico version.

theorem dist_le_range_sum_dist {α : Type u} [pseudo_metric_space α] (f : → α) (n : ) :
has_dist.dist (f 0) (f n) (finset.range n).sum (λ (i : ), has_dist.dist (f i) (f (i + 1)))

The triangle (polygon) inequality for sequences of points; finset.range version.

theorem dist_le_Ico_sum_of_dist_le {α : Type u} [pseudo_metric_space α] {f : → α} {m n : } (hmn : m n) {d : } (hd : ∀ {k : }, m kk < nhas_dist.dist (f k) (f (k + 1)) d k) :
has_dist.dist (f m) (f n) (finset.Ico m n).sum (λ (i : ), d i)

A version of dist_le_Ico_sum_dist with each intermediate distance replaced with an upper estimate.

theorem dist_le_range_sum_of_dist_le {α : Type u} [pseudo_metric_space α] {f : → α} (n : ) {d : } (hd : ∀ {k : }, k < nhas_dist.dist (f k) (f (k + 1)) d k) :
has_dist.dist (f 0) (f n) (finset.range n).sum (λ (i : ), d i)

A version of dist_le_range_sum_dist with each intermediate distance replaced with an upper estimate.

theorem abs_dist_sub_le {α : Type u} [pseudo_metric_space α] (x y z : α) :
theorem dist_nonneg {α : Type u} [pseudo_metric_space α] {x y : α} :

Extension for the positivity tactic: distances are nonnegative.

@[simp]
theorem abs_dist {α : Type u} [pseudo_metric_space α] {a b : α} :
@[class]
structure has_nndist (α : Type u_2) :
Type u_2
  • nndist : α → α → nnreal

A version of has_dist that takes value in ℝ≥0.

Instances of this typeclass
Instances of other typeclasses for has_nndist
  • has_nndist.has_sizeof_inst
@[protected, instance]

Distance as a nonnegative real number.

Equations
theorem nndist_edist {α : Type u} [pseudo_metric_space α] (x y : α) :

Express nndist in terms of edist

theorem edist_nndist {α : Type u} [pseudo_metric_space α] (x y : α) :

Express edist in terms of nndist

@[simp, norm_cast]
theorem coe_nnreal_ennreal_nndist {α : Type u} [pseudo_metric_space α] (x y : α) :
@[simp, norm_cast]
theorem edist_lt_coe {α : Type u} [pseudo_metric_space α] {x y : α} {c : nnreal} :
@[simp, norm_cast]
theorem edist_le_coe {α : Type u} [pseudo_metric_space α] {x y : α} {c : nnreal} :
theorem edist_lt_top {α : Type u_1} [pseudo_metric_space α] (x y : α) :

In a pseudometric space, the extended distance is always finite

theorem edist_ne_top {α : Type u} [pseudo_metric_space α] (x y : α) :

In a pseudometric space, the extended distance is always finite

@[simp]
theorem nndist_self {α : Type u} [pseudo_metric_space α] (a : α) :

nndist x x vanishes

theorem dist_nndist {α : Type u} [pseudo_metric_space α] (x y : α) :

Express dist in terms of nndist

@[simp, norm_cast]
theorem coe_nndist {α : Type u} [pseudo_metric_space α] (x y : α) :
@[simp, norm_cast]
theorem dist_lt_coe {α : Type u} [pseudo_metric_space α] {x y : α} {c : nnreal} :
@[simp, norm_cast]
theorem dist_le_coe {α : Type u} [pseudo_metric_space α] {x y : α} {c : nnreal} :
@[simp]
theorem edist_lt_of_real {α : Type u} [pseudo_metric_space α] {x y : α} {r : } :
@[simp]
theorem edist_le_of_real {α : Type u} [pseudo_metric_space α] {x y : α} {r : } (hr : 0 r) :
theorem nndist_dist {α : Type u} [pseudo_metric_space α] (x y : α) :

Express nndist in terms of dist

theorem nndist_comm {α : Type u} [pseudo_metric_space α] (x y : α) :
theorem nndist_triangle {α : Type u} [pseudo_metric_space α] (x y z : α) :

Triangle inequality for the nonnegative distance

theorem dist_edist {α : Type u} [pseudo_metric_space α] (x y : α) :

Express dist in terms of edist

@[simp]
theorem metric.mem_ball {α : Type u} [pseudo_metric_space α] {x y : α} {ε : } :
theorem metric.mem_ball' {α : Type u} [pseudo_metric_space α] {x y : α} {ε : } :
theorem metric.pos_of_mem_ball {α : Type u} [pseudo_metric_space α] {x y : α} {ε : } (hy : y metric.ball x ε) :
0 < ε
theorem metric.mem_ball_self {α : Type u} [pseudo_metric_space α] {x : α} {ε : } (h : 0 < ε) :
@[simp]
theorem metric.nonempty_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
@[simp]
theorem metric.ball_eq_empty {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
metric.ball x ε = ε 0
@[simp]
theorem metric.ball_zero {α : Type u} [pseudo_metric_space α] {x : α} :
theorem metric.exists_lt_mem_ball_of_mem_ball {α : Type u} [pseudo_metric_space α] {x y : α} {ε : } (h : x metric.ball y ε) :
∃ (ε' : ) (H : ε' < ε), x metric.ball y ε'

If a point belongs to an open ball, then there is a strictly smaller radius whose ball also contains it.

See also exists_lt_subset_ball.

theorem metric.ball_eq_ball {α : Type u} [pseudo_metric_space α] (ε : ) (x : α) :
uniform_space.ball x {p : α × α | has_dist.dist p.snd p.fst < ε} = metric.ball x ε
theorem metric.ball_eq_ball' {α : Type u} [pseudo_metric_space α] (ε : ) (x : α) :
uniform_space.ball x {p : α × α | has_dist.dist p.fst p.snd < ε} = metric.ball x ε
@[simp]
theorem metric.Union_ball_nat {α : Type u} [pseudo_metric_space α] (x : α) :
(⋃ (n : ), metric.ball x n) = set.univ
@[simp]
theorem metric.Union_ball_nat_succ {α : Type u} [pseudo_metric_space α] (x : α) :
(⋃ (n : ), metric.ball x (n + 1)) = set.univ
@[simp]
theorem metric.mem_closed_ball {α : Type u} [pseudo_metric_space α] {x y : α} {ε : } :
theorem metric.mem_closed_ball' {α : Type u} [pseudo_metric_space α] {x y : α} {ε : } :
@[simp]
theorem metric.mem_sphere {α : Type u} [pseudo_metric_space α] {x y : α} {ε : } :
theorem metric.mem_sphere' {α : Type u} [pseudo_metric_space α] {x y : α} {ε : } :
theorem metric.ne_of_mem_sphere {α : Type u} [pseudo_metric_space α] {x y : α} {ε : } (h : y metric.sphere x ε) (hε : ε 0) :
y x
theorem metric.sphere_eq_empty_of_subsingleton {α : Type u} [pseudo_metric_space α] {x : α} {ε : } [subsingleton α] (hε : ε 0) :
theorem metric.sphere_is_empty_of_subsingleton {α : Type u} [pseudo_metric_space α] {x : α} {ε : } [subsingleton α] (hε : ε 0) :
theorem metric.mem_closed_ball_self {α : Type u} [pseudo_metric_space α] {x : α} {ε : } (h : 0 ε) :
@[simp]
theorem metric.nonempty_closed_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
@[simp]
theorem metric.closed_ball_eq_empty {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
theorem metric.ball_subset_closed_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
theorem metric.sphere_subset_closed_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
theorem metric.closed_ball_disjoint_ball {α : Type u} [pseudo_metric_space α] {x y : α} {δ ε : } (h : δ + ε has_dist.dist x y) :
theorem metric.ball_disjoint_closed_ball {α : Type u} [pseudo_metric_space α] {x y : α} {δ ε : } (h : δ + ε has_dist.dist x y) :
theorem metric.ball_disjoint_ball {α : Type u} [pseudo_metric_space α] {x y : α} {δ ε : } (h : δ + ε has_dist.dist x y) :
theorem metric.closed_ball_disjoint_closed_ball {α : Type u} [pseudo_metric_space α] {x y : α} {δ ε : } (h : δ + ε < has_dist.dist x y) :
theorem metric.sphere_disjoint_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
@[simp]
theorem metric.ball_union_sphere {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
@[simp]
theorem metric.sphere_union_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
@[simp]
theorem metric.closed_ball_diff_sphere {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
@[simp]
theorem metric.closed_ball_diff_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
theorem metric.mem_ball_comm {α : Type u} [pseudo_metric_space α] {x y : α} {ε : } :
theorem metric.mem_closed_ball_comm {α : Type u} [pseudo_metric_space α] {x y : α} {ε : } :
theorem metric.mem_sphere_comm {α : Type u} [pseudo_metric_space α] {x y : α} {ε : } :
theorem metric.ball_subset_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε₁ ε₂ : } (h : ε₁ ε₂) :
metric.ball x ε₁ metric.ball x ε₂
theorem metric.ball_subset_ball' {α : Type u} [pseudo_metric_space α] {x y : α} {ε₁ ε₂ : } (h : ε₁ + has_dist.dist x y ε₂) :
metric.ball x ε₁ metric.ball y ε₂
theorem metric.closed_ball_subset_closed_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε₁ ε₂ : } (h : ε₁ ε₂) :
theorem metric.closed_ball_subset_closed_ball' {α : Type u} [pseudo_metric_space α] {x y : α} {ε₁ ε₂ : } (h : ε₁ + has_dist.dist x y ε₂) :
theorem metric.closed_ball_subset_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε₁ ε₂ : } (h : ε₁ < ε₂) :
theorem metric.dist_le_add_of_nonempty_closed_ball_inter_closed_ball {α : Type u} [pseudo_metric_space α] {x y : α} {ε₁ ε₂ : } (h : (metric.closed_ball x ε₁ metric.closed_ball y ε₂).nonempty) :
has_dist.dist x y ε₁ + ε₂
theorem metric.dist_lt_add_of_nonempty_closed_ball_inter_ball {α : Type u} [pseudo_metric_space α] {x y : α} {ε₁ ε₂ : } (h : (metric.closed_ball x ε₁ metric.ball y ε₂).nonempty) :
has_dist.dist x y < ε₁ + ε₂
theorem metric.dist_lt_add_of_nonempty_ball_inter_closed_ball {α : Type u} [pseudo_metric_space α] {x y : α} {ε₁ ε₂ : } (h : (metric.ball x ε₁ metric.closed_ball y ε₂).nonempty) :
has_dist.dist x y < ε₁ + ε₂
theorem metric.dist_lt_add_of_nonempty_ball_inter_ball {α : Type u} [pseudo_metric_space α] {x y : α} {ε₁ ε₂ : } (h : (metric.ball x ε₁ metric.ball y ε₂).nonempty) :
has_dist.dist x y < ε₁ + ε₂
@[simp]
theorem metric.Union_closed_ball_nat {α : Type u} [pseudo_metric_space α] (x : α) :
theorem metric.Union_inter_closed_ball_nat {α : Type u} [pseudo_metric_space α] (s : set α) (x : α) :
(⋃ (n : ), s metric.closed_ball x n) = s
theorem metric.ball_subset {α : Type u} [pseudo_metric_space α] {x y : α} {ε₁ ε₂ : } (h : has_dist.dist x y ε₂ - ε₁) :
metric.ball x ε₁ metric.ball y ε₂
theorem metric.ball_half_subset {α : Type u} [pseudo_metric_space α] {x : α} {ε : } (y : α) (h : y metric.ball x / 2)) :
theorem metric.exists_ball_subset_ball {α : Type u} [pseudo_metric_space α] {x y : α} {ε : } (h : y metric.ball x ε) :
∃ (ε' : ) (H : ε' > 0), metric.ball y ε' metric.ball x ε
theorem metric.forall_of_forall_mem_closed_ball {α : Type u} [pseudo_metric_space α] (p : α → Prop) (x : α) (H : ∃ᶠ (R : ) in filter.at_top, ∀ (y : α), y metric.closed_ball x Rp y) (y : α) :
p y

If a property holds for all points in closed balls of arbitrarily large radii, then it holds for all points.

theorem metric.forall_of_forall_mem_ball {α : Type u} [pseudo_metric_space α] (p : α → Prop) (x : α) (H : ∃ᶠ (R : ) in filter.at_top, ∀ (y : α), y metric.ball x Rp y) (y : α) :
p y

If a property holds for all points in balls of arbitrarily large radii, then it holds for all points.

theorem metric.is_bounded_iff {α : Type u} [pseudo_metric_space α] {s : set α} :
bornology.is_bounded s ∃ (C : ), ∀ ⦃x : α⦄, x s∀ ⦃y : α⦄, y shas_dist.dist x y C
theorem metric.is_bounded_iff_eventually {α : Type u} [pseudo_metric_space α] {s : set α} :
bornology.is_bounded s ∀ᶠ (C : ) in filter.at_top, ∀ ⦃x : α⦄, x s∀ ⦃y : α⦄, y shas_dist.dist x y C
theorem metric.is_bounded_iff_exists_ge {α : Type u} [pseudo_metric_space α] {s : set α} (c : ) :
bornology.is_bounded s ∃ (C : ), c C ∀ ⦃x : α⦄, x s∀ ⦃y : α⦄, y shas_dist.dist x y C
theorem metric.is_bounded_iff_nndist {α : Type u} [pseudo_metric_space α] {s : set α} :
bornology.is_bounded s ∃ (C : nnreal), ∀ ⦃x : α⦄, x s∀ ⦃y : α⦄, y shas_nndist.nndist x y C
theorem metric.uniformity_basis_dist {α : Type u} [pseudo_metric_space α] :
(uniformity α).has_basis (λ (ε : ), 0 < ε) (λ (ε : ), {p : α × α | has_dist.dist p.fst p.snd < ε})
@[protected]
theorem metric.mk_uniformity_basis {α : Type u} [pseudo_metric_space α] {β : Type u_1} {p : β → Prop} {f : β → } (hf₀ : ∀ (i : β), p i0 < f i) (hf : ∀ ⦃ε : ⦄, 0 < ε(∃ (i : β) (hi : p i), f i ε)) :
(uniformity α).has_basis p (λ (i : β), {p : α × α | has_dist.dist p.fst p.snd < f i})

Given f : β → ℝ, if f sends {i | p i} to a set of positive numbers accumulating to zero, then f i-neighborhoods of the diagonal form a basis of 𝓤 α.

For specific bases see uniformity_basis_dist, uniformity_basis_dist_inv_nat_succ, and uniformity_basis_dist_inv_nat_pos.

theorem metric.uniformity_basis_dist_inv_nat_succ {α : Type u} [pseudo_metric_space α] :
(uniformity α).has_basis (λ (_x : ), true) (λ (n : ), {p : α × α | has_dist.dist p.fst p.snd < 1 / (n + 1)})
theorem metric.uniformity_basis_dist_inv_nat_pos {α : Type u} [pseudo_metric_space α] :
(uniformity α).has_basis (λ (n : ), 0 < n) (λ (n : ), {p : α × α | has_dist.dist p.fst p.snd < 1 / n})
theorem metric.uniformity_basis_dist_pow {α : Type u} [pseudo_metric_space α] {r : } (h0 : 0 < r) (h1 : r < 1) :
(uniformity α).has_basis (λ (n : ), true) (λ (n : ), {p : α × α | has_dist.dist p.fst p.snd < r ^ n})
theorem metric.uniformity_basis_dist_lt {α : Type u} [pseudo_metric_space α] {R : } (hR : 0 < R) :
(uniformity α).has_basis (λ (r : ), 0 < r r < R) (λ (r : ), {p : α × α | has_dist.dist p.fst p.snd < r})
@[protected]
theorem metric.mk_uniformity_basis_le {α : Type u} [pseudo_metric_space α] {β : Type u_1} {p : β → Prop} {f : β → } (hf₀ : ∀ (x : β), p x0 < f x) (hf : ∀ (ε : ), 0 < ε(∃ (x : β) (hx : p x), f x ε)) :
(uniformity α).has_basis p (λ (x : β), {p : α × α | has_dist.dist p.fst p.snd f x})

Given f : β → ℝ, if f sends {i | p i} to a set of positive numbers accumulating to zero, then closed neighborhoods of the diagonal of sizes {f i | p i} form a basis of 𝓤 α.

Currently we have only one specific basis uniformity_basis_dist_le based on this constructor. More can be easily added if needed in the future.

theorem metric.uniformity_basis_dist_le {α : Type u} [pseudo_metric_space α] :
(uniformity α).has_basis (λ (ε : ), 0 < ε) (λ (ε : ), {p : α × α | has_dist.dist p.fst p.snd ε})

Contant size closed neighborhoods of the diagonal form a basis of the uniformity filter.

theorem metric.uniformity_basis_dist_le_pow {α : Type u} [pseudo_metric_space α] {r : } (h0 : 0 < r) (h1 : r < 1) :
(uniformity α).has_basis (λ (n : ), true) (λ (n : ), {p : α × α | has_dist.dist p.fst p.snd r ^ n})
theorem metric.mem_uniformity_dist {α : Type u} [pseudo_metric_space α] {s : set × α)} :
s uniformity α ∃ (ε : ) (H : ε > 0), ∀ {a b : α}, has_dist.dist a b < ε(a, b) s
theorem metric.dist_mem_uniformity {α : Type u} [pseudo_metric_space α] {ε : } (ε0 : 0 < ε) :
{p : α × α | has_dist.dist p.fst p.snd < ε} uniformity α

A constant size neighborhood of the diagonal is an entourage.

theorem metric.uniform_continuous_iff {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} :
uniform_continuous f ∀ (ε : ), ε > 0(∃ (δ : ) (H : δ > 0), ∀ {a b : α}, has_dist.dist a b < δhas_dist.dist (f a) (f b) < ε)
theorem metric.uniform_continuous_on_iff {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} {s : set α} :
uniform_continuous_on f s ∀ (ε : ), ε > 0(∃ (δ : ) (H : δ > 0), ∀ (x : α), x s∀ (y : α), y shas_dist.dist x y < δhas_dist.dist (f x) (f y) < ε)
theorem metric.uniform_continuous_on_iff_le {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} {s : set α} :
uniform_continuous_on f s ∀ (ε : ), ε > 0(∃ (δ : ) (H : δ > 0), ∀ (x : α), x s∀ (y : α), y shas_dist.dist x y δhas_dist.dist (f x) (f y) ε)
theorem metric.uniform_embedding_iff {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} :
uniform_embedding f function.injective f uniform_continuous f ∀ (δ : ), δ > 0(∃ (ε : ) (H : ε > 0), ∀ {a b : α}, has_dist.dist (f a) (f b) < εhas_dist.dist a b < δ)
theorem metric.controlled_of_uniform_embedding {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} :
uniform_embedding f((∀ (ε : ), ε > 0(∃ (δ : ) (H : δ > 0), ∀ {a b : α}, has_dist.dist a b < δhas_dist.dist (f a) (f b) < ε)) ∀ (δ : ), δ > 0(∃ (ε : ) (H : ε > 0), ∀ {a b : α}, has_dist.dist (f a) (f b) < εhas_dist.dist a b < δ))

If a map between pseudometric spaces is a uniform embedding then the distance between f x and f y is controlled in terms of the distance between x and y.

theorem metric.totally_bounded_iff {α : Type u} [pseudo_metric_space α] {s : set α} :
totally_bounded s ∀ (ε : ), ε > 0(∃ (t : set α), t.finite s ⋃ (y : α) (H : y t), metric.ball y ε)
theorem metric.totally_bounded_of_finite_discretization {α : Type u} [pseudo_metric_space α] {s : set α} (H : ∀ (ε : ), ε > 0(∃ (β : Type u) (_x : fintype β) (F : s → β), ∀ (x y : s), F x = F yhas_dist.dist x y < ε)) :

A pseudometric space is totally bounded if one can reconstruct up to any ε>0 any element of the space from finitely many data.

theorem metric.finite_approx_of_totally_bounded {α : Type u} [pseudo_metric_space α] {s : set α} (hs : totally_bounded s) (ε : ) (H : ε > 0) :
∃ (t : set α) (H : t s), t.finite s ⋃ (y : α) (H : y t), metric.ball y ε
theorem metric.tendsto_locally_uniformly_on_iff {α : Type u} {β : Type v} [pseudo_metric_space α] {ι : Type u_1} [topological_space β] {F : ι → β → α} {f : β → α} {p : filter ι} {s : set β} :
tendsto_locally_uniformly_on F f p s ∀ (ε : ), ε > 0∀ (x : β), x s(∃ (t : set β) (H : t nhds_within x s), ∀ᶠ (n : ι) in p, ∀ (y : β), y thas_dist.dist (f y) (F n y) < ε)

Expressing locally uniform convergence on a set using dist.

theorem metric.tendsto_uniformly_on_iff {α : Type u} {β : Type v} [pseudo_metric_space α] {ι : Type u_1} {F : ι → β → α} {f : β → α} {p : filter ι} {s : set β} :
tendsto_uniformly_on F f p s ∀ (ε : ), ε > 0(∀ᶠ (n : ι) in p, ∀ (x : β), x shas_dist.dist (f x) (F n x) < ε)

Expressing uniform convergence on a set using dist.

theorem metric.tendsto_locally_uniformly_iff {α : Type u} {β : Type v} [pseudo_metric_space α] {ι : Type u_1} [topological_space β] {F : ι → β → α} {f : β → α} {p : filter ι} :
tendsto_locally_uniformly F f p ∀ (ε : ), ε > 0∀ (x : β), ∃ (t : set β) (H : t nhds x), ∀ᶠ (n : ι) in p, ∀ (y : β), y thas_dist.dist (f y) (F n y) < ε

Expressing locally uniform convergence using dist.

theorem metric.tendsto_uniformly_iff {α : Type u} {β : Type v} [pseudo_metric_space α] {ι : Type u_1} {F : ι → β → α} {f : β → α} {p : filter ι} :
tendsto_uniformly F f p ∀ (ε : ), ε > 0(∀ᶠ (n : ι) in p, ∀ (x : β), has_dist.dist (f x) (F n x) < ε)

Expressing uniform convergence using dist.

@[protected]
theorem metric.cauchy_iff {α : Type u} [pseudo_metric_space α] {f : filter α} :
cauchy f f.ne_bot ∀ (ε : ), ε > 0(∃ (t : set α) (H : t f), ∀ (x : α), x t∀ (y : α), y thas_dist.dist x y < ε)
theorem metric.nhds_basis_ball {α : Type u} [pseudo_metric_space α] {x : α} :
(nhds x).has_basis (λ (ε : ), 0 < ε) (metric.ball x)
theorem metric.mem_nhds_iff {α : Type u} [pseudo_metric_space α] {x : α} {s : set α} :
s nhds x ∃ (ε : ) (H : ε > 0), metric.ball x ε s
theorem metric.eventually_nhds_iff {α : Type u} [pseudo_metric_space α] {x : α} {p : α → Prop} :
(∀ᶠ (y : α) in nhds x, p y) ∃ (ε : ) (H : ε > 0), ∀ ⦃y : α⦄, has_dist.dist y x < εp y
theorem metric.eventually_nhds_iff_ball {α : Type u} [pseudo_metric_space α] {x : α} {p : α → Prop} :
(∀ᶠ (y : α) in nhds x, p y) ∃ (ε : ) (H : ε > 0), ∀ (y : α), y metric.ball x εp y
theorem metric.nhds_basis_closed_ball {α : Type u} [pseudo_metric_space α] {x : α} :
(nhds x).has_basis (λ (ε : ), 0 < ε) (metric.closed_ball x)
theorem metric.nhds_basis_ball_inv_nat_succ {α : Type u} [pseudo_metric_space α] {x : α} :
(nhds x).has_basis (λ (_x : ), true) (λ (n : ), metric.ball x (1 / (n + 1)))
theorem metric.nhds_basis_ball_inv_nat_pos {α : Type u} [pseudo_metric_space α] {x : α} :
(nhds x).has_basis (λ (n : ), 0 < n) (λ (n : ), metric.ball x (1 / n))
theorem metric.nhds_basis_ball_pow {α : Type u} [pseudo_metric_space α] {x : α} {r : } (h0 : 0 < r) (h1 : r < 1) :
(nhds x).has_basis (λ (n : ), true) (λ (n : ), metric.ball x (r ^ n))
theorem metric.nhds_basis_closed_ball_pow {α : Type u} [pseudo_metric_space α] {x : α} {r : } (h0 : 0 < r) (h1 : r < 1) :
(nhds x).has_basis (λ (n : ), true) (λ (n : ), metric.closed_ball x (r ^ n))
theorem metric.is_open_iff {α : Type u} [pseudo_metric_space α] {s : set α} :
is_open s ∀ (x : α), x s(∃ (ε : ) (H : ε > 0), metric.ball x ε s)
theorem metric.is_open_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
theorem metric.ball_mem_nhds {α : Type u} [pseudo_metric_space α] (x : α) {ε : } (ε0 : 0 < ε) :
theorem metric.closed_ball_mem_nhds {α : Type u} [pseudo_metric_space α] (x : α) {ε : } (ε0 : 0 < ε) :
theorem metric.closed_ball_mem_nhds_of_mem {α : Type u} [pseudo_metric_space α] {x c : α} {ε : } (h : x metric.ball c ε) :
theorem metric.nhds_within_basis_ball {α : Type u} [pseudo_metric_space α] {x : α} {s : set α} :
(nhds_within x s).has_basis (λ (ε : ), 0 < ε) (λ (ε : ), metric.ball x ε s)
theorem metric.mem_nhds_within_iff {α : Type u} [pseudo_metric_space α] {x : α} {s t : set α} :
s nhds_within x t ∃ (ε : ) (H : ε > 0), metric.ball x ε t s
theorem metric.tendsto_nhds_within_nhds_within {α : Type u} {β : Type v} [pseudo_metric_space α] {s : set α} [pseudo_metric_space β] {t : set β} {f : α → β} {a : α} {b : β} :
filter.tendsto f (nhds_within a s) (nhds_within b t) ∀ (ε : ), ε > 0(∃ (δ : ) (H : δ > 0), ∀ {x : α}, x shas_dist.dist x a < δf x t has_dist.dist (f x) b < ε)
theorem metric.tendsto_nhds_within_nhds {α : Type u} {β : Type v} [pseudo_metric_space α] {s : set α} [pseudo_metric_space β] {f : α → β} {a : α} {b : β} :
filter.tendsto f (nhds_within a s) (nhds b) ∀ (ε : ), ε > 0(∃ (δ : ) (H : δ > 0), ∀ {x : α}, x shas_dist.dist x a < δhas_dist.dist (f x) b < ε)
theorem metric.tendsto_nhds_nhds {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} {a : α} {b : β} :
filter.tendsto f (nhds a) (nhds b) ∀ (ε : ), ε > 0(∃ (δ : ) (H : δ > 0), ∀ {x : α}, has_dist.dist x a < δhas_dist.dist (f x) b < ε)
theorem metric.continuous_at_iff {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} {a : α} :
continuous_at f a ∀ (ε : ), ε > 0(∃ (δ : ) (H : δ > 0), ∀ {x : α}, has_dist.dist x a < δhas_dist.dist (f x) (f a) < ε)
theorem metric.continuous_within_at_iff {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} {a : α} {s : set α} :
continuous_within_at f s a ∀ (ε : ), ε > 0(∃ (δ : ) (H : δ > 0), ∀ {x : α}, x shas_dist.dist x a < δhas_dist.dist (f x) (f a) < ε)
theorem metric.continuous_on_iff {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} {s : set α} :
continuous_on f s ∀ (b : α), b s∀ (ε : ), ε > 0(∃ (δ : ) (H : δ > 0), ∀ (a : α), a shas_dist.dist a b < δhas_dist.dist (f a) (f b) < ε)
theorem metric.continuous_iff {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} :
continuous f ∀ (b : α) (ε : ), ε > 0(∃ (δ : ) (H : δ > 0), ∀ (a : α), has_dist.dist a b < δhas_dist.dist (f a) (f b) < ε)
theorem metric.tendsto_nhds {α : Type u} {β : Type v} [pseudo_metric_space α] {f : filter β} {u : β → α} {a : α} :
filter.tendsto u f (nhds a) ∀ (ε : ), ε > 0(∀ᶠ (x : β) in f, has_dist.dist (u x) a < ε)
theorem metric.continuous_at_iff' {α : Type u} {β : Type v} [pseudo_metric_space α] [topological_space β] {f : β → α} {b : β} :
continuous_at f b ∀ (ε : ), ε > 0(∀ᶠ (x : β) in nhds b, has_dist.dist (f x) (f b) < ε)
theorem metric.continuous_within_at_iff' {α : Type u} {β : Type v} [pseudo_metric_space α] [topological_space β] {f : β → α} {b : β} {s : set β} :
continuous_within_at f s b ∀ (ε : ), ε > 0(∀ᶠ (x : β) in nhds_within b s, has_dist.dist (f x) (f b) < ε)
theorem metric.continuous_on_iff' {α : Type u} {β : Type v} [pseudo_metric_space α] [topological_space β] {f : β → α} {s : set β} :
continuous_on f s ∀ (b : β), b s∀ (ε : ), ε > 0(∀ᶠ (x : β) in nhds_within b s, has_dist.dist (f x) (f b) < ε)
theorem metric.continuous_iff' {α : Type u} {β : Type v} [pseudo_metric_space α] [topological_space β] {f : β → α} :
continuous f ∀ (a : β) (ε : ), ε > 0(∀ᶠ (x : β) in nhds a, has_dist.dist (f x) (f a) < ε)
theorem metric.tendsto_at_top {α : Type u} {β : Type v} [pseudo_metric_space α] [nonempty β] [semilattice_sup β] {u : β → α} {a : α} :
filter.tendsto u filter.at_top (nhds a) ∀ (ε : ), ε > 0(∃ (N : β), ∀ (n : β), n Nhas_dist.dist (u n) a < ε)
theorem metric.tendsto_at_top' {α : Type u} {β : Type v} [pseudo_metric_space α] [nonempty β] [semilattice_sup β] [no_max_order β] {u : β → α} {a : α} :
filter.tendsto u filter.at_top (nhds a) ∀ (ε : ), ε > 0(∃ (N : β), ∀ (n : β), n > Nhas_dist.dist (u n) a < ε)

A variant of tendsto_at_top that uses ∃ N, ∀ n > N, ... rather than ∃ N, ∀ n ≥ N, ...

theorem metric.is_open_singleton_iff {α : Type u_1} [pseudo_metric_space α] {x : α} :
is_open {x} ∃ (ε : ) (H : ε > 0), ∀ (y : α), has_dist.dist y x < εy = x
theorem metric.exists_ball_inter_eq_singleton_of_mem_discrete {α : Type u} [pseudo_metric_space α] {s : set α} [discrete_topology s] {x : α} (hx : x s) :
∃ (ε : ) (H : ε > 0), metric.ball x ε s = {x}

Given a point x in a discrete subset s of a pseudometric space, there is an open ball centered at x and intersecting s only at x.

theorem metric.exists_closed_ball_inter_eq_singleton_of_discrete {α : Type u} [pseudo_metric_space α] {s : set α} [discrete_topology s] {x : α} (hx : x s) :
∃ (ε : ) (H : ε > 0), metric.closed_ball x ε s = {x}

Given a point x in a discrete subset s of a pseudometric space, there is a closed ball of positive radius centered at x and intersecting s only at x.

theorem dense.exists_dist_lt {α : Type u} [pseudo_metric_space α] {s : set α} (hs : dense s) (x : α) {ε : } (hε : 0 < ε) :
∃ (y : α) (H : y s), has_dist.dist x y < ε
theorem dense_range.exists_dist_lt {α : Type u} [pseudo_metric_space α] {β : Type u_1} {f : β → α} (hf : dense_range f) (x : α) {ε : } (hε : 0 < ε) :
∃ (y : β), has_dist.dist x (f y) < ε
@[protected]
theorem pseudo_metric.uniformity_basis_edist {α : Type u} [pseudo_metric_space α] :
(uniformity α).has_basis (λ (ε : ennreal), 0 < ε) (λ (ε : ennreal), {p : α × α | has_edist.edist p.fst p.snd < ε})

Expressing the uniformity in terms of edist

theorem metric.uniformity_edist {α : Type u} [pseudo_metric_space α] :
uniformity α = ⨅ (ε : ennreal) (H : ε > 0), filter.principal {p : α × α | has_edist.edist p.fst p.snd < ε}
theorem metric.eball_top_eq_univ {α : Type u} [pseudo_metric_space α] (x : α) :

In a pseudometric space, an open ball of infinite radius is the whole space

@[simp]
theorem metric.emetric_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :

Balls defined using the distance or the edistance coincide

@[simp]
theorem metric.emetric_ball_nnreal {α : Type u} [pseudo_metric_space α] {x : α} {ε : nnreal} :

Balls defined using the distance or the edistance coincide

theorem metric.emetric_closed_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε : } (h : 0 ε) :

Closed balls defined using the distance or the edistance coincide

@[simp]

Closed balls defined using the distance or the edistance coincide

@[simp]
theorem metric.emetric_ball_top {α : Type u} [pseudo_metric_space α] (x : α) :
theorem metric.inseparable_iff {α : Type u} [pseudo_metric_space α] {x y : α} :

Build a new pseudometric space from an old one where the bundled uniform structure is provably (but typically non-definitionaly) equal to some given uniform structure. See Note [forgetful inheritance].

Equations
@[reducible]

Build a new pseudo metric space from an old one where the bundled topological structure is provably (but typically non-definitionaly) equal to some given topological structure. See Note [forgetful inheritance].

Equations
def pseudo_emetric_space.to_pseudo_metric_space_of_dist {α : Type u} [e : pseudo_emetric_space α] (dist : α → α → ) (edist_ne_top : ∀ (x y : α), has_edist.edist x y ) (h : ∀ (x y : α), dist x y = (has_edist.edist x y).to_real) :

One gets a pseudometric space from an emetric space if the edistance is everywhere finite, by pushing the edistance to reals. We set it up so that the edist and the uniformity are defeq in the pseudometric space and the pseudoemetric space. In this definition, the distance is given separately, to be able to prescribe some expression which is not defeq to the push-forward of the edistance to reals.

Equations

One gets a pseudometric space from an emetric space if the edistance is everywhere finite, by pushing the edistance to reals. We set it up so that the edist and the uniformity are defeq in the pseudometric space and the emetric space.

Equations

Build a new pseudometric space from an old one where the bundled bornology structure is provably (but typically non-definitionaly) equal to some given bornology structure. See Note [forgetful inheritance].

Equations
theorem metric.complete_of_convergent_controlled_sequences {α : Type u} [pseudo_metric_space α] (B : ) (hB : ∀ (n : ), 0 < B n) (H : ∀ (u : → α), (∀ (N n m : ), N nN mhas_dist.dist (u n) (u m) < B N)(∃ (x : α), filter.tendsto u filter.at_top (nhds x))) :

A very useful criterion to show that a space is complete is to show that all sequences which satisfy a bound of the form dist (u n) (u m) < B N for all n m ≥ N are converging. This is often applied for B N = 2^{-N}, i.e., with a very fast convergence to 0, which makes it possible to use arguments of converging series, while this is impossible to do in general for arbitrary Cauchy sequences.

theorem metric.complete_of_cauchy_seq_tendsto {α : Type u} [pseudo_metric_space α] :
(∀ (u : → α), cauchy_seq u(∃ (a : α), filter.tendsto u filter.at_top (nhds a)))complete_space α
@[protected, instance]

Instantiate the reals as a pseudometric space.

Equations
theorem real.dist_eq (x y : ) :
theorem real.nndist_eq (x y : ) :
theorem real.dist_0_eq_abs (x : ) :
theorem real.dist_le_of_mem_interval {x y x' y' : } (hx : x set.interval x' y') (hy : y set.interval x' y') :
theorem real.dist_le_of_mem_Icc {x y x' y' : } (hx : x set.Icc x' y') (hy : y set.Icc x' y') :
has_dist.dist x y y' - x'
theorem real.dist_le_of_mem_Icc_01 {x y : } (hx : x set.Icc 0 1) (hy : y set.Icc 0 1) :
@[protected, instance]
theorem real.ball_eq_Ioo (x r : ) :
metric.ball x r = set.Ioo (x - r) (x + r)
theorem real.closed_ball_eq_Icc {x r : } :
metric.closed_ball x r = set.Icc (x - r) (x + r)
theorem real.Ioo_eq_ball (x y : ) :
set.Ioo x y = metric.ball ((x + y) / 2) ((y - x) / 2)
theorem real.Icc_eq_closed_ball (x y : ) :
set.Icc x y = metric.closed_ball ((x + y) / 2) ((y - x) / 2)
theorem totally_bounded_Icc {α : Type u} [pseudo_metric_space α] [preorder α] [compact_Icc_space α] (a b : α) :
theorem totally_bounded_Ico {α : Type u} [pseudo_metric_space α] [preorder α] [compact_Icc_space α] (a b : α) :
theorem totally_bounded_Ioc {α : Type u} [pseudo_metric_space α] [preorder α] [compact_Icc_space α] (a b : α) :
theorem totally_bounded_Ioo {α : Type u} [pseudo_metric_space α] [preorder α] [compact_Icc_space α] (a b : α) :
theorem squeeze_zero' {α : Type u_1} {f g : α → } {t₀ : filter α} (hf : ∀ᶠ (t : α) in t₀, 0 f t) (hft : ∀ᶠ (t : α) in t₀, f t g t) (g0 : filter.tendsto g t₀ (nhds 0)) :
filter.tendsto f t₀ (nhds 0)

Special case of the sandwich theorem; see tendsto_of_tendsto_of_tendsto_of_le_of_le' for the general case.

theorem squeeze_zero {α : Type u_1} {f g : α → } {t₀ : filter α} (hf : ∀ (t : α), 0 f t) (hft : ∀ (t : α), f t g t) (g0 : filter.tendsto g t₀ (nhds 0)) :
filter.tendsto f t₀ (nhds 0)

Special case of the sandwich theorem; see tendsto_of_tendsto_of_tendsto_of_le_of_le and tendsto_of_tendsto_of_tendsto_of_le_of_le' for the general case.

theorem cauchy_seq_iff_tendsto_dist_at_top_0 {α : Type u} {β : Type v} [pseudo_metric_space α] [nonempty β] [semilattice_sup β] {u : β → α} :
cauchy_seq u filter.tendsto (λ (n : β × β), has_dist.dist (u n.fst) (u n.snd)) filter.at_top (nhds 0)
theorem tendsto_uniformity_iff_dist_tendsto_zero {α : Type u} [pseudo_metric_space α] {ι : Type u_1} {f : ι → α × α} {p : filter ι} :
filter.tendsto f p (uniformity α) filter.tendsto (λ (x : ι), has_dist.dist (f x).fst (f x).snd) p (nhds 0)
theorem filter.tendsto.congr_dist {α : Type u} [pseudo_metric_space α] {ι : Type u_1} {f₁ f₂ : ι → α} {p : filter ι} {a : α} (h₁ : filter.tendsto f₁ p (nhds a)) (h : filter.tendsto (λ (x : ι), has_dist.dist (f₁ x) (f₂ x)) p (nhds 0)) :
filter.tendsto f₂ p (nhds a)
theorem tendsto_of_tendsto_of_dist {α : Type u} [pseudo_metric_space α] {ι : Type u_1} {f₁ f₂ : ι → α} {p : filter ι} {a : α} (h₁ : filter.tendsto f₁ p (nhds a)) (h : filter.tendsto (λ (x : ι), has_dist.dist (f₁ x) (f₂ x)) p (nhds 0)) :
filter.tendsto f₂ p (nhds a)

Alias of filter.tendsto.congr_dist.

theorem tendsto_iff_of_dist {α : Type u} [pseudo_metric_space α] {ι : Type u_1} {f₁ f₂ : ι → α} {p : filter ι} {a : α} (h : filter.tendsto (λ (x : ι), has_dist.dist (f₁ x) (f₂ x)) p (nhds 0)) :
filter.tendsto f₁ p (nhds a) filter.tendsto f₂ p (nhds a)
theorem eventually_closed_ball_subset {α : Type u} [pseudo_metric_space α] {x : α} {u : set α} (hu : u nhds x) :
∀ᶠ (r : ) in nhds 0, metric.closed_ball x r u

If u is a neighborhood of x, then for small enough r, the closed ball closed_ball x r is contained in u.

@[nolint]
theorem metric.cauchy_seq_iff {α : Type u} {β : Type v} [pseudo_metric_space α] [nonempty β] [semilattice_sup β] {u : β → α} :
cauchy_seq u ∀ (ε : ), ε > 0(∃ (N : β), ∀ (m : β), m N∀ (n : β), n Nhas_dist.dist (u m) (u n) < ε)

In a pseudometric space, Cauchy sequences are characterized by the fact that, eventually, the distance between its elements is arbitrarily small

theorem metric.cauchy_seq_iff' {α : Type u} {β : Type v} [pseudo_metric_space α] [nonempty β] [semilattice_sup β] {u : β → α} :
cauchy_seq u ∀ (ε : ), ε > 0(∃ (N : β), ∀ (n : β), n Nhas_dist.dist (u n) (u N) < ε)

A variation around the pseudometric characterization of Cauchy sequences

@[nolint]
theorem metric.uniform_cauchy_seq_on_iff {α : Type u} {β : Type v} [pseudo_metric_space α] [nonempty β] [semilattice_sup β] {γ : Type u_1} {F : β → γ → α} {s : set γ} :
uniform_cauchy_seq_on F filter.at_top s ∀ (ε : ), ε > 0(∃ (N : β), ∀ (m : β), m N∀ (n : β), n N∀ (x : γ), x shas_dist.dist (F m x) (F n x) < ε)

In a pseudometric space, unifom Cauchy sequences are characterized by the fact that, eventually, the distance between all its elements is uniformly, arbitrarily small

theorem cauchy_seq_of_le_tendsto_0' {α : Type u} {β : Type v} [pseudo_metric_space α] [nonempty β] [semilattice_sup β] {s : β → α} (b : β → ) (h : ∀ (n m : β), n mhas_dist.dist (s n) (s m) b n) (h₀ : filter.tendsto b filter.at_top (nhds 0)) :

If the distance between s n and s m, n ≤ m is bounded above by b n and b converges to zero, then s is a Cauchy sequence.

theorem cauchy_seq_of_le_tendsto_0 {α : Type u} {β : Type v} [pseudo_metric_space α] [nonempty β] [semilattice_sup β] {s : β → α} (b : β → ) (h : ∀ (n m N : β), N nN mhas_dist.dist (s n) (s m) b N) (h₀ : filter.tendsto b filter.at_top (nhds 0)) :

If the distance between s n and s m, n, m ≥ N is bounded above by b N and b converges to zero, then s is a Cauchy sequence.

theorem cauchy_seq_bdd {α : Type u} [pseudo_metric_space α] {u : → α} (hu : cauchy_seq u) :
∃ (R : ) (H : R > 0), ∀ (m n : ), has_dist.dist (u m) (u n) < R

A Cauchy sequence on the natural numbers is bounded.

theorem cauchy_seq_iff_le_tendsto_0 {α : Type u} [pseudo_metric_space α] {s : → α} :
cauchy_seq s ∃ (b : ), (∀ (n : ), 0 b n) (∀ (n m N : ), N nN mhas_dist.dist (s n) (s m) b N) filter.tendsto b filter.at_top (nhds 0)

Yet another metric characterization of Cauchy sequences on integers. This one is often the most efficient.

def pseudo_metric_space.induced {α : Type u_1} {β : Type u_2} (f : α → β) (m : pseudo_metric_space β) :

Pseudometric space structure pulled back by a function.

Equations
def inducing.comap_pseudo_metric_space {α : Type u_1} {β : Type u_2} [topological_space α] [pseudo_metric_space β] {f : α → β} (hf : inducing f) :

Pull back a pseudometric space structure by an inducing map. This is a version of pseudo_metric_space.induced useful in case if the domain already has a topological_space structure.

Equations
def uniform_inducing.comap_pseudo_metric_space {α : Type u_1} {β : Type u_2} [uniform_space α] [pseudo_metric_space β] (f : α → β) (h : uniform_inducing f) :

Pull back a pseudometric space structure by a uniform inducing map. This is a version of pseudo_metric_space.induced useful in case if the domain already has a uniform_space structure.

Equations
@[protected, instance]
def subtype.pseudo_metric_space {α : Type u} [pseudo_metric_space α] {p : α → Prop} :
Equations
theorem subtype.dist_eq {α : Type u} [pseudo_metric_space α] {p : α → Prop} (x y : subtype p) :
theorem subtype.nndist_eq {α : Type u} [pseudo_metric_space α] {p : α → Prop} (x y : subtype p) :
@[simp]
@[simp]
theorem nnreal.dist_eq (a b : nnreal) :
theorem nnreal.nndist_eq (a b : nnreal) :
@[simp]
theorem ulift.dist_eq {β : Type v} [pseudo_metric_space β] (x y : ulift β) :
theorem ulift.nndist_eq {β : Type v} [pseudo_metric_space β] (x y : ulift β) :
@[simp]
theorem ulift.dist_up_up {β : Type v} [pseudo_metric_space β] (x y : β) :
@[simp]
theorem ulift.nndist_up_up {β : Type v} [pseudo_metric_space β] (x y : β) :
@[protected, instance]
noncomputable def prod.pseudo_metric_space_max {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] :
Equations
theorem prod.dist_eq {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {x y : α × β} :
@[simp]
theorem dist_prod_same_left {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {x : α} {y₁ y₂ : β} :
has_dist.dist (x, y₁) (x, y₂) = has_dist.dist y₁ y₂
@[simp]
theorem dist_prod_same_right {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {x₁ x₂ : α} {y : β} :
has_dist.dist (x₁, y) (x₂, y) = has_dist.dist x₁ x₂
theorem ball_prod_same {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] (x : α) (y : β) (r : ) :
theorem closed_ball_prod_same {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] (x : α) (y : β) (r : ) :
theorem uniform_continuous_dist {α : Type u} [pseudo_metric_space α] :
uniform_continuous (λ (p : α × α), has_dist.dist p.fst p.snd)
theorem uniform_continuous.dist {α : Type u} {β : Type v} [pseudo_metric_space α] [uniform_space β] {f g : β → α} (hf : uniform_continuous f) (hg : uniform_continuous g) :
uniform_continuous (λ (b : β), has_dist.dist (f b) (g b))
@[continuity]
theorem continuous_dist {α : Type u} [pseudo_metric_space α] :
continuous (λ (p : α × α), has_dist.dist p.fst p.snd)
@[continuity]
theorem continuous.dist {α : Type u} {β : Type v} [pseudo_metric_space α] [topological_space β] {f g : β → α} (hf : continuous f) (hg : continuous g) :
continuous (λ (b : β), has_dist.dist (f b) (g b))
theorem filter.tendsto.dist {α : Type u} {β : Type v} [pseudo_metric_space α] {f g : β → α} {x : filter β} {a b : α} (hf : filter.tendsto f x (nhds a)) (hg : filter.tendsto g x (nhds b)) :
filter.tendsto (λ (x : β), has_dist.dist (f x) (g x)) x (nhds (has_dist.dist a b))
theorem nhds_comap_dist {α : Type u} [pseudo_metric_space α] (a : α) :
filter.comap (λ (a' : α), has_dist.dist a' a) (nhds 0) = nhds a
theorem tendsto_iff_dist_tendsto_zero {α : Type u} {β : Type v} [pseudo_metric_space α] {f : β → α} {x : filter β} {a : α} :
filter.tendsto f x (nhds a) filter.tendsto (λ (b : β), has_dist.dist (f b) a) x (nhds 0)
theorem uniform_continuous.nndist {α : Type u} {β : Type v} [pseudo_metric_space α] [uniform_space β] {f g : β → α} (hf : uniform_continuous f) (hg : uniform_continuous g) :
uniform_continuous (λ (b : β), has_nndist.nndist (f b) (g b))
theorem continuous_nndist {α : Type u} [pseudo_metric_space α] :
continuous (λ (p : α × α), has_nndist.nndist p.fst p.snd)
theorem continuous.nndist {α : Type u} {β : Type v} [pseudo_metric_space α] [topological_space β] {f g : β → α} (hf : continuous f) (hg : continuous g) :
continuous (λ (b : β), has_nndist.nndist (f b) (g b))
theorem filter.tendsto.nndist {α : Type u} {β : Type v} [pseudo_metric_space α] {f g : β → α} {x : filter β} {a b : α} (hf : filter.tendsto f x (nhds a)) (hg : filter.tendsto g x (nhds b)) :
filter.tendsto (λ (x : β), has_nndist.nndist (f x) (g x)) x (nhds (has_nndist.nndist a b))
theorem metric.is_closed_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
theorem metric.is_closed_sphere {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
@[simp]
theorem metric.closure_closed_ball {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
theorem metric.frontier_ball_subset_sphere {α : Type u} [pseudo_metric_space α] {x : α} {ε : } :
theorem metric.mem_closure_iff {α : Type u} [pseudo_metric_space α] {s : set α} {a : α} :
a closure s ∀ (ε : ), ε > 0(∃ (b : α) (H : b s), has_dist.dist a b < ε)

ε-characterization of the closure in pseudometric spaces

theorem metric.mem_closure_range_iff {α : Type u} {β : Type v} [pseudo_metric_space α] {e : β → α} {a : α} :
a closure (set.range e) ∀ (ε : ), ε > 0(∃ (k : β), has_dist.dist a (e k) < ε)
theorem metric.mem_closure_range_iff_nat {α : Type u} {β : Type v} [pseudo_metric_space α] {e : β → α} {a : α} :
a closure (set.range e) ∀ (n : ), ∃ (k : β), has_dist.dist a (e k) < 1 / (n + 1)
theorem metric.mem_of_closed' {α : Type u} [pseudo_metric_space α] {s : set α} (hs : is_closed s) {a : α} :
a s ∀ (ε : ), ε > 0(∃ (b : α) (H : b s), has_dist.dist a b < ε)
theorem metric.closed_ball_zero' {α : Type u} [pseudo_metric_space α] (x : α) :
theorem metric.dense_iff {α : Type u} [pseudo_metric_space α] {s : set α} :
dense s ∀ (x : α) (r : ), r > 0(metric.ball x r s).nonempty
theorem metric.dense_range_iff {α : Type u} {β : Type v} [pseudo_metric_space α] {f : β → α} :
dense_range f ∀ (x : α) (r : ), r > 0(∃ (y : β), has_dist.dist x (f y) < r)

If a set s is separable, then the corresponding subtype is separable in a metric space. This is not obvious, as the countable set whose closure covers s does not need in general to be contained in s.

@[protected]
theorem inducing.is_separable_preimage {α : Type u} {β : Type v} [pseudo_metric_space α] {f : β → α} [topological_space β] (hf : inducing f) {s : set α} (hs : topological_space.is_separable s) :

The preimage of a separable set by an inducing map is separable.

@[protected]
theorem embedding.is_separable_preimage {α : Type u} {β : Type v} [pseudo_metric_space α] {f : β → α} [topological_space β] (hf : embedding f) {s : set α} (hs : topological_space.is_separable s) :
theorem continuous_on.is_separable_image {α : Type u} {β : Type v} [pseudo_metric_space α] [topological_space β] {f : α → β} {s : set α} (hf : continuous_on f s) (hs : topological_space.is_separable s) :

If a map is continuous on a separable set s, then the image of s is also separable.

@[protected, instance]
noncomputable def pseudo_metric_space_pi {β : Type v} {π : β → Type u_2} [fintype β] [Π (b : β), pseudo_metric_space (π b)] :
pseudo_metric_space (Π (b : β), π b)

A finite product of pseudometric spaces is a pseudometric space, with the sup distance.

Equations
theorem nndist_pi_def {β : Type v} {π : β → Type u_2} [fintype β] [Π (b : β), pseudo_metric_space (π b)] (f g : Π (b : β), π b) :
has_nndist.nndist f g = finset.univ.sup (λ (b : β), has_nndist.nndist (f b) (g b))
theorem dist_pi_def {β : Type v} {π : β → Type u_2} [fintype β] [Π (b : β), pseudo_metric_space (π b)] (f g : Π (b : β), π b) :
has_dist.dist f g = (finset.univ.sup (λ (b : β), has_nndist.nndist (f b) (g b)))
@[simp]
theorem dist_pi_const {α : Type u} {β : Type v} [pseudo_metric_space α] [fintype β] [nonempty β] (a b : α) :
has_dist.dist (λ (x : β), a) (λ (_x : β), b) = has_dist.dist a b
@[simp]
theorem nndist_pi_const {α : Type u} {β : Type v} [pseudo_metric_space α] [fintype β] [nonempty β] (a b : α) :
has_nndist.nndist (λ (x : β), a) (λ (_x : β), b) = has_nndist.nndist a b
theorem nndist_pi_le_iff {β : Type v} {π : β → Type u_2} [fintype β] [Π (b : β), pseudo_metric_space (π b)] {f g : Π (b : β), π b} {r : nnreal} :
has_nndist.nndist f g r ∀ (b : β), has_nndist.nndist (f b) (g b) r
theorem dist_pi_lt_iff {β : Type v} {π : β → Type u_2} [fintype β] [Π (b : β), pseudo_metric_space (π b)] {f g : Π (b : β), π b} {r : } (hr : 0 < r) :
has_dist.dist f g < r ∀ (b : β), has_dist.dist (f b) (g b) < r
theorem dist_pi_le_iff {β : Type v} {π : β → Type u_2} [fintype β] [Π (b : β), pseudo_metric_space (π b)] {f g : Π (b : β), π b} {r : } (hr : 0 r) :
has_dist.dist f g r ∀ (b : β), has_dist.dist (f b) (g b) r
theorem nndist_le_pi_nndist {β : Type v} {π : β → Type u_2} [fintype β] [Π (b : β), pseudo_metric_space (π b)] (f g : Π (b : β), π b) (b : β) :
theorem dist_le_pi_dist {β : Type v} {π : β → Type u_2} [fintype β] [Π (b : β), pseudo_metric_space (π b)] (f g : Π (b : β), π b) (b : β) :
theorem ball_pi {β : Type v} {π : β → Type u_2} [fintype β] [Π (b : β), pseudo_metric_space (π b)] (x : Π (b : β), π b) {r : } (hr : 0 < r) :
metric.ball x r = set.univ.pi (λ (b : β), metric.ball (x b) r)

An open ball in a product space is a product of open balls. See also metric.ball_pi' for a version assuming nonempty β instead of 0 < r.

theorem ball_pi' {β : Type v} {π : β → Type u_2} [fintype β] [Π (b : β), pseudo_metric_space (π b)] [nonempty β] (x : Π (b : β), π b) (r : ) :
metric.ball x r = set.univ.pi (λ (b : β), metric.ball (x b) r)

An open ball in a product space is a product of open balls. See also metric.ball_pi for a version assuming 0 < r instead of nonempty β.

theorem closed_ball_pi {β : Type v} {π : β → Type u_2} [fintype β] [Π (b : β), pseudo_metric_space (π b)] (x : Π (b : β), π b) {r : } (hr : 0 r) :

A closed ball in a product space is a product of closed balls. See also metric.closed_ball_pi' for a version assuming nonempty β instead of 0 ≤ r.

theorem closed_ball_pi' {β : Type v} {π : β → Type u_2} [fintype β] [Π (b : β), pseudo_metric_space (π b)] [nonempty β] (x : Π (b : β), π b) (r : ) :

A closed ball in a product space is a product of closed balls. See also metric.closed_ball_pi for a version assuming 0 ≤ r instead of nonempty β.

@[simp]
theorem fin.nndist_insert_nth_insert_nth {n : } {α : fin (n + 1)Type u_1} [Π (i : fin (n + 1)), pseudo_metric_space (α i)] (i : fin (n + 1)) (x y : α i) (f g : Π (j : fin n), α ((i.succ_above) j)) :
@[simp]
theorem fin.dist_insert_nth_insert_nth {n : } {α : fin (n + 1)Type u_1} [Π (i : fin (n + 1)), pseudo_metric_space (α i)] (i : fin (n + 1)) (x y : α i) (f g : Π (j : fin n), α ((i.succ_above) j)) :
theorem real.dist_le_of_mem_pi_Icc {β : Type v} [fintype β] {x y x' y' : β → } (hx : x set.Icc x' y') (hy : y set.Icc x' y') :
theorem finite_cover_balls_of_compact {α : Type u} [pseudo_metric_space α] {s : set α} (hs : is_compact s) {e : } (he : 0 < e) :
∃ (t : set α) (H : t s), t.finite s ⋃ (x : α) (H : x t), metric.ball x e

Any compact set in a pseudometric space can be covered by finitely many balls of a given positive radius

theorem is_compact.finite_cover_balls {α : Type u} [pseudo_metric_space α] {s : set α} (hs : is_compact s) {e : } (he : 0 < e) :
∃ (t : set α) (H : t s), t.finite s ⋃ (x : α) (H : x t), metric.ball x e

Alias of finite_cover_balls_of_compact.

@[class]
structure proper_space (α : Type u) [pseudo_metric_space α] :
Prop

A pseudometric space is proper if all closed balls are compact.

Instances of this typeclass
theorem is_compact_sphere {α : Type u_1} [pseudo_metric_space α] [proper_space α] (x : α) (r : ) :

In a proper pseudometric space, all spheres are compact.

@[protected, instance]
def metric.sphere.compact_space {α : Type u_1} [pseudo_metric_space α] [proper_space α] (x : α) (r : ) :

In a proper pseudometric space, any sphere is a compact_space when considered as a subtype.

@[protected, instance]

A proper pseudo metric space is sigma compact, and therefore second countable.

theorem proper_space_of_compact_closed_ball_of_le {α : Type u} [pseudo_metric_space α] (R : ) (h : ∀ (x : α) (r : ), R ris_compact (metric.closed_ball x r)) :

If all closed balls of large enough radius are compact, then the space is proper. Especially useful when the lower bound for the radius is 0.

@[protected, instance]
@[protected, instance]

A proper space is locally compact

@[protected, instance]

A proper space is complete

@[protected, instance]
def pi_proper_space {β : Type v} {π : β → Type u_1} [fintype β] [Π (b : β), pseudo_metric_space (π b)] [h : ∀ (b : β), proper_space (π b)] :
proper_space (Π (b : β), π b)

A finite product of proper spaces is proper.

theorem exists_pos_lt_subset_ball {α : Type u} [pseudo_metric_space α] [proper_space α] {x : α} {r : } {s : set α} (hr : 0 < r) (hs : is_closed s) (h : s metric.ball x r) :
∃ (r' : ) (H : r' set.Ioo 0 r), s metric.ball x r'

If a nonempty ball in a proper space includes a closed set s, then there exists a nonempty ball with the same center and a strictly smaller radius that includes s.

theorem exists_lt_subset_ball {α : Type u} [pseudo_metric_space α] [proper_space α] {x : α} {r : } {s : set α} (hs : is_closed s) (h : s metric.ball x r) :
∃ (r' : ) (H : r' < r), s metric.ball x r'

If a ball in a proper space includes a closed set s, then there exists a ball with the same center and a strictly smaller radius that includes s.

theorem metric.second_countable_of_almost_dense_set {α : Type u} [pseudo_metric_space α] (H : ∀ (ε : ), ε > 0(∃ (s : set α), s.countable ∀ (x : α), ∃ (y : α) (H : y s), has_dist.dist x y ε)) :

A pseudometric space is second countable if, for every ε > 0, there is a countable set which is ε-dense.

theorem lebesgue_number_lemma_of_metric {α : Type u} [pseudo_metric_space α] {s : set α} {ι : Sort u_1} {c : ι → set α} (hs : is_compact s) (hc₁ : ∀ (i : ι), is_open (c i)) (hc₂ : s ⋃ (i : ι), c i) :
∃ (δ : ) (H : δ > 0), ∀ (x : α), x s(∃ (i : ι), metric.ball x δ c i)
theorem lebesgue_number_lemma_of_metric_sUnion {α : Type u} [pseudo_metric_space α] {s : set α} {c : set (set α)} (hs : is_compact s) (hc₁ : ∀ (t : set α), t cis_open t) (hc₂ : s ⋃₀ c) :
∃ (δ : ) (H : δ > 0), ∀ (x : α), x s(∃ (t : set α) (H : t c), metric.ball x δ t)
def metric.bounded {α : Type u} [pseudo_metric_space α] (s : set α) :
Prop

Boundedness of a subset of a pseudometric space. We formulate the definition to work even in the empty space.

Equations
@[simp]
theorem metric.bounded_iff_mem_bounded {α : Type u} [pseudo_metric_space α] {s : set α} :
metric.bounded s ∀ (x : α), x smetric.bounded s
theorem metric.bounded.mono {α : Type u} [pseudo_metric_space α] {s t : set α} (incl : s t) :

Subsets of a bounded set are also bounded

theorem metric.bounded_closed_ball {α : Type u} [pseudo_metric_space α] {x : α} {r : } :

Closed balls are bounded

theorem metric.bounded_ball {α : Type u} [pseudo_metric_space α] {x : α} {r : } :

Open balls are bounded

theorem metric.bounded_sphere {α : Type u} [pseudo_metric_space α] {x : α} {r : } :

Spheres are bounded

theorem metric.bounded_iff_subset_ball {α : Type u} [pseudo_metric_space α] {s : set α} (c : α) :

Given a point, a bounded subset is included in some ball around this point

theorem metric.bounded.subset_ball {α : Type u} [pseudo_metric_space α] {s : set α} (h : metric.bounded s) (c : α) :
∃ (r : ), s metric.closed_ball c r
theorem metric.bounded.subset_ball_lt {α : Type u} [pseudo_metric_space α] {s : set α} (h : metric.bounded s) (a : ) (c : α) :
∃ (r : ), a < r s metric.closed_ball c r
@[simp]
theorem metric.bounded.union {α : Type u} [pseudo_metric_space α] {s t : set α} (hs : metric.bounded s) (ht : metric.bounded t) :

The union of two bounded sets is bounded.

@[simp]

The union of two sets is bounded iff each of the sets is bounded.

theorem metric.bounded_bUnion {α : Type u} {β : Type v} [pseudo_metric_space α] {I : set β} {s : β → set α} (H : I.finite) :
metric.bounded (⋃ (i : β) (H : i I), s i) ∀ (i : β), i Imetric.bounded (s i)

A finite union of bounded sets is bounded

@[protected]
theorem metric.bounded.prod {α : Type u} {β : Type v} [pseudo_metric_space α] [pseudo_metric_space β] {s : set α} {t : set β} (hs : metric.bounded s) (ht : metric.bounded t) :
theorem totally_bounded.bounded {α : Type u} [pseudo_metric_space α] {s : set α} (h : totally_bounded s) :

A totally bounded set is bounded

theorem is_compact.bounded {α : Type u} [pseudo_metric_space α] {s : set α} (h : is_compact s) :

A compact set is bounded

theorem metric.bounded_of_finite {α : Type u} [pseudo_metric_space α] {s : set α} (h : s.finite) :

A finite set is bounded

theorem set.finite.bounded {α : Type u} [pseudo_metric_space α] {s : set α} (h : s.finite) :

Alias of metric.bounded_of_finite.

theorem metric.bounded_singleton {α : Type u} [pseudo_metric_space α] {x : α} :

A singleton is bounded

theorem metric.bounded_range_iff {α : Type u} {β : Type v} [pseudo_metric_space α] {f : β → α} :
metric.bounded (set.range f) ∃ (C : ), ∀ (x y : β), has_dist.dist (f x) (f y) C

Characterization of the boundedness of the range of a function

theorem metric.bounded_range_of_cauchy_map_cofinite {α : Type u} {β : Type v} [pseudo_metric_space α] {f : β → α} (hf : cauchy (filter.map f filter.cofinite)) :
theorem cauchy_seq.bounded_range {α : Type u} [pseudo_metric_space α] {f : → α} (hf : cauchy_seq f) :
theorem metric.bounded_range_of_tendsto_cofinite {α : Type u} {β : Type v} [pseudo_metric_space α] {f : β → α} {a : α} (hf : filter.tendsto f filter.cofinite (nhds a)) :

In a compact space, all sets are bounded

theorem metric.bounded_range_of_tendsto {α : Type u_1} [pseudo_metric_space α] (u : → α) {x : α} (hu : filter.tendsto u filter.at_top (nhds x)) :
theorem metric.is_compact_of_is_closed_bounded {α : Type u} [pseudo_metric_space α] {s : set α} [proper_space α] (hc : is_closed s) (hb : metric.bounded s) :

The Heine–Borel theorem: In a proper space, a closed bounded set is compact.

The Heine–Borel theorem: In a proper space, the closure of a bounded set is compact.

The Heine–Borel theorem: In a proper Hausdorff space, a set is compact if and only if it is closed and bounded.

theorem metric.bounded_Icc {α : Type u} [pseudo_metric_space α] [preorder α] [compact_Icc_space α] (a b : α) :
theorem metric.bounded_Ico {α : Type u} [pseudo_metric_space α] [preorder α] [compact_Icc_space α] (a b : α) :
theorem metric.bounded_Ioc {α : Type u} [pseudo_metric_space α] [preorder α] [compact_Icc_space α] (a b : α) :
theorem metric.bounded_Ioo {α : Type u} [pseudo_metric_space α] [preorder α] [compact_Icc_space α] (a b : α) :
theorem metric.bounded_of_bdd_above_of_bdd_below {α : Type u} [pseudo_metric_space α] [preorder α] [compact_Icc_space α] {s : set α} (h₁ : bdd_above s) (h₂ : bdd_below s) :

In a pseudo metric space with a conditionally complete linear order such that the order and the metric structure give the same topology, any order-bounded set is metric-bounded.

noncomputable def metric.diam {α : Type u} [pseudo_metric_space α] (s : set α) :

The diameter of a set in a metric space. To get controllable behavior even when the diameter should be infinite, we express it in terms of the emetric.diameter

Equations
theorem metric.diam_nonneg {α : Type u} [pseudo_metric_space α] {s : set α} :

The diameter of a set is always nonnegative

theorem metric.diam_subsingleton {α : Type u} [pseudo_metric_space α] {s : set α} (hs : s.subsingleton) :
@[simp]
theorem metric.diam_empty {α : Type u} [pseudo_metric_space α] :

The empty set has zero diameter

@[simp]
theorem metric.diam_singleton {α : Type u} [pseudo_metric_space α] {x : α} :

A singleton has zero diameter

theorem metric.diam_pair {α : Type u} [pseudo_metric_space α] {x y : α} :
theorem metric.diam_triple {α : Type u} [pseudo_metric_space α] {x y z : α} :
theorem metric.ediam_le_of_forall_dist_le {α : Type u} [pseudo_metric_space α] {s : set α} {C : } (h : ∀ (x : α), x s∀ (y : α), y shas_dist.dist x y C) :

If the distance between any two points in a set is bounded by some constant C, then ennreal.of_real C bounds the emetric diameter of this set.

theorem metric.diam_le_of_forall_dist_le {α : Type u} [pseudo_metric_space α] {s : set α} {C : } (h₀ : 0 C) (h : ∀ (x : α), x s∀ (y : α), y shas_dist.dist x y C) :

If the distance between any two points in a set is bounded by some non-negative constant, this constant bounds the diameter.

theorem metric.diam_le_of_forall_dist_le_of_nonempty {α : Type u} [pseudo_metric_space α] {s : set α} (hs : s.nonempty) {C : } (h : ∀ (x : α), x s∀ (y : α), y shas_dist.dist x y C) :

If the distance between any two points in a nonempty set is bounded by some constant, this constant bounds the diameter.

theorem metric.dist_le_diam_of_mem' {α : Type u} [pseudo_metric_space α] {s : set α} {x y : α} (h : emetric.diam s ) (hx : x s) (hy : y s) :

The distance between two points in a set is controlled by the diameter of the set.

Characterize the boundedness of a set in terms of the finiteness of its emetric.diameter.

theorem metric.bounded.ediam_ne_top {α : Type u} [pseudo_metric_space α] {s : set α} (h : metric.bounded s) :
theorem metric.dist_le_diam_of_mem {α : Type u} [pseudo_metric_space α] {s : set α} {x y : α} (h : metric.bounded s) (hx : x s) (hy : y s) :

The distance between two points in a set is controlled by the diameter of the set.

theorem metric.ediam_of_unbounded {α : Type u} [pseudo_metric_space α] {s : set α} (h : ¬metric.bounded s) :
theorem metric.diam_eq_zero_of_unbounded {α : Type u} [pseudo_metric_space α] {s : set α} (h : ¬metric.bounded s) :

An unbounded set has zero diameter. If you would prefer to get the value ∞, use emetric.diam. This lemma makes it possible to avoid side conditions in some situations

theorem metric.diam_mono {α : Type u} [pseudo_metric_space α] {s t : set α} (h : s t) (ht : metric.bounded t) :

If s ⊆ t, then the diameter of s is bounded by that of t, provided t is bounded.

theorem metric.diam_union {α : Type u} [pseudo_metric_space α] {s : set α} {x y : α} {t : set α} (xs : x s) (yt : y t) :

The diameter of a union is controlled by the sum of the diameters, and the distance between any two points in each of the sets. This lemma is true without any side condition, since it is obviously true if s ∪ t is unbounded.

theorem metric.diam_union' {α : Type u} [pseudo_metric_space α] {s t : set α} (h : (s t).nonempty) :

If two sets intersect, the diameter of the union is bounded by the sum of the diameters.

theorem metric.diam_le_of_subset_closed_ball {α : Type u} [pseudo_metric_space α] {s : set α} {x : α} {r : } (hr : 0 r) (h : s metric.closed_ball x r) :
theorem metric.diam_closed_ball {α : Type u} [pseudo_metric_space α] {x : α} {r : } (h : 0 r) :

The diameter of a closed ball of radius r is at most 2 r.

theorem metric.diam_ball {α : Type u} [pseudo_metric_space α] {x : α} {r : } (h : 0 r) :

The diameter of a ball of radius r is at most 2 r.

theorem is_complete.nonempty_Inter_of_nonempty_bInter {α : Type u} [pseudo_metric_space α] {s : set α} (h0 : is_complete (s 0)) (hs : ∀ (n : ), is_closed (s n)) (h's : ∀ (n : ), metric.bounded (s n)) (h : ∀ (N : ), (⋂ (n : ) (H : n N), s n).nonempty) (h' : filter.tendsto (λ (n : ), metric.diam (s n)) filter.at_top (nhds 0)) :
(⋂ (n : ), s n).nonempty

If a family of complete sets with diameter tending to 0 is such that each finite intersection is nonempty, then the total intersection is also nonempty.

theorem metric.nonempty_Inter_of_nonempty_bInter {α : Type u} [pseudo_metric_space α] [complete_space α] {s : set α} (hs : ∀ (n : ), is_closed (s n)) (h's : ∀ (n : ), metric.bounded (s n)) (h : ∀ (N : ), (⋂ (n : ) (H : n N), s n).nonempty) (h' : filter.tendsto (λ (n : ), metric.diam (s n)) filter.at_top (nhds 0)) :
(⋂ (n : ), s n).nonempty

In a complete space, if a family of closed sets with diameter tending to 0 is such that each finite intersection is nonempty, then the total intersection is also nonempty.

theorem tendsto_cocompact_of_tendsto_dist_comp_at_top {α : Type u} {β : Type v} [pseudo_metric_space α] {f : β → α} {l : filter β} (x : α) (h : filter.tendsto (λ (y : β), has_dist.dist (f y) x) l filter.at_top) :

Under the coercion from to , inverse images of compact sets are finite.

@[class]
structure metric_space (α : Type u) :
Type u

We now define metric_space, extending pseudo_metric_space.

Instances of this typeclass
Instances of other typeclasses for metric_space
  • metric_space.has_sizeof_inst
@[ext]

Two metric space structures with the same distance coincide.

def metric_space.of_metrizable {α : Type u_1} [topological_space α] (dist : α → α → ) (dist_self : ∀ (x : α), dist x x = 0) (dist_comm : ∀ (x y : α), dist x y = dist y x) (dist_triangle : ∀ (x y z : α), dist x z dist x y + dist y z) (H : ∀ (s : set α), is_open s ∀ (x : α), x s(∃ (ε : ) (H : ε > 0), ∀ (y : α), dist x y < εy s)) (eq_of_dist_eq_zero : ∀ (x y : α), dist x y = 0x = y) :

Construct a metric space structure whose underlying topological space structure (definitionally) agrees which a pre-existing topology which is compatible with a given distance function.

Equations
theorem eq_of_dist_eq_zero {γ : Type w} [metric_space γ] {x y : γ} :
has_dist.dist x y = 0x = y
@[simp]
theorem dist_eq_zero {γ : Type w} [metric_space γ] {x y : γ} :
has_dist.dist x y = 0 x = y
@[simp]
theorem zero_eq_dist {γ : Type w} [metric_space γ] {x y : γ} :
0 = has_dist.dist x y x = y
theorem dist_ne_zero {γ : Type w} [metric_space γ] {x y : γ} :
@[simp]
theorem dist_le_zero {γ : Type w} [metric_space γ] {x y : γ} :
@[simp]
theorem dist_pos {γ : Type w} [metric_space γ] {x y : γ} :
theorem eq_of_forall_dist_le {γ : Type w} [metric_space γ] {x y : γ} (h : ∀ (ε : ), ε > 0has_dist.dist x y ε) :
x = y
theorem eq_of_nndist_eq_zero {γ : Type w} [metric_space γ] {x y : γ} :
has_nndist.nndist x y = 0x = y

Deduce the equality of points with the vanishing of the nonnegative distance

@[simp]
theorem nndist_eq_zero {γ : Type w} [metric_space γ] {x y : γ} :

Characterize the equality of points with the vanishing of the nonnegative distance

@[simp]
theorem zero_eq_nndist {γ : Type w} [metric_space γ] {x y : γ} :
@[simp]
theorem metric.closed_ball_zero {γ : Type w} [metric_space γ] {x : γ} :
@[simp]
theorem metric.sphere_zero {γ : Type w} [metric_space γ] {x : γ} :
theorem metric.subsingleton_closed_ball {γ : Type w} [metric_space γ] (x : γ) {r : } (hr : r 0) :
theorem metric.subsingleton_sphere {γ : Type w} [metric_space γ] (x : γ) {r : } (hr : r 0) :
theorem metric.uniform_embedding_iff' {β : Type v} {γ : Type w} [metric_space γ] [metric_space β] {f : γ → β} :
uniform_embedding f (∀ (ε : ), ε > 0(∃ (δ : ) (H : δ > 0), ∀ {a b : γ}, has_dist.dist a b < δhas_dist.dist (f a) (f b) < ε)) ∀ (δ : ), δ > 0(∃ (ε : ) (H : ε > 0), ∀ {a b : γ}, has_dist.dist (f a) (f b) < εhas_dist.dist a b < δ)

A map between metric spaces is a uniform embedding if and only if the distance between f x and f y is controlled in terms of the distance between x and y and conversely.

@[protected, instance]
@[protected, instance]

A metric space induces an emetric space

Equations
theorem metric.is_closed_of_pairwise_le_dist {γ : Type w} [metric_space γ] {s : set γ} {ε : } (hε : 0 < ε) (hs : s.pairwise (λ (x y : γ), ε has_dist.dist x y)) :
theorem metric.closed_embedding_of_pairwise_le_dist {γ : Type w} [metric_space γ] {α : Type u_1} [topological_space α] [discrete_topology α] {ε : } (hε : 0 < ε) {f : α → γ} (hf : pairwise (λ (x y : α), ε has_dist.dist (f x) (f y))) :
theorem metric.uniform_embedding_bot_of_pairwise_le_dist {α : Type u} [pseudo_metric_space α] {β : Type u_1} {ε : } (hε : 0 < ε) {f : β → α} (hf : pairwise (λ (x y : β), ε has_dist.dist (f x) (f y))) :

If f : β → α sends any two distinct points to points at distance at least ε > 0, then f is a uniform embedding with respect to the discrete uniformity on β.

theorem metric_space.replace_uniformity_eq {γ : Type u_1} [U : uniform_space γ] (m : metric_space γ) (H : uniformity γ = uniformity γ) :
@[reducible]

Build a new metric space from an old one where the bundled topological structure is provably (but typically non-definitionaly) equal to some given topological structure. See Note [forgetful inheritance].

Equations
def emetric_space.to_metric_space_of_dist {α : Type u} [e : emetric_space α] (dist : α → α → ) (edist_ne_top : ∀ (x y : α), has_edist.edist x y ) (h : ∀ (x y : α), dist x y = (has_edist.edist x y).to_real) :

One gets a metric space from an emetric space if the edistance is everywhere finite, by pushing the edistance to reals. We set it up so that the edist and the uniformity are defeq in the metric space and the emetric space. In this definition, the distance is given separately, to be able to prescribe some expression which is not defeq to the push-forward of the edistance to reals.

Equations
def emetric_space.to_metric_space {α : Type u} [e : emetric_space α] (h : ∀ (x y : α), has_edist.edist x y ) :

One gets a metric space from an emetric space if the edistance is everywhere finite, by pushing the edistance to reals. We set it up so that the edist and the uniformity are defeq in the metric space and the emetric space.

Equations
theorem metric_space.replace_bornology_eq {α : Type u_1} [m : metric_space α] [B : bornology α] (H : ∀ (s : set α), bornology.is_bounded s bornology.is_bounded s) :
@[reducible]
def uniform_embedding.comap_metric_space {α : Type u_1} {β : Type u_2} [uniform_space α] [metric_space β] (f : α → β) (h : uniform_embedding f) :

Pull back a metric space structure by a uniform embedding. This is a version of metric_space.induced useful in case if the domain already has a uniform_space structure.

Equations
@[reducible]
def embedding.comap_metric_space {α : Type u_1} {β : Type u_2} [topological_space α] [metric_space β] (f : α → β) (h : embedding f) :

Pull back a metric space structure by an embedding. This is a version of metric_space.induced useful in case if the domain already has a topological_space structure.

Equations
@[protected, instance]
def subtype.metric_space {α : Type u_1} {p : α → Prop} [metric_space α] :
Equations
@[protected, instance]
Equations
@[protected, instance]
Equations
@[protected, instance]
noncomputable def nnreal.metric_space  :
Equations
@[protected, instance]
def ulift.metric_space {β : Type v} [metric_space β] :
Equations
@[protected, instance]
noncomputable def metric_space_pi {β : Type v} {π : β → Type u_2} [fintype β] [Π (b : β), metric_space (π b)] :
metric_space (Π (b : β), π b)

A finite product of metric spaces is a metric space, with the sup distance.

Equations
theorem metric.second_countable_of_countable_discretization {α : Type u} [metric_space α] (H : ∀ (ε : ), ε > 0(∃ (β : Type u_1) (_x : encodable β) (F : α → β), ∀ (x y : α), F x = F yhas_dist.dist x y ε)) :

A metric space is second countable if one can reconstruct up to any ε>0 any element of the space from countably many data.

The canonical equivalence relation on a pseudometric space.

Equations
@[reducible]
def pseudo_metric_quot (α : Type u) [pseudo_metric_space α] :
Type u

The canonical quotient of a pseudometric space, identifying points at distance 0.

Equations
@[protected, instance]
Equations
@[protected, instance]
Equations

additive, multiplicative #

The distance on those type synonyms is inherited without change.

@[protected, instance]
def additive.has_dist {X : Type u_1} [has_dist X] :
Equations
@[protected, instance]
Equations
@[simp]
theorem dist_of_mul {X : Type u_1} [has_dist X] (a b : X) :
@[simp]
@[simp]
theorem dist_to_mul {X : Type u_1} [has_dist X] (a b : additive X) :
@[protected, instance]
Equations
@[protected, instance]
Equations
@[protected, instance]
Equations
@[protected, instance]
@[protected, instance]

Order dual #

The distance on this type synonym is inherited without change.

@[protected, instance]
def order_dual.has_dist {X : Type u_1} [has_dist X] :
Equations
@[simp]
theorem dist_to_dual {X : Type u_1} [has_dist X] (a b : X) :
@[simp]
@[protected, instance]
Equations
@[protected, instance]
Equations
@[protected, instance]