The category of Compact Hausdorff Spaces #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
We construct the category of compact Hausdorff spaces.
The type of compact Hausdorff spaces is denoted CompHaus
, and it is endowed with a category
instance making it a full subcategory of Top
.
The fully faithful functor CompHaus ⥤ Top
is denoted CompHaus_to_Top
.
Note: The file topology/category/Compactum.lean
provides the equivalence between Compactum
,
which is defined as the category of algebras for the ultrafilter monad, and CompHaus
.
Compactum_to_CompHaus
is the functor from Compactum
to CompHaus
which is proven to be an
equivalence of categories in Compactum_to_CompHaus.is_equivalence
.
See topology/category/Compactum.lean
for a more detailed discussion where these definitions are
introduced.
The type of Compact Hausdorff topological spaces.
Instances for CompHaus
Equations
- CompHaus.inhabited = {default := {to_Top := {α := pempty, str := pempty.topological_space}, is_compact := CompHaus.inhabited._proof_1, is_hausdorff := CompHaus.inhabited._proof_2}}
A constructor for objects of the category CompHaus
,
taking a type, and bundling the compact Hausdorff topology
found by typeclass inference.
Equations
- CompHaus.of X = CompHaus.mk (Top.of X)
Instances for CompHaus.of
Any continuous function on compact Hausdorff spaces is a closed map.
Any continuous bijection of compact Hausdorff spaces is an isomorphism.
Any continuous bijection of compact Hausdorff spaces induces an isomorphism.
Equations
- CompHaus.iso_of_bijective f bij = let _inst : category_theory.is_iso f := _ in category_theory.as_iso f
(Implementation) The object part of the compactification functor from topological spaces to compact Hausdorff spaces.
Equations
- StoneCech_obj X = CompHaus.of (stone_cech ↥X)
(Implementation) The bijection of homsets to establish the reflective adjunction of compact Hausdorff spaces in topological spaces.
Equations
- stone_cech_equivalence X Y = {to_fun := λ (f : StoneCech_obj X ⟶ Y), {to_fun := ⇑f ∘ stone_cech_unit, continuous_to_fun := _}, inv_fun := λ (f : X ⟶ CompHaus_to_Top.obj Y), {to_fun := stone_cech_extend _, continuous_to_fun := _}, left_inv := _, right_inv := _}
The Stone-Cech compactification functor from topological spaces to compact Hausdorff spaces, left adjoint to the inclusion functor.
Equations
- Top_to_CompHaus = category_theory.adjunction.left_adjoint_of_equiv stone_cech_equivalence Top_to_CompHaus._proof_1
The category of compact Hausdorff spaces is reflective in the category of topological spaces.
An explicit limit cone for a functor F : J ⥤ CompHaus
, defined in terms of
Top.limit_cone
.
Equations
- CompHaus.limit_cone F = {X := {to_Top := (Top.limit_cone (F ⋙ CompHaus_to_Top)).X, is_compact := _, is_hausdorff := _}, π := {app := λ (j : J), (Top.limit_cone (F ⋙ CompHaus_to_Top)).π.app j, naturality' := _}}
The limit cone CompHaus.limit_cone F
is indeed a limit cone.
Equations
- CompHaus.limit_cone_is_limit F = {lift := λ (S : category_theory.limits.cone F), (Top.limit_cone_is_limit (F ⋙ CompHaus_to_Top)).lift (CompHaus_to_Top.map_cone S), fac' := _, uniq' := _}