Skip to content

Commit 71b1296

Browse files
bipointed and paths (#1400)
* bipointed and paths --------- Co-authored-by: Reynald Affeldt <[email protected]>
1 parent 33c087a commit 71b1296

File tree

8 files changed

+236
-0
lines changed

8 files changed

+236
-0
lines changed

CHANGELOG_UNRELEASED.md

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -112,6 +112,13 @@
112112
- in `boolp.v`:
113113
+ lemmas `existT_inj1`, `surjective_existT`
114114

115+
- in file `homotopy_theory/path.v`,
116+
+ new definitions `reparameterize`, `mk_path`, and `chain_path`.
117+
+ new lemmas `path_eqP`, and `chain_path_cts_point`.
118+
- in file `homotopy_theory/wedge_sigT.v`,
119+
+ new definition `bpwedge_shared_pt`.
120+
+ new notations `bpwedge`, and `bpwedge_lift`.
121+
115122
### Changed
116123

117124
- in file `normedtype.v`,

_CoqProject

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -63,6 +63,7 @@ theories/topology_theory/sigT_topology.v
6363

6464
theories/homotopy_theory/homotopy.v
6565
theories/homotopy_theory/wedge_sigT.v
66+
theories/homotopy_theory/path.v
6667

6768
theories/separation_axioms.v
6869
theories/function_spaces.v

classical/classical_sets.v

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2485,6 +2485,16 @@ Proof. by apply/eqP => /seteqP[] /(_ point) /(_ Logic.I). Qed.
24852485

24862486
End PointedTheory.
24872487

2488+
HB.mixin Record isBiPointed (X : Type) of Equality X := {
2489+
zero : X;
2490+
one : X;
2491+
zero_one_neq : zero != one;
2492+
}.
2493+
2494+
#[short(type="biPointedType")]
2495+
HB.structure Definition BiPointed :=
2496+
{ X of Choice X & isBiPointed X }.
2497+
24882498
Variant squashed T : Prop := squash (x : T).
24892499
Arguments squash {T} x.
24902500
Notation "$| T |" := (squashed T) : form_scope.

theories/Make

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -31,6 +31,7 @@ topology_theory/sigT_topology.v
3131

3232
homotopy_theory/homotopy.v
3333
homotopy_theory/wedge_sigT.v
34+
homotopy_theory/path.v
3435

3536
separation_axioms.v
3637
function_spaces.v
Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1,2 +1,3 @@
11
(* mathcomp analysis (c) 2024 Inria and AIST. License: CeCILL-C. *)
22
From mathcomp Require Export wedge_sigT.
3+
From mathcomp Require Export path.

theories/homotopy_theory/path.v

Lines changed: 181 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,181 @@
1+
(* mathcomp analysis (c) 2024 Inria and AIST. License: CeCILL-C. *)
2+
From HB Require Import structures.
3+
From mathcomp Require Import all_ssreflect all_algebra finmap generic_quotient.
4+
From mathcomp Require Import mathcomp_extra boolp classical_sets functions.
5+
From mathcomp Require Import cardinality fsbigop reals signed topology.
6+
From mathcomp Require Import function_spaces wedge_sigT.
7+
8+
(**md**************************************************************************)
9+
(* # Paths *)
10+
(* Paths from biPointed spaces. *)
11+
(* *)
12+
(* ``` *)
13+
(* mk_path ctsf f0 f1 == constructs a value in `pathType x y` given proofs *)
14+
(* that the endpoints of `f` are `x` and `y` *)
15+
(* reparameterize f phi == the path `f` with a different parameterization *)
16+
(* chain_path f g == the path which follows f, then follows g *)
17+
(* Only makes sense when `f one = g zero`. The *)
18+
(* domain is the wedge of domains of `f` and `g`. *)
19+
(* ``` *)
20+
(* The type `{path i from x to y in T}` is the continuous functions on the *)
21+
(* bipointed space i that go from x to y staying in T. It is endowed with *)
22+
(* - Topology via the compact-open topology *)
23+
(* *)
24+
(******************************************************************************)
25+
26+
Set Implicit Arguments.
27+
Unset Strict Implicit.
28+
Unset Printing Implicit Defensive.
29+
30+
Reserved Notation "{ 'path' i 'from' x 'to' y }" (
31+
at level 0, i at level 69, x at level 69, y at level 69,
32+
only parsing,
33+
format "{ 'path' i 'from' x 'to' y }").
34+
Reserved Notation "{ 'path' i 'from' x 'to' y 'in' T }" (
35+
at level 0, i at level 69, x at level 69, y at level 69, T at level 69,
36+
format "{ 'path' i 'from' x 'to' y 'in' T }").
37+
38+
Local Open Scope classical_set_scope.
39+
Local Open Scope ring_scope.
40+
Local Open Scope quotient_scope.
41+
42+
HB.mixin Record isPath {i : bpTopologicalType} {T : topologicalType} (x y : T)
43+
(f : i -> T) of isContinuous i T f := {
44+
path_zero : f zero = x;
45+
path_one : f one = y;
46+
}.
47+
48+
#[short(type="pathType")]
49+
HB.structure Definition Path {i : bpTopologicalType} {T : topologicalType}
50+
(x y : T) := {f of isPath i T x y f & isContinuous i T f}.
51+
52+
Notation "{ 'path' i 'from' x 'to' y }" := (pathType i x y) : type_scope.
53+
Notation "{ 'path' i 'from' x 'to' y 'in' T }" :=
54+
(@pathType i T x y) : type_scope.
55+
56+
HB.instance Definition _ {i : bpTopologicalType}
57+
{T : topologicalType} (x y : T) := gen_eqMixin {path i from x to y}.
58+
HB.instance Definition _ {i : bpTopologicalType}
59+
{T : topologicalType} (x y : T) := gen_choiceMixin {path i from x to y}.
60+
HB.instance Definition _ {i : bpTopologicalType}
61+
{T : topologicalType} (x y : T) :=
62+
Topological.copy {path i from x to y}
63+
(@weak_topology {path i from x to y} {compact-open, i -> T} id).
64+
65+
Section path_eq.
66+
Context {T : topologicalType} {i : bpTopologicalType} (x y : T).
67+
68+
Lemma path_eqP (a b : {path i from x to y}) : a = b <-> a =1 b.
69+
Proof.
70+
split=> [->//|].
71+
move: a b => [/= f [[+ +]]] [/= g [[+ +]]] fgE.
72+
move/funext : fgE => -> /= a1 [b1 c1] a2 [b2 c2]; congr (_ _).
73+
rewrite (Prop_irrelevance a1 a2) (Prop_irrelevance b1 b2).
74+
by rewrite (Prop_irrelevance c1 c2).
75+
Qed.
76+
77+
End path_eq.
78+
79+
Section cst_path.
80+
Context {T : topologicalType} {i : bpTopologicalType} (x: T).
81+
82+
HB.instance Definition _ := @isPath.Build i T x x (cst x) erefl erefl.
83+
End cst_path.
84+
85+
Section path_domain_path.
86+
Context {i : bpTopologicalType}.
87+
HB.instance Definition _ := @isPath.Build i i zero one idfun erefl erefl.
88+
End path_domain_path.
89+
90+
Section path_compose.
91+
Context {T U : topologicalType} (i: bpTopologicalType) (x y : T).
92+
Context (f : continuousType T U) (p : {path i from x to y}).
93+
94+
Local Lemma fp_zero : (f \o p) zero = f x.
95+
Proof. by rewrite /= !path_zero. Qed.
96+
97+
Local Lemma fp_one : (f \o p) one = f y.
98+
Proof. by rewrite /= !path_one. Qed.
99+
100+
HB.instance Definition _ := isPath.Build i U (f x) (f y) (f \o p)
101+
fp_zero fp_one.
102+
103+
End path_compose.
104+
105+
Section path_reparameterize.
106+
Context {T : topologicalType} (i j: bpTopologicalType) (x y : T).
107+
Context (f : {path i from x to y}) (phi : {path j from zero to one in i}).
108+
109+
(* we use reparameterize, as opposed to just `\o` because we can simplify the
110+
endpoints. So we don't end up with `f \o p : {path j from f zero to f one}`
111+
but rather `{path j from zero to one}`
112+
*)
113+
Definition reparameterize := f \o phi.
114+
115+
Local Lemma fphi_zero : reparameterize zero = x.
116+
Proof. by rewrite /reparameterize /= ?path_zero. Qed.
117+
118+
Local Lemma fphi_one : reparameterize one = y.
119+
Proof. by rewrite /reparameterize /= ?path_one. Qed.
120+
121+
Local Lemma fphi_cts : continuous reparameterize.
122+
Proof. by move=> ?; apply: continuous_comp; apply: cts_fun. Qed.
123+
124+
HB.instance Definition _ := isContinuous.Build _ _ reparameterize fphi_cts.
125+
126+
HB.instance Definition _ := isPath.Build j T x y reparameterize
127+
fphi_zero fphi_one.
128+
129+
End path_reparameterize.
130+
131+
Section mk_path.
132+
Context {i : bpTopologicalType} {T : topologicalType}.
133+
Context {x y : T} (f : i -> T) (ctsf : continuous f).
134+
Context (f0 : f zero = x) (f1 : f one = y).
135+
136+
Definition mk_path : i -> T := let _ := (ctsf, f0, f1) in f.
137+
138+
HB.instance Definition _ := isContinuous.Build i T mk_path ctsf.
139+
HB.instance Definition _ := @isPath.Build i T x y mk_path f0 f1.
140+
End mk_path.
141+
142+
Definition chain_path {i j : bpTopologicalType} {T : topologicalType}
143+
(f : i -> T) (g : j -> T) : bpwedge i j -> T :=
144+
wedge_fun (fun b => if b return (if b then i else j) -> T then f else g).
145+
146+
Lemma chain_path_cts_point {i j : bpTopologicalType} {T : topologicalType}
147+
(f : i -> T) (g : j -> T) :
148+
continuous f ->
149+
continuous g ->
150+
f one = g zero ->
151+
continuous (chain_path f g).
152+
Proof. by move=> cf cg fgE; apply: wedge_fun_continuous => // [[]|[] []]. Qed.
153+
154+
Section chain_path.
155+
Context {T : topologicalType} {i j : bpTopologicalType} (x y z: T).
156+
Context (p : {path i from x to y}) (q : {path j from y to z}).
157+
158+
Local Lemma chain_path_zero : chain_path p q zero = x.
159+
Proof.
160+
rewrite /chain_path /= wedge_lift_funE ?path_one ?path_zero //.
161+
by case => //= [][] //=; rewrite ?path_one ?path_zero.
162+
Qed.
163+
164+
Local Lemma chain_path_one : chain_path p q one = z.
165+
Proof.
166+
rewrite /chain_path /= wedge_lift_funE ?path_zero ?path_one //.
167+
by case => //= [][] //=; rewrite ?path_one ?path_zero.
168+
Qed.
169+
170+
Local Lemma chain_path_cts : continuous (chain_path p q).
171+
Proof.
172+
apply: chain_path_cts_point; [exact: cts_fun..|].
173+
by rewrite path_zero path_one.
174+
Qed.
175+
176+
HB.instance Definition _ := @isContinuous.Build _ T (chain_path p q)
177+
chain_path_cts.
178+
HB.instance Definition _ := @isPath.Build _ T x z (chain_path p q)
179+
chain_path_zero chain_path_one.
180+
181+
End chain_path.

theories/homotopy_theory/wedge_sigT.v

Lines changed: 31 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -24,6 +24,10 @@ From mathcomp Require Import separation_axioms function_spaces.
2424
(* wedge_prod == the mapping from the wedge as a quotient of sums to *)
2525
(* the wedge as a subspace of the product topology. *)
2626
(* It's an embedding when the index is finite. *)
27+
(* bpwedge_shared_pt b == the shared point in the bpwedge. Either zero or one *)
28+
(* depending on `b`. *)
29+
(* bpwedge == wedge of two bipointed spaces gluing zero to one *)
30+
(* bpwedge_lift == wedge_lift specialized to the bipointed wedge *)
2731
(* ``` *)
2832
(* *)
2933
(* The type `wedge p0` is endowed with the structures of: *)
@@ -35,6 +39,10 @@ From mathcomp Require Import separation_axioms function_spaces.
3539
(* - quotient *)
3640
(* - pointed *)
3741
(* *)
42+
(* The type `bpwedge` is endowed with the structures of: *)
43+
(* - topology via `quotient_topology` *)
44+
(* - quotient *)
45+
(* - bipointed *)
3846
(******************************************************************************)
3947

4048
Set Implicit Arguments.
@@ -379,3 +387,26 @@ HB.instance Definition _ := Quotient.on pwedge.
379387
HB.instance Definition _ := isPointed.Build pwedge pwedge_point.
380388

381389
End pwedge.
390+
391+
Section bpwedge.
392+
Context (X Y : bpTopologicalType).
393+
394+
Definition bpwedge_shared_pt b :=
395+
if b return (if b then X else Y) then @one X else @zero Y.
396+
Local Notation bpwedge := (@wedge bool _ bpwedge_shared_pt).
397+
Local Notation bpwedge_lift := (@wedge_lift bool _ bpwedge_shared_pt).
398+
399+
Local Lemma wedge_neq : @bpwedge_lift true zero != @bpwedge_lift false one.
400+
Proof.
401+
by apply/eqP => /eqmodP/predU1P[//|/andP[/= + _]]; exact/negP/zero_one_neq.
402+
Qed.
403+
404+
Local Lemma bpwedgeE : @bpwedge_lift true one = @bpwedge_lift false zero .
405+
Proof. by apply/eqmodP/orP; rewrite !eqxx; right. Qed.
406+
407+
HB.instance Definition _ := @isBiPointed.Build
408+
bpwedge (@bpwedge_lift true zero) (@bpwedge_lift false one) wedge_neq.
409+
End bpwedge.
410+
411+
Notation bpwedge X Y := (@wedge bool _ (bpwedge_shared_pt X Y)).
412+
Notation bpwedge_lift X Y := (@wedge_lift bool _ (bpwedge_shared_pt X Y)).

theories/topology_theory/topology_structure.v

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -105,6 +105,10 @@ HB.structure Definition Topological :=
105105
HB.structure Definition PointedTopological :=
106106
{T of PointedNbhs T & Nbhs_isTopological T}.
107107

108+
#[short(type="bpTopologicalType")]
109+
HB.structure Definition BiPointedTopological :=
110+
{ X of BiPointed X & Topological X }.
111+
108112
Section Topological1.
109113
Context {T : topologicalType}.
110114

0 commit comments

Comments
 (0)