Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit ad2baee

Browse files
committed
feat(topology/separation): t0_space and t1_space for α × β and Π i, α i (#14418)
1 parent f13e5df commit ad2baee

File tree

3 files changed

+31
-19
lines changed

3 files changed

+31
-19
lines changed

src/analysis/normed_space/pi_Lp.lean

Lines changed: 1 addition & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -206,9 +206,7 @@ instance [Π i, pseudo_emetric_space (β i)] : pseudo_emetric_space (pi_Lp p β)
206206
/-- emetric space instance on the product of finitely many emetric spaces, using the `L^p`
207207
edistance, and having as uniformity the product uniformity. -/
208208
instance [Π i, emetric_space (α i)] : emetric_space (pi_Lp p α) :=
209-
@emetric.of_t0_pseudo_emetric_space _ _ $
210-
-- TODO: add `Pi.t0_space`
211-
by { haveI : t2_space (pi_Lp p α) := Pi.t2_space, apply_instance }
209+
@emetric.of_t0_pseudo_emetric_space (pi_Lp p α) _ pi.t0_space
212210

213211
variables {p β}
214212
lemma edist_eq [Π i, pseudo_emetric_space (β i)] (x y : pi_Lp p β) :

src/data/set/prod.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -350,6 +350,9 @@ by { ext, simp [pi] }
350350

351351
lemma singleton_pi' (i : ι) (t : Π i, set (α i)) : pi {i} t = {x | x i ∈ t i} := singleton_pi i t
352352

353+
lemma univ_pi_singleton (f : Π i, α i) : pi univ (λ i, {f i}) = ({f} : set (Π i, α i)) :=
354+
ext $ λ g, by simp [funext_iff]
355+
353356
lemma pi_if {p : ι → Prop} [h : decidable_pred p] (s : set ι) (t₁ t₂ : Π i, set (α i)) :
354357
pi s (λ i, if p i then t₁ i else t₂ i) = pi {i ∈ s | p i} t₁ ∩ pi {i ∈ s | ¬ p i} t₂ :=
355358
begin

src/topology/separation.lean

Lines changed: 27 additions & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -176,6 +176,11 @@ lemma indistinguishable_iff_nhds_eq {x y : α} : indistinguishable x y ↔ 𝓝
176176

177177
alias indistinguishable_iff_nhds_eq ↔ indistinguishable.nhds_eq _
178178

179+
lemma indistinguishable.map [topological_space β] {x y : α} {f : α → β}
180+
(h : indistinguishable x y) (hf : continuous f) :
181+
indistinguishable (f x) (f y) :=
182+
λ U hU, h (f ⁻¹' U) (hU.preimage hf)
183+
179184
lemma t0_space_iff_distinguishable (α : Type u) [topological_space α] :
180185
t0_space α ↔ ∀ (x y : α), x ≠ y → ¬ indistinguishable x y :=
181186
begin
@@ -185,6 +190,10 @@ begin
185190
simp_rw xor_iff_not_iff,
186191
end
187192

193+
lemma t0_space_iff_indistinguishable (α : Type u) [topological_space α] :
194+
t0_space α ↔ ∀ (x y : α), indistinguishable x y → x = y :=
195+
(t0_space_iff_distinguishable α).trans $ forall₂_congr $ λ a b, not_imp_not
196+
188197
@[simp] lemma nhds_eq_nhds_iff [t0_space α] {a b : α} : 𝓝 a = 𝓝 b ↔ a = b :=
189198
function.injective.eq_iff $ λ x y h, of_not_not $
190199
λ hne, (t0_space_iff_distinguishable α).mp ‹_› x y hne (indistinguishable_iff_nhds_eq.mpr h)
@@ -288,19 +297,16 @@ instance subtype.t0_space [t0_space α] {p : α → Prop} : t0_space (subtype p)
288297
embedding_subtype_coe.t0_space
289298

290299
theorem t0_space_iff_or_not_mem_closure (α : Type u) [topological_space α] :
291-
t0_space α ↔ (∀ a b : α, (a ≠ b) → (a ∉ closure ({b} : set α) ∨ b ∉ closure ({a} : set α))) :=
292-
begin
293-
simp only [← not_and_distrib, t0_space_def, not_and],
294-
refine forall₃_congr (λ a b _, ⟨_, λ h, _⟩),
295-
{ rintro ⟨s, h₁, (⟨h₂, h₃ : b ∈ sᶜ⟩|⟨h₂, h₃ : a ∈ sᶜ⟩)⟩ ha hb; rw ← is_closed_compl_iff at h₁,
296-
{ exact (is_closed.closure_subset_iff h₁).mpr (set.singleton_subset_iff.mpr h₃) ha h₂ },
297-
{ exact (is_closed.closure_subset_iff h₁).mpr (set.singleton_subset_iff.mpr h₃) hb h₂ } },
298-
{ by_cases h' : a ∈ closure ({b} : set α),
299-
{ exact ⟨(closure {a})ᶜ, is_closed_closure.1,
300-
or.inr ⟨h h', not_not.mpr (subset_closure (set.mem_singleton a))⟩⟩ },
301-
{ exact ⟨(closure {b})ᶜ, is_closed_closure.1,
302-
or.inl ⟨h', not_not.mpr (subset_closure (set.mem_singleton b))⟩⟩ } }
303-
end
300+
t0_space α ↔ (∀ a b : α, a ≠ b → (a ∉ closure ({b} : set α) ∨ b ∉ closure ({a} : set α))) :=
301+
by simp only [t0_space_iff_distinguishable, indistinguishable_iff_closure, not_and_distrib]
302+
303+
instance [topological_space β] [t0_space α] [t0_space β] : t0_space (α × β) :=
304+
(t0_space_iff_indistinguishable _).2 $
305+
λ x y h, prod.ext (h.map continuous_fst).eq (h.map continuous_snd).eq
306+
307+
instance {ι : Type*} {π : ι → Type*} [Π i, topological_space (π i)] [Π i, t0_space (π i)] :
308+
t0_space (Π i, π i) :=
309+
(t0_space_iff_indistinguishable _).2 $ λ x y h, funext $ λ i, (h.map (continuous_apply i)).eq
304310

305311
/-- A T₁ space, also known as a Fréchet space, is a topological space
306312
where every singleton set is closed. Equivalently, for every pair
@@ -477,6 +483,13 @@ instance subtype.t1_space {α : Type u} [topological_space α] [t1_space α] {p
477483
t1_space (subtype p) :=
478484
embedding_subtype_coe.t1_space
479485

486+
instance [topological_space β] [t1_space α] [t1_space β] : t1_space (α × β) :=
487+
⟨λ ⟨a, b⟩, @singleton_prod_singleton _ _ a b ▸ is_closed_singleton.prod is_closed_singleton⟩
488+
489+
instance {ι : Type*} {π : ι → Type*} [Π i, topological_space (π i)] [Π i, t1_space (π i)] :
490+
t1_space (Π i, π i) :=
491+
⟨λ f, univ_pi_singleton f ▸ is_closed_set_pi (λ i hi, is_closed_singleton)⟩
492+
480493
@[priority 100] -- see Note [lower instance priority]
481494
instance t1_space.t0_space [t1_space α] : t0_space α :=
482495
⟨λ x y h, ⟨{z | z ≠ y}, is_open_ne, or.inl ⟨h, not_not_intro rfl⟩⟩⟩
@@ -501,9 +514,7 @@ hs.induction_on (by simp) $ λ x, by simp
501514

502515
lemma is_closed_map_const {α β} [topological_space α] [topological_space β] [t1_space β] {y : β} :
503516
is_closed_map (function.const α y) :=
504-
begin
505-
apply is_closed_map.of_nonempty, intros s hs h2s, simp_rw [h2s.image_const, is_closed_singleton]
506-
end
517+
is_closed_map.of_nonempty $ λ s hs h2s, by simp_rw [h2s.image_const, is_closed_singleton]
507518

508519
lemma bInter_basis_nhds [t1_space α] {ι : Sort*} {p : ι → Prop} {s : ι → set α} {x : α}
509520
(h : (𝓝 x).has_basis p s) : (⋂ i (h : p i), s i) = {x} :=

0 commit comments

Comments
 (0)