Skip to content

Commit 296aaf3

Browse files
committed
feat: port Topology.Category.Top.OpenNhds (#3834)
1 parent 828c813 commit 296aaf3

File tree

2 files changed

+179
-0
lines changed

2 files changed

+179
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1971,6 +1971,7 @@ import Mathlib.Topology.Category.Top.Adjunctions
19711971
import Mathlib.Topology.Category.Top.Basic
19721972
import Mathlib.Topology.Category.Top.EpiMono
19731973
import Mathlib.Topology.Category.Top.Limits.Basic
1974+
import Mathlib.Topology.Category.Top.OpenNhds
19741975
import Mathlib.Topology.Category.Top.Opens
19751976
import Mathlib.Topology.CompactOpen
19761977
import Mathlib.Topology.Connected
Lines changed: 178 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,178 @@
1+
/-
2+
Copyright (c) 2019 Scott Morrison. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Scott Morrison
5+
6+
! This file was ported from Lean 3 source module topology.category.Top.open_nhds
7+
! leanprover-community/mathlib commit 1ec4876214bf9f1ddfbf97ae4b0d777ebd5d6938
8+
! Please do not edit these lines, except to modify the commit id
9+
! if you have ported upstream changes.
10+
-/
11+
import Mathlib.Topology.Category.Top.Opens
12+
13+
/-!
14+
# The category of open neighborhoods of a point
15+
16+
Given an object `X` of the category `TopCat` of topological spaces and a point `x : X`, this file
17+
builds the type `OpenNhds x` of open neighborhoods of `x` in `X` and endows it with the partial
18+
order given by inclusion and the corresponding category structure (as a full subcategory of the
19+
poset category `Set X`). This is used in `Topology.Sheaves.Stalks` to build the stalk of a sheaf
20+
at `x` as a limit over `OpenNhds x`.
21+
22+
## Main declarations
23+
24+
Besides `OpenNhds`, the main constructions here are:
25+
26+
* `inclusion (x : X)`: the obvious functor `OpenNhds x ⥤ Opens X`
27+
* `functorNhds`: An open map `f : X ⟶ Y` induces a functor `OpenNhds x ⥤ OpenNhds (f x)`
28+
* `adjunctionNhds`: An open map `f : X ⟶ Y` induces an adjunction between `OpenNhds x` and
29+
`OpenNhds (f x)`.
30+
-/
31+
32+
33+
open CategoryTheory
34+
35+
open TopologicalSpace
36+
37+
open Opposite
38+
39+
universe u
40+
41+
variable {X Y : TopCat.{u}} (f : X ⟶ Y)
42+
43+
namespace TopologicalSpace
44+
45+
/-- The type of open neighbourhoods of a point `x` in a (bundled) topological space. -/
46+
def OpenNhds (x : X) :=
47+
FullSubcategory fun U : Opens X => x ∈ U
48+
#align topological_space.open_nhds TopologicalSpace.OpenNhds
49+
50+
namespace OpenNhds
51+
52+
instance partialOrder (x : X) : PartialOrder (OpenNhds x) where
53+
le U V := U.1 ≤ V.1
54+
le_refl _ := by dsimp [LE.le]; exact le_rfl
55+
le_trans _ _ _ := by dsimp [LE.le]; exact le_trans
56+
le_antisymm _ _ i j := FullSubcategory.ext _ _ <| le_antisymm i j
57+
58+
instance (x : X) : Lattice (OpenNhds x) :=
59+
{ OpenNhds.partialOrder x with
60+
inf := fun U V => ⟨U.1 ⊓ V.1, ⟨U.2, V.2⟩⟩
61+
le_inf := fun U V W => @le_inf _ _ U.1.1 V.1.1 W.1.1
62+
inf_le_left := fun U V => @inf_le_left _ _ U.1.1 V.1.1
63+
inf_le_right := fun U V => @inf_le_right _ _ U.1.1 V.1.1
64+
sup := fun U V => ⟨U.1 ⊔ V.1, Set.mem_union_left V.1.1 U.2
65+
sup_le := fun U V W => @sup_le _ _ U.1.1 V.1.1 W.1.1
66+
le_sup_left := fun U V => @le_sup_left _ _ U.1.1 V.1.1
67+
le_sup_right := fun U V => @le_sup_right _ _ U.1.1 V.1.1 }
68+
69+
instance (x : X) : OrderTop (OpenNhds x) where
70+
top := ⟨⊤, trivial⟩
71+
le_top _ := by dsimp [LE.le]; exact le_top
72+
73+
instance (x : X) : Inhabited (OpenNhds x) :=
74+
⟨⊤⟩
75+
76+
instance openNhdsCategory (x : X) : Category.{u} (OpenNhds x) := inferInstance
77+
#align topological_space.open_nhds.open_nhds_category TopologicalSpace.OpenNhds.openNhdsCategory
78+
79+
instance opensNhdsHomHasCoeToFun {x : X} {U V : OpenNhds x} : CoeFun (U ⟶ V) fun _ => U.1 → V.1 :=
80+
fun f x => ⟨x, f.le x.2⟩⟩
81+
#align topological_space.open_nhds.opens_nhds_hom_has_coe_to_fun TopologicalSpace.OpenNhds.opensNhdsHomHasCoeToFun
82+
83+
/-- The inclusion `U ⊓ V ⟶ U` as a morphism in the category of open sets. -/
84+
def infLeLeft {x : X} (U V : OpenNhds x) : U ⊓ V ⟶ U :=
85+
homOfLE inf_le_left
86+
#align topological_space.open_nhds.inf_le_left TopologicalSpace.OpenNhds.infLeLeft
87+
88+
/-- The inclusion `U ⊓ V ⟶ V` as a morphism in the category of open sets. -/
89+
def infLeRight {x : X} (U V : OpenNhds x) : U ⊓ V ⟶ V :=
90+
homOfLE inf_le_right
91+
#align topological_space.open_nhds.inf_le_right TopologicalSpace.OpenNhds.infLeRight
92+
93+
/-- The inclusion functor from open neighbourhoods of `x`
94+
to open sets in the ambient topological space. -/
95+
def inclusion (x : X) : OpenNhds x ⥤ Opens X :=
96+
fullSubcategoryInclusion _
97+
#align topological_space.open_nhds.inclusion TopologicalSpace.OpenNhds.inclusion
98+
99+
@[simp]
100+
theorem inclusion_obj (x : X) (U) (p) : (inclusion x).obj ⟨U, p⟩ = U :=
101+
rfl
102+
#align topological_space.open_nhds.inclusion_obj TopologicalSpace.OpenNhds.inclusion_obj
103+
104+
theorem openEmbedding {x : X} (U : OpenNhds x) : OpenEmbedding U.1.inclusion :=
105+
U.1.openEmbedding
106+
#align topological_space.open_nhds.open_embedding TopologicalSpace.OpenNhds.openEmbedding
107+
108+
/-- The preimage functor from neighborhoods of `f x` to neighborhoods of `x`. -/
109+
def map (x : X) : OpenNhds (f x) ⥤ OpenNhds x where
110+
obj U := ⟨(Opens.map f).obj U.1, U.2
111+
map i := (Opens.map f).map i
112+
#align topological_space.open_nhds.map TopologicalSpace.OpenNhds.map
113+
114+
-- Porting note: Changed `⟨(Opens.map f).obj U, by tidy⟩` to `⟨(Opens.map f).obj U, q⟩`
115+
@[simp]
116+
theorem map_obj (x : X) (U) (q) : (map f x).obj ⟨U, q⟩ = ⟨(Opens.map f).obj U, q⟩ :=
117+
rfl
118+
#align topological_space.open_nhds.map_obj TopologicalSpace.OpenNhds.map_obj
119+
120+
@[simp]
121+
theorem map_id_obj (x : X) (U) : (map (𝟙 X) x).obj U = U := rfl
122+
#align topological_space.open_nhds.map_id_obj TopologicalSpace.OpenNhds.map_id_obj
123+
124+
@[simp]
125+
theorem map_id_obj' (x : X) (U) (p) (q) : (map (𝟙 X) x).obj ⟨⟨U, p⟩, q⟩ = ⟨⟨U, p⟩, q⟩ :=
126+
rfl
127+
#align topological_space.open_nhds.map_id_obj' TopologicalSpace.OpenNhds.map_id_obj'
128+
129+
@[simp]
130+
theorem map_id_obj_unop (x : X) (U : (OpenNhds x)ᵒᵖ) : (map (𝟙 X) x).obj (unop U) = unop U := by
131+
simp
132+
#align topological_space.open_nhds.map_id_obj_unop TopologicalSpace.OpenNhds.map_id_obj_unop
133+
134+
@[simp]
135+
theorem op_map_id_obj (x : X) (U : (OpenNhds x)ᵒᵖ) : (map (𝟙 X) x).op.obj U = U := by simp
136+
#align topological_space.open_nhds.op_map_id_obj TopologicalSpace.OpenNhds.op_map_id_obj
137+
138+
/-- `Opens.map f` and `OpenNhds.map f` form a commuting square (up to natural isomorphism)
139+
with the inclusion functors into `Opens X`. -/
140+
def inclusionMapIso (x : X) : inclusion (f x) ⋙ Opens.map f ≅ map f x ⋙ inclusion x :=
141+
NatIso.ofComponents (fun U => by constructor; rfl; rfl; exact 𝟙 _; exact 𝟙 _) (by simp)
142+
#align topological_space.open_nhds.inclusion_map_iso TopologicalSpace.OpenNhds.inclusionMapIso
143+
144+
@[simp]
145+
theorem inclusionMapIso_hom (x : X) : (inclusionMapIso f x).hom = 𝟙 _ :=
146+
rfl
147+
#align topological_space.open_nhds.inclusion_map_iso_hom TopologicalSpace.OpenNhds.inclusionMapIso_hom
148+
149+
@[simp]
150+
theorem inclusionMapIso_inv (x : X) : (inclusionMapIso f x).inv = 𝟙 _ :=
151+
rfl
152+
#align topological_space.open_nhds.inclusion_map_iso_inv TopologicalSpace.OpenNhds.inclusionMapIso_inv
153+
154+
end OpenNhds
155+
156+
end TopologicalSpace
157+
158+
namespace IsOpenMap
159+
160+
open TopologicalSpace
161+
162+
variable {f}
163+
164+
/-- An open map `f : X ⟶ Y` induces a functor `OpenNhds x ⥤ OpenNhds (f x)`. -/
165+
@[simps]
166+
def functorNhds (h : IsOpenMap f) (x : X) : OpenNhds x ⥤ OpenNhds (f x) where
167+
obj U := ⟨h.functor.obj U.1, ⟨x, U.2, rfl⟩⟩
168+
map i := h.functor.map i
169+
#align is_open_map.functor_nhds IsOpenMap.functorNhds
170+
171+
/-- An open map `f : X ⟶ Y` induces an adjunction between `OpenNhds x` and `OpenNhds (f x)`. -/
172+
def adjunctionNhds (h : IsOpenMap f) (x : X) : IsOpenMap.functorNhds h x ⊣ OpenNhds.map f x :=
173+
Adjunction.mkOfUnitCounit
174+
{ unit := { app := fun U => homOfLE fun x hxU => ⟨x, hxU, rfl⟩ }
175+
counit := { app := fun V => homOfLE fun y ⟨_, hfxV, hxy⟩ => hxy ▸ hfxV } }
176+
#align is_open_map.adjunction_nhds IsOpenMap.adjunctionNhds
177+
178+
end IsOpenMap

0 commit comments

Comments
 (0)