@@ -243,6 +243,11 @@ Require Import reals signed.
243243(* component x == the connected component of point x *)
244244(* perfect_set A == A is closed, and is every point in A *)
245245(* is a limit point of A. *)
246+ (* totally_disconnected A == The only connected subsets of A are *)
247+ (* empty or singletons. *)
248+ (* zero_dimensional T == Points are separable by a clopen set. *)
249+ (* *)
250+ (* *)
246251(* [locally P] := forall a, A a -> G (within A (nbhs x)) *)
247252(* if P is convertible to G (globally A) *)
248253(* *)
@@ -3617,6 +3622,18 @@ move=> Axy; apply/seteqP; split => z; apply: connected_component_trans => //.
36173622by apply: connected_component_sym.
36183623Qed .
36193624
3625+ Lemma component_closed A x : closed A -> closed (connected_component A x).
3626+ Proof .
3627+ move=> clA; have [Ax|Ax] := pselect (A x); first last.
3628+ by rewrite connected_component_out //; exact: closed0.
3629+ rewrite closure_id eqEsubset; split; first exact: subset_closure.
3630+ move=> z Axz; exists (closure (connected_component A x)) => //.
3631+ split; first exact/subset_closure/connected_component_refl.
3632+ rewrite [X in _ `<=` X](closure_id A).1//.
3633+ by apply: closure_subset; exact: connected_component_sub.
3634+ by apply: connected_closure; exact: component_connected.
3635+ Qed .
3636+
36203637Lemma clopen_separatedP A : clopen A <-> separated A (~` A).
36213638Proof .
36223639split=> [[oA cA]|[] /[!(@disjoints_subset T)] /[!(@setCK T)] clAA AclA].
@@ -3736,6 +3753,70 @@ Qed.
37363753
37373754End perfect_sets.
37383755
3756+ Section totally_disconnected.
3757+ Implicit Types T : topologicalType.
3758+
3759+ Definition totally_disconnected {T} (A : set T) :=
3760+ forall x, A x -> connected_component A x = [set x].
3761+
3762+ Definition zero_dimensional T :=
3763+ (forall x y, x != y -> exists U : set T, [/\ clopen U, U x & ~ U y]).
3764+
3765+ Lemma zero_dimension_prod (I : choiceType) (T : I -> topologicalType) :
3766+ (forall i, zero_dimensional (T i)) ->
3767+ zero_dimensional (prod_topology T).
3768+ Proof .
3769+ move=> dctTI x y /eqP xneqy.
3770+ have [i/eqP/dctTI [U [clU Ux nUy]]] : exists i, x i <> y i.
3771+ by apply/existsNP=> W; exact/xneqy/functional_extensionality_dep.
3772+ exists (proj i @^-1` U); split => //; apply: clopen_comp => //.
3773+ exact/proj_continuous.
3774+ Qed .
3775+
3776+ Lemma discrete_zero_dimension {T} : discrete_space T -> zero_dimensional T.
3777+ Proof .
3778+ move=> dctT x y xny; exists [set x]; split => //; last exact/nesym/eqP.
3779+ by split; [exact: discrete_open | exact: discrete_closed].
3780+ Qed .
3781+
3782+ Lemma zero_dimension_totally_disconnected {T} :
3783+ zero_dimensional T -> totally_disconnected [set: T].
3784+ Proof .
3785+ move=> zdA x _; rewrite eqEsubset.
3786+ split=> [z [R [Rx _ ctdR Rz]]|_ ->]; last exact: connected_component_refl.
3787+ apply: contrapT => /eqP znx; have [U [[oU cU] Uz Ux]] := zdA _ _ znx.
3788+ suff : R `&` U = R by move: Rx => /[swap] <- [].
3789+ by apply: ctdR; [exists z|exists U|exists U].
3790+ Qed .
3791+
3792+ Lemma totally_disconnected_cvg {T : topologicalType} (x : T) :
3793+ hausdorff_space T -> zero_dimensional T -> compact [set: T] ->
3794+ filter_from [set D : set T | D x /\ clopen D] id --> x.
3795+ Proof .
3796+ pose F := filter_from [set D : set T | D x /\ clopen D] id.
3797+ have FF : Filter F.
3798+ apply: filter_from_filter; first by exists setT; split => //; exact: clopenT.
3799+ by move=> A B [? ?] [? ?]; exists (A `&` B) => //; split=> //; exact: clopenI.
3800+ have PF : ProperFilter F by apply: filter_from_proper; move=> ? [? _]; exists x.
3801+ move=> hsdfT zdT cmpT U Ux; rewrite nbhs_simpl -/F.
3802+ wlog oU : U Ux / open U.
3803+ move: Ux; rewrite /= nbhsE => -[] V [? ?] /filterS + /(_ V) P.
3804+ by apply; apply: P => //; exists V.
3805+ have /(iffLR (compact_near_coveringP _)) : compact (~` U).
3806+ by apply: (subclosed_compact _ cmpT) => //; exact: open_closedC.
3807+ move=> /(_ _ _ setC (powerset_filter_from_filter PF))[].
3808+ move=> y nUy; have /zdT [C [[oC cC] Cx Cy]] : x != y.
3809+ by apply: contra_notN nUy => /eqP <-; exact: nbhs_singleton.
3810+ exists (~` C, [set U | U `<=` C]); first split.
3811+ - by apply: open_nbhs_nbhs; split => //; exact: closed_openC.
3812+ - apply/near_powerset_filter_fromP; first by move=> ? ?; exact: subset_trans.
3813+ by exists C => //; exists C.
3814+ - by case=> i j [? /subsetC]; apply.
3815+ by move=> D [DF _ [C DC]]/(_ _ DC)/subsetC2/filterS; apply; exact: DF.
3816+ Qed .
3817+
3818+ End totally_disconnected.
3819+
37393820(** * Uniform spaces *)
37403821
37413822Local Notation "A ^-1" := ([set xy | A (xy.2, xy.1)]) : classical_set_scope.
@@ -6650,6 +6731,22 @@ move/nbhs_singleton: nbhsU; move: x; apply/in_setP.
66506731by rewrite -continuous_open_subspace.
66516732Unshelve. end_near. Qed .
66526733
6734+ Lemma totally_disconnected_prod (I : choiceType)
6735+ (T : I -> topologicalType) (A : forall i, set (T i)) :
6736+ (forall i, totally_disconnected (A i)) ->
6737+ @totally_disconnected (prod_topology T)
6738+ (fun f => forall i, A i (f i)).
6739+ Proof .
6740+ move=> dsctAi x /= Aix; rewrite eqEsubset; split; last first.
6741+ by move=> ? ->; exact: connected_component_refl.
6742+ move=> f /= [C /= [Cx CA ctC Cf]]; apply/functional_extensionality_dep => i.
6743+ suff : proj i @` C `<=` [set x i] by apply; exists f.
6744+ rewrite -(dsctAi i) // => Ti ?; exists (proj i @` C) => //.
6745+ split; [by exists x | by move=> ? [r Cr <-]; exact: CA |].
6746+ apply/(connected_continuous_connected ctC)/continuous_subspaceT.
6747+ exact: proj_continuous.
6748+ Qed .
6749+
66536750Section UniformPointwise.
66546751Context {U : topologicalType} {V : uniformType}.
66556752
0 commit comments