Skip to content

Commit 3800c53

Browse files
committed
able to separate the discrete stuff more cleanly
1 parent b74a117 commit 3800c53

File tree

12 files changed

+284
-234
lines changed

12 files changed

+284
-234
lines changed

CHANGELOG_UNRELEASED.md

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -93,6 +93,9 @@
9393
+ new lemmas `min_continuous`, `min_fun_continuous`, `max_continuous`, and
9494
`max_fun_continuous`.
9595

96+
- in file `bool_topology.v`,
97+
+ new lemma `bool_compact`.
98+
9699
### Changed
97100

98101
- in file `normedtype.v`,
@@ -146,6 +149,16 @@
146149

147150
### Changed
148151

152+
- moved from `topology_structure.v` to `discrete_topology.v`:
153+
`discrete_open`, `discrete_set1`, `discrete_closed`, and `discrete_cvg`.
154+
155+
- moved from `pseudometric_structure.v` to `discrete_topology.v`:
156+
`discrete_ent`, `discrete_ball`, and `discrete_topology`.
157+
158+
- in file `cantor.v`, `cantor_space` now defined in terms of `bool`.
159+
- in file `separation_axioms.v`, updated `discrete_hausdorff`, and
160+
`discrete_zero_dimension` to take a `discreteTopologicalType`.
161+
149162
### Renamed
150163

151164
- in `normedtype.v`:
@@ -178,6 +191,10 @@
178191
+ lemma `cst_mfun_subproof` (use lemma `measurable_cst` instead)
179192
+ lemma `cst_nnfun_subproof` (turned into a `Let`)
180193
+ lemma `indic_mfun_subproof` (use lemma `measurable_fun_indic` instead)
194+
- in file `topology_structure.v`, removed `discrete_sing`, `discrete_nbhs`, and `discrete_space`.
195+
- in file `nat_topology.v`, removed `discrete_nat`.
196+
- in file `pseudometric_structure.v`, removed `discrete_ball_center`, `discrete_topology_type`, and
197+
`discrete_space_discrete`.
181198

182199
### Infrastructure
183200

_CoqProject

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -60,6 +60,7 @@ theories/topology_theory/num_topology.v
6060
theories/topology_theory/quotient_topology.v
6161
theories/topology_theory/one_point_compactification.v
6262
theories/topology_theory/sigT_topology.v
63+
theories/topology_theory/discrete_topology.v
6364

6465
theories/homotopy_theory/homotopy.v
6566
theories/homotopy_theory/wedge_sigT.v

theories/Make

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -28,6 +28,7 @@ topology_theory/num_topology.v
2828
topology_theory/quotient_topology.v
2929
topology_theory/one_point_compactification.v
3030
topology_theory/sigT_topology.v
31+
topology_theory/discrete_topology.v
3132

3233
homotopy_theory/homotopy.v
3334
homotopy_theory/wedge_sigT.v

theories/cantor.v

Lines changed: 26 additions & 28 deletions
Original file line numberDiff line numberDiff line change
@@ -43,12 +43,13 @@ Import numFieldTopology.Exports.
4343

4444
Local Open Scope classical_set_scope.
4545

46-
Definition cantor_space :=
47-
prod_topology (fun _ : nat => discrete_topology discrete_bool).
46+
Definition cantor_space : Type :=
47+
prod_topology (fun _ : nat => bool).
4848

4949
HB.instance Definition _ := Pointed.on cantor_space.
5050
HB.instance Definition _ := Nbhs.on cantor_space.
5151
HB.instance Definition _ := Topological.on cantor_space.
52+
HB.instance Definition _ := Uniform.on cantor_space.
5253

5354
Definition cantor_like (T : topologicalType) :=
5455
[/\ perfect_set [set: T],
@@ -58,20 +59,18 @@ Definition cantor_like (T : topologicalType) :=
5859

5960
Lemma cantor_space_compact : compact [set: cantor_space].
6061
Proof.
61-
have := @tychonoff _ (fun _ : nat => _) _ (fun=> discrete_bool_compact).
62+
have := @tychonoff _ (fun _ : nat => _) _ (fun=> bool_compact).
6263
by congr (compact _); rewrite eqEsubset.
6364
Qed.
6465

6566
Lemma cantor_space_hausdorff : hausdorff_space cantor_space.
6667
Proof.
6768
apply: hausdorff_product => ?; apply: discrete_hausdorff.
68-
exact: discrete_space_discrete.
6969
Qed.
7070

7171
Lemma cantor_zero_dimensional : zero_dimensional cantor_space.
7272
Proof.
7373
apply: zero_dimension_prod => _; apply: discrete_zero_dimension.
74-
exact: discrete_space_discrete.
7574
Qed.
7675

7776
Lemma cantor_perfect : perfect_set [set: cantor_space].
@@ -102,13 +101,12 @@ Qed.
102101
(* *)
103102
(******************************************************************************)
104103
Section topological_trees.
105-
Context {K : nat -> ptopologicalType} {X : ptopologicalType}
104+
Context {K : nat -> pdiscreteTopologicalType} {X : ptopologicalType}
106105
(refine_apx : forall n, set X -> K n -> set X)
107106
(tree_invariant : set X -> Prop).
108107

109108
Hypothesis cmptX : compact [set: X].
110109
Hypothesis hsdfX : hausdorff_space X.
111-
Hypothesis discreteK : forall n, discrete_space (K n).
112110
Hypothesis refine_cover : forall n U, U = \bigcup_e @refine_apx n U e.
113111
Hypothesis refine_invar : forall n U e,
114112
tree_invariant U -> tree_invariant (@refine_apx n U e).
@@ -122,7 +120,7 @@ Hypothesis refine_separates: forall x y : X, x != y ->
122120
Let refine_subset n U e : @refine_apx n U e `<=` U.
123121
Proof. by rewrite [X in _ `<=` X](refine_cover n); exact: bigcup_sup. Qed.
124122

125-
Let T := prod_topology K.
123+
Let T := prod_topology (K : nat -> ptopologicalType).
126124

127125
Local Fixpoint branch_apx (b : T) n :=
128126
if n is m.+1 then refine_apx (branch_apx b m) (b m) else [set: X].
@@ -193,7 +191,6 @@ near=> z => i; rewrite leq_eqVlt => /predU1P[|iSn]; last by rewrite (near IH z).
193191
move=> [->]; near: z; exists (proj n @^-1` [set b n]).
194192
split => //; suff : @open T (proj n @^-1` [set b n]) by [].
195193
apply: open_comp; [move=> + _; exact: proj_continuous| apply: discrete_open].
196-
exact: discreteK.
197194
Unshelve. all: end_near. Qed.
198195

199196
Let apx_prefix b c n :
@@ -293,9 +290,7 @@ Local Lemma cantor_map : exists f : cantor_space -> T,
293290
set_surj [set: cantor_space] [set: T] f &
294291
set_inj [set: cantor_space] f ].
295292
Proof.
296-
have [] := @tree_map_props
297-
(fun=> discrete_topology discrete_bool) T c_ind c_invar cmptT hsdfT.
298-
- by move=> ?; exact: discrete_space_discrete.
293+
have [] := @tree_map_props (fun=> bool) T c_ind c_invar cmptT hsdfT.
299294
- move=> n V; rewrite eqEsubset; split => [t Vt|t [? ? []]//].
300295
have [?|?] := pselect (U_ n `&` V !=set0 /\ ~` U_ n `&` V !=set0).
301296
+ have [Unt|Unt] := pselect (U_ n t).
@@ -350,38 +345,40 @@ Qed.
350345

351346
End TreeStructure.
352347

348+
Section cantor.
349+
Context {R : realType}.
350+
353351
(**md**************************************************************************)
354352
(* ## Part 3: Finitely branching trees are Cantor-like *)
355353
(******************************************************************************)
356354
Section FinitelyBranchingTrees.
357355

358356
Definition tree_of (T : nat -> pointedType) : Type :=
359-
prod_topology (fun n => discrete_topology_type (T n)).
357+
prod_topology (fun n => discrete_topology (T n)).
360358

361359
HB.instance Definition _ (T : nat -> pointedType) : Pointed (tree_of T):=
362360
Pointed.on (tree_of T).
363361

364362
HB.instance Definition _ (T : nat -> pointedType) := Uniform.on (tree_of T).
365363

366-
HB.instance Definition _ {R : realType} (T : nat -> pointedType) :
364+
HB.instance Definition _ (T : nat -> pointedType) :
367365
@PseudoMetric R _ :=
368366
@PseudoMetric.on (tree_of T).
369367

370368
Lemma cantor_like_finite_prod (T : nat -> ptopologicalType) :
371-
(forall n, finite_set [set: discrete_topology_type (T n)]) ->
369+
(forall n, finite_set [set: discrete_topology (T n)]) ->
372370
(forall n, (exists xy : T n * T n, xy.1 != xy.2)) ->
373371
cantor_like (tree_of T).
374372
Proof.
375373
move=> finiteT twoElems; split.
376-
- exact/(@perfect_diagonal (discrete_topology_type \o T))/twoElems.
374+
- exact/(@perfect_diagonal (discrete_topology \o T))/twoElems.
377375
- have := tychonoff (fun n => finite_compact (finiteT n)).
378376
set A := (X in compact X -> _).
379377
suff : A = [set: tree_of (fun x : nat => T x)] by move=> ->.
380378
by rewrite eqEsubset.
381-
- apply: (@hausdorff_product _ (discrete_topology_type \o T)) => n.
379+
- apply: (@hausdorff_product _ (discrete_topology \o T)) => n.
382380
by apply: discrete_hausdorff; exact: discrete_space_discrete.
383381
- apply: zero_dimension_prod => ?; apply: discrete_zero_dimension.
384-
exact: discrete_space_discrete.
385382
Qed.
386383

387384
End FinitelyBranchingTrees.
@@ -390,7 +387,7 @@ End FinitelyBranchingTrees.
390387
(* ## Part 4: Building a finitely branching tree to cover `T` *)
391388
(******************************************************************************)
392389
Section alexandroff_hausdorff.
393-
Context {R : realType} {T : pseudoPMetricType R}.
390+
Context {T : pseudoPMetricType R}.
394391

395392
Hypothesis cptT : compact [set: T].
396393
Hypothesis hsdfT : hausdorff_space T.
@@ -469,9 +466,14 @@ HB.instance Definition _ n := gen_eqMixin (K' n).
469466
HB.instance Definition _ n := gen_choiceMixin (K' n).
470467
HB.instance Definition _ n := isPointed.Build (K' n) (K'p n).
471468

472-
Let K n := K' n.
469+
Let K n := discrete_topology (K' n).
473470
Let Tree := @tree_of K.
474471

472+
HB.instance Definition _ n :=
473+
DiscreteTopology.on (K n).
474+
HB.instance Definition _ n :=
475+
Pointed.on (K n).
476+
475477
Let embed_refine n (U : set T) (k : K n) :=
476478
(if pselect (projT1 k `&` U !=set0)
477479
then projT1 k
@@ -486,17 +488,12 @@ case: e => W; have [//| _ _ _ _] := projT2 (cid (ent_balls' (count_unif n))).
486488
exact.
487489
Qed.
488490

489-
Let discrete_subproof (P : choiceType) :
490-
discrete_space (principal_filter_type P).
491-
Proof. by []. Qed.
492-
493491
Local Lemma cantor_surj_pt1 : exists2 f : Tree -> T,
494492
continuous f & set_surj [set: Tree] [set: T] f.
495493
Proof.
496494
pose entn n := projT2 (cid (ent_balls' (count_unif n))).
497-
have [//| | |? []//| |? []// | |] := @tree_map_props
498-
(discrete_topology_type \o K) T (embed_refine) (embed_invar) cptT hsdfT.
499-
- by move=> n; exact: discrete_space_discrete.
495+
have [ | |? []//| |? []// | |] := @tree_map_props
496+
K T (embed_refine) (embed_invar) cptT hsdfT.
500497
- move=> n U; rewrite eqEsubset; split=> [t Ut|t [? ? []]//].
501498
have [//|_ _ _ + _] := entn n; rewrite -subTset.
502499
move=> /(_ t I)[W cbW Wt]; exists (existT _ W cbW) => //.
@@ -535,7 +532,7 @@ Local Lemma cantor_surj_pt2 :
535532
exists f : {surj [set: cantor_space] >-> [set: Tree]}, continuous f.
536533
Proof.
537534
have [|f [ctsf _]] := @homeomorphism_cantor_like R Tree; last by exists f.
538-
apply: (@cantor_like_finite_prod (discrete_topology_type \o K)) => [n /=|n].
535+
apply: (@cantor_like_finite_prod (discrete_topology \o K)) => [n /=|n].
539536
have [//| fs _ _ _ _] := projT2 (cid (ent_balls' (count_unif n))).
540537
suff -> : [set: {classic K' n}] =
541538
(@projT1 (set T) _) @^-1` (projT1 (cid (ent_balls' (count_unif n)))).
@@ -569,3 +566,4 @@ by exists f; rewrite -cstf; exact: cst_continuous.
569566
Qed.
570567

571568
End alexandroff_hausdorff.
569+
End cantor.

theories/separation_axioms.v

Lines changed: 8 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -111,11 +111,6 @@ have [q [Aq clsAp_q]] := !! Aco _ _ pA; rewrite (hT p q) //.
111111
by apply: cvg_cluster clsAp_q; apply: cvg_within.
112112
Qed.
113113

114-
Lemma discrete_hausdorff {dsc : discrete_space T} : hausdorff_space.
115-
Proof.
116-
by move=> p q /(_ _ _ (discrete_set1 p) (discrete_set1 q)) [] // x [] -> ->.
117-
Qed.
118-
119114
Lemma compact_cluster_set1 (x : T) F V :
120115
hausdorff_space -> compact V -> nbhs x V ->
121116
ProperFilter F -> F V -> cluster F = [set x] -> F --> x.
@@ -211,6 +206,12 @@ move=> [[P Q]] /= [Px Qx] /= [/open_subspaceW oP /open_subspaceW oQ].
211206
by move=> ?; exists (P, Q); split => //=; [exact: oP | exact: oQ].
212207
Qed.
213208

209+
Lemma discrete_hausdorff {T : discreteTopologicalType} : hausdorff_space T.
210+
Proof.
211+
by move=> p q /(_ _ _ (discrete_set1 p) (discrete_set1 q)) [] // x [] -> ->.
212+
Qed.
213+
214+
214215
Lemma order_hausdorff {d} {T : orderTopologicalType d} : hausdorff_space T.
215216
Proof.
216217
rewrite open_hausdorff=> p q; wlog : p q / (p < q)%O.
@@ -567,9 +568,9 @@ Definition totally_disconnected {T} (A : set T) :=
567568
Definition zero_dimensional T :=
568569
(forall x y, x != y -> exists U : set T, [/\ clopen U, U x & ~ U y]).
569570

570-
Lemma discrete_zero_dimension {T} : discrete_space T -> zero_dimensional T.
571+
Lemma discrete_zero_dimension {T : discreteTopologicalType} : zero_dimensional T.
571572
Proof.
572-
move=> dctT x y xny; exists [set x]; split => //; last exact/nesym/eqP.
573+
move=> x y xny; exists [set x]; split => //; last exact/nesym/eqP.
573574
by split; [exact: discrete_open | exact: discrete_closed].
574575
Qed.
575576

theories/sequences.v

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1314,9 +1314,9 @@ Lemma cvg_nseries_near (u : nat^nat) : cvgn (nseries u) ->
13141314
\forall n \near \oo, u n = 0%N.
13151315
Proof.
13161316
move=> /cvg_ex[l ul]; have /ul[a _ aul] : nbhs l [set l].
1317-
by exists [set l]; split=> //; exists [set l] => //; rewrite bigcup_set1.
1317+
by rewrite nbhs_principalE.
13181318
have /ul[b _ bul] : nbhs l [set l.-1; l].
1319-
by exists [set l]; split => //; exists [set l] => //; rewrite bigcup_set1.
1319+
by rewrite nbhs_principalE ; apply/principal_filterP => /=; right.
13201320
exists (maxn a b) => // n /= abn.
13211321
rewrite (_ : u = fun n => nseries u n.+1 - nseries u n)%N; last first.
13221322
by rewrite funeqE => i; rewrite /nseries big_nat_recr//= addnC addnK.

theories/topology_theory/bool_topology.v

Lines changed: 7 additions & 19 deletions
Original file line numberDiff line numberDiff line change
@@ -3,6 +3,7 @@ From HB Require Import structures.
33
From mathcomp Require Import all_ssreflect all_algebra all_classical.
44
From mathcomp Require Import signed reals topology_structure uniform_structure.
55
From mathcomp Require Import pseudometric_structure order_topology compact.
6+
From mathcomp Require Import discrete_topology.
67

78
(**md**************************************************************************)
89
(* # Topology for boolean numbers *)
@@ -15,25 +16,19 @@ Import Order.TTheory GRing.Theory Num.Theory.
1516
Local Open Scope classical_set_scope.
1617
Local Open Scope ring_scope.
1718

18-
HB.instance Definition _ := Nbhs_isNbhsTopological.Build bool
19-
principal_filter_proper discrete_sing discrete_nbhs.
20-
21-
Lemma discrete_bool : discrete_space bool.
22-
Proof. by []. Qed.
19+
HB.instance Definition _ := hasNbhs.Build bool principal_filter.
20+
HB.instance Definition _ := Discrete_ofNbhs.Build bool erefl.
21+
HB.instance Definition _ := DiscreteUniform_ofNbhs.Build bool.
2322

2423
Lemma bool_compact : compact [set: bool].
2524
Proof. by rewrite setT_bool; apply/compactU; exact: compact_set1. Qed.
2625

27-
Section bool_ord_topology.
28-
Local Open Scope classical_set_scope.
29-
Local Open Scope order_scope.
30-
3126
Local Lemma bool_nbhs_itv (b : bool) :
3227
nbhs b = filter_from
3328
(fun i => itv_open_ends i /\ b \in i)
3429
(fun i => [set` i]).
3530
Proof.
36-
rewrite discrete_bool eqEsubset; split=> U; first last.
31+
rewrite nbhs_principalE eqEsubset; split=> U; first last.
3732
by case => V [_ Vb] VU; apply/principal_filterP/VU; apply: Vb.
3833
move/principal_filterP; case: b.
3934
move=> Ut; exists `]false, +oo[; first split => //; first by left.
@@ -43,12 +38,5 @@ by move=> r /=; rewrite in_itv /=; case: r.
4338
Qed.
4439

4540
HB.instance Definition _ := Order_isNbhs.Build _ bool bool_nbhs_itv.
46-
End bool_ord_topology.
47-
48-
Lemma discrete_bool_compact : compact [set: discrete_topology discrete_bool].
49-
Proof. by rewrite setT_bool; apply/compactU; exact: compact_set1. Qed.
50-
51-
Definition pseudoMetric_bool {R : realType} :=
52-
[the pseudoMetricType R of discrete_topology discrete_bool : Type].
53-
54-
#[global] Hint Resolve discrete_bool : core.
41+
HB.instance Definition _ {R : numDomainType} :=
42+
@DiscretePseudoMetric_ofUniform.Build R bool.

0 commit comments

Comments
 (0)