Skip to content

Commit

Permalink
Homeomorphisms for (X*Y)*Z -> X*(Y*Z) and X*Y -> Y*X (#1367)
Browse files Browse the repository at this point in the history
* homeomorphims for products

---------

Co-authored-by: Reynald Affeldt <[email protected]>
  • Loading branch information
zstone1 and affeldt-aist authored Nov 10, 2024
1 parent fa2bfc6 commit e792f1a
Show file tree
Hide file tree
Showing 3 changed files with 58 additions and 1 deletion.
5 changes: 5 additions & 0 deletions CHANGELOG_UNRELEASED.md
Original file line number Diff line number Diff line change
Expand Up @@ -71,6 +71,11 @@
+ lemma `countable_measurable`
- in `realfun.v`:
+ lemma `cvgr_dnbhsP`
+ new definitions `prodA`, and `prodAr`.
+ new lemmas `prodAK`, `prodArK`, and `swapK`.
- in file `product_topology.v`,
+ new lemmas `swap_continuous`, `prodA_continuous`, and
`prodAr_continuous`.

### Changed

Expand Down
24 changes: 23 additions & 1 deletion classical/mathcomp_extra.v
Original file line number Diff line number Diff line change
Expand Up @@ -18,6 +18,8 @@ From mathcomp Require Import finset interval.
(* {in A &, {mono f : x y /~ x <= y}} *)
(* sigT_fun f := lifts a family of functions f into a function on *)
(* the dependent sum *)
(* prodA x := sends (X * Y) * Z to X * (Y * Z) *)
(* prodAr x := sends X * (Y * Z) to (X * Y) * Z *)
(* ``` *)
(* *)
(******************************************************************************)
Expand Down Expand Up @@ -373,7 +375,27 @@ Lemma real_ltr_distlC [R : numDomainType] [x y : R] (e : R) :
x - y \is Num.real -> (`|x - y| < e) = (x - e < y < x + e).
Proof. by move=> ?; rewrite distrC real_ltr_distl// -rpredN opprB. Qed.

Definition swap (T1 T2 : Type) (x : T1 * T2) := (x.2, x.1).
Definition swap {T1 T2 : Type} (x : T1 * T2) := (x.2, x.1).

Section reassociate_products.
Context {X Y Z : Type}.

Definition prodA (xyz : (X * Y) * Z) : X * (Y * Z) :=
(xyz.1.1, (xyz.1.2, xyz.2)).

Definition prodAr (xyz : X * (Y * Z)) : (X * Y) * Z :=
((xyz.1, xyz.2.1), xyz.2.2).

Lemma prodAK : cancel prodA prodAr.
Proof. by case; case. Qed.

Lemma prodArK : cancel prodAr prodA.
Proof. by case => ? []. Qed.

End reassociate_products.

Lemma swapK {T1 T2 : Type} : cancel (@swap T1 T2) (@swap T2 T1).
Proof. by case=> ? ?. Qed.

Definition map_pair {S U : Type} (f : S -> U) (x : (S * S)) : (U * U) :=
(f x.1, f x.2).
Expand Down
30 changes: 30 additions & 0 deletions theories/topology_theory/product_topology.v
Original file line number Diff line number Diff line change
Expand Up @@ -231,3 +231,33 @@ by apply: (ngPr (b, a, i)); split => //; exact: N1N2N.
Unshelve. all: by end_near. Qed.
#[deprecated(since="mathcomp-analysis 0.6.0", note="renamed to `compact_setX`")]
Notation compact_setM := compact_setX (only parsing).

Lemma swap_continuous {X Y : topologicalType} : continuous (@swap X Y).
Proof.
case=> a b W /= [[U V]][] /= Ub Va UVW; exists (V, U); first by split.
by case => //= ? ? [] ? ?; exact: UVW.
Qed.

Section reassociate_continuous.
Context {X Y Z : topologicalType}.

Lemma prodA_continuous : continuous (@prodA X Y Z).
Proof.
move=> [] [a b] c U [[/= P V]] [Pa] [][/= Q R][ Qb Rc] QRV PVU.
exists ((P `*` Q), R); first split.
- exists (P, Q); first by [].
by case=> x y /= [Px Qy].
- by [].
- by move=> [[x y] z] [] [] ? ? ?; apply: PVU; split => //; exact: QRV.
Qed.

Lemma prodAr_continuous : continuous (@prodAr X Y Z).
Proof.
case=> a [b c] U [[V R]] [/= [[P Q /=]]] [Pa Qb] PQV Rc VRU.
exists (P, (Q `*` R)); first split => //.
- exists (Q, R); first by [].
by case=> ? ? /= [].
- by case=> x [y z] [? [? ?]]; apply: VRU; split => //; exact: PQV.
Qed.

End reassociate_continuous.

0 comments on commit e792f1a

Please sign in to comment.