@@ -7,6 +7,7 @@ Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébas
7
7
import data.int.interval
8
8
import topology.algebra.order.compact
9
9
import topology.metric_space.emetric_space
10
+ import topology.uniform_space.complete_separated
10
11
11
12
/-!
12
13
# Metric spaces
@@ -85,7 +86,7 @@ uniform_space.of_core (uniform_space.core_of_dist dist dist_self dist_comm dist_
85
86
86
87
/-- The distance function (given an ambient metric space on `α`), which returns
87
88
a nonnegative real number `dist x y` given `x y : α`. -/
88
- class has_dist (α : Type *) := (dist : α → α → ℝ)
89
+ @[ext] class has_dist (α : Type *) := (dist : α → α → ℝ)
89
90
90
91
export has_dist (dist)
91
92
@@ -126,6 +127,22 @@ class pseudo_metric_space (α : Type u) extends has_dist α : Type u :=
126
127
(to_uniform_space : uniform_space α := uniform_space_of_dist dist dist_self dist_comm dist_triangle)
127
128
(uniformity_dist : 𝓤 α = ⨅ ε>0 , 𝓟 {p:α×α | dist p.1 p.2 < ε} . control_laws_tac)
128
129
130
+ /-- Two pseudo metric space structures with the same distance function coincide. -/
131
+ @[ext] lemma pseudo_metric_space.ext {α : Type *} {m m' : pseudo_metric_space α}
132
+ (h : m.to_has_dist = m'.to_has_dist) : m = m' :=
133
+ begin
134
+ unfreezingI { rcases m, rcases m' },
135
+ dsimp at h,
136
+ unfreezingI { subst h },
137
+ congr,
138
+ { ext x y : 2 ,
139
+ dsimp at m_edist_dist m'_edist_dist,
140
+ simp [m_edist_dist, m'_edist_dist] },
141
+ { dsimp at m_uniformity_dist m'_uniformity_dist,
142
+ rw ← m'_uniformity_dist at m_uniformity_dist,
143
+ exact uniform_space_eq m_uniformity_dist }
144
+ end
145
+
129
146
variables [pseudo_metric_space α]
130
147
131
148
@[priority 100 ] -- see Note [lower instance priority]
@@ -879,6 +896,18 @@ lemma exists_closed_ball_inter_eq_singleton_of_discrete [discrete_topology s] {x
879
896
∃ ε > 0 , metric.closed_ball x ε ∩ s = {x} :=
880
897
nhds_basis_closed_ball.exists_inter_eq_singleton_of_mem_discrete hx
881
898
899
+ lemma _root_.dense.exists_dist_lt {s : set α} (hs : dense s) (x : α) {ε : ℝ} (hε : 0 < ε) :
900
+ ∃ y ∈ s, dist x y < ε :=
901
+ begin
902
+ have : (ball x ε).nonempty, by simp [hε],
903
+ simpa only [mem_ball'] using hs.exists_mem_open is_open_ball this
904
+ end
905
+
906
+ lemma _root_.dense_range.exists_dist_lt {β : Type *} {f : β → α} (hf : dense_range f)
907
+ (x : α) {ε : ℝ} (hε : 0 < ε) :
908
+ ∃ y, dist x (f y) < ε :=
909
+ exists_range_iff.1 (hf.exists_dist_lt x hε)
910
+
882
911
end metric
883
912
884
913
open metric
@@ -967,6 +996,12 @@ def pseudo_metric_space.replace_uniformity {α} [U : uniform_space α] (m : pseu
967
996
to_uniform_space := U,
968
997
uniformity_dist := H.trans pseudo_metric_space.uniformity_dist }
969
998
999
+ lemma pseudo_metric_space.replace_uniformity_eq {α} [U : uniform_space α]
1000
+ (m : pseudo_metric_space α)
1001
+ (H : @uniformity _ U = @uniformity _ pseudo_emetric_space.to_uniform_space') :
1002
+ m.replace_uniformity H = m :=
1003
+ by { ext, refl }
1004
+
970
1005
/-- One gets a pseudometric space from an emetric space if the edistance
971
1006
is everywhere finite, by pushing the edistance to reals. We set it up so that the edist and the
972
1007
uniformity are defeq in the pseudometric space and the pseudoemetric space. In this definition, the
@@ -1170,17 +1205,23 @@ theorem metric.cauchy_seq_iff' {u : β → α} :
1170
1205
cauchy_seq u ↔ ∀ε>0 , ∃N, ∀n≥N, dist (u n) (u N) < ε :=
1171
1206
uniformity_basis_dist.cauchy_seq_iff'
1172
1207
1208
+ /-- If the distance between `s n` and `s m`, `n ≤ m` is bounded above by `b n`
1209
+ and `b` converges to zero, then `s` is a Cauchy sequence. -/
1210
+ lemma cauchy_seq_of_le_tendsto_0' {s : β → α} (b : β → ℝ)
1211
+ (h : ∀ n m : β, n ≤ m → dist (s n) (s m) ≤ b n) (h₀ : tendsto b at_top (𝓝 0 )) :
1212
+ cauchy_seq s :=
1213
+ metric.cauchy_seq_iff'.2 $ λ ε ε0 ,
1214
+ (h₀.eventually (gt_mem_nhds ε0 )).exists.imp $ λ N hN n hn,
1215
+ calc dist (s n) (s N) = dist (s N) (s n) : dist_comm _ _
1216
+ ... ≤ b N : h _ _ hn
1217
+ ... < ε : hN
1218
+
1173
1219
/-- If the distance between `s n` and `s m`, `n, m ≥ N` is bounded above by `b N`
1174
1220
and `b` converges to zero, then `s` is a Cauchy sequence. -/
1175
1221
lemma cauchy_seq_of_le_tendsto_0 {s : β → α} (b : β → ℝ)
1176
- (h : ∀ n m N : β, N ≤ n → N ≤ m → dist (s n) (s m) ≤ b N) (h₀ : tendsto b at_top (nhds 0 )) :
1222
+ (h : ∀ n m N : β, N ≤ n → N ≤ m → dist (s n) (s m) ≤ b N) (h₀ : tendsto b at_top (𝓝 0 )) :
1177
1223
cauchy_seq s :=
1178
- metric.cauchy_seq_iff.2 $ λ ε ε0 ,
1179
- (metric.tendsto_at_top.1 h₀ ε ε0 ).imp $ λ N hN m hm n hn,
1180
- calc dist (s m) (s n) ≤ b N : h m n N hm hn
1181
- ... ≤ |b N| : le_abs_self _
1182
- ... = dist (b N) 0 : by rw real.dist_0_eq_abs; refl
1183
- ... < ε : (hN _ (le_refl N))
1224
+ cauchy_seq_of_le_tendsto_0' b (λ n m hnm, h _ _ _ le_rfl hnm) h₀
1184
1225
1185
1226
/-- A Cauchy sequence on the natural numbers is bounded. -/
1186
1227
theorem cauchy_seq_bdd {u : ℕ → α} (hu : cauchy_seq u) :
@@ -2086,6 +2127,41 @@ diam_le_of_subset_closed_ball h subset.rfl
2086
2127
lemma diam_ball {r : ℝ} (h : 0 ≤ r) : diam (ball x r) ≤ 2 * r :=
2087
2128
diam_le_of_subset_closed_ball h ball_subset_closed_ball
2088
2129
2130
+ /-- If a family of complete sets with diameter tending to `0` is such that each finite intersection
2131
+ is nonempty, then the total intersection is also nonempty. -/
2132
+ lemma _root_.is_complete.nonempty_Inter_of_nonempty_bInter {s : ℕ → set α} (h0 : is_complete (s 0 ))
2133
+ (hs : ∀ n, is_closed (s n)) (h's : ∀ n, bounded (s n)) (h : ∀ N, (⋂ n ≤ N, s n).nonempty)
2134
+ (h' : tendsto (λ n, diam (s n)) at_top (𝓝 0 )) :
2135
+ (⋂ n, s n).nonempty :=
2136
+ begin
2137
+ let u := λ N, (h N).some,
2138
+ have I : ∀ n N, n ≤ N → u N ∈ s n,
2139
+ { assume n N hn,
2140
+ apply mem_of_subset_of_mem _ ((h N).some_spec),
2141
+ assume x hx,
2142
+ simp only [mem_Inter] at hx,
2143
+ exact hx n hn },
2144
+ have : ∀ n, u n ∈ s 0 := λ n, I 0 n (zero_le _),
2145
+ have : cauchy_seq u,
2146
+ { apply cauchy_seq_of_le_tendsto_0 _ _ h',
2147
+ assume m n N hm hn,
2148
+ exact dist_le_diam_of_mem (h's N) (I _ _ hm) (I _ _ hn) },
2149
+ obtain ⟨x, hx, xlim⟩ : ∃ (x : α) (H : x ∈ s 0 ), tendsto (λ (n : ℕ), u n) at_top (𝓝 x) :=
2150
+ cauchy_seq_tendsto_of_is_complete h0 (λ n, I 0 n (zero_le _)) this ,
2151
+ refine ⟨x, mem_Inter.2 (λ n, _)⟩,
2152
+ apply (hs n).mem_of_tendsto xlim,
2153
+ filter_upwards [Ici_mem_at_top n] with p hp,
2154
+ exact I n p hp,
2155
+ end
2156
+
2157
+ /-- In a complete space, if a family of closed sets with diameter tending to `0` is such that each
2158
+ finite intersection is nonempty, then the total intersection is also nonempty. -/
2159
+ lemma nonempty_Inter_of_nonempty_bInter [complete_space α] {s : ℕ → set α}
2160
+ (hs : ∀ n, is_closed (s n)) (h's : ∀ n, bounded (s n)) (h : ∀ N, (⋂ n ≤ N, s n).nonempty)
2161
+ (h' : tendsto (λ n, diam (s n)) at_top (𝓝 0 )) :
2162
+ (⋂ n, s n).nonempty :=
2163
+ (hs 0 ).is_complete.nonempty_Inter_of_nonempty_bInter hs h's h h'
2164
+
2089
2165
end diam
2090
2166
2091
2167
end metric
@@ -2130,6 +2206,16 @@ end int
2130
2206
class metric_space (α : Type u) extends pseudo_metric_space α : Type u :=
2131
2207
(eq_of_dist_eq_zero : ∀ {x y : α}, dist x y = 0 → x = y)
2132
2208
2209
+ /-- Two metric space structures with the same distance coincide. -/
2210
+ @[ext] lemma metric_space.ext {α : Type *} {m m' : metric_space α}
2211
+ (h : m.to_has_dist = m'.to_has_dist) : m = m' :=
2212
+ begin
2213
+ have h' : m.to_pseudo_metric_space = m'.to_pseudo_metric_space := pseudo_metric_space.ext h,
2214
+ unfreezingI { rcases m, rcases m' },
2215
+ dsimp at h',
2216
+ unfreezingI { subst h' },
2217
+ end
2218
+
2133
2219
/-- Construct a metric space structure whose underlying topological space structure
2134
2220
(definitionally) agrees which a pre-existing topology which is compatible with a given distance
2135
2221
function. -/
@@ -2213,7 +2299,7 @@ instance metric_space.to_separated : separated_space γ :=
2213
2299
separated_def.2 $ λ x y h, eq_of_forall_dist_le $
2214
2300
λ ε ε0 , le_of_lt (h _ (dist_mem_uniformity ε0 ))
2215
2301
2216
- /-- If a `pseudo_metric_space` is separated, then it is a `metric_space`. -/
2302
+ /-- If a `pseudo_metric_space` is separated, then it is a `metric_space`. -/
2217
2303
def of_t2_pseudo_metric_space {α : Type *} [pseudo_metric_space α]
2218
2304
(h : separated_space α) : metric_space α :=
2219
2305
{ eq_of_dist_eq_zero := λ x y hdist,
@@ -2257,6 +2343,30 @@ def metric_space.replace_uniformity {γ} [U : uniform_space γ] (m : metric_spac
2257
2343
{ eq_of_dist_eq_zero := @eq_of_dist_eq_zero _ _,
2258
2344
..pseudo_metric_space.replace_uniformity m.to_pseudo_metric_space H, }
2259
2345
2346
+ lemma metric_space.replace_uniformity_eq {γ} [U : uniform_space γ] (m : metric_space γ)
2347
+ (H : @uniformity _ U = @uniformity _ emetric_space.to_uniform_space') :
2348
+ m.replace_uniformity H = m :=
2349
+ by { ext, refl }
2350
+
2351
+ /-- Build a new metric space from an old one where the bundled topological structure is provably
2352
+ (but typically non-definitionaly) equal to some given topological structure.
2353
+ See Note [forgetful inheritance].
2354
+ -/
2355
+ @[reducible] def metric_space.replace_topology {γ} [U : topological_space γ] (m : metric_space γ)
2356
+ (H : U = m.to_pseudo_metric_space.to_uniform_space.to_topological_space) :
2357
+ metric_space γ :=
2358
+ begin
2359
+ let t := m.to_pseudo_metric_space.to_uniform_space.replace_topology H,
2360
+ letI : uniform_space γ := t,
2361
+ have : @uniformity _ t = @uniformity _ m.to_pseudo_metric_space.to_uniform_space := rfl,
2362
+ exact m.replace_uniformity this
2363
+ end
2364
+
2365
+ lemma metric_space.replace_topology_eq {γ} [U : topological_space γ] (m : metric_space γ)
2366
+ (H : U = m.to_pseudo_metric_space.to_uniform_space.to_topological_space) :
2367
+ m.replace_topology H = m :=
2368
+ by { ext, refl }
2369
+
2260
2370
/-- One gets a metric space from an emetric space if the edistance
2261
2371
is everywhere finite, by pushing the edistance to reals. We set it up so that the edist and the
2262
2372
uniformity are defeq in the metric space and the emetric space. In this definition, the distance
@@ -2288,29 +2398,53 @@ def metric_space.induced {γ β} (f : γ → β) (hf : function.injective f)
2288
2398
2289
2399
/-- Pull back a metric space structure by a uniform embedding. This is a version of
2290
2400
`metric_space.induced` useful in case if the domain already has a `uniform_space` structure. -/
2291
- def uniform_embedding.comap_metric_space {α β} [uniform_space α] [metric_space β] (f : α → β)
2292
- (h : uniform_embedding f) : metric_space α :=
2401
+ @[reducible] def uniform_embedding.comap_metric_space
2402
+ {α β} [uniform_space α] [metric_space β] (f : α → β) (h : uniform_embedding f) :
2403
+ metric_space α :=
2293
2404
(metric_space.induced f h.inj ‹_›).replace_uniformity h.comap_uniformity.symm
2294
2405
2406
+ /-- Pull back a metric space structure by an embedding. This is a version of
2407
+ `metric_space.induced` useful in case if the domain already has a `topological_space` structure. -/
2408
+ @[reducible] def embedding.comap_metric_space
2409
+ {α β} [topological_space α] [metric_space β] (f : α → β) (h : embedding f) :
2410
+ metric_space α :=
2411
+ begin
2412
+ letI : uniform_space α := embedding.comap_uniform_space f h,
2413
+ exact uniform_embedding.comap_metric_space f (h.to_uniform_embedding f),
2414
+ end
2415
+
2295
2416
instance subtype.metric_space {α : Type *} {p : α → Prop } [t : metric_space α] :
2296
2417
metric_space (subtype p) :=
2297
2418
metric_space.induced coe (λ x y, subtype.ext) t
2298
2419
2299
2420
theorem subtype.dist_eq {p : α → Prop } (x y : subtype p) : dist x y = dist (x : α) y := rfl
2300
2421
2422
+ local attribute [instance] filter.unique
2423
+
2301
2424
instance : metric_space empty :=
2302
2425
{ dist := λ _ _, 0 ,
2303
2426
dist_self := λ _, rfl,
2304
2427
dist_comm := λ _ _, rfl,
2305
2428
eq_of_dist_eq_zero := λ _ _ _, subsingleton.elim _ _,
2306
- dist_triangle := λ _ _ _, show (0 :ℝ) ≤ 0 + 0 , by rw add_zero, }
2429
+ dist_triangle := λ _ _ _, show (0 :ℝ) ≤ 0 + 0 , by rw add_zero,
2430
+ to_uniform_space := empty.uniform_space,
2431
+ uniformity_dist := subsingleton.elim _ _ }
2307
2432
2308
- instance : metric_space punit :=
2433
+ instance : metric_space punit.{u + 1 } :=
2309
2434
{ dist := λ _ _, 0 ,
2310
2435
dist_self := λ _, rfl,
2311
2436
dist_comm := λ _ _, rfl,
2312
2437
eq_of_dist_eq_zero := λ _ _ _, subsingleton.elim _ _,
2313
- dist_triangle := λ _ _ _, show (0 :ℝ) ≤ 0 + 0 , by rw add_zero, }
2438
+ dist_triangle := λ _ _ _, show (0 :ℝ) ≤ 0 + 0 , by rw add_zero,
2439
+ to_uniform_space := punit.uniform_space,
2440
+ uniformity_dist :=
2441
+ begin
2442
+ simp only,
2443
+ haveI : ne_bot (⨅ ε > (0 : ℝ), 𝓟 {p : punit.{u + 1 } × punit.{u + 1 } | 0 < ε}),
2444
+ { exact @uniformity.ne_bot _ (uniform_space_of_dist (λ _ _, 0 ) (λ _, rfl) (λ _ _, rfl)
2445
+ (λ _ _ _, by rw zero_add)) _ },
2446
+ refine (eq_top_of_ne_bot _).trans (eq_top_of_ne_bot _).symm,
2447
+ end }
2314
2448
2315
2449
section real
2316
2450
0 commit comments