Skip to content

Commit d1d1772

Browse files
committed
feat(LightProfinite): being light is a property of a profinite space (#10391)
This PR defines the class `Profinite.IsLight` which is the property of a profinite space to have countably many clopens. We prove that such a profinite space gives rise to a `LightProfinite`, and the underlying profinite space of a `LightProfinite` is light. - [x] depends on: #10390
1 parent d424344 commit d1d1772

File tree

3 files changed

+275
-1
lines changed

3 files changed

+275
-1
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -3776,6 +3776,7 @@ import Mathlib.Topology.Category.CompHaus.Limits
37763776
import Mathlib.Topology.Category.CompHaus.Projective
37773777
import Mathlib.Topology.Category.Compactum
37783778
import Mathlib.Topology.Category.LightProfinite.Basic
3779+
import Mathlib.Topology.Category.LightProfinite.IsLight
37793780
import Mathlib.Topology.Category.Locale
37803781
import Mathlib.Topology.Category.Profinite.AsLimit
37813782
import Mathlib.Topology.Category.Profinite.Basic
Lines changed: 224 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,224 @@
1+
/-
2+
Copyright (c) 2023 Dagur Asgeirsson. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Dagur Asgeirsson
5+
-/
6+
import Mathlib.CategoryTheory.Limits.Shapes.Countable
7+
import Mathlib.Topology.Category.LightProfinite.Basic
8+
import Mathlib.Topology.Category.Profinite.AsLimit
9+
import Mathlib.Topology.Category.Profinite.CofilteredLimit
10+
import Mathlib.Topology.ClopenBox
11+
/-!
12+
# Being light is a property of profinite spaces
13+
14+
We define a class `Profinite.IsLight` which says that a profinite space `S` has countably many
15+
clopen subsets. We prove that a profinite space which `IsLight` gives rise to a `LightProfinite` and
16+
that the underlying profinite space of a `LightProfinite` is light.
17+
18+
## Main results
19+
20+
* `IsLight` is preserved under taking cartesian products
21+
22+
* Given a monomorphism whose target `IsLight`, its sourse is also light: `Profinite.isLight_of_mono`
23+
24+
## Project
25+
26+
* Prove the Stone duality theorem that `Profinite` is equivalent to the opposite category of
27+
boolean algebras. Then the property of being light says precisely that the corresponding
28+
boolean algebra is countable. Maybe constructions of limits and colimits in `LightProfinite`
29+
become easier when transporting over this equivalence.
30+
31+
-/
32+
33+
universe u
34+
35+
open CategoryTheory Limits FintypeCat Opposite TopologicalSpace
36+
37+
open scoped Classical
38+
39+
namespace Profinite
40+
41+
/-- A profinite space *is light* if it has countably many clopen subsets. -/
42+
class IsLight (S : Profinite) : Prop where
43+
/-- The set of clopens is countable -/
44+
countable_clopens : Countable (Clopens S)
45+
46+
attribute [instance] Profinite.IsLight.countable_clopens
47+
48+
instance instIsLightProd (X Y : Profinite) [X.IsLight] [Y.IsLight] :
49+
(Profinite.of (X × Y)).IsLight where
50+
countable_clopens := Clopens.countable_prod
51+
52+
instance instCountableDiscreteQuotientOfIsLight (S : Profinite) [S.IsLight] :
53+
Countable (DiscreteQuotient S) := (DiscreteQuotient.finsetClopens_inj S).countable
54+
55+
end Profinite
56+
57+
namespace LightProfinite
58+
59+
/-- A profinite space which is light gives rise to a light profinite space. -/
60+
noncomputable def ofIsLight (S : Profinite.{u}) [S.IsLight] : LightProfinite.{u} where
61+
diagram := sequentialFunctor _ ⋙ S.fintypeDiagram
62+
cone := (Functor.Initial.limitConeComp (sequentialFunctor _) S.lim).cone
63+
isLimit := (Functor.Initial.limitConeComp (sequentialFunctor _) S.lim).isLimit
64+
65+
instance (S : LightProfinite.{u}) : S.toProfinite.IsLight where
66+
countable_clopens := by
67+
refine @Countable.of_equiv _ _ ?_ (LocallyConstant.equivClopens (X := S.toProfinite))
68+
refine @Function.Surjective.countable
69+
(Σ (n : ℕ), LocallyConstant ((S.diagram ⋙ FintypeCat.toProfinite).obj ⟨n⟩) (Fin 2)) _ ?_ ?_ ?_
70+
· apply @instCountableSigma _ _ _ ?_
71+
intro n
72+
refine @Finite.to_countable _ ?_
73+
refine @Finite.of_injective _ ((S.diagram ⋙ FintypeCat.toProfinite).obj ⟨n⟩ → (Fin 2)) ?_ _
74+
LocallyConstant.coe_injective
75+
refine @Pi.finite _ _ ?_ _
76+
simp only [Functor.comp_obj, toProfinite_obj_toCompHaus_toTop_α]
77+
infer_instance
78+
· exact fun a ↦ a.snd.comap (S.cone.π.app ⟨a.fst⟩).1
79+
· intro a
80+
obtain ⟨n, g, h⟩ := Profinite.exists_locallyConstant S.cone S.isLimit a
81+
exact ⟨⟨unop n, g⟩, h.symm⟩
82+
83+
instance (S : LightProfinite.{u}) : (lightToProfinite.obj S).IsLight :=
84+
(inferInstance : S.toProfinite.IsLight)
85+
86+
end LightProfinite
87+
88+
section Mono
89+
/-!
90+
Given a monomorphism `f : X ⟶ Y` in `Profinite`, such that `Y.IsLight`, we want to prove that
91+
`X.IsLight`. We do this by regarding `Y` as a `LightProfinite`, and constructing a `LightProfinite`
92+
with `X` as its underlying `Profinite`. The diagram is just the image of `f` in the diagram of `Y`.
93+
-/
94+
95+
namespace Profinite
96+
97+
variable {X : Profinite} {Y : LightProfinite} (f : X ⟶ Y.toProfinite)
98+
99+
/-- The image of `f` in the diagram of `Y` -/
100+
noncomputable def lightProfiniteDiagramOfHom :
101+
ℕᵒᵖ ⥤ FintypeCat where
102+
obj := fun n ↦ FintypeCat.of (Set.range (f ≫ Y.cone.π.app n) : Set (Y.diagram.obj n))
103+
map := fun h ⟨x, hx⟩ ↦ ⟨Y.diagram.map h x, (by
104+
obtain ⟨y, hx⟩ := hx
105+
rw [← hx]
106+
use y
107+
have := Y.cone.π.naturality h
108+
simp only [Functor.const_obj_obj, Functor.comp_obj, Functor.const_obj_map, Category.id_comp,
109+
Functor.comp_map] at this
110+
rw [this]
111+
rfl )⟩
112+
map_id := by -- `aesop` can handle it but is a bit slow
113+
intro
114+
simp only [Functor.comp_obj, id_eq, Functor.const_obj_obj, Functor.const_obj_map,
115+
Functor.comp_map, eq_mp_eq_cast, cast_eq, eq_mpr_eq_cast, CategoryTheory.Functor.map_id,
116+
FintypeCat.id_apply]
117+
rfl
118+
map_comp := by -- `aesop` can handle it but is a bit slow
119+
intros
120+
simp only [Functor.comp_obj, id_eq, Functor.const_obj_obj, Functor.const_obj_map,
121+
Functor.comp_map, eq_mp_eq_cast, cast_eq, eq_mpr_eq_cast, Functor.map_comp,
122+
FintypeCat.comp_apply]
123+
rfl
124+
125+
/-- An auxiliary definition to prove continuity in `lightProfiniteConeOfHom_π_app`. -/
126+
def lightProfiniteConeOfHom_π_app' (n : ℕᵒᵖ) :
127+
C(X, Set.range (f ≫ Y.cone.π.app n)) where
128+
toFun := fun x ↦ ⟨Y.cone.π.app n (f x), ⟨x, rfl⟩⟩
129+
continuous_toFun := Continuous.subtype_mk ((Y.cone.π.app n).continuous.comp f.continuous) _
130+
131+
/-- An auxiliary definition for `lightProfiniteConeOfHom`. -/
132+
def lightProfiniteConeOfHom_π_app (n : ℕᵒᵖ) :
133+
X ⟶ (lightProfiniteDiagramOfHom f ⋙ FintypeCat.toProfinite).obj n where
134+
toFun x := ⟨Y.cone.π.app n (f x), ⟨x, rfl⟩⟩
135+
continuous_toFun := by
136+
convert (lightProfiniteConeOfHom_π_app' f n).continuous_toFun
137+
change ⊥ = _
138+
ext U
139+
rw [isOpen_induced_iff]
140+
have := discreteTopology_bot
141+
refine ⟨fun _ ↦ ⟨Subtype.val '' U, isOpen_discrete _,
142+
Function.Injective.preimage_image Subtype.val_injective _⟩, fun _ ↦ isOpen_discrete U⟩
143+
-- This is annoying
144+
145+
/-- The cone on `lightProfiniteDiagramOfHom` -/
146+
def lightProfiniteConeOfHom :
147+
Cone (lightProfiniteDiagramOfHom f ⋙ FintypeCat.toProfinite) where
148+
pt := X
149+
π := {
150+
app := fun n ↦ lightProfiniteConeOfHom_π_app f n
151+
naturality := by
152+
intro n m h
153+
have := Y.cone.π.naturality h
154+
simp only [Functor.const_obj_obj, Functor.comp_obj, Functor.const_obj_map, Category.id_comp,
155+
Functor.comp_map] at this
156+
simp only [Functor.comp_obj, toProfinite_obj_toCompHaus_toTop_α, Functor.const_obj_obj,
157+
Functor.const_obj_map, lightProfiniteConeOfHom_π_app, this, CategoryTheory.comp_apply,
158+
Category.id_comp, Functor.comp_map]
159+
rfl }
160+
161+
instance [Mono f] : IsIso ((Profinite.limitConeIsLimit ((lightProfiniteDiagramOfHom f) ⋙
162+
FintypeCat.toProfinite)).lift (lightProfiniteConeOfHom f)) := by
163+
apply Profinite.isIso_of_bijective
164+
refine ⟨fun a b h ↦ ?_, fun a ↦ ?_⟩
165+
· have hf : Function.Injective f := by rwa [← Profinite.mono_iff_injective]
166+
suffices f a = f b by exact hf this
167+
apply LightProfinite.ext
168+
intro n
169+
apply_fun fun f : (Profinite.limitCone ((lightProfiniteDiagramOfHom f) ⋙
170+
FintypeCat.toProfinite)).pt ↦ f.val n at h
171+
erw [ContinuousMap.coe_mk, Subtype.ext_iff] at h
172+
exact h
173+
· suffices ∃ x, ∀ n, lightProfiniteConeOfHom_π_app f (op n) x = a.val (op n) by
174+
obtain ⟨x, h⟩ := this
175+
use x
176+
apply Subtype.ext
177+
apply funext
178+
intro n
179+
exact h (unop n)
180+
have : Set.Nonempty
181+
(⋂ (n : ℕ), (lightProfiniteConeOfHom_π_app f (op n)) ⁻¹' {a.val (op n)}) := by
182+
refine IsCompact.nonempty_iInter_of_directed_nonempty_isCompact_isClosed
183+
(fun n ↦ (lightProfiniteConeOfHom_π_app f (op n)) ⁻¹' {a.val (op n)})
184+
(directed_of_isDirected_le ?_)
185+
(fun _ ↦ (Set.singleton_nonempty _).preimage fun ⟨a, ⟨b, hb⟩⟩ ↦ ⟨b, Subtype.ext hb⟩)
186+
(fun _ ↦ (IsClosed.preimage (lightProfiniteConeOfHom_π_app _ _).continuous
187+
(T1Space.t1 _)).isCompact)
188+
(fun _ ↦ IsClosed.preimage (lightProfiniteConeOfHom_π_app _ _).continuous (T1Space.t1 _))
189+
intro i j h x hx
190+
simp only [Functor.comp_obj, profiniteToCompHaus_obj, compHausToTop_obj,
191+
toProfinite_obj_toCompHaus_toTop_α, Functor.comp_map, profiniteToCompHaus_map,
192+
compHausToTop_map, Set.mem_setOf_eq, Set.mem_preimage, Set.mem_singleton_iff] at hx ⊢
193+
have := (lightProfiniteConeOfHom f).π.naturality (homOfLE h).op
194+
simp only [lightProfiniteConeOfHom, Functor.const_obj_obj, Functor.comp_obj,
195+
Functor.const_obj_map, Category.id_comp, Functor.comp_map] at this
196+
rw [this]
197+
simp only [CategoryTheory.comp_apply]
198+
rw [hx, ← a.prop (homOfLE h).op]
199+
rfl
200+
obtain ⟨x, hx⟩ := this
201+
exact ⟨x, Set.mem_iInter.1 hx⟩
202+
203+
204+
/-- When `f` is a monomorphism the cone `lightProfiniteConeOfHom` is limiting. -/
205+
noncomputable
206+
def lightProfiniteIsLimitOfMono [Mono f] : IsLimit (lightProfiniteConeOfHom f) :=
207+
Limits.IsLimit.ofIsoLimit (Profinite.limitConeIsLimit _)
208+
(Limits.Cones.ext (asIso ((Profinite.limitConeIsLimit ((lightProfiniteDiagramOfHom f) ⋙
209+
FintypeCat.toProfinite)).lift (lightProfiniteConeOfHom f))) fun _ => rfl).symm
210+
211+
/--
212+
Putting together all of the above, we get a `LightProfinite` whose underlying `Profinite` is `X`
213+
-/
214+
noncomputable def lightProfiniteOfMono [Mono f] : LightProfinite :=
215+
⟨lightProfiniteDiagramOfHom f, lightProfiniteConeOfHom f, lightProfiniteIsLimitOfMono f⟩
216+
217+
/-- `lightProfiniteOfMono` phrased in terms of `Profinite.IsLight`. -/
218+
theorem isLight_of_mono {X Y : Profinite} [Y.IsLight] (f : X ⟶ Y) [Mono f] : X.IsLight := by
219+
change (lightProfiniteOfMono (Y := LightProfinite.ofIsLight Y) f).toProfinite.IsLight
220+
infer_instance
221+
222+
end Profinite
223+
224+
end Mono

Mathlib/Topology/DiscreteQuotient.lean

Lines changed: 50 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3,6 +3,7 @@ Copyright (c) 2021 Adam Topaz. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Calle Sönne, Adam Topaz
55
-/
6+
import Mathlib.Data.Setoid.Partition
67
import Mathlib.Topology.Separation
78
import Mathlib.Topology.LocallyConstant.Basic
89

@@ -62,7 +63,7 @@ of finite discrete spaces.
6263
-/
6364

6465

65-
open Set Function
66+
open Set Function TopologicalSpace
6667

6768
variable {α X Y Z : Type*} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z]
6869

@@ -394,6 +395,54 @@ instance [CompactSpace X] : Finite S := by
394395
have : CompactSpace S := Quotient.compactSpace
395396
rwa [← isCompact_univ_iff, isCompact_iff_finite, finite_univ_iff] at this
396397

398+
variable (X)
399+
400+
open Classical in
401+
/--
402+
If `X` is a compact space, then we associate to any discrete quotient on `X` a finite set of
403+
clopen subsets of `X`, given by the fibers of `proj`.
404+
405+
TODO: prove that these form a partition of `X` 
406+
-/
407+
noncomputable def finsetClopens [CompactSpace X]
408+
(d : DiscreteQuotient X) : Finset (Clopens X) := have : Fintype d := Fintype.ofFinite _
409+
(Set.range (fun (x : d) ↦ ⟨_, d.isClopen_preimage {x}⟩) : Set (Clopens X)).toFinset
410+
411+
/-- A helper lemma to prove that `finsetClopens X` is injective, see `finsetClopens_inj`. -/
412+
lemma comp_finsetClopens [CompactSpace X] :
413+
(Set.image (fun (t : Clopens X) ↦ t.carrier) ∘ Finset.toSet) ∘
414+
finsetClopens X = fun ⟨f, _⟩ ↦ f.classes := by
415+
ext d
416+
simp only [Setoid.classes, Setoid.Rel, Set.mem_setOf_eq, Function.comp_apply,
417+
finsetClopens, Set.coe_toFinset, Set.mem_image, Set.mem_range,
418+
exists_exists_eq_and]
419+
constructor
420+
· refine fun ⟨y, h⟩ ↦ ⟨Quotient.out (s := d.toSetoid) y, ?_⟩
421+
ext
422+
simpa [← h] using Quotient.mk_eq_iff_out (s := d.toSetoid)
423+
· exact fun ⟨y, h⟩ ↦ ⟨d.proj y, by ext; simp [h, proj]⟩
424+
425+
/-- `finsetClopens X` is injective. -/
426+
theorem finsetClopens_inj [CompactSpace X] :
427+
(finsetClopens X).Injective := by
428+
apply Function.Injective.of_comp (f := Set.image (fun (t : Clopens X) ↦ t.carrier) ∘ Finset.toSet)
429+
rw [comp_finsetClopens]
430+
intro ⟨_, _⟩ ⟨_, _⟩ h
431+
congr
432+
rw [Setoid.classes_inj]
433+
exact h
434+
435+
/--
436+
The discrete quotients of a compact space are in bijection with a subtype of the type of
437+
`Finset (Clopens X)`.
438+
439+
TODO: show that this is precisely those finsets of clopens which form a partition of `X`.
440+
-/
441+
noncomputable
442+
def equivFinsetClopens [CompactSpace X] := Equiv.ofInjective _ (finsetClopens_inj X)
443+
444+
variable {X}
445+
397446
end DiscreteQuotient
398447

399448
namespace LocallyConstant

0 commit comments

Comments
 (0)