The Gromov-Hausdorff distance is realized #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
In this file, we construct of a good coupling between nonempty compact metric spaces, minimizing their Hausdorff distance. This construction is instrumental to study the Gromov-Hausdorff distance between nonempty compact metric spaces.
Given two nonempty compact metric spaces X
and Y
, we define optimal_GH_coupling X Y
as a
compact metric space, together with two isometric embeddings optimal_GH_injl
and optimal_GH_injr
respectively of X
and Y
into optimal_GH_coupling X Y
. The main property of the optimal
coupling is that the Hausdorff distance between X
and Y
in optimal_GH_coupling X Y
is smaller
than the corresponding distance in any other coupling. We do not prove completely this fact in this
file, but we show a good enough approximation of this fact in Hausdorff_dist_optimal_le_HD
, that
will suffice to obtain the full statement once the Gromov-Hausdorff distance is properly defined,
in Hausdorff_dist_optimal
.
The key point in the construction is that the set of possible distances coming from isometric
embeddings of X
and Y
in metric spaces is a set of equicontinuous functions. By Arzela-Ascoli,
it is compact, and one can find such a distance which is minimal. This distance defines a premetric
space structure on X ⊕ Y
. The corresponding metric quotient is optimal_GH_coupling X Y
.
The set of functions on X ⊕ Y
that are candidates distances to realize the
minimum of the Hausdorff distances between X
and Y
in a coupling
Equations
- Gromov_Hausdorff.candidates X Y = {f : prod_space_fun X Y | (((((∀ (x y : X), f (sum.inl x, sum.inl y) = has_dist.dist x y) ∧ ∀ (x y : Y), f (sum.inr x, sum.inr y) = has_dist.dist x y) ∧ ∀ (x y : X ⊕ Y), f (x, y) = f (y, x)) ∧ ∀ (x y z : X ⊕ Y), f (x, z) ≤ f (x, y) + f (y, z)) ∧ ∀ (x : X ⊕ Y), f (x, x) = 0) ∧ ∀ (x y : X ⊕ Y), f (x, y) ≤ ↑(max_var X Y)}
candidates give rise to elements of bounded_continuous_functions
Equations
The distance on X ⊕ Y
as a candidate
Equations
- Gromov_Hausdorff.candidates_b_dist X Y = Gromov_Hausdorff.candidates_b_of_candidates (λ (p : (X ⊕ Y) × (X ⊕ Y)), has_dist.dist p.fst p.snd) _
We will then choose the candidate minimizing the Hausdorff distance. Except that we are not in a metric space setting, so we need to define our custom version of Hausdorff distance, called HD, and prove its basic properties.
Explicit bound on HD (dist)
. This means that when looking for minimizers it will
be sufficient to look for functions with HD(f)
bounded by this bound.
With the optimal candidate, construct a premetric space structure on X ⊕ Y
, on which the
predistance is given by the candidate. Then, we will identify points at 0
predistance
to obtain a genuine metric space
Equations
- Gromov_Hausdorff.premetric_optimal_GH_dist X Y = {to_has_dist := {dist := λ (p q : X ⊕ Y), ⇑(optimal_GH_dist X Y) (p, q)}, dist_self := _, dist_comm := _, dist_triangle := _, edist := λ (x y : X ⊕ Y), ↑⟨(λ (p q : X ⊕ Y), ⇑(optimal_GH_dist X Y) (p, q)) x y, _⟩, edist_dist := _, to_uniform_space := uniform_space_of_dist (λ (p q : X ⊕ Y), ⇑(optimal_GH_dist X Y) (p, q)) _ _ _, uniformity_dist := _, to_bornology := bornology.of_dist (λ (p q : X ⊕ Y), ⇑(optimal_GH_dist X Y) (p, q)) _ _ _, cobounded_sets := _}
A metric space which realizes the optimal coupling between X
and Y
Equations
Instances for Gromov_Hausdorff.optimal_GH_coupling
Injection of X
in the optimal coupling between X
and Y
Equations
- Gromov_Hausdorff.optimal_GH_injl X Y x = quotient.mk' (sum.inl x)
The injection of X
in the optimal coupling between X
and Y
is an isometry.
Injection of Y
in the optimal coupling between X
and Y
Equations
- Gromov_Hausdorff.optimal_GH_injr X Y y = quotient.mk' (sum.inr y)
The injection of Y
in the optimal coupling between X
and Y
is an isometry.
The optimal coupling between two compact spaces X
and Y
is still a compact space
For any candidate f
, HD(f)
is larger than or equal to the Hausdorff distance in the
optimal coupling. This follows from the fact that HD of the optimal candidate is exactly
the Hausdorff distance in the optimal coupling, although we only prove here the inequality
we need.