Skip to content

Commit d4f568d

Browse files
committed
feat(Topology/UniformSpace): define the Hausdorff uniformity
This PR defines the Hausdorff uniformity on `Closeds`, `Compacts` and `NonemptyCompacts`. Since `Closeds` and `NonemptyCompacts` already have metrics, they are changed to use the newly defined uniformity.
1 parent 3e58cbe commit d4f568d

File tree

6 files changed

+328
-31
lines changed

6 files changed

+328
-31
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -7088,6 +7088,7 @@ import Mathlib.Topology.UniformSpace.AbstractCompletion
70887088
import Mathlib.Topology.UniformSpace.Ascoli
70897089
import Mathlib.Topology.UniformSpace.Basic
70907090
import Mathlib.Topology.UniformSpace.Cauchy
7091+
import Mathlib.Topology.UniformSpace.Closeds
70917092
import Mathlib.Topology.UniformSpace.Compact
70927093
import Mathlib.Topology.UniformSpace.CompactConvergence
70937094
import Mathlib.Topology.UniformSpace.CompareReals

Mathlib/Data/Rel.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -219,9 +219,15 @@ def preimage : Set α := {a | ∃ b ∈ t, a ~[R] b}
219219
@[gcongr] lemma image_subset_image (hs : s₁ ⊆ s₂) : image R s₁ ⊆ image R s₂ :=
220220
fun _ ⟨a, ha, hab⟩ ↦ ⟨a, hs ha, hab⟩
221221

222+
@[gcongr] lemma image_subset_image_left (hR : R₁ ⊆ R₂) : image R₁ s ⊆ image R₂ s :=
223+
fun _ ⟨a, ha, hab⟩ ↦ ⟨a, ha, hR hab⟩
224+
222225
@[gcongr] lemma preimage_subset_preimage (ht : t₁ ⊆ t₂) : preimage R t₁ ⊆ preimage R t₂ :=
223226
fun _ ⟨a, ha, hab⟩ ↦ ⟨a, ht ha, hab⟩
224227

228+
@[gcongr] lemma preimage_subset_preimage_left (hR : R₁ ⊆ R₂) : preimage R₁ t ⊆ preimage R₂ t :=
229+
fun _ ⟨a, ha, hab⟩ ↦ ⟨a, ha, hR hab⟩
230+
225231
variable (R t) in
226232
@[simp] lemma image_inv : R.inv.image t = preimage R t := rfl
227233

@@ -381,6 +387,12 @@ lemma subset_iterate_comp [R.IsRefl] {S : SetRel α β} : ∀ {n}, S ⊆ (R ○
381387
| 0 => .rfl
382388
| _n + 1 => right_subset_comp.trans subset_iterate_comp
383389

390+
lemma self_subset_image [R.IsRefl] (s : Set α) : s ⊆ R.image s :=
391+
fun x hx => ⟨x, hx, R.rfl⟩
392+
393+
lemma self_subset_preimage [R.IsRefl] (s : Set α) : s ⊆ R.preimage s :=
394+
fun x hx => ⟨x, hx, R.rfl⟩
395+
384396
lemma exists_eq_singleton_of_prod_subset_id {s t : Set α} (hs : s.Nonempty) (ht : t.Nonempty)
385397
(hst : s ×ˢ t ⊆ SetRel.id) : ∃ x, s = {x} ∧ t = {x} := by
386398
obtain ⟨a, ha⟩ := hs

Mathlib/Topology/MetricSpace/Closeds.lean

Lines changed: 63 additions & 31 deletions
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,7 @@ Authors: Sébastien Gouëzel
55
-/
66
import Mathlib.Analysis.SpecificLimits.Basic
77
import Mathlib.Topology.MetricSpace.HausdorffDistance
8-
import Mathlib.Topology.Sets.Compacts
8+
import Mathlib.Topology.UniformSpace.Closeds
99

1010
/-!
1111
# Closed subsets
@@ -30,15 +30,62 @@ namespace EMetric
3030

3131
section
3232

33+
variable {α : Type*} [PseudoEMetricSpace α]
34+
35+
theorem mem_hausdorffEntourage_of_hausdorffEdist_lt {s t : Set α} {δ : ℝ≥0∞}
36+
(h : hausdorffEdist s t < δ) : (s, t) ∈ hausdorffEntourage {p | edist p.1 p.2 < δ} := by
37+
rw [hausdorffEdist, max_lt_iff] at h
38+
rw [hausdorffEntourage, Set.mem_setOf]
39+
conv => enter [2, 2, 1, 1, _]; rw [edist_comm]
40+
have {s t : Set α} (h : ⨆ x ∈ s, infEdist x t < δ) :
41+
s ⊆ SetRel.preimage {p | edist p.1 p.2 < δ} t := by
42+
intro x hx
43+
have := (le_iSup₂ x hx).trans_lt h
44+
simp_rw [infEdist, iInf_lt_iff, exists_prop] at this
45+
exact this
46+
exact ⟨this h.1, this h.2
47+
48+
theorem hausdorffEdist_le_of_mem_hausdorffEntourage {s t : Set α} {δ : ℝ≥0∞}
49+
(h : (s, t) ∈ hausdorffEntourage {p | edist p.1 p.2 ≤ δ}) : hausdorffEdist s t ≤ δ := by
50+
rw [hausdorffEdist, max_le_iff]
51+
rw [hausdorffEntourage, Set.mem_setOf] at h
52+
conv at h => enter [2, 2, 1, 1, _]; rw [edist_comm]
53+
have {s t : Set α} (h : s ⊆ SetRel.preimage {p | edist p.1 p.2 ≤ δ} t) :
54+
⨆ x ∈ s, infEdist x t ≤ δ := by
55+
rw [iSup₂_le_iff]
56+
intro x hx
57+
obtain ⟨y, hy, hxy⟩ := h hx
58+
refine iInf₂_le_of_le y hy hxy
59+
exact ⟨this h.1, this h.2
60+
61+
/-- The Hausdorff pseudo emetric on the powerset of a pseudo emetric space.
62+
See note [reducible non-instances]. -/
63+
protected abbrev _root_.PseudoEMetricSpace.hausdorff : PseudoEMetricSpace (Set α) where
64+
edist s t := hausdorffEdist s t
65+
edist_self _ := hausdorffEdist_self
66+
edist_comm _ _ := hausdorffEdist_comm
67+
edist_triangle _ _ _ := hausdorffEdist_triangle
68+
toUniformSpace := .hausdorff
69+
uniformity_edist := by
70+
refine le_antisymm
71+
(le_iInf₂ fun ε hε => Filter.le_principal_iff.mpr ?_)
72+
(uniformity_basis_edist.lift' monotone_hausdorffEntourage |>.ge_iff.mpr fun ε hε =>
73+
Filter.mem_iInf_of_mem ε <| Filter.mem_iInf_of_mem hε fun _ =>
74+
mem_hausdorffEntourage_of_hausdorffEdist_lt)
75+
obtain ⟨δ, hδ, hδε⟩ := exists_between hε
76+
filter_upwards [Filter.mem_lift' (uniformity_basis_edist_le.mem_of_mem hδ)]
77+
with _ h using hδε.trans_le' <| hausdorffEdist_le_of_mem_hausdorffEntourage h
78+
79+
end
80+
81+
section
82+
3383
variable {α β : Type*} [EMetricSpace α] [EMetricSpace β] {s : Set α}
3484

3585
/-- In emetric spaces, the Hausdorff edistance defines an emetric space structure
3686
on the type of closed subsets -/
3787
instance Closeds.emetricSpace : EMetricSpace (Closeds α) where
38-
edist s t := hausdorffEdist (s : Set α) t
39-
edist_self _ := hausdorffEdist_self
40-
edist_comm _ _ := hausdorffEdist_comm
41-
edist_triangle _ _ _ := hausdorffEdist_triangle
88+
__ := PseudoEMetricSpace.hausdorff.induced SetLike.coe
4289
eq_of_edist_eq_zero {s t} h :=
4390
Closeds.ext <| (hausdorffEdist_zero_iff_eq_of_closed s.isClosed t.isClosed).1 h
4491

@@ -59,20 +106,11 @@ theorem continuous_infEdist_hausdorffEdist :
59106
_ = infEdist y t + 2 * edist (x, s) (y, t) := by rw [← mul_two, mul_comm]
60107

61108
/-- Subsets of a given closed subset form a closed set -/
62-
theorem Closeds.isClosed_subsets_of_isClosed (hs : IsClosed s) :
63-
IsClosed { t : Closeds α | (t : Set α) ⊆ s } := by
64-
refine isClosed_of_closure_subset fun
65-
(t : Closeds α) (ht : t ∈ closure {t : Closeds α | (t : Set α) ⊆ s}) (x : α) (hx : x ∈ t) => ?_
66-
have : x ∈ closure s := by
67-
refine mem_closure_iff.2 fun ε εpos => ?_
68-
obtain ⟨u : Closeds α, hu : u ∈ {t : Closeds α | (t : Set α) ⊆ s}, Dtu : edist t u < ε⟩ :=
69-
mem_closure_iff.1 ht ε εpos
70-
obtain ⟨y : α, hy : y ∈ u, Dxy : edist x y < ε⟩ := exists_edist_lt_of_hausdorffEdist_lt hx Dtu
71-
exact ⟨y, hu hy, Dxy⟩
72-
rwa [hs.closure_eq] at this
109+
@[deprecated (since := "2025-11-05")]
110+
alias Closeds.isClosed_subsets_of_isClosed := TopologicalSpace.Closeds.isClosed_subsets_of_isClosed
73111

74112
@[deprecated (since := "2025-08-20")]
75-
alias isClosed_subsets_of_isClosed := Closeds.isClosed_subsets_of_isClosed
113+
alias isClosed_subsets_of_isClosed := TopologicalSpace.Closeds.isClosed_subsets_of_isClosed
76114

77115
/-- By definition, the edistance on `Closeds α` is given by the Hausdorff edistance -/
78116
theorem Closeds.edist_eq {s t : Closeds α} : edist s t = hausdorffEdist (s : Set α) t :=
@@ -231,10 +269,7 @@ namespace NonemptyCompacts
231269
/-- In an emetric space, the type of non-empty compact subsets is an emetric space,
232270
where the edistance is the Hausdorff edistance -/
233271
instance emetricSpace : EMetricSpace (NonemptyCompacts α) where
234-
edist s t := hausdorffEdist (s : Set α) t
235-
edist_self _ := hausdorffEdist_self
236-
edist_comm _ _ := hausdorffEdist_comm
237-
edist_triangle _ _ _ := hausdorffEdist_triangle
272+
__ := PseudoEMetricSpace.hausdorff.induced SetLike.coe
238273
eq_of_edist_eq_zero {s t} h := NonemptyCompacts.ext <| by
239274
have : closure (s : Set α) = closure t := hausdorffEdist_zero_iff_closure_eq_closure.1 h
240275
rwa [s.isCompact.isClosed.closure_eq, t.isCompact.isClosed.closure_eq] at this
@@ -244,21 +279,18 @@ theorem isometry_toCloseds : Isometry (@NonemptyCompacts.toCloseds α _ _) :=
244279
fun _ _ => rfl
245280

246281
/-- `NonemptyCompacts.toCloseds` is a uniform embedding (as it is an isometry) -/
247-
theorem isUniformEmbedding_toCloseds :
248-
IsUniformEmbedding (@NonemptyCompacts.toCloseds α _ _) :=
249-
isometry_toCloseds.isUniformEmbedding
282+
@[deprecated (since := "2025-11-05")]
283+
alias isUniformEmbedding_toCloseds := TopologicalSpace.NonemptyCompacts.isUniformEmbedding_toCloseds
250284

251285
@[deprecated (since := "2025-08-20")]
252-
alias ToCloseds.isUniformEmbedding := isUniformEmbedding_toCloseds
286+
alias ToCloseds.isUniformEmbedding := TopologicalSpace.NonemptyCompacts.isUniformEmbedding_toCloseds
253287

254288
/-- `NonemptyCompacts.toCloseds` is continuous (as it is an isometry) -/
255-
@[fun_prop]
256-
theorem continuous_toCloseds : Continuous (@NonemptyCompacts.toCloseds α _ _) :=
257-
isometry_toCloseds.continuous
289+
@[deprecated (since := "2025-11-05")]
290+
alias continuous_toCloseds := TopologicalSpace.NonemptyCompacts.continuous_toCloseds
258291

259-
lemma isClosed_subsets_of_isClosed (hs : IsClosed s) :
260-
IsClosed {A : NonemptyCompacts α | (A : Set α) ⊆ s} :=
261-
(Closeds.isClosed_subsets_of_isClosed hs).preimage continuous_toCloseds
292+
@[deprecated (since := "2025-11-05")]
293+
alias isClosed_subsets_of_isClosed := TopologicalSpace.NonemptyCompacts.isClosed_subsets_of_isClosed
262294

263295
/-- The range of `NonemptyCompacts.toCloseds` is closed in a complete space -/
264296
theorem isClosed_in_closeds [CompleteSpace α] :

Mathlib/Topology/Order.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -809,6 +809,10 @@ theorem isOpen_induced_eq {s : Set α} :
809809
theorem isOpen_induced {s : Set β} (h : IsOpen s) : IsOpen[induced f t] (f ⁻¹' s) :=
810810
⟨s, h, rfl⟩
811811

812+
theorem isClosed_induced {s : Set β} (h : IsClosed s) : IsClosed[induced f t] (f ⁻¹' s) := by
813+
simp_rw [← isOpen_compl_iff]
814+
exact isOpen_induced h.isOpen_compl
815+
812816
theorem map_nhds_induced_eq (a : α) : map f (@nhds α (induced f t) a) = 𝓝[range f] f a := by
813817
rw [nhds_induced, Filter.map_comap, nhdsWithin]
814818

Mathlib/Topology/Sets/Compacts.lean

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -239,10 +239,16 @@ protected theorem isCompact (s : NonemptyCompacts α) : IsCompact (s : Set α) :
239239
protected theorem nonempty (s : NonemptyCompacts α) : (s : Set α).Nonempty :=
240240
s.nonempty'
241241

242+
theorem toCompacts_injective : Function.Injective (toCompacts (α := α)) :=
243+
.of_comp (f := SetLike.coe) SetLike.coe_injective
244+
242245
/-- Reinterpret a nonempty compact as a closed set. -/
243246
def toCloseds [T2Space α] (s : NonemptyCompacts α) : Closeds α :=
244247
⟨s, s.isCompact.isClosed⟩
245248

249+
theorem toCloseds_injective [T2Space α] : Function.Injective (toCloseds (α := α)) :=
250+
.of_comp (f := SetLike.coe) SetLike.coe_injective
251+
246252
@[ext]
247253
protected theorem ext {s t : NonemptyCompacts α} (h : (s : Set α) = t) : s = t :=
248254
SetLike.ext' h

0 commit comments

Comments
 (0)