Skip to content

Commit 283d192

Browse files
committed
homeomorphims for products
1 parent 57b7b2d commit 283d192

File tree

3 files changed

+64
-2
lines changed

3 files changed

+64
-2
lines changed

CHANGELOG_UNRELEASED.md

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -57,6 +57,13 @@
5757
`locally_compact_completely_regular`, and
5858
`completely_regular_regular`.
5959

60+
- in file `mathcomp_extra.v`,
61+
+ new definitions `left_assoc_prod`, and `right_assoc_prod`.
62+
+ new lemmas `left_assoc_prodK`, `right_assoc_prodK`, and `swapK`.
63+
- in file `product_topology.v`,
64+
+ new lemmas `swap_continuous`, `left_assoc_prod_continuous`, and
65+
`right_assoc_prod_continuous`.
66+
6067
### Changed
6168

6269
- in file `normedtype.v`,

classical/mathcomp_extra.v

Lines changed: 22 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -13,9 +13,11 @@ From mathcomp Require Import finset interval.
1313
(* dfwith f x == fun j => x if j = i, and f j otherwise *)
1414
(* given x : T i *)
1515
(* swap x := (x.2, x.1) *)
16-
(* map_pair f x := (f x.1, f x.2) *)
16+
(* map_pair f x := (f x.1, f x.2) *)
1717
(* monotonous A f := {in A &, {mono f : x y / x <= y}} \/ *)
1818
(* {in A &, {mono f : x y /~ x <= y}} *)
19+
(* left_assoc_prod x := sends (X * Y) * Z to X * (Y * Z) *)
20+
(* right_assoc_prod x := sends X * (Y * Z) to (X * Y) * Z *)
1921
(* ``` *)
2022
(* *)
2123
(******************************************************************************)
@@ -369,7 +371,25 @@ Lemma real_ltr_distlC [R : numDomainType] [x y : R] (e : R) :
369371
x - y \is Num.real -> (`|x - y| < e) = (x - e < y < x + e).
370372
Proof. by move=> ?; rewrite distrC real_ltr_distl// -rpredN opprB. Qed.
371373

372-
Definition swap (T1 T2 : Type) (x : T1 * T2) := (x.2, x.1).
374+
Definition swap {T1 T2 : Type} (x : T1 * T2) := (x.2, x.1).
375+
376+
Section reassociate_products.
377+
Context {X Y Z : Type}.
378+
Definition left_assoc_prod (xyz : (X * Y) * Z) : X * (Y * Z) :=
379+
(xyz.1.1,(xyz.1.2,xyz.2)).
380+
381+
Definition right_assoc_prod (xyz : X * (Y * Z)) : (X * Y) * Z :=
382+
((xyz.1,xyz.2.1),xyz.2.2).
383+
384+
Lemma left_assoc_prodK : cancel left_assoc_prod right_assoc_prod.
385+
Proof. by case;case. Qed.
386+
387+
Lemma right_assoc_prodK : cancel right_assoc_prod left_assoc_prod.
388+
Proof. by case => ? []. Qed.
389+
End reassociate_products.
390+
391+
Lemma swapK {T1 T2 : Type} : cancel (@swap T1 T2) (@swap T2 T1).
392+
Proof. by case=> ? ?. Qed.
373393

374394
Definition map_pair {S U : Type} (f : S -> U) (x : (S * S)) : (U * U) :=
375395
(f x.1, f x.2).

theories/topology_theory/product_topology.v

Lines changed: 35 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -231,3 +231,38 @@ by apply: (ngPr (b, a, i)); split => //; exact: N1N2N.
231231
Unshelve. all: by end_near. Qed.
232232
#[deprecated(since="mathcomp-analysis 0.6.0", note="renamed to `compact_setX`")]
233233
Notation compact_setM := compact_setX (only parsing).
234+
235+
Lemma swap_continuous {X Y : topologicalType} : continuous (@swap X Y).
236+
Proof.
237+
case=> a b W /=[[U V]][] /= Ua Vb UVW; exists (V, U); first by split.
238+
by case => //= ? ? [] ? ?; apply: UVW.
239+
Qed.
240+
241+
Section reassociate_continuous.
242+
Context {X Y Z : topologicalType}.
243+
244+
Lemma left_assoc_prod_continuous : continuous (@left_assoc_prod X Y Z).
245+
Proof.
246+
case;case=> a b c U [[/= P V]] [Pa] [][/= Q R][ Qb Rc] QRV PVU.
247+
exists ((P `*` Q), R); first split.
248+
- exists (P, Q); first split => //=.
249+
by case=> x y /= [Px Qy]; split => //.
250+
- done.
251+
- case; case=> x y z /= [][] ? ? ?; apply: PVU; split => //.
252+
exact: QRV.
253+
Qed.
254+
255+
Lemma right_assoc_prod_continuous : continuous (@right_assoc_prod X Y Z).
256+
Proof.
257+
case=> a [b c] U [[V R]] [/= [[P Q /=]]] [Pa Qb] PQV Rc VRU.
258+
exists (P, (Q `*` R)); first split.
259+
- done.
260+
- exists (Q, R); first split => //=.
261+
by case=> ? ? /= [? ?]; split.
262+
- case=> x [y z] [? [? ?]]; apply: VRU; split => //.
263+
exact: PQV.
264+
Qed.
265+
266+
End reassociate_continuous.
267+
268+

0 commit comments

Comments
 (0)