Extended metric spaces #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This file is devoted to the definition and study of emetric_spaces
, i.e., metric
spaces in which the distance is allowed to take the value ∞. This extended distance is
called edist
, and takes values in ℝ≥0∞
.
Many definitions and theorems expected on emetric spaces are already introduced on uniform spaces and topological spaces. For example: open and closed sets, compactness, completeness, continuity and uniform continuity.
The class emetric_space
therefore extends uniform_space
(and topological_space
).
Since a lot of elementary properties don't require eq_of_edist_eq_zero
we start setting up the
theory of pseudo_emetric_space
, where we don't require edist x y = 0 → x = y
and we specialize
to emetric_space
at the end.
Characterizing uniformities associated to a (generalized) distance function D
in terms of the elements of the uniformity.
has_edist α
means that α
is equipped with an extended distance.
Instances of this typeclass
Instances of other typeclasses for has_edist
- has_edist.has_sizeof_inst
Creating a uniform space from an extended distance.
Equations
- uniform_space_of_edist edist edist_self edist_comm edist_triangle = uniform_space.of_fun edist edist_self edist_comm edist_triangle uniform_space_of_edist._proof_1
- to_has_edist : has_edist α
- edist_self : ∀ (x : α), has_edist.edist x x = 0
- edist_comm : ∀ (x y : α), has_edist.edist x y = has_edist.edist y x
- edist_triangle : ∀ (x y z : α), has_edist.edist x z ≤ has_edist.edist x y + has_edist.edist y z
- to_uniform_space : uniform_space α
- uniformity_edist : (uniformity α = ⨅ (ε : ennreal) (H : ε > 0), filter.principal {p : α × α | has_edist.edist p.fst p.snd < ε}) . "control_laws_tac"
Extended (pseudo) metric spaces, with an extended distance edist
possibly taking the
value ∞
Each pseudo_emetric 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 pseudo_emetric_space
structure, the uniformity fields are not necessary, they
will be filled in by default. There is a default value for the uniformity, that can be substituted
in cases of interest, for instance when instantiating a pseudo_emetric_space
structure
on a product.
Continuity of edist
is proved in topology.instances.ennreal
Instances of this typeclass
- emetric_space.to_pseudo_emetric_space
- pseudo_metric_space.to_pseudo_emetric_space
- subtype.pseudo_emetric_space
- mul_opposite.pseudo_emetric_space
- add_opposite.pseudo_emetric_space
- ulift.pseudo_emetric_space
- prod.pseudo_emetric_space_max
- pseudo_emetric_space_pi
- additive.pseudo_emetric_space
- multiplicative.pseudo_emetric_space
- order_dual.pseudo_emetric_space
- pi_Lp.pseudo_emetric_space
Instances of other typeclasses for pseudo_emetric_space
- pseudo_emetric_space.has_sizeof_inst
Triangle inequality for the extended distance
The triangle (polygon) inequality for sequences of points; finset.Ico
version.
The triangle (polygon) inequality for sequences of points; finset.range
version.
A version of edist_le_Ico_sum_edist
with each intermediate distance replaced
with an upper estimate.
A version of edist_le_range_sum_edist
with each intermediate distance replaced
with an upper estimate.
Reformulation of the uniform structure in terms of the extended distance
Characterization of the elements of the uniformity in terms of the extended distance
Given f : β → ℝ≥0∞
, 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_edist
, uniformity_basis_edist'
,
uniformity_basis_edist_nnreal
, and uniformity_basis_edist_inv_nat
.
Given f : β → ℝ≥0∞
, if f
sends {i | p i}
to a set of positive numbers
accumulating to zero, then closed f i
-neighborhoods of the diagonal form a basis of 𝓤 α
.
For specific bases see uniformity_basis_edist_le
and uniformity_basis_edist_le'
.
Fixed size neighborhoods of the diagonal belong to the uniform structure
ε-δ characterization of uniform continuity on a set for pseudoemetric spaces
ε-δ characterization of uniform continuity on pseudoemetric spaces
ε-δ characterization of uniform embeddings on pseudoemetric spaces
If a map between pseudoemetric spaces is a uniform embedding then the edistance between f x
and f y
is controlled in terms of the distance between x
and y
.
ε-δ characterization of Cauchy sequences on pseudoemetric spaces
A very useful criterion to show that a space is complete is to show that all sequences
which satisfy a bound of the form edist (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.
A sequentially complete pseudoemetric space is complete.
Expressing locally uniform convergence on a set using edist
.
Expressing uniform convergence on a set using edist
.
Expressing locally uniform convergence using edist
.
Expressing uniform convergence using edist
.
Auxiliary function to replace the uniformity on a pseudoemetric space with a uniformity which is equal to the original one, but maybe not defeq. This is useful if one wants to construct a pseudoemetric space with a specified uniformity. See Note [forgetful inheritance] explaining why having definitionally the right uniformity is often important.
Equations
- m.replace_uniformity H = {to_has_edist := {edist := has_edist.edist pseudo_emetric_space.to_has_edist}, edist_self := _, edist_comm := _, edist_triangle := _, to_uniform_space := U, uniformity_edist := _}
The extended pseudometric induced by a function taking values in a pseudoemetric space.
Equations
- pseudo_emetric_space.induced f m = {to_has_edist := {edist := λ (x y : α), has_edist.edist (f x) (f y)}, edist_self := _, edist_comm := _, edist_triangle := _, to_uniform_space := uniform_space.comap f pseudo_emetric_space.to_uniform_space, uniformity_edist := _}
Pseudoemetric space instance on subsets of pseudoemetric spaces
Equations
The extended psuedodistance on a subset of a pseudoemetric space is the restriction of the original pseudodistance, by definition
Pseudoemetric space instance on the additive opposite of a pseudoemetric space.
Equations
Pseudoemetric space instance on the multiplicative opposite of a pseudoemetric space.
Equations
Equations
The product of two pseudoemetric spaces, with the max distance, is an extended pseudometric spaces. We make sure that the uniform structure thus constructed is the one corresponding to the product of uniform spaces, to avoid diamond problems.
Equations
- prod.pseudo_emetric_space_max = {to_has_edist := {edist := λ (x y : α × β), has_edist.edist x.fst y.fst ⊔ has_edist.edist x.snd y.snd}, edist_self := _, edist_comm := _, edist_triangle := _, to_uniform_space := prod.uniform_space pseudo_emetric_space.to_uniform_space, uniformity_edist := _}
The product of a finite number of pseudoemetric spaces, with the max distance, is still a pseudoemetric space. This construction would also work for infinite products, but it would not give rise to the product topology. Hence, we only formalize it in the good situation of finitely many spaces.
Equations
- pseudo_emetric_space_pi = {to_has_edist := {edist := λ (f g : Π (b : β), π b), finset.univ.sup (λ (b : β), has_edist.edist (f b) (g b))}, edist_self := _, edist_comm := _, edist_triangle := _, to_uniform_space := Pi.uniform_space (λ (b : β), π b) (λ (i : β), pseudo_emetric_space.to_uniform_space), uniformity_edist := _}
emetric.ball x ε
is the set of all points y
with edist y x < ε
Equations
- emetric.ball x ε = {y : α | has_edist.edist y x < ε}
emetric.closed_ball x ε
is the set of all points y
with edist y x ≤ ε
Equations
- emetric.closed_ball x ε = {y : α | has_edist.edist y x ≤ ε}
Relation “two points are at a finite edistance” is an equivalence relation.
Equations
- emetric.edist_lt_top_setoid = {r := λ (x y : α), has_edist.edist x y < ⊤, iseqv := _}
In a pseudoemetric space, Cauchy sequences are characterized by the fact that, eventually, the pseudoedistance between its elements is arbitrarily small
A variation around the emetric characterization of Cauchy sequences
A variation of the emetric characterization of Cauchy sequences that deals with
ℝ≥0
upper bounds.
For a set s
in a pseudo emetric space, if for every ε > 0
there exists a countable
set that is ε
-dense in s
, then there exists a countable subset t ⊆ s
that is dense in s
.
A compact set in a pseudo emetric space is separable, i.e., it is a subset of the closure of a countable set.
A sigma compact pseudo emetric space has second countable topology. This is not an instance
to avoid a loop with sigma_compact_space_of_locally_compact_second_countable
.
The diameter of a set in a pseudoemetric space, named emetric.diam
Equations
- emetric.diam s = ⨆ (x : α) (H : x ∈ s) (y : α) (H : y ∈ s), has_edist.edist x y
If two points belong to some set, their edistance is bounded by the diameter of the set
If the distance between any two points in a set is bounded by some constant, this constant bounds the diameter.
The diameter of a subsingleton vanishes.
The diameter of the empty set vanishes
The diameter of a singleton vanishes
The diameter is monotonous with respect to inclusion
The diameter of a union is controlled by the diameter of the sets, and the edistance between two points in the sets.
- to_pseudo_emetric_space : pseudo_emetric_space α
- eq_of_edist_eq_zero : ∀ {x y : α}, has_edist.edist x y = 0 → x = y
We now define emetric_space
, extending pseudo_emetric_space
.
Instances of this typeclass
- metric_space.to_emetric_space
- subtype.emetric_space
- mul_opposite.emetric_space
- add_opposite.emetric_space
- ulift.emetric_space
- prod.emetric_space_max
- emetric_space_pi
- uniform_space.separation_quotient.emetric_space
- additive.emetric_space
- multiplicative.emetric_space
- order_dual.emetric_space
- pi_Lp.emetric_space
- emetric.closeds.emetric_space
- emetric.nonempty_compacts.emetric_space
Instances of other typeclasses for emetric_space
- emetric_space.has_sizeof_inst
Characterize the equality of points by the vanishing of their extended distance
Two points coincide if their distance is < ε
for all positive ε
An emetric space is separated
A map between emetric spaces is a uniform embedding if and only if the edistance between f x
and f y
is controlled in terms of the distance between x
and y
and conversely.
If a pseudo_emetric_space
is a T₀ space, then it is an emetric_space
.
Equations
- emetric_space.of_t0_pseudo_emetric_space α = {to_pseudo_emetric_space := {to_has_edist := pseudo_emetric_space.to_has_edist _inst_3, edist_self := _, edist_comm := _, edist_triangle := _, to_uniform_space := pseudo_emetric_space.to_uniform_space _inst_3, uniformity_edist := _}, eq_of_edist_eq_zero := _}
Auxiliary function to replace the uniformity on an emetric space with a uniformity which is equal to the original one, but maybe not defeq. This is useful if one wants to construct an emetric space with a specified uniformity. See Note [forgetful inheritance] explaining why having definitionally the right uniformity is often important.
Equations
- m.replace_uniformity H = {to_pseudo_emetric_space := {to_has_edist := {edist := has_edist.edist pseudo_emetric_space.to_has_edist}, edist_self := _, edist_comm := _, edist_triangle := _, to_uniform_space := U, uniformity_edist := _}, eq_of_edist_eq_zero := _}
The extended metric induced by an injective function taking values in a emetric space.
Equations
- emetric_space.induced f hf m = {to_pseudo_emetric_space := {to_has_edist := {edist := λ (x y : γ), has_edist.edist (f x) (f y)}, edist_self := _, edist_comm := _, edist_triangle := _, to_uniform_space := uniform_space.comap f pseudo_emetric_space.to_uniform_space, uniformity_edist := _}, eq_of_edist_eq_zero := _}
Emetric space instance on subsets of emetric spaces
Equations
Emetric space instance on the multiplicative opposite of an emetric space.
Emetric space instance on the additive opposite of an emetric space.
Equations
The product of two emetric spaces, with the max distance, is an extended metric spaces. We make sure that the uniform structure thus constructed is the one corresponding to the product of uniform spaces, to avoid diamond problems.
Equations
- prod.emetric_space_max = {to_pseudo_emetric_space := {to_has_edist := pseudo_emetric_space.to_has_edist prod.pseudo_emetric_space_max, edist_self := _, edist_comm := _, edist_triangle := _, to_uniform_space := pseudo_emetric_space.to_uniform_space prod.pseudo_emetric_space_max, uniformity_edist := _}, eq_of_edist_eq_zero := _}
Reformulation of the uniform structure in terms of the extended distance
The product of a finite number of emetric spaces, with the max distance, is still an emetric space. This construction would also work for infinite products, but it would not give rise to the product topology. Hence, we only formalize it in the good situation of finitely many spaces.
Equations
- emetric_space_pi = {to_pseudo_emetric_space := {to_has_edist := pseudo_emetric_space.to_has_edist pseudo_emetric_space_pi, edist_self := _, edist_comm := _, edist_triangle := _, to_uniform_space := pseudo_emetric_space.to_uniform_space pseudo_emetric_space_pi, uniformity_edist := _}, eq_of_edist_eq_zero := _}
A compact set in an emetric space is separable, i.e., it is the closure of a countable set.
Separation quotient #
Equations
- uniform_space.separation_quotient.has_edist = {edist := λ (x y : uniform_space.separation_quotient X), quotient.lift_on₂' x y has_edist.edist uniform_space.separation_quotient.has_edist._proof_1}
additive
, multiplicative
#
The distance on those type synonyms is inherited without change.
Equations
- additive.has_edist = _inst_3
Equations
- multiplicative.has_edist = _inst_3
Equations
- additive.pseudo_emetric_space = _inst_3
Equations
- multiplicative.pseudo_emetric_space = _inst_3
Equations
- additive.emetric_space = _inst_3
Equations
- multiplicative.emetric_space = _inst_3
Order dual #
The distance on this type synonym is inherited without change.
Equations
- order_dual.has_edist = _inst_3
Equations
- order_dual.pseudo_emetric_space = _inst_3
Equations
- order_dual.emetric_space = _inst_3