Skip to content

Commit

Permalink
feat(topology/metric_space): use weaker TC assumptions (#14316)
Browse files Browse the repository at this point in the history
Assume `t0_space` instead of `separated_space` in `metric.of_t0_pseudo_metric_space` and `emetric.of_t0_pseudo_emetric_space` (both definition used to have `t2` in their names).
  • Loading branch information
urkud committed May 22, 2022
1 parent 9e9a2c9 commit 005df45
Show file tree
Hide file tree
Showing 3 changed files with 24 additions and 32 deletions.
8 changes: 4 additions & 4 deletions src/analysis/normed/group/basic.lean
Expand Up @@ -1028,18 +1028,18 @@ lemma semi_normed_group.mem_closure_iff {s : set E} {x : E} :
x ∈ closure s ↔ ∀ ε > 0, ∃ y ∈ s, ∥x - y∥ < ε :=
by simp [metric.mem_closure_iff, dist_eq_norm]

lemma norm_le_zero_iff' [separated_space E] {g : E} :
lemma norm_le_zero_iff' [t0_space E] {g : E} :
∥g∥ ≤ 0 ↔ g = 0 :=
begin
letI : normed_group E := { to_metric_space := of_t2_pseudo_metric_space ‹_›,
letI : normed_group E := { to_metric_space := metric.of_t0_pseudo_metric_space E,
.. ‹semi_normed_group E› },
rw [← dist_zero_right], exact dist_le_zero
end

lemma norm_eq_zero_iff' [separated_space E] {g : E} : ∥g∥ = 0 ↔ g = 0 :=
lemma norm_eq_zero_iff' [t0_space E] {g : E} : ∥g∥ = 0 ↔ g = 0 :=
(norm_nonneg g).le_iff_eq.symm.trans norm_le_zero_iff'

lemma norm_pos_iff' [separated_space E] {g : E} : 0 < ∥g∥ ↔ g ≠ 0 :=
lemma norm_pos_iff' [t0_space E] {g : E} : 0 < ∥g∥ ↔ g ≠ 0 :=
by rw [← not_le, norm_le_zero_iff']

lemma cauchy_seq_sum_of_eventually_eq {u v : ℕ → E} {N : ℕ} (huv : ∀ n ≥ N, u n = v n)
Expand Down
24 changes: 11 additions & 13 deletions src/topology/metric_space/basic.lean
Expand Up @@ -207,7 +207,7 @@ def pseudo_metric_space.of_metrizable {α : Type*} [topological_space α] (dist
(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 ∈ s, ∃ ε > 0, ∀ y, dist x y < ε → y ∈ s) :
pseudo_metric_space α :=
pseudo_metric_space α :=
{ dist := dist,
dist_self := dist_self,
dist_comm := dist_comm,
Expand Down Expand Up @@ -1022,6 +1022,10 @@ by { convert metric.emetric_closed_ball ε.2, simp }
@[simp] lemma metric.emetric_ball_top (x : α) : emetric.ball x ⊤ = univ :=
eq_univ_of_forall $ λ y, edist_lt_top _ _

lemma metric.indistinguishable_iff {x y : α} : indistinguishable x y ↔ dist x y = 0 :=
by rw [emetric.indistinguishable_iff, edist_nndist, dist_nndist, ennreal.coe_eq_zero,
nnreal.coe_eq_zero]

/-- 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].
Expand Down Expand Up @@ -2475,26 +2479,20 @@ begin
end

@[priority 100] -- see Note [lower instance priority]
instance metric_space.to_separated : separated_space γ :=
instance _root_.metric_space.to_separated : separated_space γ :=
separated_def.2 $ λ x y h, eq_of_forall_dist_le $
λ ε ε0, le_of_lt (h _ (dist_mem_uniformity ε0))

/-- If a `pseudo_metric_space` is separated, then it is a `metric_space`. -/
def of_t2_pseudo_metric_space {α : Type*} [pseudo_metric_space α]
(h : separated_space α) : metric_space α :=
{ eq_of_dist_eq_zero := λ x y hdist,
begin
refine separated_def.1 h x y (λ s hs, _),
obtain ⟨ε, hε, H⟩ := mem_uniformity_dist.1 hs,
exact H (show dist x y < ε, by rwa [hdist])
end
/-- If a `pseudo_metric_space` is a T₀ space, then it is a `metric_space`. -/
def of_t0_pseudo_metric_space (α : Type*) [pseudo_metric_space α] [t0_space α] :
metric_space α :=
{ eq_of_dist_eq_zero := λ x y hdist, indistinguishable.eq $ metric.indistinguishable_iff.2 hdist,
..‹pseudo_metric_space α› }

/-- A metric space induces an emetric space -/
@[priority 100] -- see Note [lower instance priority]
instance metric_space.to_emetric_space : emetric_space γ :=
{ eq_of_edist_eq_zero := assume x y h, by simpa [edist_dist] using h,
..pseudo_metric_space.to_pseudo_emetric_space, }
emetric.of_t0_pseudo_emetric_space γ

lemma is_closed_of_pairwise_le_dist {s : set γ} {ε : ℝ} (hε : 0 < ε)
(hs : s.pairwise (λ x y, ε ≤ dist x y)) : is_closed s :=
Expand Down
24 changes: 9 additions & 15 deletions src/topology/metric_space/emetric_space.lean
Expand Up @@ -613,6 +613,9 @@ theorem tendsto_at_top [nonempty β] [semilattice_sup β] {u : β → α} {a :
(at_top_basis.tendsto_iff nhds_basis_eball).trans $
by simp only [exists_prop, true_and, mem_Ici, mem_ball]

theorem indistinguishable_iff : indistinguishable x y ↔ edist x y = 0 :=
by simp [indistinguishable_iff_closure, mem_closure_iff, edist_comm, forall_lt_iff_le']

/-- In a pseudoemetric space, Cauchy sequences are characterized by the fact that, eventually,
the pseudoedistance between its elements is arbitrarily small -/
@[nolint ge_or_gt] -- see Note [nolint_ge]
Expand Down Expand Up @@ -698,6 +701,7 @@ to avoid a loop with `sigma_compact_space_of_locally_compact_second_countable`.
lemma second_countable_of_sigma_compact [sigma_compact_space α] :
second_countable_topology α :=
begin

suffices : separable_space α, by exactI uniform_space.second_countable_of_separable α,
choose T hTsub hTc hsubT
using λ n, subset_countable_closure_of_compact (is_compact_compact_covering α n),
Expand Down Expand Up @@ -883,15 +887,10 @@ instance to_separated : separated_space γ :=
separated_def.2 $ λ x y h, eq_of_forall_edist_le $
λ ε ε0, le_of_lt (h _ (edist_mem_uniformity ε0))

/-- If a `pseudo_emetric_space` is separated, then it is an `emetric_space`. -/
def emetric_of_t2_pseudo_emetric_space {α : Type*} [pseudo_emetric_space α]
(h : separated_space α) : emetric_space α :=
{ eq_of_edist_eq_zero := λ x y hdist,
begin
refine separated_def.1 h x y (λ s hs, _),
obtain ⟨ε, hε, H⟩ := mem_uniformity_edist.1 hs,
exact H (show edist x y < ε, by rwa [hdist])
end
/-- If a `pseudo_emetric_space` is a T₀ space, then it is an `emetric_space`. -/
def emetric.of_t0_pseudo_emetric_space (α : Type*) [pseudo_emetric_space α] [t0_space α] :
emetric_space α :=
{ eq_of_edist_eq_zero := λ x y hdist, indistinguishable.eq $ emetric.indistinguishable_iff.2 hdist,
..‹pseudo_emetric_space α› }

/-- Auxiliary function to replace the uniformity on an emetric space with
Expand Down Expand Up @@ -998,12 +997,7 @@ lemma diam_eq_zero_iff : diam s = 0 ↔ s.subsingleton :=
⟨λ h x hx y hy, edist_le_zero.1 $ h ▸ edist_le_diam_of_mem hx hy, diam_subsingleton⟩

lemma diam_pos_iff : 0 < diam s ↔ ∃ (x ∈ s) (y ∈ s), x ≠ y :=
begin
have := not_congr (@diam_eq_zero_iff _ _ s),
dunfold set.subsingleton at this,
push_neg at this,
simpa only [pos_iff_ne_zero, exists_prop] using this
end
by simp only [pos_iff_ne_zero, ne.def, diam_eq_zero_iff, set.subsingleton, not_forall]

end diam

Expand Down

0 comments on commit 005df45

Please sign in to comment.