Skip to content

Commit 048fe93

Browse files
authored
feat: Urysohn lemma in regular locally compact spaces (#8124)
Also change a few theorems to instances, as instance loops are not a problem any more. --- <!-- The text above the `---` will become the commit message when your PR is merged. Please leave a blank newline before the `---`, otherwise GitHub will format the text above it as a title. To indicate co-authors, include lines at the bottom of the commit message (that is, before the `---`) using the following format: Co-authored-by: Author Name <author@email.com> Any other comments you want to keep out of the PR commit should go below the `---`, and placed outside this HTML comment, or else they will be invisible to reviewers. If this PR depends on other PRs, please list them below this comment, using the following format: - [ ] depends on: #abc [optional extra text] - [ ] depends on: #xyz [optional extra text] --> [![Open in Gitpod](https://gitpod.io/button/open-in-gitpod.svg)](https://gitpod.io/from-referrer/)
1 parent b6ec745 commit 048fe93

File tree

4 files changed

+155
-57
lines changed

4 files changed

+155
-57
lines changed

Mathlib/Topology/EMetricSpace/Basic.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -852,9 +852,9 @@ open TopologicalSpace
852852

853853
variable (α)
854854

855-
/-- A sigma compact pseudo emetric space has second countable topology. This is not an instance
856-
to avoid a loop with `sigmaCompactSpace_of_locally_compact_second_countable`. -/
857-
theorem secondCountable_of_sigmaCompact [SigmaCompactSpace α] : SecondCountableTopology α := by
855+
/-- A sigma compact pseudo emetric space has second countable topology. -/
856+
instance (priority := 90) secondCountable_of_sigmaCompact [SigmaCompactSpace α] :
857+
SecondCountableTopology α := by
858858
suffices SeparableSpace α by exact UniformSpace.secondCountable_of_separable α
859859
choose T _ hTc hsubT using fun n =>
860860
subset_countable_closure_of_compact (isCompact_compactCovering α n)

Mathlib/Topology/Metrizable/Urysohn.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -109,7 +109,8 @@ theorem exists_inducing_l_infty : ∃ f : X → ℕ →ᵇ ℝ, Inducing f := by
109109
/-- *Urysohn's metrization theorem* (Tychonoff's version):
110110
a regular topological space with second countable topology `X` is metrizable,
111111
i.e., there exists a pseudometric space structure that generates the same topology. -/
112-
lemma PseudoMetrizableSpace.of_regularSpace_secondCountableTopology : PseudoMetrizableSpace X :=
112+
instance (priority := 90) PseudoMetrizableSpace.of_regularSpace_secondCountableTopology :
113+
PseudoMetrizableSpace X :=
113114
let ⟨_, hf⟩ := exists_inducing_l_infty X
114115
hf.pseudoMetrizableSpace
115116

@@ -124,8 +125,7 @@ theorem exists_embedding_l_infty : ∃ f : X → ℕ →ᵇ ℝ, Embedding f :=
124125
/-- *Urysohn's metrization theorem* (Tychonoff's version): a T₃ topological space with second
125126
countable topology `X` is metrizable, i.e., there exists a metric space structure that generates the
126127
same topology. -/
127-
theorem metrizableSpace_of_t3_second_countable : MetrizableSpace X :=
128+
instance (priority := 90) metrizableSpace_of_t3_second_countable : MetrizableSpace X :=
128129
let ⟨_, hf⟩ := exists_embedding_l_infty X
129130
hf.metrizableSpace
130131
#align topological_space.metrizable_space_of_t3_second_countable TopologicalSpace.metrizableSpace_of_t3_second_countable
131-

Mathlib/Topology/Separation.lean

Lines changed: 14 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -1731,8 +1731,8 @@ section ClosableCompactSubsetOpenSpace
17311731

17321732
/-- A class of topological spaces in which, given a compact set included inside an open set, then
17331733
the closure of the compact set is also included in the open set.
1734-
Satisfied notably for T2 spaces and regular spaces, and useful
1735-
when discussing classes of regular measures. -/
1734+
Satisfied notably for T2 spaces and regular spaces, and useful when discussing classes of
1735+
regular measures. Equivalent to regularity among locally compact spaces. -/
17361736
class ClosableCompactSubsetOpenSpace (X : Type*) [TopologicalSpace X] : Prop :=
17371737
closure_subset_of_isOpen : ∀ (K U : Set X), IsCompact K → IsOpen U → K ⊆ U → closure K ⊆ U
17381738

@@ -1741,12 +1741,12 @@ theorem IsCompact.closure_subset_of_isOpen [ClosableCompactSubsetOpenSpace X]
17411741
closure s ⊆ u :=
17421742
ClosableCompactSubsetOpenSpace.closure_subset_of_isOpen s u hs hu h
17431743

1744-
instance [T2Space X] : ClosableCompactSubsetOpenSpace X :=
1744+
instance (priority := 150) [T2Space X] : ClosableCompactSubsetOpenSpace X :=
17451745
fun K _U K_comp _U_open KU ↦ by rwa [K_comp.isClosed.closure_eq]⟩
17461746

17471747
/-- In a (possibly non-Hausdorff) regular space, if a compact set `s` is contained in an
17481748
open set `u`, then its closure is also contained in `u`. -/
1749-
instance [RegularSpace X] : ClosableCompactSubsetOpenSpace X := by
1749+
instance (priority := 150) [RegularSpace X] : ClosableCompactSubsetOpenSpace X := by
17501750
refine ⟨fun s u hs hu h ↦ ?_⟩
17511751
obtain ⟨F, sF, F_closed, Fu⟩ : ∃ F, s ⊆ F ∧ IsClosed F ∧ F ⊆ u := by
17521752
apply hs.induction_on (p := fun t ↦ ∃ F, t ⊆ F ∧ IsClosed F ∧ F ⊆ u)
@@ -1779,6 +1779,16 @@ theorem exists_compact_closed_between [LocallyCompactSpace X] [ClosableCompactSu
17791779
· apply M_comp.closure_subset_of_isOpen hU
17801780
exact ML.trans (interior_subset.trans LU)
17811781

1782+
/-- A locally compact space with the `ClosableCompactSubsetOpenSpace` is `Regular`. -/
1783+
instance (priority := 80) [LocallyCompactSpace X] [ClosableCompactSubsetOpenSpace X] :
1784+
RegularSpace X := by
1785+
apply RegularSpace.ofExistsMemNhdsIsClosedSubset (fun x s hx ↦ ?_)
1786+
rcases _root_.mem_nhds_iff.1 hx with ⟨u, us, u_open, xu⟩
1787+
rcases exists_compact_closed_between (isCompact_singleton (a := x)) u_open (by simpa using xu)
1788+
with ⟨t, -, t_closed, xt, tu⟩
1789+
have : interior t ∈ 𝓝 x := isOpen_interior.mem_nhds (by simpa using xt)
1790+
exact ⟨t, interior_mem_nhds.mp this, t_closed, tu.trans us⟩
1791+
17821792
protected theorem IsCompact.closure [WeaklyLocallyCompactSpace X] [ClosableCompactSubsetOpenSpace X]
17831793
{K : Set X} (hK : IsCompact K) : IsCompact (closure K) := by
17841794
rcases exists_compact_superset hK with ⟨L, L_comp, hL⟩

0 commit comments

Comments
 (0)