diff --git a/src/foundation-core/commuting-squares-of-homotopies.lagda.md b/src/foundation-core/commuting-squares-of-homotopies.lagda.md
index 7c3e1c6877..27936d2400 100644
--- a/src/foundation-core/commuting-squares-of-homotopies.lagda.md
+++ b/src/foundation-core/commuting-squares-of-homotopies.lagda.md
@@ -37,7 +37,9 @@ homotopy is called a
{{#concept "coherence" Disambiguation="commuting square of homotopies" Agda=coherence-square-homotopies}}
of the square.
-## Definition
+## Definitions
+
+### Commuting squares of homotopies
```agda
module _
diff --git a/src/foundation-core/constant-maps.lagda.md b/src/foundation-core/constant-maps.lagda.md
index 420d04b577..ee7650e01c 100644
--- a/src/foundation-core/constant-maps.lagda.md
+++ b/src/foundation-core/constant-maps.lagda.md
@@ -12,9 +12,24 @@ open import foundation.universe-levels
+## Idea
+
+The {{#concept "constant map" Agda=const}} from `A` to `B` at an element `b : B`
+is the function `x ↦ b` mapping every element `x : A` to the given element
+`b : B`. In common type theoretic notation, the constant map at `b` is the
+function
+
+```text
+ λ x → b.
+```
+
## Definition
```agda
const : {l1 l2 : Level} (A : UU l1) (B : UU l2) → B → A → B
const A B b x = b
```
+
+## See also
+
+- [Constant pointed maps](structured-types.constant-pointed-maps.md)
diff --git a/src/foundation-core/homotopies.lagda.md b/src/foundation-core/homotopies.lagda.md
index 9c53f25ce0..2159d361ea 100644
--- a/src/foundation-core/homotopies.lagda.md
+++ b/src/foundation-core/homotopies.lagda.md
@@ -274,6 +274,18 @@ inv-nat-htpy :
ap f p ∙ H y = H x ∙ ap g p
inv-nat-htpy H p = inv (nat-htpy H p)
+nat-refl-htpy :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B)
+ {x y : A} (p : x = y) →
+ nat-htpy (refl-htpy' f) p = inv right-unit
+nat-refl-htpy f refl = refl
+
+inv-nat-refl-htpy :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B)
+ {x y : A} (p : x = y) →
+ inv-nat-htpy (refl-htpy' f) p = right-unit
+inv-nat-refl-htpy f refl = refl
+
nat-htpy-id :
{l : Level} {A : UU l} {f : A → A} (H : f ~ id)
{x y : A} (p : x = y) → H x ∙ p = ap f p ∙ H y
diff --git a/src/foundation-core/identity-types.lagda.md b/src/foundation-core/identity-types.lagda.md
index e36fc57851..8a4289d598 100644
--- a/src/foundation-core/identity-types.lagda.md
+++ b/src/foundation-core/identity-types.lagda.md
@@ -231,6 +231,8 @@ use of the fact that the identity types are defined _for all types_. In other
words, since identity types are themselves types, we can consider identity types
of identity types, and so on.
+#### Associators
+
```agda
module _
{l : Level} {A : UU l}
@@ -264,6 +266,21 @@ module _
{l : Level} {A : UU l}
where
+ double-assoc :
+ {x y z w v : A} (p : x = y) (q : y = z) (r : z = w) (s : w = v) →
+ (((p ∙ q) ∙ r) ∙ s) = p ∙ (q ∙ (r ∙ s))
+ double-assoc refl q r s = assoc q r s
+
+ triple-assoc :
+ {x y z w v u : A}
+ (p : x = y) (q : y = z) (r : z = w) (s : w = v) (t : v = u) →
+ ((((p ∙ q) ∙ r) ∙ s) ∙ t) = p ∙ (q ∙ (r ∙ (s ∙ t)))
+ triple-assoc refl q r s t = double-assoc q r s t
+```
+
+#### Unit laws
+
+```agda
left-unit : {x y : A} {p : x = y} → refl ∙ p = p
left-unit = refl
@@ -387,11 +404,21 @@ module _
p ∙ q = r → q = inv p ∙ r
left-transpose-eq-concat refl q r s = s
+ inv-left-transpose-eq-concat :
+ {x y : A} (p : x = y) {z : A} (q : y = z) (r : x = z) →
+ q = inv p ∙ r → p ∙ q = r
+ inv-left-transpose-eq-concat refl q r s = s
+
right-transpose-eq-concat :
{x y : A} (p : x = y) {z : A} (q : y = z) (r : x = z) →
p ∙ q = r → p = r ∙ inv q
right-transpose-eq-concat p refl r s = (inv right-unit ∙ s) ∙ inv right-unit
+ inv-right-transpose-eq-concat :
+ {x y : A} (p : x = y) {z : A} (q : y = z) (r : x = z) →
+ p = r ∙ inv q → p ∙ q = r
+ inv-right-transpose-eq-concat p refl r s = right-unit ∙ (s ∙ right-unit)
+
double-transpose-eq-concat :
{x y u v : A} (r : x = u) (p : x = y) (s : u = v) (q : y = v) →
p ∙ q = r ∙ s → (inv r) ∙ p = s ∙ (inv q)
diff --git a/src/foundation/commuting-triangles-of-identifications.lagda.md b/src/foundation/commuting-triangles-of-identifications.lagda.md
index 071dbaa96f..5e15c81f17 100644
--- a/src/foundation/commuting-triangles-of-identifications.lagda.md
+++ b/src/foundation/commuting-triangles-of-identifications.lagda.md
@@ -8,6 +8,9 @@ module foundation.commuting-triangles-of-identifications where
```agda
open import foundation.action-on-identifications-functions
+open import foundation.binary-equivalences
+open import foundation.commuting-squares-of-identifications
+open import foundation.dependent-pair-types
open import foundation.identity-types
open import foundation.path-algebra
open import foundation.universe-levels
@@ -33,8 +36,7 @@ A triangle of [identifications](foundation-core.identity-types.md)
z
```
-is said to **commute** if there is a higher identification between the `x = z`
-and the concatenated identification `x = y = z`.
+is said to **commute** if there is an identification `p = q ∙ r`.
## Definitions
@@ -47,11 +49,23 @@ module _
coherence-triangle-identifications :
(left : x = z) (right : y = z) (top : x = y) → UU l
- coherence-triangle-identifications left right top = left = (top ∙ right)
+ coherence-triangle-identifications left right top = left = top ∙ right
coherence-triangle-identifications' :
(left : x = z) (right : y = z) (top : x = y) → UU l
- coherence-triangle-identifications' left right top = (top ∙ right) = left
+ coherence-triangle-identifications' left right top = top ∙ right = left
+```
+
+### The horizontally constant triangle of identifications
+
+```agda
+module _
+ {l : Level} {A : UU l} {x y : A}
+ where
+
+ horizontal-refl-coherence-triangle-identifications :
+ (p : x = y) → coherence-triangle-identifications p p refl
+ horizontal-refl-coherence-triangle-identifications p = refl
```
## Properties
@@ -168,7 +182,7 @@ module _
coherence-triangle-identifications left right top ≃
coherence-triangle-identifications (p ∙ left) right (p ∙ top)
equiv-left-whisker-concat-coherence-triangle-identifications =
- ( inv-equiv (equiv-concat-assoc' (p ∙ left) p top right)) ∘e
+ ( equiv-inv-concat' _ (assoc p top right)) ∘e
( equiv-left-whisker-concat p)
left-whisker-concat-coherence-triangle-identifications :
@@ -183,11 +197,18 @@ module _
left-unwhisker-concat-coherence-triangle-identifications =
map-inv-equiv equiv-left-whisker-concat-coherence-triangle-identifications
+ is-equiv-left-whisker-concat-coherence-triangle-identifications :
+ is-equiv
+ ( left-whisker-concat-coherence-triangle-identifications)
+ is-equiv-left-whisker-concat-coherence-triangle-identifications =
+ is-equiv-map-equiv
+ equiv-left-whisker-concat-coherence-triangle-identifications
+
equiv-left-whisker-concat-coherence-triangle-identifications' :
coherence-triangle-identifications' left right top ≃
coherence-triangle-identifications' (p ∙ left) right (p ∙ top)
equiv-left-whisker-concat-coherence-triangle-identifications' =
- ( inv-equiv (equiv-concat-assoc p top right (p ∙ left))) ∘e
+ ( equiv-concat (assoc p top right) _) ∘e
( equiv-left-whisker-concat p)
left-whisker-concat-coherence-triangle-identifications' :
@@ -201,6 +222,12 @@ module _
coherence-triangle-identifications' left right top
left-unwhisker-concat-coherence-triangle-identifications' =
map-inv-equiv equiv-left-whisker-concat-coherence-triangle-identifications'
+
+ is-equiv-left-whisker-concat-coherence-triangle-identifications' :
+ is-equiv left-whisker-concat-coherence-triangle-identifications'
+ is-equiv-left-whisker-concat-coherence-triangle-identifications' =
+ is-equiv-map-equiv
+ equiv-left-whisker-concat-coherence-triangle-identifications'
```
#### Right whiskering commuting squares of identifications
@@ -243,6 +270,12 @@ module _
right-unwhisker-concat-coherence-triangle-identifications =
map-inv-equiv equiv-right-whisker-concat-coherence-triangle-identifications
+ is-equiv-right-whisker-concat-coherence-triangle-identifications :
+ is-equiv right-whisker-concat-coherence-triangle-identifications
+ is-equiv-right-whisker-concat-coherence-triangle-identifications =
+ is-equiv-map-equiv
+ equiv-right-whisker-concat-coherence-triangle-identifications
+
equiv-right-whisker-concat-coherence-triangle-identifications' :
coherence-triangle-identifications' left right top ≃
coherence-triangle-identifications' (left ∙ p) (right ∙ p) top
@@ -261,6 +294,12 @@ module _
coherence-triangle-identifications' left right top
right-unwhisker-concat-coherence-triangle-identifications' =
map-inv-equiv equiv-right-whisker-concat-coherence-triangle-identifications'
+
+ is-equiv-right-whisker-concat-coherence-triangle-identifications' :
+ is-equiv right-whisker-concat-coherence-triangle-identifications'
+ is-equiv-right-whisker-concat-coherence-triangle-identifications' =
+ is-equiv-map-equiv
+ equiv-right-whisker-concat-coherence-triangle-identifications'
```
#### Splicing a pair of mutual inverse identifications in a commuting triangle of identifications
@@ -415,3 +454,951 @@ module _
refl = map-coherence-triangle-identifications p p refl refl
compute-refl-top-map-coherence-triangle-identifications p = refl
```
+
+### The action of functions on commuting triangles of identifications
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B)
+ {x y z : A} (left : x = z) (right : y = z) (top : x = y)
+ where
+
+ action-function-coherence-triangle-identifications :
+ coherence-triangle-identifications left right top →
+ coherence-triangle-identifications (ap f left) (ap f right) (ap f top)
+ action-function-coherence-triangle-identifications s =
+ ap (ap f) s ∙ ap-concat f top right
+```
+
+### Inverting one side of a commuting triangle of identifications
+
+```agda
+module _
+ {l1 : Level} {A : UU l1}
+ where
+
+ transpose-top-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : b = c) (top : b = a)
+ {top' : a = b} (α : inv top = top') →
+ coherence-triangle-identifications right left top →
+ coherence-triangle-identifications left right top'
+ transpose-top-coherence-triangle-identifications left right top refl t =
+ left-transpose-eq-concat _ _ _ (inv t)
+
+ equiv-transpose-top-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : b = c) (top : b = a) →
+ coherence-triangle-identifications right left top ≃
+ coherence-triangle-identifications left right (inv top)
+ equiv-transpose-top-coherence-triangle-identifications left right top =
+ equiv-left-transpose-eq-concat _ _ _ ∘e equiv-inv _ _
+
+ equiv-higher-transpose-top-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : b = c) (top : b = a)
+ {left' : a = c} (p : left = left')
+ {top' : a = b} (α : inv top = top') →
+ (s : coherence-triangle-identifications right left top) →
+ (t : coherence-triangle-identifications right left' top) →
+ coherence-triangle-identifications t (left-whisker-concat top p) s ≃
+ coherence-triangle-identifications
+ ( transpose-top-coherence-triangle-identifications left right top α s)
+ ( transpose-top-coherence-triangle-identifications left' right top α t)
+ ( p)
+ equiv-higher-transpose-top-coherence-triangle-identifications
+ left right top refl refl s t =
+ ( equiv-ap
+ ( equiv-transpose-top-coherence-triangle-identifications left right top)
+ ( _)
+ ( _)) ∘e
+ ( equiv-inv _ _) ∘e
+ ( equiv-concat' _ right-unit)
+
+ higher-transpose-top-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : b = c) (top : b = a) →
+ {left' : a = c} (p : left = left')
+ {top' : a = b} (q : inv top = top') →
+ (s : coherence-triangle-identifications right left top) →
+ (t : coherence-triangle-identifications right left' top) →
+ coherence-triangle-identifications t (left-whisker-concat top p) s →
+ coherence-triangle-identifications
+ ( transpose-top-coherence-triangle-identifications left right top q s)
+ ( transpose-top-coherence-triangle-identifications left' right top q t)
+ ( p)
+ higher-transpose-top-coherence-triangle-identifications
+ left right top p q s t =
+ map-equiv
+ ( equiv-higher-transpose-top-coherence-triangle-identifications
+ left right top p q s t)
+
+ transpose-right-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : c = b) (top : a = b)
+ {right' : b = c} (p : inv right = right') →
+ coherence-triangle-identifications top right left →
+ coherence-triangle-identifications left right' top
+ transpose-right-coherence-triangle-identifications left right top refl t =
+ right-transpose-eq-concat _ _ _ (inv t)
+
+ equiv-transpose-right-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : c = b) (top : a = b) →
+ coherence-triangle-identifications top right left ≃
+ coherence-triangle-identifications left (inv right) top
+ equiv-transpose-right-coherence-triangle-identifications left right top =
+ equiv-right-transpose-eq-concat _ _ _ ∘e equiv-inv _ _
+
+ equiv-higher-transpose-right-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : c = b) (top : a = b) →
+ {left' : a = c} (p : left = left')
+ {right' : b = c} (q : inv right = right') →
+ (s : coherence-triangle-identifications top right left) →
+ (t : coherence-triangle-identifications top right left') →
+ coherence-triangle-identifications t (right-whisker-concat p right) s ≃
+ coherence-triangle-identifications
+ ( transpose-right-coherence-triangle-identifications left right top q s)
+ ( transpose-right-coherence-triangle-identifications left' right top q t)
+ ( p)
+ equiv-higher-transpose-right-coherence-triangle-identifications
+ left right top refl refl s t =
+ ( equiv-ap
+ ( equiv-transpose-right-coherence-triangle-identifications left right top)
+ ( _)
+ ( _)) ∘e
+ ( equiv-inv _ _) ∘e
+ ( equiv-concat' t right-unit)
+
+ higher-transpose-right-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : c = b) (top : a = b) →
+ {left' : a = c} (p : left = left')
+ {right' : b = c} (q : inv right = right') →
+ (s : coherence-triangle-identifications top right left) →
+ (t : coherence-triangle-identifications top right left') →
+ coherence-triangle-identifications t (right-whisker-concat p right) s →
+ coherence-triangle-identifications
+ ( transpose-right-coherence-triangle-identifications left right top q s)
+ ( transpose-right-coherence-triangle-identifications left' right top q t)
+ ( p)
+ higher-transpose-right-coherence-triangle-identifications
+ left right top p q s t =
+ map-equiv
+ ( equiv-higher-transpose-right-coherence-triangle-identifications
+ ( left)
+ ( right)
+ ( top)
+ ( p)
+ ( q)
+ ( s)
+ ( t))
+```
+
+### Concatenating identifications on edges with coherences of commuting triangles of identifications
+
+```agda
+module _
+ {l1 : Level} {A : UU l1}
+ where
+
+ equiv-concat-top-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : b = c) (top : a = b)
+ {top' : a = b} (p : top' = top) →
+ coherence-triangle-identifications left right top' ≃
+ coherence-triangle-identifications left right top
+ equiv-concat-top-coherence-triangle-identifications left right top p =
+ equiv-concat' left (right-whisker-concat p right)
+
+ concat-top-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : b = c) (top : a = b)
+ {top' : a = b} (p : top' = top) →
+ coherence-triangle-identifications left right top' →
+ coherence-triangle-identifications left right top
+ concat-top-coherence-triangle-identifications left right top p =
+ map-equiv
+ ( equiv-concat-top-coherence-triangle-identifications left right top p)
+
+ equiv-concat-right-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : b = c) (top : a = b)
+ {right' : b = c} (p : right' = right) →
+ coherence-triangle-identifications left right' top ≃
+ coherence-triangle-identifications left right top
+ equiv-concat-right-coherence-triangle-identifications left right top p =
+ equiv-concat' left (left-whisker-concat top p)
+
+ concat-right-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : b = c) (top : a = b)
+ {right' : b = c} (p : right' = right) →
+ coherence-triangle-identifications left right' top →
+ coherence-triangle-identifications left right top
+ concat-right-coherence-triangle-identifications left right top p =
+ map-equiv
+ ( equiv-concat-right-coherence-triangle-identifications left right top p)
+
+ equiv-concat-left-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : b = c) (top : a = b)
+ {left' : a = c} (p : left = left') →
+ coherence-triangle-identifications left' right top ≃
+ coherence-triangle-identifications left right top
+ equiv-concat-left-coherence-triangle-identifications left right top p =
+ equiv-concat p (top ∙ right)
+
+ concat-left-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : b = c) (top : a = b)
+ {left' : a = c} (p : left = left') →
+ coherence-triangle-identifications left' right top →
+ coherence-triangle-identifications left right top
+ concat-left-coherence-triangle-identifications left right top p =
+ map-equiv
+ ( equiv-concat-left-coherence-triangle-identifications left right top p)
+```
+
+### Horizontal pasting of commuting triangles of identifications
+
+Consider a commuting diagram of identifications of the form
+
+```text
+ top-left top-right
+ a ---> b ---> c
+ \ | /
+ left \ |m / right
+ \ | /
+ ∨ ∨ ∨
+ d
+```
+
+Then the outer triangle commutes too. Indeed, an identification
+`left = top-left ∙ middle` is given. Then, an identification
+
+```text
+ top-left ∙ middle = (top-left ∙ top-right) ∙ right
+```
+
+is obtained immediately by left whiskering the right triangle with the
+identification `top-left`. Note that this direct construction of the coherence
+of the outer commuting triangle of identifications avoids any use of
+associativity.
+
+```agda
+module _
+ {l : Level} {A : UU l}
+ {a b c d : A} (left : a = d) (middle : b = d) (right : c = d)
+ (top-left : a = b) (top-right : b = c)
+ where
+
+ horizontal-pasting-coherence-triangle-identifications :
+ coherence-triangle-identifications left middle top-left →
+ coherence-triangle-identifications middle right top-right →
+ coherence-triangle-identifications left right (top-left ∙ top-right)
+ horizontal-pasting-coherence-triangle-identifications s t =
+ ( s) ∙
+ ( left-whisker-concat-coherence-triangle-identifications
+ ( top-left)
+ ( middle)
+ ( right)
+ ( top-right)
+ ( t))
+```
+
+### Horizontal pasting of commuting triangles of identifications is a binary equivalence
+
+```agda
+module _
+ {l : Level} {A : UU l}
+ {a b c d : A} (left : a = d) (middle : b = d) (right : c = d)
+ (top-left : a = b) (top-right : b = c)
+ where
+
+ abstract
+ is-equiv-horizontal-pasting-coherence-triangle-identifications :
+ (s : coherence-triangle-identifications left middle top-left) →
+ is-equiv
+ ( horizontal-pasting-coherence-triangle-identifications
+ ( left)
+ ( middle)
+ ( right)
+ ( top-left)
+ ( top-right)
+ ( s))
+ is-equiv-horizontal-pasting-coherence-triangle-identifications s =
+ is-equiv-comp _ _
+ ( is-equiv-left-whisker-concat-coherence-triangle-identifications
+ ( top-left)
+ ( middle)
+ ( right)
+ ( top-right))
+ ( is-equiv-concat s _)
+
+ equiv-horizontal-pasting-coherence-triangle-identifications :
+ (s : coherence-triangle-identifications left middle top-left) →
+ coherence-triangle-identifications middle right top-right ≃
+ coherence-triangle-identifications left right (top-left ∙ top-right)
+ pr1 (equiv-horizontal-pasting-coherence-triangle-identifications s) =
+ horizontal-pasting-coherence-triangle-identifications _ _ _ _ _ s
+ pr2 (equiv-horizontal-pasting-coherence-triangle-identifications s) =
+ is-equiv-horizontal-pasting-coherence-triangle-identifications s
+
+ abstract
+ is-equiv-horizontal-pasting-coherence-triangle-identifications' :
+ (t : coherence-triangle-identifications middle right top-right) →
+ is-equiv
+ ( λ s →
+ horizontal-pasting-coherence-triangle-identifications
+ ( left)
+ ( middle)
+ ( right)
+ ( top-left)
+ ( top-right)
+ ( s)
+ ( t))
+ is-equiv-horizontal-pasting-coherence-triangle-identifications' t =
+ is-equiv-concat' _
+ ( left-whisker-concat-coherence-triangle-identifications
+ ( top-left)
+ ( middle)
+ ( right)
+ ( top-right)
+ ( t))
+
+ equiv-horizontal-pasting-coherence-triangle-identifications' :
+ (t : coherence-triangle-identifications middle right top-right) →
+ coherence-triangle-identifications left middle top-left ≃
+ coherence-triangle-identifications left right (top-left ∙ top-right)
+ pr1 (equiv-horizontal-pasting-coherence-triangle-identifications' t) s =
+ horizontal-pasting-coherence-triangle-identifications
+ ( left)
+ ( middle)
+ ( right)
+ ( top-left)
+ ( top-right)
+ ( s)
+ ( t)
+ pr2 (equiv-horizontal-pasting-coherence-triangle-identifications' t) =
+ is-equiv-horizontal-pasting-coherence-triangle-identifications' t
+
+ is-binary-equiv-horizontal-pasting-coherence-triangle-identifications :
+ is-binary-equiv
+ ( horizontal-pasting-coherence-triangle-identifications
+ ( left)
+ ( middle)
+ ( right)
+ ( top-left)
+ ( top-right))
+ pr1 is-binary-equiv-horizontal-pasting-coherence-triangle-identifications =
+ is-equiv-horizontal-pasting-coherence-triangle-identifications'
+ pr2 is-binary-equiv-horizontal-pasting-coherence-triangle-identifications =
+ is-equiv-horizontal-pasting-coherence-triangle-identifications
+```
+
+### The left unit law for horizontal pastinf of commuting triangles of identifications
+
+Consider a commuting triangle of identifications
+
+```text
+ top
+ a ------> b
+ \ /
+ left \ / right
+ \ /
+ ∨ ∨
+ c
+```
+
+with `t : left = top ∙ right`. Then we have an identification
+
+```text
+ horizontal-pasting left left right refl top refl t = t
+```
+
+```agda
+module _
+ {l : Level} {A : UU l} {a b c : A}
+ where
+
+ left-unit-law-horizontal-pasting-coherence-triangle-identifications :
+ (left : a = c) (right : b = c) (top : a = b) →
+ (t : coherence-triangle-identifications left right top) →
+ horizontal-pasting-coherence-triangle-identifications
+ ( left)
+ ( left)
+ ( right)
+ ( refl)
+ ( top)
+ ( refl)
+ ( t) =
+ t
+ left-unit-law-horizontal-pasting-coherence-triangle-identifications
+ ._ right top refl =
+ refl
+```
+
+### The left unit law for horizontal pastinf of commuting triangles of identifications
+
+Consider a commuting triangle of identifications
+
+```text
+ top
+ a ------> b
+ \ /
+ left \ / right
+ \ /
+ ∨ ∨
+ c
+```
+
+with `t : left = top ∙ right`. Then we have a commuting triangle of
+identifications
+
+```text
+ horizontal-pasting t refl
+ left ----------------------> (top ∙ refl) ∙ right
+ \ /
+ \ /
+ t \ / right-whisker right-unit right
+ \ /
+ \ /
+ ∨ ∨
+ top ∙ right
+```
+
+```agda
+module _
+ {l : Level} {A : UU l} {a b c : A}
+ where
+
+ right-unit-law-horizontal-pasting-coherence-triangle-identifications :
+ (left : a = c) (right : b = c) (top : a = b) →
+ (t : coherence-triangle-identifications left right top) →
+ coherence-triangle-identifications
+ ( t)
+ ( right-whisker-concat right-unit right)
+ ( horizontal-pasting-coherence-triangle-identifications
+ ( left)
+ ( right)
+ ( right)
+ ( top)
+ ( refl)
+ ( t)
+ ( refl))
+ right-unit-law-horizontal-pasting-coherence-triangle-identifications
+ ._ right refl refl = refl
+```
+
+### Associativity of horizontal pasting of coherences of commuting triangles of identifications
+
+```agda
+module _
+ {l : Level} {A : UU l} {a b c d e : A}
+ where
+
+ associative-horizontal-pasting-coherence-triangle-identifications :
+ (left : a = e) (mid-left : b = e) (mid-right : c = e) (right : d = e)
+ (top-left : a = b) (top-middle : b = c) (top-right : c = d)
+ (r : coherence-triangle-identifications left mid-left top-left) →
+ (s : coherence-triangle-identifications mid-left mid-right top-middle) →
+ (t : coherence-triangle-identifications mid-right right top-right) →
+ coherence-triangle-identifications
+ ( horizontal-pasting-coherence-triangle-identifications
+ ( left)
+ ( mid-left)
+ ( right)
+ ( top-left)
+ ( top-middle ∙ top-right)
+ ( r)
+ ( horizontal-pasting-coherence-triangle-identifications
+ ( mid-left)
+ ( mid-right)
+ ( right)
+ ( top-middle)
+ ( top-right)
+ ( s)
+ ( t)))
+ ( right-whisker-concat (assoc top-left top-middle top-right) right)
+ ( horizontal-pasting-coherence-triangle-identifications
+ ( left)
+ ( mid-right)
+ ( right)
+ ( top-left ∙ top-middle)
+ ( top-right)
+ ( horizontal-pasting-coherence-triangle-identifications
+ ( left)
+ ( mid-left)
+ ( mid-right)
+ ( top-left)
+ ( top-middle)
+ ( r)
+ ( s))
+ ( t))
+ associative-horizontal-pasting-coherence-triangle-identifications
+ refl .refl .refl .refl refl refl refl refl refl refl =
+ refl
+```
+
+### Left whiskering of horizontal pasting of commuting triangles of identifications
+
+Consider two commuting triangles of identifications as in the diagram
+
+```text
+ s t
+ a ----> b ----> c
+ \ | /
+ \ L | R /
+ l \ |m / r
+ \ | /
+ ∨ ∨ ∨
+ d,
+```
+
+and consider furthermore a commuting triangle of identifications
+
+```text
+ t'
+ b ----> c
+ | /
+ | R' /
+ m | / r
+ | /
+ ∨ ∨
+ d
+```
+
+where the identifications `m : b = d` and `r : c = d` are the same as in the
+previous diagram. Finally, consider an identification `p : t = t'` and an
+identification `q` witnessing that the triangle
+
+```text
+ R
+ m ------> t ∙ r
+ \ /
+ R' \ / right-whisker p r
+ \ /
+ ∨ ∨
+ t' ∙ r
+```
+
+commutes. Then the triangle
+
+```text
+ horizontal-pasting L R
+ l ----------------> (s ∙ t) ∙ r
+ \ /
+ \ /
+ horizontal-pasting L R' \ / right-whisker (left-whisker s p) r
+ \ /
+ ∨ ∨
+ (s ∙ t') ∙ r
+```
+
+commutes.
+
+```agda
+module _
+ {l : Level} {A : UU l} {a b c d : A}
+ where
+
+ left-whisker-horizontal-pasting-coherence-triangle-identifications :
+ (left : a = d) (middle : b = d) (right : c = d)
+ (top-left : a = b) (top-right top-right' : b = c) →
+ (L : coherence-triangle-identifications left middle top-left)
+ (R : coherence-triangle-identifications middle right top-right)
+ (R' : coherence-triangle-identifications middle right top-right')
+ (p : top-right = top-right') →
+ coherence-triangle-identifications R' (right-whisker-concat p right) R →
+ coherence-triangle-identifications
+ ( horizontal-pasting-coherence-triangle-identifications
+ ( left)
+ ( middle)
+ ( right)
+ ( top-left)
+ ( top-right')
+ ( L)
+ ( R'))
+ ( right-whisker-concat (left-whisker-concat top-left p) right)
+ ( horizontal-pasting-coherence-triangle-identifications
+ ( left)
+ ( middle)
+ ( right)
+ ( top-left)
+ ( top-right)
+ ( L)
+ ( R))
+ left-whisker-horizontal-pasting-coherence-triangle-identifications
+ ._ ._ refl refl refl .refl refl refl ._ refl refl =
+ refl
+```
+
+### Right whiskering of horizontal pasting of commuting triangles of identifications
+
+Consider two commuting triangles of identifications as in the diagram
+
+```text
+ s t
+ a ----> b ----> c
+ \ | /
+ \ L | R /
+ l \ |m / r
+ \ | /
+ ∨ ∨ ∨
+ d,
+```
+
+and consider furthermore a commuting triangle of identifications
+
+```text
+ s'
+ a ----> b
+ \ |
+ \ L' |
+ l \ |m
+ \ |
+ ∨ ∨
+ d,
+```
+
+where the identifications `m : b = d` and `l : a = d` are the same as in the
+previous diagram. Finally, consider an identification `p : s = s'` and an
+identification `q` witnessing that the triangle
+
+```text
+ L
+ l ------> s ∙ m
+ \ /
+ L' \ / right-whisker p m
+ \ /
+ ∨ ∨
+ s' ∙ r
+```
+
+commutes. Then the triangle
+
+```text
+ horizontal-pasting L R
+ l ----------------> (s ∙ t) ∙ r
+ \ /
+ \ /
+ horizontal-pasting L R' \ / right-whisker (right-whisker p t) r
+ \ /
+ ∨ ∨
+ (s' ∙ t) ∙ r
+```
+
+commutes.
+
+```agda
+module _
+ {l : Level} {A : UU l} {a b c d : A}
+ where
+
+ right-whisker-horizontal-pasting-coherence-triangle-identifications :
+ (left : a = d) (middle : b = d) (right : c = d)
+ (top-left top-left' : a = b) (top-right : b = c) →
+ (L : coherence-triangle-identifications left middle top-left)
+ (L' : coherence-triangle-identifications left middle top-left')
+ (R : coherence-triangle-identifications middle right top-right)
+ (p : top-left = top-left') →
+ coherence-triangle-identifications L' (right-whisker-concat p middle) L →
+ coherence-triangle-identifications
+ ( horizontal-pasting-coherence-triangle-identifications
+ ( left)
+ ( middle)
+ ( right)
+ ( top-left')
+ ( top-right)
+ ( L')
+ ( R))
+ ( right-whisker-concat (right-whisker-concat p top-right) right)
+ ( horizontal-pasting-coherence-triangle-identifications
+ ( left)
+ ( middle)
+ ( right)
+ ( top-left)
+ ( top-right)
+ ( L)
+ ( R))
+ right-whisker-horizontal-pasting-coherence-triangle-identifications
+ refl .refl .refl refl .refl refl refl ._ refl refl refl =
+ refl
+```
+
+### Right pasting of commuting triangles of identifications
+
+Consider a commuting diagram of identifications of the form
+
+```text
+ top
+ a ------------> b
+ \ ⎻⎽ /
+ \ ⎺⎽ mid / top-right
+ \ ⎺⎽ ∨
+ left \ ⎺> c
+ \ /
+ \ / bottom-right
+ ∨ ∨
+ d
+```
+
+Then the outer triangle commutes too. Indeed, an identification
+`left = mid ∙ bottom-right` is given. Then, an identification
+
+```text
+ mid ∙ bottom-right = top ∙ (top-right ∙ bottom-right)
+```
+
+is obtained immediately by right whiskering the upper triangle with the
+identification `bottom-right`. Note that this direct construction of the
+coherence of the outer commuting triangle of identifications avoids any use of
+associativity.
+
+```agda
+module _
+ {l : Level} {A : UU l} {a b c d : A}
+ (left : a = d) (top-right : b = c) (bottom-right : c = d)
+ (middle : a = c) (top : a = b)
+ where
+
+ right-pasting-coherence-triangle-identifications :
+ (s : coherence-triangle-identifications left bottom-right middle) →
+ (t : coherence-triangle-identifications middle top-right top) →
+ coherence-triangle-identifications left (top-right ∙ bottom-right) top
+ right-pasting-coherence-triangle-identifications s t =
+ ( s) ∙
+ ( right-whisker-concat-coherence-triangle-identifications
+ ( middle)
+ ( top-right)
+ ( top)
+ ( bottom-right)
+ ( t))
+```
+
+### Right pasting commuting triangles of identifications is a binary equivalence
+
+```agda
+module _
+ {l : Level} {A : UU l} {a b c d : A}
+ (left : a = d) (top-right : b = c) (bottom-right : c = d)
+ (middle : a = c) (top : a = b)
+ where
+
+ abstract
+ is-equiv-right-pasting-coherence-triangle-identifications :
+ (s : coherence-triangle-identifications left bottom-right middle) →
+ is-equiv
+ ( right-pasting-coherence-triangle-identifications
+ ( left)
+ ( top-right)
+ ( bottom-right)
+ ( middle)
+ ( top)
+ ( s))
+ is-equiv-right-pasting-coherence-triangle-identifications s =
+ is-equiv-comp _ _
+ ( is-equiv-right-whisker-concat-coherence-triangle-identifications
+ ( middle)
+ ( top-right)
+ ( top)
+ ( bottom-right))
+ ( is-equiv-concat s _)
+
+ equiv-right-pasting-coherence-triangle-identifications :
+ (s : coherence-triangle-identifications left bottom-right middle) →
+ coherence-triangle-identifications middle top-right top ≃
+ coherence-triangle-identifications left (top-right ∙ bottom-right) top
+ pr1 (equiv-right-pasting-coherence-triangle-identifications s) =
+ right-pasting-coherence-triangle-identifications
+ ( left)
+ ( top-right)
+ ( bottom-right)
+ ( middle)
+ ( top)
+ ( s)
+ pr2 (equiv-right-pasting-coherence-triangle-identifications s) =
+ is-equiv-right-pasting-coherence-triangle-identifications s
+
+ abstract
+ is-equiv-right-pasting-coherence-triangle-identifications' :
+ (t : coherence-triangle-identifications middle top-right top) →
+ is-equiv
+ ( λ (s : coherence-triangle-identifications left bottom-right middle) →
+ right-pasting-coherence-triangle-identifications
+ ( left)
+ ( top-right)
+ ( bottom-right)
+ ( middle)
+ ( top)
+ ( s)
+ ( t))
+ is-equiv-right-pasting-coherence-triangle-identifications' t =
+ is-equiv-concat' _
+ ( right-whisker-concat-coherence-triangle-identifications
+ ( middle)
+ ( top-right)
+ ( top)
+ ( bottom-right)
+ ( t))
+
+ equiv-right-pasting-coherence-triangle-identifications' :
+ (t : coherence-triangle-identifications middle top-right top) →
+ coherence-triangle-identifications left bottom-right middle ≃
+ coherence-triangle-identifications left (top-right ∙ bottom-right) top
+ pr1 (equiv-right-pasting-coherence-triangle-identifications' t) s =
+ right-pasting-coherence-triangle-identifications
+ ( left)
+ ( top-right)
+ ( bottom-right)
+ ( middle)
+ ( top)
+ ( s)
+ ( t)
+ pr2 (equiv-right-pasting-coherence-triangle-identifications' t) =
+ is-equiv-right-pasting-coherence-triangle-identifications' t
+
+ is-binary-equiv-right-pasting-coherence-triangle-identifications :
+ is-binary-equiv
+ ( right-pasting-coherence-triangle-identifications
+ ( left)
+ ( top-right)
+ ( bottom-right)
+ ( middle)
+ ( top))
+ pr1 is-binary-equiv-right-pasting-coherence-triangle-identifications =
+ is-equiv-right-pasting-coherence-triangle-identifications'
+ pr2 is-binary-equiv-right-pasting-coherence-triangle-identifications =
+ is-equiv-right-pasting-coherence-triangle-identifications
+```
+
+### Left pasting of commuting triangles of identifications
+
+**Note.** For left pasting there are two potential constructions. One takes a
+commuting diagram of identifications of the form
+
+```text
+ top
+ a ------------> b
+ \ ⎽⎻ /
+ top-left \ m ⎽⎺ /
+ ∨ ⎽⎺ /
+ c <⎺ / right
+ \ /
+ bottom-left \ /
+ ∨ ∨
+ d
+```
+
+and returns an identification witnessing that the outer triangle commutes. In
+this case the top triangle is an ordinary commuting triangle of identifications,
+and the bottom triangle is inverted along the top edge `m`.
+
+The other left pasting of commuting triangles of identifications takes a
+commuting diagram of identifications of the form
+
+```text
+ top
+ a ------------> b
+ \ ⎽-> /
+ top-left \ m ⎽⎺ /
+ ∨ ⎽⎺ /
+ c ⎺ / right
+ \ /
+ bottom-left \ /
+ ∨ ∨
+ d
+```
+
+and returns an identification witnessing that the outer rectangle commutes. In
+this case the bottom triangle of identifications is an ordinary commuting
+triangle of identifications, and the top triangle is inverted along the right
+edge `m`.
+
+Both constructions have yet to be formalized.
+
+### Vertically pasting commuting squares and commuting triangles of identifications
+
+Consider a diagram of the form
+
+```text
+ top
+ a ------------> b
+ \ /
+ top-left \ / top-right
+ ∨ mid ∨
+ c ----> d
+ \ /
+ bottom-left \ / bottom-right
+ ∨ ∨
+ e
+```
+
+with `s : top-left ∙ mid = top ∙ top-right` witnessing that the square
+commutes, and with `t : bottom-left = mid ∙ bottom-right` witnessing that the
+triangle commutes. Then the outer triangle commutes.
+
+```agda
+module _
+ {l : Level} {A : UU l}
+ where
+
+ vertical-pasting-coherence-square-coherence-triangle-identifications :
+ {a b c d e : A}
+ (top : a = b) (top-left : a = c) (top-right : b = d) (mid : c = d)
+ (bottom-left : c = e) (bottom-right : d = e) →
+ coherence-square-identifications top top-left top-right mid →
+ coherence-triangle-identifications bottom-left bottom-right mid →
+ coherence-triangle-identifications
+ ( top-left ∙ bottom-left)
+ ( top-right ∙ bottom-right)
+ ( top)
+ vertical-pasting-coherence-square-coherence-triangle-identifications
+ top top-left top-right mid bottom-left bottom-right s t =
+ ( left-whisker-concat top-left t) ∙
+ ( right-whisker-concat-coherence-square-identifications
+ ( top)
+ ( top-left)
+ ( top-right)
+ ( mid)
+ ( s)
+ ( bottom-right))
+```
+
+### Vertical pasting of horizontally constant commuting squares of identifications and commuting triangles of identifications
+
+```agda
+module _
+ {l : Level} {A : UU l}
+ where
+
+ vertical-pasting-coherence-horizontally-constant-square-coherence-triangle-identifications :
+ {a c e : A} (p : a = c)
+ (bottom-left : c = e) (bottom-right : c = e) →
+ (t : coherence-triangle-identifications bottom-left bottom-right refl) →
+ ( vertical-pasting-coherence-square-coherence-triangle-identifications
+ ( refl)
+ ( p)
+ ( p)
+ ( refl)
+ ( bottom-left)
+ ( bottom-right)
+ ( horizontal-refl-coherence-square-identifications p)
+ ( t)) =
+ ( left-whisker-concat p t)
+ vertical-pasting-coherence-horizontally-constant-square-coherence-triangle-identifications
+ refl refl .refl refl =
+ refl
+```
+
+### Vertical pasting of verticaly constant commuting squares of identifications and commuting triangles of identifications
+
+```agda
+module _
+ {l : Level} {A : UU l}
+ where
+
+ vertical-pasting-coherence-vertically-constant-square-coherence-triangle-identifications :
+ {a b c : A} (left : a = c) (right : b = c) (top : a = b) →
+ (t : coherence-triangle-identifications left right top) →
+ ( vertical-pasting-coherence-square-coherence-triangle-identifications
+ ( top)
+ ( refl)
+ ( refl)
+ ( top)
+ ( left)
+ ( right)
+ ( vertical-refl-coherence-square-identifications top)
+ ( t)) =
+ t
+ vertical-pasting-coherence-vertically-constant-square-coherence-triangle-identifications
+ ._ refl refl refl = refl
+```
diff --git a/src/foundation/constant-maps.lagda.md b/src/foundation/constant-maps.lagda.md
index b02dab9244..55c8368ea0 100644
--- a/src/foundation/constant-maps.lagda.md
+++ b/src/foundation/constant-maps.lagda.md
@@ -193,3 +193,7 @@ htpy-ap-diagonal-htpy-eq-diagonal-Id :
htpy-ap-diagonal-htpy-eq-diagonal-Id A x y =
inv-htpy (htpy-diagonal-Id-ap-diagonal-htpy-eq A x y)
```
+
+## See also
+
+- [Constant pointed maps](structured-types.constant-pointed-maps.md)
diff --git a/src/foundation/higher-homotopies-morphisms-arrows.lagda.md b/src/foundation/higher-homotopies-morphisms-arrows.lagda.md
index ef383f8263..842b83be66 100644
--- a/src/foundation/higher-homotopies-morphisms-arrows.lagda.md
+++ b/src/foundation/higher-homotopies-morphisms-arrows.lagda.md
@@ -36,7 +36,7 @@ Consider two [morphisms of arrows](foundation.morphisms-arrows.md)
`α := (i , j , H)` and `α' := (i' , j' , H')` and two
[homotopies of morphisms of arrows](foundation.homotopies-morphisms-arrows.md)
`β := (I , J , K)` and `β' : (I' , J' , K')` between them. A
-{{#concept "2-homotopy of morphisms of arrows" Agda=htpy-htpy-hom-arrow}} is a
+{{#concept "`2`-homotopy of morphisms of arrows" Agda=htpy-htpy-hom-arrow}} is a
triple `(γ₀, γ₁ , γ₂)` consisting of [homotopies](foundation-core.homotopies.md)
```text
diff --git a/src/foundation/identity-types.lagda.md b/src/foundation/identity-types.lagda.md
index 471f7d66b7..80997d9a37 100644
--- a/src/foundation/identity-types.lagda.md
+++ b/src/foundation/identity-types.lagda.md
@@ -82,11 +82,25 @@ module _
( is-section-inv-concat p)
( is-retraction-inv-concat p)
+ abstract
+ is-equiv-inv-concat :
+ {x y : A} (p : x = y) (z : A) → is-equiv (inv-concat p z)
+ is-equiv-inv-concat p z =
+ is-equiv-is-invertible
+ ( concat p z)
+ ( is-retraction-inv-concat p)
+ ( is-section-inv-concat p)
+
equiv-concat :
{x y : A} (p : x = y) (z : A) → (y = z) ≃ (x = z)
pr1 (equiv-concat p z) = concat p z
pr2 (equiv-concat p z) = is-equiv-concat p z
+ equiv-inv-concat :
+ {x y : A} (p : x = y) (z : A) → (x = z) ≃ (y = z)
+ pr1 (equiv-inv-concat p z) = inv-concat p z
+ pr2 (equiv-inv-concat p z) = is-equiv-inv-concat p z
+
map-equiv-concat-equiv :
{x x' : A} → ((y : A) → (x = y) ≃ (x' = y)) → (x' = x)
map-equiv-concat-equiv {x} e = map-equiv (e x) refl
@@ -124,11 +138,25 @@ module _
( is-section-inv-concat' q)
( is-retraction-inv-concat' q)
+ abstract
+ is-equiv-inv-concat' :
+ (x : A) {y z : A} (q : y = z) → is-equiv (inv-concat' x q)
+ is-equiv-inv-concat' x q =
+ is-equiv-is-invertible
+ ( concat' x q)
+ ( is-retraction-inv-concat' q)
+ ( is-section-inv-concat' q)
+
equiv-concat' :
(x : A) {y z : A} (q : y = z) → (x = y) ≃ (x = z)
pr1 (equiv-concat' x q) = concat' x q
pr2 (equiv-concat' x q) = is-equiv-concat' x q
+ equiv-inv-concat' :
+ (x : A) {y z : A} (q : y = z) → (x = z) ≃ (x = y)
+ pr1 (equiv-inv-concat' x q) = inv-concat' x q
+ pr2 (equiv-inv-concat' x q) = is-equiv-inv-concat' x q
+
is-binary-equiv-concat :
{l : Level} {A : UU l} {x y z : A} →
is-binary-equiv (λ (p : x = y) (q : y = z) → p ∙ q)
diff --git a/src/foundation/morphisms-twisted-arrows.lagda.md b/src/foundation/morphisms-twisted-arrows.lagda.md
index 82a89e820f..0a7a0dd0fb 100644
--- a/src/foundation/morphisms-twisted-arrows.lagda.md
+++ b/src/foundation/morphisms-twisted-arrows.lagda.md
@@ -49,9 +49,13 @@ module _
(f : A → B) (g : X → Y)
where
+ coherence-hom-twisted-arrow :
+ (X → A) → (B → Y) → UU (l3 ⊔ l4)
+ coherence-hom-twisted-arrow i j = j ∘ f ∘ i ~ g
+
hom-twisted-arrow : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4)
hom-twisted-arrow =
- Σ (X → A) (λ i → Σ (B → Y) (λ j → j ∘ f ∘ i ~ g))
+ Σ (X → A) (λ i → Σ (B → Y) (coherence-hom-twisted-arrow i))
module _
(α : hom-twisted-arrow)
diff --git a/src/foundation/precomposition-functions.lagda.md b/src/foundation/precomposition-functions.lagda.md
index 3f21781779..eba63757e6 100644
--- a/src/foundation/precomposition-functions.lagda.md
+++ b/src/foundation/precomposition-functions.lagda.md
@@ -143,7 +143,7 @@ module _
compute-fiber-precomp g = compute-coherence-triangle-fiber-precomp (g ∘ f)
compute-total-fiber-precomp :
- Σ ( B → X) (λ g → fiber (precomp f X) (g ∘ f)) ≃
+ Σ (B → X) (λ g → fiber (precomp f X) (g ∘ f)) ≃
Σ (B → X) (λ u → Σ (B → X) (λ v → u ∘ f ~ v ∘ f))
compute-total-fiber-precomp = equiv-tot compute-fiber-precomp
diff --git a/src/foundation/spans.lagda.md b/src/foundation/spans.lagda.md
index 385e0b9da2..b3756a7b23 100644
--- a/src/foundation/spans.lagda.md
+++ b/src/foundation/spans.lagda.md
@@ -84,3 +84,4 @@ module _
- [Cospans](foundation.cospans.md)
- [Span diagrams](foundation.span-diagrams.md)
- [Spans of families of types](foundation.spans-families-of-types.md)
+- [Spans of pointed types](structured-types.pointed-spans.md)
diff --git a/src/foundation/whiskering-higher-homotopies-composition.lagda.md b/src/foundation/whiskering-higher-homotopies-composition.lagda.md
index 27db8cd6d1..577f7ba83c 100644
--- a/src/foundation/whiskering-higher-homotopies-composition.lagda.md
+++ b/src/foundation/whiskering-higher-homotopies-composition.lagda.md
@@ -27,9 +27,9 @@ family of maps `h : (x : A) → B x → C x`. Then we obtain a map
```
This operation is called the
-{{#concept "left whiskering" Disambiguation="2-homotopies with respect to composition" Agda=left-whisker-comp²}}.
-Alternatively the left whiskering operation of 2-homotopies can be defined using
-the
+{{#concept "left whiskering" Disambiguation="`2`-homotopies with respect to composition" Agda=left-whisker-comp²}}.
+Alternatively the left whiskering operation of `2`-homotopies can be defined
+using the
[action on higher identifications of functions](foundation.action-on-higher-identifications-functions.md)
by
diff --git a/src/higher-group-theory/equivalences-higher-groups.lagda.md b/src/higher-group-theory/equivalences-higher-groups.lagda.md
index a8f1a9f4e2..bc8cbe7ebe 100644
--- a/src/higher-group-theory/equivalences-higher-groups.lagda.md
+++ b/src/higher-group-theory/equivalences-higher-groups.lagda.md
@@ -21,6 +21,7 @@ open import higher-group-theory.higher-groups
open import higher-group-theory.homomorphisms-higher-groups
open import structured-types.pointed-equivalences
+open import structured-types.pointed-isomorphisms
open import structured-types.pointed-types
```
@@ -63,7 +64,7 @@ module _
where
is-iso-∞-Group : hom-∞-Group G H → UU (l1 ⊔ l2)
- is-iso-∞-Group = is-iso-pointed-map
+ is-iso-∞-Group = is-pointed-iso
```
## Properties
diff --git a/src/higher-group-theory/homomorphisms-higher-groups.lagda.md b/src/higher-group-theory/homomorphisms-higher-groups.lagda.md
index e40ff05191..cd94013ef2 100644
--- a/src/higher-group-theory/homomorphisms-higher-groups.lagda.md
+++ b/src/higher-group-theory/homomorphisms-higher-groups.lagda.md
@@ -75,7 +75,7 @@ module _
where
htpy-hom-∞-Group : (g : hom-∞-Group G H) → UU (l1 ⊔ l2)
- htpy-hom-∞-Group = htpy-pointed-map f
+ htpy-hom-∞-Group = pointed-htpy f
extensionality-hom-∞-Group :
(g : hom-∞-Group G H) → (f = g) ≃ htpy-hom-∞-Group g
diff --git a/src/structured-types.lagda.md b/src/structured-types.lagda.md
index 11ac977476..d2e68002b1 100644
--- a/src/structured-types.lagda.md
+++ b/src/structured-types.lagda.md
@@ -7,15 +7,18 @@ module structured-types where
open import structured-types.cartesian-products-types-equipped-with-endomorphisms public
open import structured-types.central-h-spaces public
+open import structured-types.commuting-squares-of-pointed-homotopies public
open import structured-types.commuting-squares-of-pointed-maps public
+open import structured-types.commuting-triangles-of-pointed-maps public
open import structured-types.conjugation-pointed-types public
-open import structured-types.constant-maps-pointed-types public
+open import structured-types.constant-pointed-maps public
open import structured-types.contractible-pointed-types public
open import structured-types.cyclic-types public
open import structured-types.dependent-products-h-spaces public
open import structured-types.dependent-products-pointed-types public
open import structured-types.dependent-products-wild-monoids public
open import structured-types.dependent-types-equipped-with-automorphisms public
+open import structured-types.equivalences-pointed-arrows public
open import structured-types.equivalences-types-equipped-with-automorphisms public
open import structured-types.equivalences-types-equipped-with-endomorphisms public
open import structured-types.faithful-pointed-maps public
@@ -34,28 +37,43 @@ open import structured-types.magmas public
open import structured-types.mere-equivalences-types-equipped-with-endomorphisms public
open import structured-types.morphisms-h-spaces public
open import structured-types.morphisms-magmas public
+open import structured-types.morphisms-pointed-arrows public
+open import structured-types.morphisms-twisted-pointed-arrows public
open import structured-types.morphisms-types-equipped-with-automorphisms public
open import structured-types.morphisms-types-equipped-with-endomorphisms public
open import structured-types.morphisms-wild-monoids public
open import structured-types.noncoherent-h-spaces public
+open import structured-types.opposite-pointed-spans public
+open import structured-types.pointed-2-homotopies public
open import structured-types.pointed-cartesian-product-types public
open import structured-types.pointed-dependent-functions public
open import structured-types.pointed-dependent-pair-types public
open import structured-types.pointed-equivalences public
open import structured-types.pointed-families-of-types public
open import structured-types.pointed-homotopies public
+open import structured-types.pointed-isomorphisms public
open import structured-types.pointed-maps public
+open import structured-types.pointed-retractions public
open import structured-types.pointed-sections public
+open import structured-types.pointed-span-diagrams public
+open import structured-types.pointed-spans public
open import structured-types.pointed-types public
open import structured-types.pointed-types-equipped-with-automorphisms public
open import structured-types.pointed-unit-type public
open import structured-types.pointed-universal-property-contractible-types public
+open import structured-types.postcomposition-pointed-maps public
+open import structured-types.precomposition-pointed-maps public
open import structured-types.sets-equipped-with-automorphisms public
open import structured-types.symmetric-elements-involutive-types public
open import structured-types.symmetric-h-spaces public
+open import structured-types.transposition-pointed-span-diagrams public
open import structured-types.types-equipped-with-automorphisms public
open import structured-types.types-equipped-with-endomorphisms public
+open import structured-types.uniform-pointed-homotopies public
+open import structured-types.universal-property-pointed-equivalences public
open import structured-types.unpointed-maps public
+open import structured-types.whiskering-pointed-2-homotopies-concatenation public
+open import structured-types.whiskering-pointed-homotopies-composition public
open import structured-types.wild-groups public
open import structured-types.wild-loops public
open import structured-types.wild-monoids public
diff --git a/src/structured-types/commuting-squares-of-pointed-homotopies.lagda.md b/src/structured-types/commuting-squares-of-pointed-homotopies.lagda.md
new file mode 100644
index 0000000000..450d973a09
--- /dev/null
+++ b/src/structured-types/commuting-squares-of-pointed-homotopies.lagda.md
@@ -0,0 +1,64 @@
+# Commuting squares of pointed homotopies
+
+```agda
+module structured-types.commuting-squares-of-pointed-homotopies where
+```
+
+Imports
+
+```agda
+open import foundation.universe-levels
+
+open import structured-types.pointed-2-homotopies
+open import structured-types.pointed-dependent-functions
+open import structured-types.pointed-families-of-types
+open import structured-types.pointed-homotopies
+open import structured-types.pointed-types
+```
+
+
+
+## Idea
+
+A square of [pointed homotopies](structured-types.pointed-homotopies.md)
+
+```text
+ top
+ f ------> g
+ | |
+ left | | right
+ v v
+ h ------> i
+ bottom
+```
+
+is said to be a
+{{#concept "commuting square" Disambiguation="pointed homotopies" Agda=coherence-square-pointed-homotopies}}
+of pointed homotopies if there is a pointed homotopy
+`left ∙h bottom ~∗ top ∙h right `. Such a pointed homotopy is called a
+{{#concept "coherence" Disambiguation="commuting square of homotopies" Agda=coherence-square-pointed-homotopies}}
+of the square.
+
+## Definitions
+
+### Commuting squares of pointed homotopies
+
+```agda
+module _
+ {l1 l2 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Fam l2 A} {f g h i : pointed-Π A B}
+ (top : f ~∗ g) (left : f ~∗ h) (right : g ~∗ i) (bottom : h ~∗ i)
+ where
+
+ coherence-square-pointed-homotopies : UU (l1 ⊔ l2)
+ coherence-square-pointed-homotopies =
+ pointed-2-htpy
+ ( concat-pointed-htpy left bottom)
+ ( concat-pointed-htpy top right)
+
+ coherence-square-pointed-homotopies' : UU (l1 ⊔ l2)
+ coherence-square-pointed-homotopies' =
+ pointed-2-htpy
+ ( concat-pointed-htpy top right)
+ ( concat-pointed-htpy left bottom)
+```
diff --git a/src/structured-types/commuting-squares-of-pointed-maps.lagda.md b/src/structured-types/commuting-squares-of-pointed-maps.lagda.md
index 1a296ce4ff..ed8779ac44 100644
--- a/src/structured-types/commuting-squares-of-pointed-maps.lagda.md
+++ b/src/structured-types/commuting-squares-of-pointed-maps.lagda.md
@@ -7,33 +7,53 @@ module structured-types.commuting-squares-of-pointed-maps where
Imports
```agda
+open import foundation.action-on-identifications-functions
+open import foundation.commuting-squares-of-identifications
open import foundation.commuting-squares-of-maps
+open import foundation.dependent-pair-types
+open import foundation.function-types
+open import foundation.identity-types
open import foundation.universe-levels
+open import foundation.whiskering-homotopies-composition
open import structured-types.pointed-homotopies
open import structured-types.pointed-maps
open import structured-types.pointed-types
+open import structured-types.whiskering-pointed-homotopies-composition
```
## Idea
-A coherence square of pointed maps
+Consider a square of [pointed maps](structured-types.pointed-maps.md)
```text
- A ------> X
- | |
- | |
- | |
- V V
- B ------> Y
+ top
+ A --------> X
+ | |
+ left | | right
+ ∨ ∨
+ B --------> Y.
+ bottom
```
-is a coherence of the underlying square of (unpointed) maps, plus a coherence
-between the pointedness preservation proofs.
+Such a square is said to be a
+{{#concept "commuting square" Disambiguation="pointed maps" Agda=coherence-square-pointed-maps}}
+of pointed maps if there is a
+[pointed homotopy](structured-types.pointed-homotopies.md)
-## Definition
+```text
+ bottom ∘∗ left ~∗ right ∘∗ top.
+```
+
+Such a homotopy is referred to as the
+{{#concept "coherence" Disambiguation="commuting squares of pointed maps" Agda=coherence-square-pointed-maps}}
+of the commuting square of pointed maps.
+
+## Definitions
+
+### Coherences of commuting squares of pointed maps
```agda
module _
@@ -55,5 +75,285 @@ module _
( map-pointed-map right)
( map-pointed-map bottom)
coherence-square-maps-coherence-square-pointed-maps =
- htpy-pointed-htpy (bottom ∘∗ left) (right ∘∗ top)
+ htpy-pointed-htpy
+```
+
+## Operations
+
+### Left whiskering of coherences of commuting squares of pointed maps
+
+Consider a commuting square of pointed maps
+
+```text
+ top
+ A --------> X
+ | |
+ left | | right
+ ∨ ∨
+ B --------> Y
+ bottom
+```
+
+and consider a pointed map `f : Y →∗ Z`. Then the square
+
+```text
+ top
+ A -------------> X
+ | |
+ left | | f ∘∗ right
+ ∨ ∨
+ B -------------> Z
+ f ∘∗ bottom
+```
+
+also commutes.
+
+```agda
+module _
+ {l1 l2 l3 l4 l5 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4} {Z : Pointed-Type l5}
+ (f : Y →∗ Z)
+ (top : A →∗ X) (left : A →∗ B) (right : X →∗ Y) (bottom : B →∗ Y)
+ (s : coherence-square-pointed-maps top left right bottom)
+ where
+
+ left-whisker-comp-coherence-square-pointed-maps :
+ coherence-square-pointed-maps top left (f ∘∗ right) (f ∘∗ bottom)
+ left-whisker-comp-coherence-square-pointed-maps =
+ concat-pointed-htpy
+ ( associative-comp-pointed-map f bottom left)
+ ( concat-pointed-htpy
+ ( left-whisker-comp-pointed-htpy f _ _ s)
+ ( inv-associative-comp-pointed-map f right top))
+```
+
+### Left whiskering of coherences of commuting squares of pointed maps
+
+Consider a commuting square of pointed maps
+
+```text
+ top
+ A --------> X
+ | |
+ left | | right
+ ∨ ∨
+ B --------> Y
+ bottom
+```
+
+and consider a pointed map `f : Z →∗ A`. Then the square
+
+```text
+ f ∘∗ top
+ A ----------> X
+ | |
+ left ∘∗ f | | right
+ ∨ ∨
+ B ----------> Z
+ bottom
+```
+
+also commutes.
+
+```agda
+module _
+ {l1 l2 l3 l4 l5 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4} {Z : Pointed-Type l5}
+ (top : A →∗ X) (left : A →∗ B) (right : X →∗ Y) (bottom : B →∗ Y)
+ (s : coherence-square-pointed-maps top left right bottom)
+ (f : Z →∗ A)
+ where
+
+ right-whisker-comp-coherence-square-pointed-maps :
+ coherence-square-pointed-maps (top ∘∗ f) (left ∘∗ f) right bottom
+ right-whisker-comp-coherence-square-pointed-maps =
+ concat-pointed-htpy
+ ( inv-associative-comp-pointed-map bottom left f)
+ ( concat-pointed-htpy
+ ( right-whisker-comp-pointed-htpy _ _ s f)
+ ( associative-comp-pointed-map right top f))
+```
+
+### Horizontal pasting of coherences of commuting squares of pointed maps
+
+Consider two commuting squares of pointed maps, as in the diagram
+
+```text
+ top-left top-right
+ A -------------> B --------------> C
+ | | |
+ left | | middle | right
+ ∨ ∨ ∨
+ D -------------> E --------------> F
+ bottom-left bottom-right
+```
+
+with pointed homotopies
+
+```text
+ H : bottom-left ∘∗ left ~∗ middle ∘∗ top
+ K : bottom-right ∘∗ middle ~∗ right ∘∗ top-right.
+```
+
+The
+{{#concept "horizontal pasting" Disambiguation="commuting squares of pointed maps" Agda=horizontal-pasting-coherence-square-pointed-maps}}
+of these coherences of commuting squares of pointed maps is the coherence of the
+commuting square
+
+```text
+ top-right ∘∗ top-left
+ A -----------------------------> C
+ | |
+ left | | right
+ ∨ ∨
+ D -----------------------------> F
+ bottom-right ∘∗ bottom-left
+```
+
+obtained by concatenation of the following three pointed homotopies:
+
+```text
+ (bottom-right ∘∗ bottom-left) ∘∗ left
+ ~∗ (bottom-right ∘∗ middle) ∘∗ top-left
+ ~∗ bottom-right ∘∗ (middle ∘∗ top-left)
+ ~∗ right ∘∗ (top-right ∘∗ top-left).
+```
+
+The first and third homotopy in this concatenation are the whiskerings of
+coherences of
+[commuting triangles of pointed maps](structured-types.commuting-triangles-of-pointed-maps.md).
+
+```agda
+module _
+ {l1 l2 l3 l4 l5 l6 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4}
+ {U : Pointed-Type l5} {V : Pointed-Type l6}
+ (top-left : A →∗ X) (top-right : X →∗ U)
+ (left : A →∗ B) (middle : X →∗ Y) (right : U →∗ V)
+ (bottom-left : B →∗ Y) (bottom-right : Y →∗ V)
+ (left-square :
+ coherence-square-pointed-maps top-left left middle bottom-left)
+ (right-square :
+ coherence-square-pointed-maps top-right middle right bottom-right)
+ where
+
+ horizontal-pasting-coherence-square-pointed-maps :
+ coherence-square-pointed-maps
+ ( top-right ∘∗ top-left)
+ ( left)
+ ( right)
+ ( bottom-right ∘∗ bottom-left)
+ horizontal-pasting-coherence-square-pointed-maps =
+ concat-pointed-htpy
+ ( left-whisker-comp-coherence-square-pointed-maps
+ ( bottom-right)
+ ( top-left)
+ ( left)
+ ( middle)
+ ( bottom-left)
+ ( left-square))
+ ( concat-pointed-htpy
+ ( associative-comp-pointed-map bottom-right middle top-left)
+ ( right-whisker-comp-coherence-square-pointed-maps
+ ( top-right)
+ ( middle)
+ ( right)
+ ( bottom-right)
+ ( right-square)
+ ( top-left)))
+```
+
+### Vertical pasting of coherences of commuting squares of pointed maps
+
+Consider two commuting squares of pointed maps, as in the diagram
+
+```text
+ top
+ A --------> B
+ | |
+ top-left | | top-right
+ ∨ middle ∨
+ C --------> D
+ | |
+ bottom-left | | bottom-right
+ ∨ ∨
+ E --------> F
+ bottom
+```
+
+with pointed homotopies
+
+```text
+ H : middle ∘∗ top-left ~∗ top-right ∘∗ top
+ K : bottom ∘∗ bottom-left ~∗ bottom-right ∘∗ middle.
+```
+
+The
+{{#concept "vertical pasting" Disambiguation="commuting squares of pointed maps" Agda=vertical-pasting-coherence-square-pointed-maps}}
+of these coherences of commuting squares of pointed maps is the coherence of the
+commuting square
+
+```text
+ top
+ A --------> B
+ | |
+ bottom-left ∘∗ top-left | | bottom-right ∘∗ top-right
+ ∨ ∨
+ E --------> F
+ bottom
+```
+
+obtained by concatenation of the following three pointed homotopies:
+
+```text
+ bottom ∘∗ (bottom-left ∘∗ top-left)
+ ~∗ bottom-right ∘∗ (middle ∘∗ top-left)
+ ~∗ (bottom-right ∘∗ middle) ∘∗ top-left
+ ~∗ (bottom-right ∘∗ top-right) ∘∗ top.
+```
+
+The first and third homotopy in this concatenation are the whiskerings of
+coherences of
+[commuting triangles of pointed maps](structured-types.commuting-triangles-of-pointed-maps.md).
+
+```agda
+module _
+ {l1 l2 l3 l4 l5 l6 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {C : Pointed-Type l3} {D : Pointed-Type l4}
+ {E : Pointed-Type l5} {F : Pointed-Type l6}
+ (top : A →∗ B) (top-left : A →∗ C) (top-right : B →∗ D) (middle : C →∗ D)
+ (bottom-left : C →∗ E) (bottom-right : D →∗ F) (bottom : E →∗ F)
+ (top-square : coherence-square-pointed-maps top top-left top-right middle)
+ (bottom-square :
+ coherence-square-pointed-maps middle bottom-left bottom-right bottom)
+ where
+
+ vertical-pasting-coherence-square-pointed-maps :
+ coherence-square-pointed-maps
+ ( top)
+ ( bottom-left ∘∗ top-left)
+ ( bottom-right ∘∗ top-right)
+ ( bottom)
+ vertical-pasting-coherence-square-pointed-maps =
+ concat-pointed-htpy
+ ( right-whisker-comp-coherence-square-pointed-maps
+ ( middle)
+ ( bottom-left)
+ ( bottom-right)
+ ( bottom)
+ ( bottom-square)
+ ( top-left))
+ ( concat-pointed-htpy
+ ( inv-associative-comp-pointed-map bottom-right middle top-left)
+ ( left-whisker-comp-coherence-square-pointed-maps
+ ( bottom-right)
+ ( top)
+ ( top-left)
+ ( top-right)
+ ( middle)
+ ( top-square)))
```
diff --git a/src/structured-types/commuting-triangles-of-pointed-maps.lagda.md b/src/structured-types/commuting-triangles-of-pointed-maps.lagda.md
new file mode 100644
index 0000000000..34a14fabd0
--- /dev/null
+++ b/src/structured-types/commuting-triangles-of-pointed-maps.lagda.md
@@ -0,0 +1,119 @@
+# Commuting triangles of pointed maps
+
+```agda
+module structured-types.commuting-triangles-of-pointed-maps where
+```
+
+Imports
+
+```agda
+open import foundation.universe-levels
+
+open import structured-types.pointed-homotopies
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+open import structured-types.whiskering-pointed-homotopies-composition
+```
+
+
+
+## Idea
+
+Consider a triangle of [pointed maps](structured-types.pointed-maps.md)
+
+```text
+ top
+ A ------> B
+ \ /
+ left \ / right
+ \ /
+ ∨ ∨
+ C
+```
+
+Such a triangle is said to be a
+{{#concept "commuting triangle of pointed maps" Agda=coherence-triangle-pointed-maps}}
+if there is a [pointed homotopy](structured-types.pointed-homotopies.md)
+
+```text
+ left ~∗ right ∘∗ top.
+```
+
+Such a homotopy is referred to as the
+{{#concept "coherence" Disambiguation="commuting triangles of pointed maps" Agda=coherence-triangle-pointed-maps}}
+of the commuting triangle of pointed maps.
+
+## Definitions
+
+### Coherences of commuting triangles of pointed maps
+
+```agda
+module _
+ {l1 l2 l3 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Type l2} {C : Pointed-Type l3}
+ (left : A →∗ C) (right : B →∗ C) (top : A →∗ B)
+ where
+
+ coherence-triangle-pointed-maps : UU (l1 ⊔ l3)
+ coherence-triangle-pointed-maps =
+ left ~∗ right ∘∗ top
+
+ coherence-triangle-pointed-maps' : UU (l1 ⊔ l3)
+ coherence-triangle-pointed-maps' =
+ right ∘∗ top ~∗ left
+```
+
+## Properties
+
+### Left whiskering of coherences of commuting triangles of pointed maps
+
+Consider a commuting triangle of pointed maps
+
+```text
+ top
+ A ------> B
+ \ /
+ left \ / right
+ \ /
+ ∨ ∨
+ C
+```
+
+and consider a pointed map `f : C →∗ X`. The
+{{#concept "left whiskering" Disambiguation="commuting triangles of pointed maps" Agda=left-whisker-coherence-triangle-pointed-maps}}
+is a coherence of the triangle of pointed maps
+
+```text
+ top
+ A ------> B
+ \ /
+ f ∘∗ left \ / f ∘∗ right
+ \ /
+ ∨ ∨
+ X
+```
+
+In other words, left whiskering of coherences of commuting triangles of pointed
+maps is an operation
+
+```text
+ (left ~∗ right ∘∗ top) → (f ∘∗ left ~∗ (f ∘∗ right) ∘∗ top).
+```
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Type l2} {C : Pointed-Type l3}
+ {X : Pointed-Type l4} (f : C →∗ X)
+ (left : A →∗ C) (right : B →∗ C) (top : A →∗ B)
+ where
+
+ left-whisker-coherence-triangle-pointed-maps :
+ coherence-triangle-pointed-maps left right top →
+ coherence-triangle-pointed-maps (f ∘∗ left) (f ∘∗ right) top
+ left-whisker-coherence-triangle-pointed-maps H =
+ concat-pointed-htpy
+ ( left-whisker-comp-pointed-htpy f left (right ∘∗ top) H)
+ ( inv-pointed-htpy
+ ( associative-comp-pointed-map f right top))
+```
diff --git a/src/structured-types/conjugation-pointed-types.lagda.md b/src/structured-types/conjugation-pointed-types.lagda.md
index f7b73c1199..fb95e9ed66 100644
--- a/src/structured-types/conjugation-pointed-types.lagda.md
+++ b/src/structured-types/conjugation-pointed-types.lagda.md
@@ -97,9 +97,7 @@ module _
{u : type-Pointed-Type B} (p : point-Pointed-Type B = u) →
conjugation-Ω p ~∗ action-on-loops-conjugation-Pointed-Type p
compute-action-on-loops-conjugation-Pointed-Type p =
- concat-htpy-pointed-map
- ( conjugation-Ω p) (conjugation-Ω' p)
- ( action-on-loops-conjugation-Pointed-Type p)
+ concat-pointed-htpy
( compute-conjugation-Ω p)
( compute-action-on-loops-conjugation-Pointed-Type' p)
diff --git a/src/structured-types/constant-maps-pointed-types.lagda.md b/src/structured-types/constant-maps-pointed-types.lagda.md
deleted file mode 100644
index f50eab18e4..0000000000
--- a/src/structured-types/constant-maps-pointed-types.lagda.md
+++ /dev/null
@@ -1,37 +0,0 @@
-# Constant maps of pointed types
-
-```agda
-module structured-types.constant-maps-pointed-types where
-```
-
-Imports
-
-```agda
-open import foundation.dependent-pair-types
-open import foundation.identity-types
-open import foundation.universe-levels
-
-open import structured-types.pointed-maps
-open import structured-types.pointed-types
-```
-
-
-
-## Idea
-
-Given a type `X` and a pointed type `A`, the constant map from `X` to `A` maps
-every element of `X` to the base point of `A`.
-
-## Definition
-
-```agda
-const-Pointed-Type :
- {l1 l2 : Level} (X : UU l1) (A : Pointed-Type l2) → X → type-Pointed-Type A
-const-Pointed-Type X A x = point-Pointed-Type A
-
-pointed-const-Pointed-Type :
- {l1 l2 : Level} (A : Pointed-Type l1) (B : Pointed-Type l2) → A →∗ B
-pr1 (pointed-const-Pointed-Type A B) =
- const-Pointed-Type (type-Pointed-Type A) B
-pr2 (pointed-const-Pointed-Type A B) = refl
-```
diff --git a/src/structured-types/constant-pointed-maps.lagda.md b/src/structured-types/constant-pointed-maps.lagda.md
new file mode 100644
index 0000000000..8b46613b0f
--- /dev/null
+++ b/src/structured-types/constant-pointed-maps.lagda.md
@@ -0,0 +1,65 @@
+# Constant pointed maps
+
+```agda
+module structured-types.constant-pointed-maps where
+```
+
+Imports
+
+```agda
+open import foundation.constant-maps
+open import foundation.dependent-pair-types
+open import foundation.identity-types
+open import foundation.universe-levels
+
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+```
+
+
+
+## Idea
+
+Given two [pointed types](structured-types.pointed-types.md) `A` and `B` the
+{{#concept "constant pointed map" Agda=constant-pointed-map}} from `A` to `B` is
+the [pointed map](structured-types.pointed-maps.md)
+`constant-pointed-map : A →∗ B` mapping every element in `A` to the base point
+of `B`.
+
+## Definitions
+
+### Constant pointed maps
+
+```agda
+module _
+ {l1 l2 : Level} (A : Pointed-Type l1) (B : Pointed-Type l2)
+ where
+
+ map-constant-pointed-map : type-Pointed-Type A → type-Pointed-Type B
+ map-constant-pointed-map =
+ const (type-Pointed-Type A) (type-Pointed-Type B) (point-Pointed-Type B)
+
+ preserves-point-constant-pointed-map :
+ map-constant-pointed-map (point-Pointed-Type A) = point-Pointed-Type B
+ preserves-point-constant-pointed-map = refl
+
+ constant-pointed-map : A →∗ B
+ pr1 constant-pointed-map = map-constant-pointed-map
+ pr2 constant-pointed-map = preserves-point-constant-pointed-map
+```
+
+### The pointed type of pointed maps
+
+```agda
+module _
+ {l1 l2 : Level} (A : Pointed-Type l1) (B : Pointed-Type l2)
+ where
+
+ pointed-map-Pointed-Type : Pointed-Type (l1 ⊔ l2)
+ pr1 pointed-map-Pointed-Type = A →∗ B
+ pr2 pointed-map-Pointed-Type = constant-pointed-map A B
+```
+
+## See also
+
+- [Constant maps](foundation.constant-maps.md)
diff --git a/src/structured-types/equivalences-pointed-arrows.lagda.md b/src/structured-types/equivalences-pointed-arrows.lagda.md
new file mode 100644
index 0000000000..febdc24bdb
--- /dev/null
+++ b/src/structured-types/equivalences-pointed-arrows.lagda.md
@@ -0,0 +1,131 @@
+# Equivalences of pointed arrows
+
+```agda
+module structured-types.equivalences-pointed-arrows where
+```
+
+Imports
+
+```agda
+open import foundation.dependent-pair-types
+open import foundation.equivalences
+open import foundation.equivalences-arrows
+open import foundation.function-types
+open import foundation.identity-types
+open import foundation.universe-levels
+
+open import structured-types.commuting-squares-of-pointed-maps
+open import structured-types.pointed-equivalences
+open import structured-types.pointed-homotopies
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+```
+
+
+
+## Idea
+
+An {{#concept "equivalence of pointed arrows" Agda=equiv-pointed-arrow}} from a
+[pointed map](structured-types.pointed-maps.md) `f : A →∗ B` to a pointed map
+`g : X →∗ Y` is a [triple](foundation.dependent-pair-types.md) `(i , j , H)`
+consisting of [pointed equivalences](structured-types.pointed-equivalences.md)
+`i : A ≃∗ X` and `j : B ≃∗ Y` and a
+[pointed homotopy](structured-types.pointed-homotopies.md)
+`H : j ∘∗ f ~∗ g ∘∗ i` witnessing that the square
+
+```text
+ i
+ A -----> X
+ | |
+ f | | g
+ V V
+ B -----> Y
+ j
+```
+
+[commutes](structured-types.commuting-squares-of-pointed-maps.md).
+
+## Definitions
+
+### Equivalences of pointed arrows
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4}
+ (f : A →∗ B) (g : X →∗ Y)
+ where
+
+ coherence-equiv-pointed-arrow :
+ (A ≃∗ X) → (B ≃∗ Y) → UU (l1 ⊔ l4)
+ coherence-equiv-pointed-arrow i j =
+ coherence-square-pointed-maps
+ ( pointed-map-pointed-equiv i)
+ ( f)
+ ( g)
+ ( pointed-map-pointed-equiv j)
+
+ equiv-pointed-arrow : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4)
+ equiv-pointed-arrow =
+ Σ (A ≃∗ X) (λ i → Σ (B ≃∗ Y) (coherence-equiv-pointed-arrow i))
+
+module _
+ {l1 l2 l3 l4 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4}
+ {f : A →∗ B} {g : X →∗ Y} (e : equiv-pointed-arrow f g)
+ where
+
+ pointed-equiv-domain-equiv-pointed-arrow : A ≃∗ X
+ pointed-equiv-domain-equiv-pointed-arrow = pr1 e
+
+ equiv-domain-equiv-pointed-arrow : type-Pointed-Type A ≃ type-Pointed-Type X
+ equiv-domain-equiv-pointed-arrow =
+ equiv-pointed-equiv pointed-equiv-domain-equiv-pointed-arrow
+
+ pointed-map-domain-equiv-pointed-arrow : A →∗ X
+ pointed-map-domain-equiv-pointed-arrow =
+ pointed-map-pointed-equiv pointed-equiv-domain-equiv-pointed-arrow
+
+ map-domain-equiv-pointed-arrow : type-Pointed-Type A → type-Pointed-Type X
+ map-domain-equiv-pointed-arrow =
+ map-pointed-equiv pointed-equiv-domain-equiv-pointed-arrow
+
+ preserves-point-map-domain-equiv-pointed-arrow :
+ map-domain-equiv-pointed-arrow (point-Pointed-Type A) =
+ point-Pointed-Type X
+ preserves-point-map-domain-equiv-pointed-arrow =
+ preserves-point-pointed-equiv pointed-equiv-domain-equiv-pointed-arrow
+
+ pointed-equiv-codomain-equiv-pointed-arrow : B ≃∗ Y
+ pointed-equiv-codomain-equiv-pointed-arrow = pr1 (pr2 e)
+
+ equiv-codomain-equiv-pointed-arrow : type-Pointed-Type B ≃ type-Pointed-Type Y
+ equiv-codomain-equiv-pointed-arrow =
+ equiv-pointed-equiv pointed-equiv-codomain-equiv-pointed-arrow
+
+ map-codomain-equiv-pointed-arrow : type-Pointed-Type B → type-Pointed-Type Y
+ map-codomain-equiv-pointed-arrow =
+ map-pointed-equiv pointed-equiv-codomain-equiv-pointed-arrow
+
+ preserves-point-map-codomain-equiv-pointed-arrow :
+ map-codomain-equiv-pointed-arrow (point-Pointed-Type B) =
+ point-Pointed-Type Y
+ preserves-point-map-codomain-equiv-pointed-arrow =
+ preserves-point-pointed-equiv pointed-equiv-codomain-equiv-pointed-arrow
+
+ coh-equiv-pointed-arrow :
+ coherence-equiv-pointed-arrow
+ ( f)
+ ( g)
+ ( pointed-equiv-domain-equiv-pointed-arrow)
+ ( pointed-equiv-codomain-equiv-pointed-arrow)
+ coh-equiv-pointed-arrow = pr2 (pr2 e)
+
+ htpy-coh-equiv-pointed-arrow :
+ coherence-equiv-arrow
+ ( map-pointed-map f)
+ ( map-pointed-map g)
+ ( equiv-domain-equiv-pointed-arrow)
+ ( equiv-codomain-equiv-pointed-arrow)
+ htpy-coh-equiv-pointed-arrow = htpy-pointed-htpy coh-equiv-pointed-arrow
+```
diff --git a/src/structured-types/h-spaces.lagda.md b/src/structured-types/h-spaces.lagda.md
index 28cd269fbc..ddb89b6c0a 100644
--- a/src/structured-types/h-spaces.lagda.md
+++ b/src/structured-types/h-spaces.lagda.md
@@ -19,11 +19,14 @@ open import foundation.identity-types
open import foundation.type-arithmetic-dependent-pair-types
open import foundation.unital-binary-operations
open import foundation.universe-levels
+open import foundation.whiskering-identifications-concatenation
open import foundation-core.endomorphisms
open import structured-types.magmas
open import structured-types.noncoherent-h-spaces
+open import structured-types.pointed-homotopies
+open import structured-types.pointed-maps
open import structured-types.pointed-sections
open import structured-types.pointed-types
```
@@ -158,34 +161,22 @@ module _
{l : Level} (A : Pointed-Type l)
where
+ ev-endo-Pointed-Type : endo-Pointed-Type (type-Pointed-Type A) →∗ A
+ pr1 ev-endo-Pointed-Type = ev-point-Pointed-Type A
+ pr2 ev-endo-Pointed-Type = refl
+
pointed-section-ev-point-Pointed-Type : UU l
pointed-section-ev-point-Pointed-Type =
- pointed-section-Pointed-Type
- ( endo-Pointed-Type (type-Pointed-Type A))
- ( A)
- ( pair (ev-point-Pointed-Type A) refl)
-
-compute-pointed-section-ev-point-Pointed-Type :
- {l : Level} (A : Pointed-Type l) →
- pointed-section-ev-point-Pointed-Type A ≃ coherent-unital-mul-Pointed-Type A
-compute-pointed-section-ev-point-Pointed-Type (pair A a) =
- ( equiv-tot
- ( λ μ →
- ( equiv-Σ
- ( λ p →
- Σ ( (x : A) → μ x a = x)
- ( λ q → p a = q a))
- ( equiv-funext)
- ( λ H →
- equiv-tot
- ( λ K →
- equiv-inv (K a) (htpy-eq H a) ∘e
- equiv-concat' (K a) right-unit ∘e
- equiv-concat' (K a) right-unit))))) ∘e
- ( associative-Σ
- ( A → (A → A))
- ( λ μ → μ a = id)
- ( λ μp →
- Σ ( (x : A) → pr1 μp x a = x)
- ( λ H → H a = ( ( ap (λ h → h a) (pr2 μp) ∙ refl) ∙ refl))))
+ pointed-section ev-endo-Pointed-Type
+
+ compute-pointed-section-ev-point-Pointed-Type :
+ pointed-section-ev-point-Pointed-Type ≃ coherent-unital-mul-Pointed-Type A
+ compute-pointed-section-ev-point-Pointed-Type =
+ ( equiv-tot
+ ( λ _ →
+ equiv-Σ _
+ ( equiv-funext)
+ ( λ _ →
+ equiv-tot (λ _ → inv-equiv (equiv-right-whisker-concat refl))))) ∘e
+ ( associative-Σ _ _ _)
```
diff --git a/src/structured-types/involutive-type-of-h-space-structures.lagda.md b/src/structured-types/involutive-type-of-h-space-structures.lagda.md
index 1454c26b51..deb276431d 100644
--- a/src/structured-types/involutive-type-of-h-space-structures.lagda.md
+++ b/src/structured-types/involutive-type-of-h-space-structures.lagda.md
@@ -9,6 +9,7 @@ module structured-types.involutive-type-of-h-space-structures where
```agda
open import foundation.action-on-identifications-functions
open import foundation.binary-transport
+open import foundation.constant-maps
open import foundation.contractible-types
open import foundation.dependent-pair-types
open import foundation.equality-dependent-function-types
@@ -23,7 +24,7 @@ open import foundation.symmetric-identity-types
open import foundation.torsorial-type-families
open import foundation.universe-levels
-open import structured-types.constant-maps-pointed-types
+open import structured-types.constant-pointed-maps
open import structured-types.pointed-types
open import univalent-combinatorics.2-element-types
@@ -53,7 +54,9 @@ h-space-Involutive-Type A X =
symmetric-Id
( ( X) ,
( λ x →
- ν (const-Pointed-Type _ A) x refl))))
+ ν ( map-constant-pointed-map (type-2-Element-Type X , x) A)
+ ( x)
+ ( refl)))))
```
## Properties
@@ -78,24 +81,52 @@ module _
Eq-symmetric-Id
( ( X) ,
( λ x →
- ( H (const-Pointed-Type _ A)) ∙
- ( pr1 (pr2 μ') (const-Pointed-Type _ A) x refl)))
+ ( H ( map-constant-pointed-map
+ ( type-2-Element-Type X , x)
+ ( A))) ∙
+ ( pr1
+ ( pr2 μ')
+ ( map-constant-pointed-map (type-2-Element-Type X , x) A)
+ ( x)
+ ( refl))))
( tr-symmetric-Id
( ( X) ,
- ( λ x → pr1 (pr2 μ) (const-Pointed-Type _ A) x refl))
+ ( λ x →
+ pr1
+ ( pr2 μ)
+ ( map-constant-pointed-map (type-2-Element-Type X , x) A)
+ ( x)
+ ( refl)))
( ( X) ,
( λ x →
- ( H (const-Pointed-Type _ A)) ∙
- ( pr1 (pr2 μ') (const-Pointed-Type _ A) x refl)))
+ ( H ( map-constant-pointed-map
+ ( type-2-Element-Type X , x)
+ ( A))) ∙
+ ( pr1
+ ( pr2 μ')
+ ( map-constant-pointed-map (type-2-Element-Type X , x) A)
+ ( x)
+ ( refl))))
( id-equiv)
- ( λ x → K (const-Pointed-Type _ A) x refl)
+ ( λ x →
+ K ( map-constant-pointed-map (type-2-Element-Type X , x) A)
+ ( x)
+ ( refl))
( pr2 (pr2 μ)))
( map-equiv-symmetric-Id
( equiv-concat
- ( H (const-Pointed-Type _ A))
+ ( H ( const
+ ( type-2-Element-Type X)
+ ( type-Pointed-Type A)
+ ( point-Pointed-Type A)))
( point-Pointed-Type A))
( ( X) ,
- ( λ x → pr1 (pr2 μ') (const-Pointed-Type _ A) x refl))
+ ( λ x →
+ pr1
+ ( pr2 μ')
+ ( map-constant-pointed-map (type-2-Element-Type X , x) A)
+ ( x)
+ ( refl)))
( pr2 (pr2 μ')))))
refl-htpy-h-space-Involutive-Type :
@@ -111,8 +142,18 @@ module _
( pr2
( pr2 (refl-htpy-h-space-Involutive-Type (μ , unit-laws-μ , coh-μ)))) x =
( ( compute-pr2-tr-symmetric-Id
- ( X , ( λ x → unit-laws-μ (const-Pointed-Type _ A) x refl))
- ( X , ( λ x → unit-laws-μ (const-Pointed-Type _ A) x refl))
+ ( ( X) ,
+ ( λ x →
+ unit-laws-μ
+ ( map-constant-pointed-map (type-2-Element-Type X , x) A)
+ ( x)
+ ( refl)))
+ ( ( X) ,
+ ( λ x →
+ unit-laws-μ
+ ( map-constant-pointed-map (type-2-Element-Type X , x) A)
+ ( x)
+ ( refl)))
( id-equiv)
( λ y → refl)
( λ y → pr2 coh-μ y)
@@ -133,23 +174,53 @@ module _
( ν , (λ f x p → refl))
( is-contr-equiv
( Σ ( symmetric-Id
- ( X , (λ x → ν (const-Pointed-Type _ A) x refl)))
+ ( ( X) ,
+ ( λ x →
+ ν ( map-constant-pointed-map (type-2-Element-Type X , x) A)
+ ( x)
+ ( refl))))
( Eq-symmetric-Id
- ( X , λ x → ν (const-Pointed-Type _ A) x refl)
+ ( ( X) ,
+ ( λ x →
+ ν ( map-constant-pointed-map (type-2-Element-Type X , x) A)
+ ( x)
+ ( refl)))
( ρ)))
( equiv-tot
( λ α →
equiv-binary-tr
( Eq-symmetric-Id
- ( X , (λ x → ν (const-Pointed-Type _ A) x refl)))
+ ( ( X) ,
+ ( λ x →
+ ν ( map-constant-pointed-map
+ ( type-2-Element-Type X , x)
+ ( A))
+ ( x)
+ ( refl))))
( refl-Eq-unordered-pair-tr-symmetric-Id
- ( X , (λ x → ν (const-Pointed-Type _ A) x refl))
+ ( ( X) ,
+ ( λ x →
+ ν ( map-constant-pointed-map
+ ( type-2-Element-Type X , x)
+ ( A))
+ ( x)
+ ( refl)))
( ρ))
( id-equiv-symmetric-Id
- ( X , (λ x → ν (const-Pointed-Type _ A) x refl))
+ ( ( X) ,
+ ( λ x →
+ ν ( map-constant-pointed-map
+ ( type-2-Element-Type X , x)
+ ( A))
+ ( x)
+ ( refl)))
( α))))
( is-torsorial-Eq-symmetric-Id
- ( X , (λ x → ν (const-Pointed-Type _ A) x refl))
+ ( ( X) ,
+ ( λ x →
+ ν ( map-constant-pointed-map (type-2-Element-Type X , x) A)
+ ( x)
+ ( refl)))
( ρ))))
htpy-eq-h-space-Involutive-Type :
diff --git a/src/structured-types/morphisms-pointed-arrows.lagda.md b/src/structured-types/morphisms-pointed-arrows.lagda.md
new file mode 100644
index 0000000000..d86af9b863
--- /dev/null
+++ b/src/structured-types/morphisms-pointed-arrows.lagda.md
@@ -0,0 +1,551 @@
+# Morphisms of pointed arrows
+
+```agda
+module structured-types.morphisms-pointed-arrows where
+```
+
+Imports
+
+```agda
+open import foundation.action-on-identifications-functions
+open import foundation.commuting-squares-of-homotopies
+open import foundation.commuting-squares-of-identifications
+open import foundation.commuting-triangles-of-identifications
+open import foundation.contractible-types
+open import foundation.dependent-pair-types
+open import foundation.functoriality-dependent-pair-types
+open import foundation.fundamental-theorem-of-identity-types
+open import foundation.homotopy-induction
+open import foundation.morphisms-arrows
+open import foundation.path-algebra
+open import foundation.structure-identity-principle
+open import foundation.universe-levels
+open import foundation.whiskering-homotopies-composition
+
+open import foundation-core.commuting-squares-of-maps
+open import foundation-core.equivalences
+open import foundation-core.function-types
+open import foundation-core.homotopies
+open import foundation-core.identity-types
+open import foundation-core.torsorial-type-families
+
+open import structured-types.commuting-squares-of-pointed-homotopies
+open import structured-types.commuting-squares-of-pointed-maps
+open import structured-types.pointed-2-homotopies
+open import structured-types.pointed-homotopies
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+open import structured-types.whiskering-pointed-2-homotopies-concatenation
+open import structured-types.whiskering-pointed-homotopies-composition
+```
+
+
+
+## Idea
+
+A {{#concept "morphism of pointed arrows" Agda=hom-pointed-arrow}} from a
+[pointed map](structured-types.pointed-maps.md) `f : A →∗ B` to a pointed map
+`g : X →∗ Y` is a [triple](foundation.dependent-pair-types.md) `(i , j , H)`
+consisting of pointed maps `i : A →∗ X` and `j : B →∗ Y` and a
+[pointed homotopy](structured-types.pointed-homotopies.md)
+`H : j ∘∗ f ~∗ g ∘∗ i` witnessing that the square
+
+```text
+ i
+ A -----> X
+ | |
+ f | | g
+ V V
+ B -----> Y
+ j
+```
+
+[commutes](structured-types.commuting-squares-of-pointed-maps.md). Morphisms of
+pointed arrows can be composed horizontally or vertically by pasting of squares.
+
+## Definitions
+
+### Morphisms of pointed arrows
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4}
+ (f : A →∗ B) (g : X →∗ Y)
+ where
+
+ coherence-hom-pointed-arrow :
+ (A →∗ X) → (B →∗ Y) → UU (l1 ⊔ l4)
+ coherence-hom-pointed-arrow i = coherence-square-pointed-maps i f g
+
+ hom-pointed-arrow : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4)
+ hom-pointed-arrow =
+ Σ (A →∗ X) (λ i → Σ (B →∗ Y) (coherence-hom-pointed-arrow i))
+
+ pointed-map-domain-hom-pointed-arrow :
+ hom-pointed-arrow → A →∗ X
+ pointed-map-domain-hom-pointed-arrow = pr1
+
+ map-domain-hom-pointed-arrow :
+ hom-pointed-arrow → type-Pointed-Type A → type-Pointed-Type X
+ map-domain-hom-pointed-arrow h =
+ map-pointed-map (pointed-map-domain-hom-pointed-arrow h)
+
+ preserves-point-map-domain-hom-pointed-arrow :
+ (h : hom-pointed-arrow) →
+ map-domain-hom-pointed-arrow h (point-Pointed-Type A) =
+ point-Pointed-Type X
+ preserves-point-map-domain-hom-pointed-arrow h =
+ preserves-point-pointed-map (pointed-map-domain-hom-pointed-arrow h)
+
+ pointed-map-codomain-hom-pointed-arrow :
+ hom-pointed-arrow → B →∗ Y
+ pointed-map-codomain-hom-pointed-arrow = pr1 ∘ pr2
+
+ map-codomain-hom-pointed-arrow :
+ hom-pointed-arrow → type-Pointed-Type B → type-Pointed-Type Y
+ map-codomain-hom-pointed-arrow h =
+ map-pointed-map (pointed-map-codomain-hom-pointed-arrow h)
+
+ preserves-point-map-codomain-hom-pointed-arrow :
+ (h : hom-pointed-arrow) →
+ map-codomain-hom-pointed-arrow h (point-Pointed-Type B) =
+ point-Pointed-Type Y
+ preserves-point-map-codomain-hom-pointed-arrow h =
+ preserves-point-pointed-map (pointed-map-codomain-hom-pointed-arrow h)
+
+ coh-hom-pointed-arrow :
+ (h : hom-pointed-arrow) →
+ coherence-hom-pointed-arrow
+ ( pointed-map-domain-hom-pointed-arrow h)
+ ( pointed-map-codomain-hom-pointed-arrow h)
+ coh-hom-pointed-arrow = pr2 ∘ pr2
+
+ htpy-coh-hom-pointed-arrow :
+ (h : hom-pointed-arrow) →
+ coherence-hom-arrow
+ ( map-pointed-map f)
+ ( map-pointed-map g)
+ ( map-domain-hom-pointed-arrow h)
+ ( map-codomain-hom-pointed-arrow h)
+ htpy-coh-hom-pointed-arrow h =
+ htpy-pointed-htpy
+ ( coh-hom-pointed-arrow h)
+```
+
+### Transposing morphisms of pointed arrows
+
+The
+{{#concept "transposition" Disambiguation="morphism of pointed arrows" Agda=transpose-hom-pointed-arrow}}
+of a morphism of pointed arrows
+
+```text
+ i
+ A -----> X
+ | |
+ f | | g
+ V V
+ B -----> Y
+ j
+```
+
+is the morphism of pointed arrows
+
+```text
+ f
+ A -----> B
+ | |
+ i | | j
+ V V
+ X -----> Y.
+ g
+```
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4}
+ (f : A →∗ B) (g : X →∗ Y) (α : hom-pointed-arrow f g)
+ where
+
+ pointed-map-domain-transpose-hom-pointed-arrow : A →∗ B
+ pointed-map-domain-transpose-hom-pointed-arrow = f
+
+ pointed-map-codomain-transpose-hom-pointed-arrow : X →∗ Y
+ pointed-map-codomain-transpose-hom-pointed-arrow = g
+
+ coh-transpose-hom-pointed-arrow :
+ coherence-hom-pointed-arrow
+ ( pointed-map-domain-hom-pointed-arrow f g α)
+ ( pointed-map-codomain-hom-pointed-arrow f g α)
+ ( pointed-map-domain-transpose-hom-pointed-arrow)
+ ( pointed-map-codomain-transpose-hom-pointed-arrow)
+ coh-transpose-hom-pointed-arrow =
+ inv-pointed-htpy
+ ( coh-hom-pointed-arrow f g α)
+
+ transpose-hom-pointed-arrow :
+ hom-pointed-arrow
+ ( pointed-map-domain-hom-pointed-arrow f g α)
+ ( pointed-map-codomain-hom-pointed-arrow f g α)
+ pr1 transpose-hom-pointed-arrow =
+ pointed-map-domain-transpose-hom-pointed-arrow
+ pr1 (pr2 transpose-hom-pointed-arrow) =
+ pointed-map-codomain-transpose-hom-pointed-arrow
+ pr2 (pr2 transpose-hom-pointed-arrow) = coh-transpose-hom-pointed-arrow
+```
+
+### The identity morphism of pointed arrows
+
+The identity morphism of pointed arrows is defined as
+
+```text
+ id
+ A -----> A
+ | |
+ f | | f
+ V V
+ B -----> B
+ id
+```
+
+where the pointed homotopy `id ∘∗ f ~∗ f ∘∗ id` is the concatenation of the left
+unit law pointed homotopy and the inverse pointed homotopy of the right unit law
+pointed homotopy.
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} {f : A →∗ B}
+ where
+
+ id-hom-pointed-arrow : hom-pointed-arrow f f
+ pr1 id-hom-pointed-arrow = id-pointed-map
+ pr1 (pr2 id-hom-pointed-arrow) = id-pointed-map
+ pr2 (pr2 id-hom-pointed-arrow) =
+ concat-pointed-htpy
+ ( left-unit-law-comp-pointed-map f)
+ ( inv-pointed-htpy (right-unit-law-comp-pointed-map f))
+```
+
+### Composition of morphisms of pointed arrows
+
+Consider a commuting diagram of the form
+
+```text
+ α₀ β₀
+ A -----> X -----> U
+ | | |
+ f | α g | β | h
+ V V V
+ B -----> Y -----> V.
+ α₁ β₁
+```
+
+Then the outer rectangle commutes by horizontal pasting of commuting squares of
+pointed maps. The
+{{#concept "composition" Disambiguation="morphism of pointed arrows" Agda=comp-hom-pointed-arrow}}
+of `β : g → h` with `α : f → g` is therefore defined to be
+
+```text
+ β₀ ∘ α₀
+ A ----------> U
+ | |
+ f | α □ β | h
+ V V
+ B ----------> V.
+ β₁ ∘ α₁
+```
+
+```agda
+module _
+ {l1 l2 l3 l4 l5 l6 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4}
+ {U : Pointed-Type l5} {V : Pointed-Type l6}
+ (f : A →∗ B) (g : X →∗ Y) (h : U →∗ V) (b : hom-pointed-arrow g h)
+ (a : hom-pointed-arrow f g)
+ where
+
+ pointed-map-domain-comp-hom-pointed-arrow : A →∗ U
+ pointed-map-domain-comp-hom-pointed-arrow =
+ pointed-map-domain-hom-pointed-arrow g h b ∘∗
+ pointed-map-domain-hom-pointed-arrow f g a
+
+ map-domain-comp-hom-pointed-arrow :
+ type-Pointed-Type A → type-Pointed-Type U
+ map-domain-comp-hom-pointed-arrow =
+ map-pointed-map pointed-map-domain-comp-hom-pointed-arrow
+
+ preserves-point-map-domain-comp-hom-pointed-arrow :
+ map-domain-comp-hom-pointed-arrow (point-Pointed-Type A) =
+ point-Pointed-Type U
+ preserves-point-map-domain-comp-hom-pointed-arrow =
+ preserves-point-pointed-map pointed-map-domain-comp-hom-pointed-arrow
+
+ pointed-map-codomain-comp-hom-pointed-arrow : B →∗ V
+ pointed-map-codomain-comp-hom-pointed-arrow =
+ pointed-map-codomain-hom-pointed-arrow g h b ∘∗
+ pointed-map-codomain-hom-pointed-arrow f g a
+
+ map-codomain-comp-hom-pointed-arrow :
+ type-Pointed-Type B → type-Pointed-Type V
+ map-codomain-comp-hom-pointed-arrow =
+ map-pointed-map pointed-map-codomain-comp-hom-pointed-arrow
+
+ preserves-point-map-codomain-comp-hom-pointed-arrow :
+ map-codomain-comp-hom-pointed-arrow (point-Pointed-Type B) =
+ point-Pointed-Type V
+ preserves-point-map-codomain-comp-hom-pointed-arrow =
+ preserves-point-pointed-map pointed-map-codomain-comp-hom-pointed-arrow
+
+ coh-comp-hom-pointed-arrow :
+ coherence-hom-pointed-arrow f h
+ ( pointed-map-domain-comp-hom-pointed-arrow)
+ ( pointed-map-codomain-comp-hom-pointed-arrow)
+ coh-comp-hom-pointed-arrow =
+ horizontal-pasting-coherence-square-pointed-maps
+ ( pointed-map-domain-hom-pointed-arrow f g a)
+ ( pointed-map-domain-hom-pointed-arrow g h b)
+ ( f)
+ ( g)
+ ( h)
+ ( pointed-map-codomain-hom-pointed-arrow f g a)
+ ( pointed-map-codomain-hom-pointed-arrow g h b)
+ ( coh-hom-pointed-arrow f g a)
+ ( coh-hom-pointed-arrow g h b)
+
+ comp-hom-pointed-arrow : hom-pointed-arrow f h
+ pr1 comp-hom-pointed-arrow = pointed-map-domain-comp-hom-pointed-arrow
+ pr1 (pr2 comp-hom-pointed-arrow) = pointed-map-codomain-comp-hom-pointed-arrow
+ pr2 (pr2 comp-hom-pointed-arrow) = coh-comp-hom-pointed-arrow
+```
+
+### Homotopies of morphisms of pointed arrows
+
+A
+{{#concept "homotopy of morphisms of pointed arrows" Agda=htpy-hom-pointed-arrow}}
+from `(i , j , H)` to `(i' , j' , H')` is a triple `(I , J , K)` consisting of
+pointed homotopies `I : i ~∗ i'` and `J : j ~∗ j'` and a pointed `2`-homotopy
+`K` witnessing that the
+[square of pointed homotopies](structured-types.commuting-squares-of-pointed-homotopies.md)
+
+```text
+ J ·r f
+ (j ∘∗ f) --------> (j' ∘∗ f)
+ | |
+ H | | H'
+ ∨ ∨
+ (g ∘∗ i) ---------> (g ∘∗ i')
+ g ·l I
+```
+
+commutes.
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4}
+ (f : A →∗ B) (g : X →∗ Y) (α : hom-pointed-arrow f g)
+ where
+
+ coherence-htpy-hom-pointed-arrow :
+ (β : hom-pointed-arrow f g)
+ (I :
+ pointed-map-domain-hom-pointed-arrow f g α ~∗
+ pointed-map-domain-hom-pointed-arrow f g β)
+ (J :
+ pointed-map-codomain-hom-pointed-arrow f g α ~∗
+ pointed-map-codomain-hom-pointed-arrow f g β) →
+ UU (l1 ⊔ l4)
+ coherence-htpy-hom-pointed-arrow β I J =
+ coherence-square-pointed-homotopies
+ ( right-whisker-comp-pointed-htpy _ _ J f)
+ ( coh-hom-pointed-arrow f g α)
+ ( coh-hom-pointed-arrow f g β)
+ ( left-whisker-comp-pointed-htpy g _ _ I)
+
+ htpy-hom-pointed-arrow :
+ (β : hom-pointed-arrow f g) → UU (l1 ⊔ l2 ⊔ l3 ⊔ l4)
+ htpy-hom-pointed-arrow β =
+ Σ ( pointed-map-domain-hom-pointed-arrow f g α ~∗
+ pointed-map-domain-hom-pointed-arrow f g β)
+ ( λ I →
+ Σ ( pointed-map-codomain-hom-pointed-arrow f g α ~∗
+ pointed-map-codomain-hom-pointed-arrow f g β)
+ ( coherence-htpy-hom-pointed-arrow β I))
+
+ module _
+ (β : hom-pointed-arrow f g) (η : htpy-hom-pointed-arrow β)
+ where
+
+ pointed-htpy-domain-htpy-hom-pointed-arrow :
+ pointed-map-domain-hom-pointed-arrow f g α ~∗
+ pointed-map-domain-hom-pointed-arrow f g β
+ pointed-htpy-domain-htpy-hom-pointed-arrow = pr1 η
+
+ pointed-htpy-codomain-htpy-hom-pointed-arrow :
+ pointed-map-codomain-hom-pointed-arrow f g α ~∗
+ pointed-map-codomain-hom-pointed-arrow f g β
+ pointed-htpy-codomain-htpy-hom-pointed-arrow = pr1 (pr2 η)
+
+ coh-htpy-hom-pointed-arrow :
+ coherence-htpy-hom-pointed-arrow β
+ ( pointed-htpy-domain-htpy-hom-pointed-arrow)
+ ( pointed-htpy-codomain-htpy-hom-pointed-arrow)
+ coh-htpy-hom-pointed-arrow = pr2 (pr2 η)
+```
+
+### The reflexive homotopy of pointed arrows
+
+Consider a morphism of pointed arrows
+
+```text
+ α₀
+ A -----> X
+ | |
+ (f₀ , f₁) | α₂ | (g₀ , g₁)
+ ∨ ∨
+ B -----> Y
+ α₁
+```
+
+from `f : A →∗ B` to `g : X →∗ Y`. The reflexive homotopy of morphisms of arrows
+`r := (r₀ , r₁ , r₂)` on `α := (α₀ , α₁ , α₂)` is given by
+
+```text
+ r₀ := refl-pointed-htpy : α₀ ~∗ α₀
+ r₁ := refl-pointed-htpy : α₁ ~∗ α₁
+```
+
+and a pointed `2`-homotopy `r₂` witnessing that the square of pointed homotopies
+
+```text
+ r₁ ·r f
+ (α₁ ∘ f) --------> (α₁ ∘ f)
+ | |
+ α₂ | | α₂
+ V V
+ (g ∘ α₀) --------> (g ∘ α₀)
+ g ·l r₀
+```
+
+commutes. Note that `r₁ ·r f ~∗ refl-pointed-htpy` and
+`g ·l r₀ ≐ refl-pointed-htpy`. By
+[whiskering of pointed `2`-homotopies](structured-types.whiskering-pointed-2-homotopies-concatenation.md)
+with respect to concatenation it follows that
+
+```text
+ (r₁ ·r f) ∙h α₂ ~∗ refl-pointed-htpy ∙h α₂ ~∗ α₂.
+```
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4}
+ (f : A →∗ B) (g : X →∗ Y) (α : hom-pointed-arrow f g)
+ where
+
+ pointed-htpy-domain-refl-htpy-hom-pointed-arrow :
+ pointed-map-domain-hom-pointed-arrow f g α ~∗
+ pointed-map-domain-hom-pointed-arrow f g α
+ pointed-htpy-domain-refl-htpy-hom-pointed-arrow = refl-pointed-htpy _
+
+ pointed-htpy-codomain-refl-htpy-hom-pointed-arrow :
+ pointed-map-codomain-hom-pointed-arrow f g α ~∗
+ pointed-map-codomain-hom-pointed-arrow f g α
+ pointed-htpy-codomain-refl-htpy-hom-pointed-arrow = refl-pointed-htpy _
+
+ coh-refl-htpy-hom-pointed-arrow :
+ coherence-htpy-hom-pointed-arrow f g α α
+ ( pointed-htpy-domain-refl-htpy-hom-pointed-arrow)
+ ( pointed-htpy-codomain-refl-htpy-hom-pointed-arrow)
+ coh-refl-htpy-hom-pointed-arrow =
+ concat-pointed-2-htpy
+ ( right-unit-law-concat-pointed-htpy (coh-hom-pointed-arrow f g α))
+ ( inv-pointed-2-htpy
+ ( concat-pointed-2-htpy
+ ( right-whisker-concat-pointed-2-htpy
+ ( right-whisker-comp-pointed-htpy
+ ( pointed-map-codomain-hom-pointed-arrow f g α)
+ ( pointed-map-codomain-hom-pointed-arrow f g α)
+ ( pointed-htpy-codomain-refl-htpy-hom-pointed-arrow)
+ ( f))
+ ( refl-pointed-htpy _)
+ ( compute-refl-right-whisker-comp-pointed-htpy
+ ( pointed-map-codomain-hom-pointed-arrow f g α)
+ ( f))
+ ( coh-hom-pointed-arrow f g α))
+ ( left-unit-law-concat-pointed-htpy (coh-hom-pointed-arrow f g α))))
+
+ refl-htpy-hom-pointed-arrow : htpy-hom-pointed-arrow f g α α
+ pr1 refl-htpy-hom-pointed-arrow =
+ pointed-htpy-domain-refl-htpy-hom-pointed-arrow
+ pr1 (pr2 refl-htpy-hom-pointed-arrow) =
+ pointed-htpy-codomain-refl-htpy-hom-pointed-arrow
+ pr2 (pr2 refl-htpy-hom-pointed-arrow) =
+ coh-refl-htpy-hom-pointed-arrow
+```
+
+### Characterization of the identity types of morphisms of pointed arrows
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4}
+ (f : A →∗ B) (g : X →∗ Y) (α : hom-pointed-arrow f g)
+ where
+
+ is-torsorial-htpy-hom-pointed-arrow :
+ is-torsorial (htpy-hom-pointed-arrow f g α)
+ is-torsorial-htpy-hom-pointed-arrow =
+ is-torsorial-Eq-structure
+ ( is-torsorial-pointed-htpy _)
+ ( pointed-map-domain-hom-pointed-arrow f g α , refl-pointed-htpy _)
+ ( is-torsorial-Eq-structure
+ ( is-torsorial-pointed-htpy _)
+ ( pointed-map-codomain-hom-pointed-arrow f g α , refl-pointed-htpy _)
+ ( is-contr-equiv' _
+ ( equiv-tot
+ ( λ H →
+ equiv-concat-pointed-2-htpy'
+ ( inv-pointed-2-htpy
+ ( concat-pointed-2-htpy
+ ( right-whisker-concat-pointed-2-htpy _ _
+ ( compute-refl-right-whisker-comp-pointed-htpy
+ ( pointed-map-codomain-hom-pointed-arrow f g α)
+ ( f))
+ ( H))
+ ( left-unit-law-concat-pointed-htpy H)))))
+ ( is-torsorial-pointed-2-htpy
+ ( concat-pointed-htpy
+ ( coh-hom-pointed-arrow f g α)
+ ( refl-pointed-htpy _)))))
+
+ htpy-eq-hom-pointed-arrow :
+ (β : hom-pointed-arrow f g) → α = β → htpy-hom-pointed-arrow f g α β
+ htpy-eq-hom-pointed-arrow β refl = refl-htpy-hom-pointed-arrow f g α
+
+ is-equiv-htpy-eq-hom-pointed-arrow :
+ (β : hom-pointed-arrow f g) → is-equiv (htpy-eq-hom-pointed-arrow β)
+ is-equiv-htpy-eq-hom-pointed-arrow =
+ fundamental-theorem-id
+ ( is-torsorial-htpy-hom-pointed-arrow)
+ ( htpy-eq-hom-pointed-arrow)
+
+ extensionality-hom-pointed-arrow :
+ (β : hom-pointed-arrow f g) → (α = β) ≃ htpy-hom-pointed-arrow f g α β
+ pr1 (extensionality-hom-pointed-arrow β) =
+ htpy-eq-hom-pointed-arrow β
+ pr2 (extensionality-hom-pointed-arrow β) =
+ is-equiv-htpy-eq-hom-pointed-arrow β
+
+ eq-htpy-hom-pointed-arrow :
+ (β : hom-pointed-arrow f g) → htpy-hom-pointed-arrow f g α β → α = β
+ eq-htpy-hom-pointed-arrow β =
+ map-inv-equiv (extensionality-hom-pointed-arrow β)
+```
+
+## See also
+
+- [Equivalences of pointed arrows](structured-types.equivalences-pointed-arrows.md)
+- [Morphisms of twisted pointed arrows](structured-types.morphisms-twisted-pointed-arrows.md)
diff --git a/src/structured-types/morphisms-twisted-pointed-arrows.lagda.md b/src/structured-types/morphisms-twisted-pointed-arrows.lagda.md
new file mode 100644
index 0000000000..16616f1516
--- /dev/null
+++ b/src/structured-types/morphisms-twisted-pointed-arrows.lagda.md
@@ -0,0 +1,113 @@
+# Morphisms of twisted pointed arrows
+
+```agda
+module structured-types.morphisms-twisted-pointed-arrows where
+```
+
+Imports
+
+```agda
+open import foundation.dependent-pair-types
+open import foundation.function-types
+open import foundation.identity-types
+open import foundation.morphisms-twisted-arrows
+open import foundation.universe-levels
+
+open import structured-types.pointed-homotopies
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+```
+
+
+
+A
+{{#concept "morphism of twisted pointed arrows" Agda=hom-twisted-pointed-arrow}}
+from a [pointed map](structured-types.pointed-maps.md) `f : A →∗ B` to a pointed
+map `g : X →∗ Y` is a [triple](foundation.dependent-pair-types.md) `(i , j , H)`
+consisting of pointed maps `i : X →∗ A` and `j : B →∗ Y` and a
+[pointed homotopy](structured-types.pointed-homotopies.md)
+`H : j ∘∗ f ~∗ g ∘∗ i` witnessing that the diagram
+
+```text
+ i
+ A <----- X
+ | |
+ f | | g
+ V V
+ B -----> Y
+ j
+```
+
+commutes.
+
+## Definitions
+
+### Morphisms of twisted pointed arrows
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4}
+ (f : A →∗ B) (g : X →∗ Y)
+ where
+
+ coherence-hom-twisted-pointed-arrow :
+ (X →∗ A) → (B →∗ Y) → UU (l3 ⊔ l4)
+ coherence-hom-twisted-pointed-arrow i j = ((j ∘∗ f) ∘∗ i) ~∗ g
+
+ hom-twisted-pointed-arrow : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4)
+ hom-twisted-pointed-arrow =
+ Σ (X →∗ A) (λ i → Σ (B →∗ Y) (coherence-hom-twisted-pointed-arrow i))
+
+module _
+ {l1 l2 l3 l4 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4}
+ {f : A →∗ B} {g : X →∗ Y} (h : hom-twisted-pointed-arrow f g)
+ where
+
+ pointed-map-domain-hom-twisted-pointed-arrow : X →∗ A
+ pointed-map-domain-hom-twisted-pointed-arrow = pr1 h
+
+ map-domain-hom-twisted-pointed-arrow :
+ type-Pointed-Type X → type-Pointed-Type A
+ map-domain-hom-twisted-pointed-arrow =
+ map-pointed-map pointed-map-domain-hom-twisted-pointed-arrow
+
+ preserves-point-map-domain-hom-twisted-pointed-arrow :
+ map-domain-hom-twisted-pointed-arrow (point-Pointed-Type X) =
+ point-Pointed-Type A
+ preserves-point-map-domain-hom-twisted-pointed-arrow =
+ preserves-point-pointed-map pointed-map-domain-hom-twisted-pointed-arrow
+
+ pointed-map-codomain-hom-twisted-pointed-arrow : B →∗ Y
+ pointed-map-codomain-hom-twisted-pointed-arrow = pr1 (pr2 h)
+
+ map-codomain-hom-twisted-pointed-arrow :
+ type-Pointed-Type B → type-Pointed-Type Y
+ map-codomain-hom-twisted-pointed-arrow =
+ map-pointed-map pointed-map-codomain-hom-twisted-pointed-arrow
+
+ preserves-point-map-codomain-hom-twisted-pointed-arrow :
+ map-codomain-hom-twisted-pointed-arrow (point-Pointed-Type B) =
+ point-Pointed-Type Y
+ preserves-point-map-codomain-hom-twisted-pointed-arrow =
+ preserves-point-pointed-map
+ ( pointed-map-codomain-hom-twisted-pointed-arrow)
+
+ coh-hom-twisted-pointed-arrow :
+ coherence-hom-twisted-pointed-arrow
+ ( f)
+ ( g)
+ ( pointed-map-domain-hom-twisted-pointed-arrow)
+ ( pointed-map-codomain-hom-twisted-pointed-arrow)
+ coh-hom-twisted-pointed-arrow = pr2 (pr2 h)
+
+ htpy-coh-hom-twisted-pointed-arrow :
+ coherence-hom-twisted-arrow
+ ( map-pointed-map f)
+ ( map-pointed-map g)
+ ( map-domain-hom-twisted-pointed-arrow)
+ ( map-codomain-hom-twisted-pointed-arrow)
+ htpy-coh-hom-twisted-pointed-arrow =
+ htpy-pointed-htpy coh-hom-twisted-pointed-arrow
+```
diff --git a/src/structured-types/opposite-pointed-spans.lagda.md b/src/structured-types/opposite-pointed-spans.lagda.md
new file mode 100644
index 0000000000..2c29bc256f
--- /dev/null
+++ b/src/structured-types/opposite-pointed-spans.lagda.md
@@ -0,0 +1,60 @@
+# Opposite pointed spans
+
+```agda
+module structured-types.opposite-pointed-spans where
+```
+
+Imports
+
+```agda
+open import foundation.dependent-pair-types
+open import foundation.universe-levels
+
+open import structured-types.pointed-spans
+open import structured-types.pointed-types
+```
+
+
+
+## Idea
+
+Consider a [pointed span](structured-types.pointed-spans.md) `𝒮 := (S , f , g)`
+from `A` to `B`. The
+{{#concept "opposite pointed span" Agda=opposite-pointed-span}} of
+`𝒮 := (S , f , g)` is the pointed span `(S , g , f)` from `B` to `A`. In other
+words, the opposite of a pointed span
+
+```text
+ f g
+ A <----- S -----> B
+```
+
+is the pointed span
+
+```text
+ g f
+ B <----- S -----> A.
+```
+
+## Definitions
+
+### The opposite of a pointed span
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ where
+
+ opposite-pointed-span :
+ pointed-span l3 A B → pointed-span l3 B A
+ pr1 (opposite-pointed-span s) =
+ spanning-pointed-type-pointed-span s
+ pr1 (pr2 (opposite-pointed-span s)) =
+ right-pointed-map-pointed-span s
+ pr2 (pr2 (opposite-pointed-span s)) =
+ left-pointed-map-pointed-span s
+```
+
+## See also
+
+- [Transpositions of span diagrams](foundation.transposition-span-diagrams.md)
diff --git a/src/structured-types/pointed-2-homotopies.lagda.md b/src/structured-types/pointed-2-homotopies.lagda.md
new file mode 100644
index 0000000000..7b40b04004
--- /dev/null
+++ b/src/structured-types/pointed-2-homotopies.lagda.md
@@ -0,0 +1,592 @@
+# Pointed `2`-homotopies
+
+```agda
+module structured-types.pointed-2-homotopies where
+```
+
+Imports
+
+```agda
+open import foundation.action-on-identifications-functions
+open import foundation.binary-equivalences
+open import foundation.commuting-triangles-of-identifications
+open import foundation.contractible-types
+open import foundation.dependent-pair-types
+open import foundation.equivalences
+open import foundation.functoriality-dependent-pair-types
+open import foundation.fundamental-theorem-of-identity-types
+open import foundation.homotopies
+open import foundation.homotopy-induction
+open import foundation.identity-types
+open import foundation.path-algebra
+open import foundation.structure-identity-principle
+open import foundation.torsorial-type-families
+open import foundation.universe-levels
+open import foundation.whiskering-identifications-concatenation
+
+open import structured-types.pointed-dependent-functions
+open import structured-types.pointed-families-of-types
+open import structured-types.pointed-homotopies
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+open import structured-types.uniform-pointed-homotopies
+```
+
+
+
+## Idea
+
+Consider two [pointed homotopies](structured-types.pointed-homotopies.md)
+`H := (H₀ , H₁)` and `K := (K₀ , K₁)` between two
+[pointed dependent functions](structured-types.pointed-dependent-functions.md)
+`f := (f₀ , f₁)` and `g := (g₀ , g₁)` with base point coherences
+
+```text
+ H₀ * H₀ *
+ f₀ * ------> g₀ * f₀ * ------> g₀ *
+ \ / \ ∧
+ f₁ \ H₁ / g₁ and f₁ \ H̃₁ / inv g₁
+ \ / \ /
+ ∨ ∨ ∨ /
+ * *
+```
+
+and
+
+```text
+ K₀ * K₀ *
+ f₀ * ------> g₀ * f₀ * ------> g₀ *
+ \ / \ ∧
+ f₁ \ K₁ / g₁ and f₁ \ K̃₁ / inv g₁
+ \ / \ /
+ ∨ ∨ ∨ /
+ * *,
+```
+
+where
+
+```text
+ H̃₁ := coherence-triangle-inv-right f₁ g₁ (H₀ *) H₁
+ K̃₁ := coherence-triangle-inv-right f₁ g₁ (K₀ *) K₁
+```
+
+A {{#concept "pointed `2`-homotopy" Agda=_~²∗_}} `H ~²∗ K` then consists of an
+unpointed [homotopy](foundation-core.homotopies.md) `α₀ : H₀ ~ K₀` and an
+[identification](foundation-core.identity-types.md) witnessing that the triangle
+
+```text
+ H₁
+ f₁ ------> (H₀ *) ∙ g₁
+ \ /
+ K₁ \ / right-whisker-concat (α₀ *) g₁
+ \ /
+ ∨ ∨
+ (K₀ *) ∙ g₁
+```
+
+[commutes](foundation.commuting-triangles-of-identifications.md). Equivalently,
+following the [equivalence](foundation-core.equivalences.md) of pointed
+homotopies and
+[uniform pointed homotopies](structured-types.uniform-pointed-homotopies.md), a
+uniform pointed `2`-homotopy consists of an unpointed homotopy `α₀ : H₀ ~ K₀`
+and an identification witnessing that `α₀` preserves the base point, i.e.,
+witnessing that the triangle
+
+```text
+ α₀ *
+ H₀ * ------> K₀ *
+ \ ∧
+ H̃₁ \ / inv K̃₁
+ \ /
+ ∨ /
+ f₁ ∙ inv g₁
+```
+
+commutes. Note that such identifications are often much harder to construct. Our
+preferred definition of pointed `2`-homotopies is therefore the non-uniform
+definition described first.
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B} (H K : pointed-htpy f g)
+ where
+
+ unpointed-htpy-pointed-htpy : UU (l1 ⊔ l2)
+ unpointed-htpy-pointed-htpy =
+ htpy-pointed-htpy H ~ htpy-pointed-htpy K
+
+ coherence-point-unpointed-htpy-pointed-htpy :
+ unpointed-htpy-pointed-htpy → UU l2
+ coherence-point-unpointed-htpy-pointed-htpy α =
+ coherence-triangle-identifications
+ ( coherence-point-pointed-htpy K)
+ ( right-whisker-concat
+ ( α (point-Pointed-Type A))
+ ( preserves-point-function-pointed-Π g))
+ ( coherence-point-pointed-htpy H)
+
+ pointed-2-htpy : UU (l1 ⊔ l2)
+ pointed-2-htpy =
+ Σ unpointed-htpy-pointed-htpy coherence-point-unpointed-htpy-pointed-htpy
+
+ infix 6 _~²∗_
+
+ _~²∗_ : UU (l1 ⊔ l2)
+ _~²∗_ = pointed-2-htpy
+
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B} {H K : pointed-htpy f g}
+ (α : pointed-2-htpy H K)
+ where
+
+ htpy-pointed-2-htpy : unpointed-htpy-pointed-htpy H K
+ htpy-pointed-2-htpy = pr1 α
+
+ coherence-point-pointed-2-htpy :
+ coherence-point-unpointed-htpy-pointed-htpy H K htpy-pointed-2-htpy
+ coherence-point-pointed-2-htpy = pr2 α
+
+ preserves-point-pointed-2-htpy :
+ preserves-point-unpointed-htpy-pointed-Π
+ ( make-uniform-pointed-htpy
+ ( htpy-pointed-htpy H)
+ ( coherence-point-pointed-htpy H))
+ ( make-uniform-pointed-htpy
+ ( htpy-pointed-htpy K)
+ ( coherence-point-pointed-htpy K))
+ ( htpy-pointed-2-htpy)
+ preserves-point-pointed-2-htpy =
+ transpose-right-coherence-triangle-identifications
+ ( htpy-pointed-2-htpy (point-Pointed-Type A))
+ ( preserves-point-pointed-htpy K)
+ ( preserves-point-pointed-htpy H)
+ ( refl)
+ ( higher-transpose-right-coherence-triangle-identifications
+ ( htpy-pointed-htpy H (point-Pointed-Type A))
+ ( preserves-point-function-pointed-Π g)
+ ( preserves-point-function-pointed-Π f)
+ ( htpy-pointed-2-htpy (point-Pointed-Type A))
+ ( refl)
+ ( coherence-point-pointed-htpy H)
+ ( coherence-point-pointed-htpy K)
+ ( coherence-point-pointed-2-htpy))
+
+ uniform-pointed-htpy-pointed-2-htpy :
+ uniform-pointed-htpy
+ ( uniform-pointed-htpy-pointed-htpy H)
+ ( uniform-pointed-htpy-pointed-htpy K)
+ pr1 uniform-pointed-htpy-pointed-2-htpy =
+ htpy-pointed-2-htpy
+ pr2 uniform-pointed-htpy-pointed-2-htpy =
+ preserves-point-pointed-2-htpy
+```
+
+### The reflexive pointed `2`-homotopy
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B} (H : f ~∗ g)
+ where
+
+ htpy-refl-pointed-2-htpy : unpointed-htpy-pointed-htpy H H
+ htpy-refl-pointed-2-htpy = refl-htpy
+
+ coherence-point-refl-pointed-2-htpy :
+ coherence-point-unpointed-htpy-pointed-htpy H H
+ ( htpy-refl-pointed-2-htpy)
+ coherence-point-refl-pointed-2-htpy = inv right-unit
+
+ refl-pointed-2-htpy : H ~²∗ H
+ pr1 refl-pointed-2-htpy = htpy-refl-pointed-2-htpy
+ pr2 refl-pointed-2-htpy = coherence-point-refl-pointed-2-htpy
+```
+
+### Concatenation of pointed `2`-homotopies
+
+Consider two pointed dependent functions `f := (f₀ , f₁)` and `g := (g₀ , g₁)`
+and three pointed homotopies `H := (H₀ , H₁)`, `K := (K₀ , K₁)`, and
+`L := (L₀ , L₁)` between them. Furthermore, consider two pointed `2`-homotopies
+`α := (α₀ , α₁) : H ~²∗ K` and `β := (β₀ , β₁) : K ~²∗ L`. The underlying
+homotopy of the concatenation `α ∙h β` is simply the concatenation of homotopies
+
+```text
+ (α ∙h β)₀ := α₀ ∙h β₀.
+```
+
+The base point coherence `(α ∙h β)₁` is an identification witnessing that the
+triangle
+
+```text
+ H₁
+ f₁ ------> (H₀ *) ∙ h₁
+ \ /
+ L₁ \ / right-whisker-concat ((α₀ *) ∙h (β₀ *)) g₁
+ \ /
+ ∨ ∨
+ (L₀ *) ∙ h₁
+```
+
+commutes. Note that right whiskering of identifications with respect to
+concatenation distributes over concatenation. The identification witnessing the
+commutativity of the above triangle can therefore be constructed by constructing
+an identification witnessing that the triangle
+
+```text
+ H₁
+ f₁ ----------> (H₀ *) ∙ h₁
+ \ /
+ \ / right-whisker-concat (α₀ *) g₁
+ \ ∨
+ L₁ \ (K₀ *) ∙ h₁
+ \ /
+ \ / right-whisker-concat (β₀ *) g₁
+ ∨ ∨
+ (L₀ *) ∙ h₁
+```
+
+commutes. This triangle commutes by right pasting of commuting triangles of
+identifications.
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B} {H K L : f ~∗ g} (α : H ~²∗ K) (β : K ~²∗ L)
+ where
+
+ htpy-concat-pointed-2-htpy :
+ unpointed-htpy-pointed-htpy H L
+ htpy-concat-pointed-2-htpy =
+ htpy-pointed-2-htpy α ∙h htpy-pointed-2-htpy β
+
+ coherence-point-concat-pointed-2-htpy :
+ coherence-point-unpointed-htpy-pointed-htpy H L htpy-concat-pointed-2-htpy
+ coherence-point-concat-pointed-2-htpy =
+ ( right-pasting-coherence-triangle-identifications _ _ _ _
+ ( coherence-point-pointed-htpy H)
+ ( coherence-point-pointed-2-htpy β)
+ ( coherence-point-pointed-2-htpy α)) ∙
+ ( inv
+ ( left-whisker-concat
+ ( coherence-point-pointed-htpy H)
+ ( distributive-right-whisker-concat-concat
+ ( htpy-pointed-2-htpy α _)
+ ( htpy-pointed-2-htpy β _)
+ ( preserves-point-function-pointed-Π g))))
+
+ concat-pointed-2-htpy : H ~²∗ L
+ pr1 concat-pointed-2-htpy = htpy-concat-pointed-2-htpy
+ pr2 concat-pointed-2-htpy = coherence-point-concat-pointed-2-htpy
+```
+
+### Inverses of pointed `2`-homotopies
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B} {H K : f ~∗ g} (α : H ~²∗ K)
+ where
+
+ htpy-inv-pointed-2-htpy :
+ unpointed-htpy-pointed-htpy K H
+ htpy-inv-pointed-2-htpy = inv-htpy (htpy-pointed-2-htpy α)
+
+ coherence-point-inv-pointed-2-htpy :
+ coherence-point-unpointed-htpy-pointed-htpy K H htpy-inv-pointed-2-htpy
+ coherence-point-inv-pointed-2-htpy =
+ transpose-right-coherence-triangle-identifications
+ ( coherence-point-pointed-htpy H)
+ ( right-whisker-concat (htpy-pointed-2-htpy α _) _)
+ ( coherence-point-pointed-htpy K)
+ ( inv (ap-inv _ (htpy-pointed-2-htpy α _)))
+ ( coherence-point-pointed-2-htpy α)
+
+ inv-pointed-2-htpy : K ~²∗ H
+ pr1 inv-pointed-2-htpy = htpy-inv-pointed-2-htpy
+ pr2 inv-pointed-2-htpy = coherence-point-inv-pointed-2-htpy
+```
+
+## Properties
+
+### Extensionality of pointed homotopies
+
+Pointed `2`-homotopies characterize identifications of pointed homotopies.
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B} (H : f ~∗ g)
+ where
+
+ is-torsorial-pointed-2-htpy :
+ is-torsorial (pointed-2-htpy H)
+ is-torsorial-pointed-2-htpy =
+ is-torsorial-Eq-structure
+ ( is-torsorial-htpy _)
+ ( htpy-pointed-htpy H , refl-htpy)
+ ( is-torsorial-Id' _)
+
+ pointed-2-htpy-eq : (K : f ~∗ g) → H = K → H ~²∗ K
+ pointed-2-htpy-eq .H refl = refl-pointed-2-htpy H
+
+ is-equiv-pointed-2-htpy-eq :
+ (K : f ~∗ g) → is-equiv (pointed-2-htpy-eq K)
+ is-equiv-pointed-2-htpy-eq =
+ fundamental-theorem-id
+ ( is-torsorial-pointed-2-htpy)
+ ( pointed-2-htpy-eq)
+
+ extensionality-pointed-htpy :
+ (K : f ~∗ g) → (H = K) ≃ (H ~²∗ K)
+ pr1 (extensionality-pointed-htpy K) = pointed-2-htpy-eq K
+ pr2 (extensionality-pointed-htpy K) = is-equiv-pointed-2-htpy-eq K
+
+ eq-pointed-2-htpy :
+ (K : f ~∗ g) → H ~²∗ K → H = K
+ eq-pointed-2-htpy K = map-inv-equiv (extensionality-pointed-htpy K)
+```
+
+### Concatenation of pointed `2`-homotopies is a binary equivalence
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B} {H K L : f ~∗ g}
+ where
+
+ abstract
+ is-equiv-concat-pointed-2-htpy :
+ (α : H ~²∗ K) → is-equiv (λ (β : K ~²∗ L) → concat-pointed-2-htpy α β)
+ is-equiv-concat-pointed-2-htpy α =
+ is-equiv-map-Σ _
+ ( is-equiv-concat-htpy (htpy-pointed-2-htpy α) _)
+ ( λ β →
+ is-equiv-comp _ _
+ ( is-equiv-right-pasting-coherence-triangle-identifications'
+ ( coherence-point-pointed-htpy L)
+ ( right-whisker-concat
+ ( htpy-pointed-2-htpy α _)
+ ( preserves-point-function-pointed-Π g))
+ ( right-whisker-concat
+ ( β _)
+ ( preserves-point-function-pointed-Π g))
+ ( coherence-point-pointed-htpy K)
+ ( coherence-point-pointed-htpy H)
+ ( coherence-point-pointed-2-htpy α))
+ ( is-equiv-concat' _
+ ( inv
+ ( left-whisker-concat
+ ( coherence-point-pointed-htpy H)
+ ( distributive-right-whisker-concat-concat
+ ( htpy-pointed-2-htpy α _)
+ ( β _)
+ ( preserves-point-function-pointed-Π g))))))
+
+ equiv-concat-pointed-2-htpy : H ~²∗ K → (K ~²∗ L) ≃ (H ~²∗ L)
+ pr1 (equiv-concat-pointed-2-htpy α) = concat-pointed-2-htpy α
+ pr2 (equiv-concat-pointed-2-htpy α) = is-equiv-concat-pointed-2-htpy α
+
+ abstract
+ is-equiv-concat-pointed-2-htpy' :
+ (β : K ~²∗ L) → is-equiv (λ (α : H ~²∗ K) → concat-pointed-2-htpy α β)
+ is-equiv-concat-pointed-2-htpy' β =
+ is-equiv-map-Σ _
+ ( is-equiv-concat-htpy' _ (htpy-pointed-2-htpy β))
+ ( λ α →
+ is-equiv-comp _ _
+ ( is-equiv-right-pasting-coherence-triangle-identifications
+ ( coherence-point-pointed-htpy L)
+ ( right-whisker-concat
+ ( α _)
+ ( preserves-point-function-pointed-Π g))
+ ( right-whisker-concat
+ ( htpy-pointed-2-htpy β _)
+ ( preserves-point-function-pointed-Π g))
+ ( coherence-point-pointed-htpy K)
+ ( coherence-point-pointed-htpy H)
+ ( coherence-point-pointed-2-htpy β))
+ ( is-equiv-concat' _
+ ( inv
+ ( left-whisker-concat
+ ( coherence-point-pointed-htpy H)
+ ( distributive-right-whisker-concat-concat
+ ( α _)
+ ( htpy-pointed-2-htpy β _)
+ ( preserves-point-function-pointed-Π g))))))
+
+ equiv-concat-pointed-2-htpy' :
+ K ~²∗ L → (H ~²∗ K) ≃ (H ~²∗ L)
+ pr1 (equiv-concat-pointed-2-htpy' β) α = concat-pointed-2-htpy α β
+ pr2 (equiv-concat-pointed-2-htpy' β) = is-equiv-concat-pointed-2-htpy' β
+
+ is-binary-equiv-concat-pointed-2-htpy :
+ is-binary-equiv (λ (α : H ~²∗ K) (β : K ~²∗ L) → concat-pointed-2-htpy α β)
+ pr1 is-binary-equiv-concat-pointed-2-htpy = is-equiv-concat-pointed-2-htpy'
+ pr2 is-binary-equiv-concat-pointed-2-htpy = is-equiv-concat-pointed-2-htpy
+```
+
+### Associativity of concatenation of pointed homotopies
+
+Associativity of concatenation of three pointed homotopies `G`, `H`, and `K` is
+a pointed `2`-homotopy
+
+```text
+ (G ∙h H) ∙h K ~²∗ G ∙h (H ∙h K).
+```
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g h k : pointed-Π A B} (G : f ~∗ g) (H : g ~∗ h) (K : h ~∗ k)
+ where
+
+ htpy-associative-concat-pointed-htpy :
+ htpy-pointed-htpy
+ ( concat-pointed-htpy (concat-pointed-htpy G H) K) ~
+ htpy-pointed-htpy
+ ( concat-pointed-htpy G (concat-pointed-htpy H K))
+ htpy-associative-concat-pointed-htpy =
+ assoc-htpy
+ ( htpy-pointed-htpy G)
+ ( htpy-pointed-htpy H)
+ ( htpy-pointed-htpy K)
+
+ coherence-associative-concat-pointed-htpy :
+ coherence-point-unpointed-htpy-pointed-htpy
+ ( concat-pointed-htpy (concat-pointed-htpy G H) K)
+ ( concat-pointed-htpy G (concat-pointed-htpy H K))
+ ( htpy-associative-concat-pointed-htpy)
+ coherence-associative-concat-pointed-htpy =
+ associative-horizontal-pasting-coherence-triangle-identifications
+ ( preserves-point-function-pointed-Π f)
+ ( preserves-point-function-pointed-Π g)
+ ( preserves-point-function-pointed-Π h)
+ ( preserves-point-function-pointed-Π k)
+ ( htpy-pointed-htpy G _)
+ ( htpy-pointed-htpy H _)
+ ( htpy-pointed-htpy K _)
+ ( coherence-point-pointed-htpy G)
+ ( coherence-point-pointed-htpy H)
+ ( coherence-point-pointed-htpy K)
+
+ associative-concat-pointed-htpy :
+ concat-pointed-htpy (concat-pointed-htpy G H) K ~²∗
+ concat-pointed-htpy G (concat-pointed-htpy H K)
+ pr1 associative-concat-pointed-htpy =
+ htpy-associative-concat-pointed-htpy
+ pr2 associative-concat-pointed-htpy =
+ coherence-associative-concat-pointed-htpy
+```
+
+### The left unit law of concatenation of pointed homotopies
+
+Consider a pointed homotopy `H := (H₀ , H₁)` between pointed dependent functions
+`f := (f₀ , f₁)` and `g := (g₀ , g₁)`. Then there is a pointed `2`-homotopy of
+type
+
+```text
+ refl-pointed-htpy ∙h H ~²∗ H.
+```
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B} (H : f ~∗ g)
+ where
+
+ htpy-left-unit-law-concat-pointed-htpy :
+ unpointed-htpy-pointed-htpy
+ ( concat-pointed-htpy (refl-pointed-htpy f) H)
+ ( H)
+ htpy-left-unit-law-concat-pointed-htpy = refl-htpy
+
+ coherence-point-left-unit-law-concat-pointed-htpy :
+ coherence-point-unpointed-htpy-pointed-htpy
+ ( concat-pointed-htpy (refl-pointed-htpy f) H)
+ ( H)
+ ( htpy-left-unit-law-concat-pointed-htpy)
+ coherence-point-left-unit-law-concat-pointed-htpy =
+ inv (right-unit ∙ right-unit ∙ ap-id (coherence-point-pointed-htpy H))
+
+ left-unit-law-concat-pointed-htpy :
+ concat-pointed-htpy (refl-pointed-htpy f) H ~²∗ H
+ pr1 left-unit-law-concat-pointed-htpy =
+ htpy-left-unit-law-concat-pointed-htpy
+ pr2 left-unit-law-concat-pointed-htpy =
+ coherence-point-left-unit-law-concat-pointed-htpy
+```
+
+### The right unit law of concatenation of pointed homotopies
+
+Consider a pointed homotopy `H := (H₀ , H₁)` between pointed dependent functions
+`f := (f₀ , f₁)` and `g := (g₀ , g₁)`. Then there is a pointed `2`-homotopy
+
+```text
+ H ∙h refl-pointed-htpy ~²∗ H.
+```
+
+The underlying homotopy of this pointed `2`-homotopy is the homotopy
+`right-unit-htpy`. The base point coherence of this homotopy is an
+identification witnessing that the triangle
+
+```text
+ (H ∙h refl)₁
+ f₁ ------------> (H₀ * ∙ refl) ∙ g₁
+ \ /
+ H₁ \ / right-whisker-concat right-unit g₁
+ \ /
+ ∨ ∨
+ (H₀ *) ∙ g₁
+```
+
+commutes. Here, the identification `(H ∙h refl)₁` is the horizontal pasting of
+commuting triangles of identifications
+
+```text
+ H₀ * refl
+ f₀ * --> g₀ * ---> g₀ *
+ \ | /
+ \ | g₁ /
+ f₁ \ | / g₁
+ \ | /
+ \ | /
+ ∨ ∨ ∨
+ *.
+```
+
+The upper triangle therefore commutes by the right unit law of horizontal
+pasting of commuting triangles of identifications.
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B} (H : f ~∗ g)
+ where
+
+ htpy-right-unit-law-concat-pointed-htpy :
+ unpointed-htpy-pointed-htpy
+ ( concat-pointed-htpy H (refl-pointed-htpy g))
+ ( H)
+ htpy-right-unit-law-concat-pointed-htpy = right-unit-htpy
+
+ coherence-point-right-unit-law-concat-pointed-htpy :
+ coherence-point-unpointed-htpy-pointed-htpy
+ ( concat-pointed-htpy H (refl-pointed-htpy g))
+ ( H)
+ ( htpy-right-unit-law-concat-pointed-htpy)
+ coherence-point-right-unit-law-concat-pointed-htpy =
+ right-unit-law-horizontal-pasting-coherence-triangle-identifications
+ ( preserves-point-function-pointed-Π f)
+ ( preserves-point-function-pointed-Π g)
+ ( htpy-pointed-htpy H _)
+ ( coherence-point-pointed-htpy H)
+
+ right-unit-law-concat-pointed-htpy :
+ concat-pointed-htpy H (refl-pointed-htpy g) ~²∗ H
+ pr1 right-unit-law-concat-pointed-htpy =
+ htpy-right-unit-law-concat-pointed-htpy
+ pr2 right-unit-law-concat-pointed-htpy =
+ coherence-point-right-unit-law-concat-pointed-htpy
+```
diff --git a/src/structured-types/pointed-equivalences.lagda.md b/src/structured-types/pointed-equivalences.lagda.md
index f45422806f..0ec64334e1 100644
--- a/src/structured-types/pointed-equivalences.lagda.md
+++ b/src/structured-types/pointed-equivalences.lagda.md
@@ -9,6 +9,7 @@ module structured-types.pointed-equivalences where
```agda
open import foundation.action-on-identifications-functions
open import foundation.cartesian-product-types
+open import foundation.commuting-squares-of-identifications
open import foundation.contractible-maps
open import foundation.contractible-types
open import foundation.dependent-pair-types
@@ -19,45 +20,171 @@ open import foundation.function-types
open import foundation.functoriality-dependent-pair-types
open import foundation.homotopies
open import foundation.identity-types
+open import foundation.injective-maps
+open import foundation.path-algebra
open import foundation.propositions
+open import foundation.retractions
+open import foundation.sections
open import foundation.structure-identity-principle
open import foundation.torsorial-type-families
+open import foundation.transposition-identifications-along-equivalences
open import foundation.type-arithmetic-dependent-pair-types
open import foundation.univalence
open import foundation.universe-levels
+open import foundation.whiskering-identifications-concatenation
open import structured-types.pointed-homotopies
open import structured-types.pointed-maps
+open import structured-types.pointed-retractions
+open import structured-types.pointed-sections
open import structured-types.pointed-types
+open import structured-types.postcomposition-pointed-maps
+open import structured-types.precomposition-pointed-maps
+open import structured-types.universal-property-pointed-equivalences
+open import structured-types.whiskering-pointed-homotopies-composition
```
## Idea
-A **pointed equivalence** is an equivalence in the category of pointed spaces.
-Equivalently, a pointed equivalence is a
+A {{#concept "pointed equivalence" Agda=_≃∗_}} `e : A ≃∗ B` consists of an
+[equivalence](foundation-core.equivalences.md) `e : A ≃ B` equipped with an
+[identification](foundation-core.identity-types.md) `p : e * = *` witnessing
+that the underlying map of `e` preserves the base point of `A`.
+
+The notion of pointed equivalence is described equivalently as a
[pointed map](structured-types.pointed-maps.md) of which the underlying function
-is an [equivalence](foundation-core.equivalences.md).
+is an [equivalence](foundation-core.equivalences.md), i.e.,
+
+```text
+ (A ≃∗ B) ≃ Σ (f : A →∗ B), is-equiv (map-pointed-map f)
+```
+
+Furthermore, a pointed equivalence can also be described equivalently as an
+equivalence in the category of
+[pointed types](structured-types.pointed-types.md).
## Definitions
-### Pointed equivalences
+### The predicate of being a pointed equivalence
```agda
module _
- {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
where
- is-equiv-pointed-map : (A →∗ B) → UU (l1 ⊔ l2)
- is-equiv-pointed-map f = is-equiv (map-pointed-map f)
+ is-pointed-equiv : UU (l1 ⊔ l2)
+ is-pointed-equiv = is-equiv (map-pointed-map f)
+
+ is-prop-is-pointed-equiv : is-prop is-pointed-equiv
+ is-prop-is-pointed-equiv = is-property-is-equiv (map-pointed-map f)
- is-prop-is-equiv-pointed-map : (f : A →∗ B) → is-prop (is-equiv-pointed-map f)
- is-prop-is-equiv-pointed-map = is-property-is-equiv ∘ map-pointed-map
+ is-pointed-equiv-Prop : Prop (l1 ⊔ l2)
+ is-pointed-equiv-Prop = is-equiv-Prop (map-pointed-map f)
- is-equiv-pointed-map-Prop : (A →∗ B) → Prop (l1 ⊔ l2)
- is-equiv-pointed-map-Prop = is-equiv-Prop ∘ map-pointed-map
+ module _
+ (H : is-pointed-equiv)
+ where
+ map-inv-is-pointed-equiv : type-Pointed-Type B → type-Pointed-Type A
+ map-inv-is-pointed-equiv = map-inv-is-equiv H
+
+ is-section-map-inv-is-pointed-equiv :
+ is-section (map-pointed-map f) map-inv-is-pointed-equiv
+ is-section-map-inv-is-pointed-equiv = is-section-map-inv-is-equiv H
+
+ is-retraction-map-inv-is-pointed-equiv :
+ is-retraction (map-pointed-map f) map-inv-is-pointed-equiv
+ is-retraction-map-inv-is-pointed-equiv =
+ is-retraction-map-inv-is-equiv H
+
+ preserves-point-map-inv-is-pointed-equiv :
+ map-inv-is-pointed-equiv (point-Pointed-Type B) = point-Pointed-Type A
+ preserves-point-map-inv-is-pointed-equiv =
+ preserves-point-is-retraction-pointed-map f
+ ( map-inv-is-pointed-equiv)
+ ( is-retraction-map-inv-is-pointed-equiv)
+
+ pointed-map-inv-is-pointed-equiv : B →∗ A
+ pointed-map-inv-is-pointed-equiv =
+ pointed-map-is-retraction-pointed-map f
+ ( map-inv-is-pointed-equiv)
+ ( is-retraction-map-inv-is-pointed-equiv)
+
+ coherence-point-is-retraction-map-inv-is-pointed-equiv :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( pointed-map-inv-is-pointed-equiv ∘∗ f)
+ ( id-pointed-map)
+ ( is-retraction-map-inv-is-pointed-equiv)
+ coherence-point-is-retraction-map-inv-is-pointed-equiv =
+ coherence-point-is-retraction-pointed-map f
+ ( map-inv-is-pointed-equiv)
+ ( is-retraction-map-inv-is-pointed-equiv)
+
+ is-pointed-retraction-pointed-map-inv-is-pointed-equiv :
+ is-pointed-retraction f pointed-map-inv-is-pointed-equiv
+ is-pointed-retraction-pointed-map-inv-is-pointed-equiv =
+ is-pointed-retraction-is-retraction-pointed-map f
+ ( map-inv-is-pointed-equiv)
+ ( is-retraction-map-inv-is-pointed-equiv)
+
+ coherence-point-is-section-map-inv-is-pointed-equiv :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( f ∘∗ pointed-map-inv-is-pointed-equiv)
+ ( id-pointed-map)
+ ( is-section-map-inv-is-pointed-equiv)
+ coherence-point-is-section-map-inv-is-pointed-equiv =
+ ( right-whisker-concat
+ ( ap-concat
+ ( map-pointed-map f)
+ ( inv (ap _ (preserves-point-pointed-map f)))
+ ( _) ∙
+ ( horizontal-concat-Id²
+ ( ap-inv
+ ( map-pointed-map f)
+ ( ap _ (preserves-point-pointed-map f)) ∙
+ ( inv
+ ( ap
+ ( inv)
+ ( ap-comp
+ ( map-pointed-map f)
+ ( map-inv-is-pointed-equiv)
+ ( preserves-point-pointed-map f)))))
+ ( inv (coherence-map-inv-is-equiv H (point-Pointed-Type A)))))
+ ( preserves-point-pointed-map f)) ∙
+ ( assoc
+ ( inv
+ ( ap
+ ( map-pointed-map f ∘ map-inv-is-pointed-equiv)
+ ( preserves-point-pointed-map f)))
+ ( is-section-map-inv-is-pointed-equiv _)
+ ( preserves-point-pointed-map f)) ∙
+ ( inv
+ ( ( right-unit) ∙
+ ( left-transpose-eq-concat
+ ( ap
+ ( map-pointed-map f ∘ map-inv-is-pointed-equiv)
+ ( preserves-point-pointed-map f))
+ ( is-section-map-inv-is-pointed-equiv _)
+ ( ( is-section-map-inv-is-pointed-equiv _) ∙
+ ( preserves-point-pointed-map f))
+ ( ( inv (nat-htpy is-section-map-inv-is-pointed-equiv _)) ∙
+ ( left-whisker-concat
+ ( is-section-map-inv-is-pointed-equiv _)
+ ( ap-id (preserves-point-pointed-map f)))))))
+
+ is-pointed-section-pointed-map-inv-is-pointed-equiv :
+ is-pointed-section f pointed-map-inv-is-pointed-equiv
+ pr1 is-pointed-section-pointed-map-inv-is-pointed-equiv =
+ is-section-map-inv-is-pointed-equiv
+ pr2 is-pointed-section-pointed-map-inv-is-pointed-equiv =
+ coherence-point-is-section-map-inv-is-pointed-equiv
+```
+
+### Pointed equivalences
+
+```agda
pointed-equiv :
{l1 l2 : Level} (A : Pointed-Type l1) (B : Pointed-Type l2) → UU (l1 ⊔ l2)
pointed-equiv A B =
@@ -65,17 +192,8 @@ pointed-equiv A B =
( λ e → map-equiv e (point-Pointed-Type A) = point-Pointed-Type B)
infix 6 _≃∗_
-_≃∗_ = pointed-equiv
-
-compute-pointed-equiv :
- {l1 l2 : Level} (A : Pointed-Type l1) (B : Pointed-Type l2) →
- (A ≃∗ B) ≃ Σ (A →∗ B) (is-equiv-pointed-map {A = A} {B})
-compute-pointed-equiv A B = equiv-right-swap-Σ
-inv-compute-pointed-equiv :
- {l1 l2 : Level} (A : Pointed-Type l1) (B : Pointed-Type l2) →
- Σ (A →∗ B) (is-equiv-pointed-map {A = A} {B}) ≃ (A ≃∗ B)
-inv-compute-pointed-equiv A B = equiv-right-swap-Σ
+_≃∗_ = pointed-equiv
module _
{l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (e : A ≃∗ B)
@@ -84,23 +202,112 @@ module _
equiv-pointed-equiv : type-Pointed-Type A ≃ type-Pointed-Type B
equiv-pointed-equiv = pr1 e
- map-equiv-pointed-equiv : type-Pointed-Type A → type-Pointed-Type B
- map-equiv-pointed-equiv = map-equiv equiv-pointed-equiv
-
- is-equiv-map-equiv-pointed-equiv : is-equiv map-equiv-pointed-equiv
- is-equiv-map-equiv-pointed-equiv = is-equiv-map-equiv equiv-pointed-equiv
+ map-pointed-equiv : type-Pointed-Type A → type-Pointed-Type B
+ map-pointed-equiv = map-equiv equiv-pointed-equiv
- preserves-point-equiv-pointed-equiv :
- Id (map-equiv-pointed-equiv (point-Pointed-Type A)) (point-Pointed-Type B)
- preserves-point-equiv-pointed-equiv = pr2 e
+ preserves-point-pointed-equiv :
+ map-pointed-equiv (point-Pointed-Type A) = point-Pointed-Type B
+ preserves-point-pointed-equiv = pr2 e
pointed-map-pointed-equiv : A →∗ B
- pr1 pointed-map-pointed-equiv = map-equiv-pointed-equiv
- pr2 pointed-map-pointed-equiv = preserves-point-equiv-pointed-equiv
+ pr1 pointed-map-pointed-equiv = map-pointed-equiv
+ pr2 pointed-map-pointed-equiv = preserves-point-pointed-equiv
+
+ is-equiv-map-pointed-equiv : is-equiv map-pointed-equiv
+ is-equiv-map-pointed-equiv = is-equiv-map-equiv equiv-pointed-equiv
+
+ is-pointed-equiv-pointed-equiv :
+ is-pointed-equiv pointed-map-pointed-equiv
+ is-pointed-equiv-pointed-equiv = is-equiv-map-pointed-equiv
+
+ pointed-map-inv-pointed-equiv : B →∗ A
+ pointed-map-inv-pointed-equiv =
+ pointed-map-inv-is-pointed-equiv
+ ( pointed-map-pointed-equiv)
+ ( is-pointed-equiv-pointed-equiv)
+
+ map-inv-pointed-equiv : type-Pointed-Type B → type-Pointed-Type A
+ map-inv-pointed-equiv =
+ map-inv-is-pointed-equiv
+ ( pointed-map-pointed-equiv)
+ ( is-pointed-equiv-pointed-equiv)
+
+ preserves-point-map-inv-pointed-equiv :
+ map-inv-pointed-equiv (point-Pointed-Type B) = point-Pointed-Type A
+ preserves-point-map-inv-pointed-equiv =
+ preserves-point-map-inv-is-pointed-equiv
+ ( pointed-map-pointed-equiv)
+ ( is-pointed-equiv-pointed-equiv)
+
+ is-pointed-section-pointed-map-inv-pointed-equiv :
+ is-pointed-section
+ ( pointed-map-pointed-equiv)
+ ( pointed-map-inv-pointed-equiv)
+ is-pointed-section-pointed-map-inv-pointed-equiv =
+ is-pointed-section-pointed-map-inv-is-pointed-equiv
+ ( pointed-map-pointed-equiv)
+ ( is-pointed-equiv-pointed-equiv)
+
+ is-section-map-inv-pointed-equiv :
+ is-section
+ ( map-pointed-equiv)
+ ( map-inv-pointed-equiv)
+ is-section-map-inv-pointed-equiv =
+ is-section-map-inv-is-pointed-equiv
+ ( pointed-map-pointed-equiv)
+ ( is-pointed-equiv-pointed-equiv)
+
+ coherence-point-is-section-map-inv-pointed-equiv :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( pointed-map-pointed-equiv ∘∗ pointed-map-inv-pointed-equiv)
+ ( id-pointed-map)
+ ( is-section-map-inv-pointed-equiv)
+ coherence-point-is-section-map-inv-pointed-equiv =
+ coherence-point-is-section-map-inv-is-pointed-equiv
+ ( pointed-map-pointed-equiv)
+ ( is-pointed-equiv-pointed-equiv)
+
+ is-pointed-retraction-pointed-map-inv-pointed-equiv :
+ is-pointed-retraction
+ ( pointed-map-pointed-equiv)
+ ( pointed-map-inv-pointed-equiv)
+ is-pointed-retraction-pointed-map-inv-pointed-equiv =
+ is-pointed-retraction-pointed-map-inv-is-pointed-equiv
+ ( pointed-map-pointed-equiv)
+ ( is-pointed-equiv-pointed-equiv)
+
+ is-retraction-map-inv-pointed-equiv :
+ is-retraction
+ ( map-pointed-equiv)
+ ( map-inv-pointed-equiv)
+ is-retraction-map-inv-pointed-equiv =
+ is-retraction-map-inv-is-pointed-equiv
+ ( pointed-map-pointed-equiv)
+ ( is-pointed-equiv-pointed-equiv)
+
+ coherence-point-is-retraction-map-inv-pointed-equiv :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( pointed-map-inv-pointed-equiv ∘∗ pointed-map-pointed-equiv)
+ ( id-pointed-map)
+ ( is-retraction-map-inv-pointed-equiv)
+ coherence-point-is-retraction-map-inv-pointed-equiv =
+ coherence-point-is-retraction-map-inv-is-pointed-equiv
+ ( pointed-map-pointed-equiv)
+ ( is-pointed-equiv-pointed-equiv)
+```
+
+### The equivalence between pointed equivalences and the type of pointed maps that are pointed equivalences
+
+```agda
+module _
+ {l1 l2 : Level} (A : Pointed-Type l1) (B : Pointed-Type l2)
+ where
- is-equiv-pointed-map-pointed-equiv :
- is-equiv-pointed-map pointed-map-pointed-equiv
- is-equiv-pointed-map-pointed-equiv = is-equiv-map-equiv-pointed-equiv
+ compute-pointed-equiv : (A ≃∗ B) ≃ Σ (A →∗ B) is-pointed-equiv
+ compute-pointed-equiv = equiv-right-swap-Σ
+
+ inv-compute-pointed-equiv : Σ (A →∗ B) is-pointed-equiv ≃ (A ≃∗ B)
+ inv-compute-pointed-equiv = equiv-right-swap-Σ
```
### The identity pointed equivalence
@@ -132,24 +339,6 @@ module _
( pointed-map-pointed-equiv e)
```
-### Pointed isomorphisms
-
-```agda
-module _
- {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
- where
-
- section-pointed-map : UU (l1 ⊔ l2)
- section-pointed-map = Σ (B →∗ A) (λ g → (f ∘∗ g) ~∗ id-pointed-map)
-
- retraction-pointed-map : UU (l1 ⊔ l2)
- retraction-pointed-map =
- Σ (B →∗ A) (λ g → (g ∘∗ f) ~∗ id-pointed-map)
-
- is-iso-pointed-map : UU (l1 ⊔ l2)
- is-iso-pointed-map = section-pointed-map × retraction-pointed-map
-```
-
## Properties
### Extensionality of the universe of pointed types
@@ -164,233 +353,188 @@ module _
is-torsorial-equiv-Pointed-Type =
is-torsorial-Eq-structure
( is-torsorial-equiv (type-Pointed-Type A))
- ( pair (type-Pointed-Type A) id-equiv)
+ ( type-Pointed-Type A , id-equiv)
( is-torsorial-Id (point-Pointed-Type A))
- extensionality-Pointed-Type : (B : Pointed-Type l1) → Id A B ≃ (A ≃∗ B)
+ extensionality-Pointed-Type : (B : Pointed-Type l1) → (A = B) ≃ (A ≃∗ B)
extensionality-Pointed-Type =
extensionality-Σ
- ( λ b e → Id (map-equiv e (point-Pointed-Type A)) b)
+ ( λ b e → map-equiv e (point-Pointed-Type A) = b)
( id-equiv)
( refl)
( λ B → equiv-univalence)
( λ a → id-equiv)
- eq-pointed-equiv : (B : Pointed-Type l1) → A ≃∗ B → Id A B
+ eq-pointed-equiv : (B : Pointed-Type l1) → A ≃∗ B → A = B
eq-pointed-equiv B = map-inv-equiv (extensionality-Pointed-Type B)
```
-### Being a pointed equivalence is equivalent to being a pointed isomorphism
+### Any pointed map satisfying the universal property of pointed equivalences is a pointed equivalence
+
+In other words, any pointed map `f : A →∗ B` such that precomposing by `f`
+
+```text
+ - ∘∗ f : (B →∗ C) → (A →∗ C)
+```
+
+is an equivalence for any pointed type `C`, is an equivalence.
```agda
module _
{l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ (H : universal-property-pointed-equiv f)
where
- is-contr-section-is-equiv-pointed-map :
- is-equiv-pointed-map f → is-contr (section-pointed-map f)
- is-contr-section-is-equiv-pointed-map H =
- is-torsorial-Eq-structure
- ( is-contr-section-is-equiv H)
- ( pair (map-inv-is-equiv H) (is-section-map-inv-is-equiv H))
- ( is-contr-equiv
- ( fiber
- ( ap (map-pointed-map f))
- ( ( is-section-map-inv-is-equiv H (point-Pointed-Type B)) ∙
- ( inv (preserves-point-pointed-map f))))
- ( equiv-tot
- ( λ p →
- ( equiv-right-transpose-eq-concat
- ( ap (map-pointed-map f) p)
- ( preserves-point-pointed-map f)
- ( is-section-map-inv-is-equiv H (point-Pointed-Type B))) ∘e
- ( equiv-inv
- ( is-section-map-inv-is-equiv H (point-Pointed-Type B))
- ( ( ap (map-pointed-map f) p) ∙
- ( preserves-point-pointed-map f))) ∘e
- ( equiv-concat'
- ( is-section-map-inv-is-equiv H (point-Pointed-Type B))
- ( right-unit))))
- ( is-contr-map-is-equiv
- ( is-emb-is-equiv H
- ( map-inv-is-equiv H (point-Pointed-Type B))
- ( point-Pointed-Type A))
- ( ( is-section-map-inv-is-equiv H (point-Pointed-Type B)) ∙
- ( inv (preserves-point-pointed-map f)))))
-
- is-contr-retraction-is-equiv-pointed-map :
- is-equiv-pointed-map f → is-contr (retraction-pointed-map f)
- is-contr-retraction-is-equiv-pointed-map H =
- is-torsorial-Eq-structure
- ( is-contr-retraction-is-equiv H)
- ( pair (map-inv-is-equiv H) (is-retraction-map-inv-is-equiv H))
- ( is-contr-equiv
- ( fiber
- ( λ p →
- ap (map-inv-is-equiv H) (preserves-point-pointed-map f) ∙ p ∙ refl)
- ( is-retraction-map-inv-is-equiv H (point-Pointed-Type A)))
- ( equiv-tot (λ p → equiv-inv _ _))
- ( is-contr-map-is-equiv
- ( is-equiv-comp
- ( _∙ refl)
- ( ap (map-inv-is-equiv H) (preserves-point-pointed-map f) ∙_)
- ( is-equiv-concat
- ( ap
- ( map-inv-is-equiv H)
- ( preserves-point-pointed-map f))
- ( point-Pointed-Type A))
- ( is-equiv-concat'
- ( map-inv-is-equiv
- ( H)
- ( map-pointed-map f (point-Pointed-Type A)))
- ( refl)))
- ( is-retraction-map-inv-is-equiv H (point-Pointed-Type A))))
-
- is-contr-is-iso-is-equiv-pointed-map :
- is-equiv-pointed-map f → is-contr (is-iso-pointed-map f)
- is-contr-is-iso-is-equiv-pointed-map H =
- is-contr-product
- ( is-contr-section-is-equiv-pointed-map H)
- ( is-contr-retraction-is-equiv-pointed-map H)
-
- is-iso-is-equiv-pointed-map :
- is-equiv-pointed-map f → is-iso-pointed-map f
- is-iso-is-equiv-pointed-map H =
- center (is-contr-is-iso-is-equiv-pointed-map H)
-
- is-equiv-is-iso-pointed-map :
- is-iso-pointed-map f → is-equiv-pointed-map f
- pr1 (pr1 (is-equiv-is-iso-pointed-map H)) = pr1 (pr1 (pr1 H))
- pr2 (pr1 (is-equiv-is-iso-pointed-map H)) = pr1 (pr2 (pr1 H))
- pr1 (pr2 (is-equiv-is-iso-pointed-map H)) = pr1 (pr1 (pr2 H))
- pr2 (pr2 (is-equiv-is-iso-pointed-map H)) = pr1 (pr2 (pr2 H))
-
- is-prop-is-iso-pointed-map : is-prop (is-iso-pointed-map f)
- is-prop-is-iso-pointed-map =
- is-prop-is-proof-irrelevant
- ( λ H →
- is-contr-is-iso-is-equiv-pointed-map (is-equiv-is-iso-pointed-map H))
-
- equiv-is-iso-is-equiv-pointed-map :
- is-equiv-pointed-map f ≃ (is-iso-pointed-map f)
- pr1 equiv-is-iso-is-equiv-pointed-map = is-iso-is-equiv-pointed-map
- pr2 equiv-is-iso-is-equiv-pointed-map =
- is-equiv-is-prop
- ( is-property-is-equiv (map-pointed-map f))
- ( is-prop-is-iso-pointed-map)
- ( is-equiv-is-iso-pointed-map)
+ pointed-map-inv-is-pointed-equiv-universal-property-pointed-equiv : B →∗ A
+ pointed-map-inv-is-pointed-equiv-universal-property-pointed-equiv =
+ map-inv-is-equiv (H A) id-pointed-map
+
+ map-inv-is-pointed-equiv-universal-property-pointed-equiv :
+ type-Pointed-Type B → type-Pointed-Type A
+ map-inv-is-pointed-equiv-universal-property-pointed-equiv =
+ map-pointed-map
+ pointed-map-inv-is-pointed-equiv-universal-property-pointed-equiv
+
+ preserves-point-map-inv-is-pointed-equiv-universal-property-pointed-equiv :
+ map-inv-is-pointed-equiv-universal-property-pointed-equiv
+ ( point-Pointed-Type B) =
+ point-Pointed-Type A
+ preserves-point-map-inv-is-pointed-equiv-universal-property-pointed-equiv =
+ preserves-point-pointed-map
+ pointed-map-inv-is-pointed-equiv-universal-property-pointed-equiv
+
+ is-pointed-retraction-pointed-map-inv-is-pointed-equiv-universal-property-pointed-equiv :
+ is-pointed-retraction f
+ pointed-map-inv-is-pointed-equiv-universal-property-pointed-equiv
+ is-pointed-retraction-pointed-map-inv-is-pointed-equiv-universal-property-pointed-equiv =
+ pointed-htpy-eq _ _ (is-section-map-inv-is-equiv (H A) (id-pointed-map))
+
+ is-retraction-map-inv-is-pointed-equiv-universal-property-pointed-equiv :
+ is-retraction
+ ( map-pointed-map f)
+ ( map-inv-is-pointed-equiv-universal-property-pointed-equiv)
+ is-retraction-map-inv-is-pointed-equiv-universal-property-pointed-equiv =
+ htpy-pointed-htpy
+ ( is-pointed-retraction-pointed-map-inv-is-pointed-equiv-universal-property-pointed-equiv)
+
+ is-pointed-section-pointed-map-inv-is-pointed-equiv-universal-property-pointed-equiv :
+ is-pointed-section f
+ pointed-map-inv-is-pointed-equiv-universal-property-pointed-equiv
+ is-pointed-section-pointed-map-inv-is-pointed-equiv-universal-property-pointed-equiv =
+ pointed-htpy-eq _ _
+ ( is-injective-is-equiv (H B)
+ ( eq-pointed-htpy ((f ∘∗ _) ∘∗ f) (id-pointed-map ∘∗ f)
+ ( concat-pointed-htpy
+ ( associative-comp-pointed-map f _ f)
+ ( concat-pointed-htpy
+ ( left-whisker-comp-pointed-htpy f _ _
+ ( is-pointed-retraction-pointed-map-inv-is-pointed-equiv-universal-property-pointed-equiv))
+ ( concat-pointed-htpy
+ ( right-unit-law-comp-pointed-map f)
+ ( inv-left-unit-law-comp-pointed-map f))))))
+
+ is-section-map-inv-is-pointed-equiv-universal-property-pointed-equiv :
+ is-section
+ ( map-pointed-map f)
+ ( map-inv-is-pointed-equiv-universal-property-pointed-equiv)
+ is-section-map-inv-is-pointed-equiv-universal-property-pointed-equiv =
+ htpy-pointed-htpy
+ ( is-pointed-section-pointed-map-inv-is-pointed-equiv-universal-property-pointed-equiv)
+
+ is-pointed-equiv-universal-property-pointed-equiv : is-pointed-equiv f
+ is-pointed-equiv-universal-property-pointed-equiv =
+ is-equiv-is-invertible
+ ( map-inv-is-pointed-equiv-universal-property-pointed-equiv)
+ ( is-section-map-inv-is-pointed-equiv-universal-property-pointed-equiv)
+ ( is-retraction-map-inv-is-pointed-equiv-universal-property-pointed-equiv)
```
-### Precomposing by pointed equivalences is a pointed equivalence
+### Any pointed equivalence satisfies the universal property of pointed equivalences
```agda
module _
{l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
where
- is-equiv-is-equiv-precomp-pointed-map :
- ( {l : Level} (C : Pointed-Type l) → is-equiv (precomp-pointed-map C f)) →
- is-equiv-pointed-map f
- is-equiv-is-equiv-precomp-pointed-map H =
- is-equiv-is-invertible
- ( map-pointed-map g)
- ( htpy-eq
- ( ap pr1
- ( ap pr1
- ( eq-is-contr
- ( is-contr-map-is-equiv (H B) f)
- { pair
- ( f ∘∗ g)
- ( eq-htpy-pointed-map
- ( ( f ∘∗ g) ∘∗ f)
- ( f)
- ( concat-htpy-pointed-map
- ( ( f ∘∗ g) ∘∗ f)
- ( f ∘∗ (g ∘∗ f))
- ( f)
- ( associative-comp-pointed-map f g f)
- ( concat-htpy-pointed-map
- ( f ∘∗ (g ∘∗ f))
- ( f ∘∗ id-pointed-map)
- ( f)
- ( left-whisker-comp-pointed-map f
- ( g ∘∗ f)
- ( id-pointed-map)
- ( G))
- ( right-unit-law-comp-pointed-map f))))}
- { pair
- ( id-pointed-map)
- ( eq-htpy-pointed-map
- ( id-pointed-map ∘∗ f)
- ( f)
- ( left-unit-law-comp-pointed-map f))}))))
- ( pr1 G)
- where
- g : B →∗ A
- g = pr1 (center (is-contr-map-is-equiv (H A) id-pointed-map))
- G : (g ∘∗ f) ~∗ id-pointed-map
- G = map-equiv
- ( extensionality-pointed-map
- ( g ∘∗ f)
- ( id-pointed-map))
- ( pr2 (center (is-contr-map-is-equiv (H A) id-pointed-map)))
-
- is-equiv-precomp-is-equiv-pointed-map :
- is-equiv-pointed-map f →
- {l : Level} → (C : Pointed-Type l) → is-equiv (precomp-pointed-map C f)
- is-equiv-precomp-is-equiv-pointed-map E C =
- pair
- ( pair
- ( precomp-pointed-map C h)
- ( λ k →
- eq-htpy-pointed-map
- ( (k ∘∗ h) ∘∗ f)
- ( k)
- ( concat-htpy-pointed-map
- ( (k ∘∗ h) ∘∗ f)
- ( k ∘∗ (h ∘∗ f))
- ( k)
- ( associative-comp-pointed-map k h f)
- ( concat-htpy-pointed-map
- ( k ∘∗ (h ∘∗ f))
- ( k ∘∗ id-pointed-map)
- ( k)
- ( left-whisker-comp-pointed-map k
- ( h ∘∗ f)
- ( id-pointed-map)
- ( H))
- ( right-unit-law-comp-pointed-map k)))))
- ( pair
- ( precomp-pointed-map C g)
- ( λ k →
- eq-htpy-pointed-map
- ( (k ∘∗ f) ∘∗ g)
- ( k)
- ( concat-htpy-pointed-map
- ( (k ∘∗ f) ∘∗ g)
- ( k ∘∗ (f ∘∗ g))
- ( k)
- ( associative-comp-pointed-map k f g)
- ( concat-htpy-pointed-map
- ( k ∘∗ (f ∘∗ g))
- ( k ∘∗ id-pointed-map)
- ( k)
- ( left-whisker-comp-pointed-map k
- ( f ∘∗ g)
- ( id-pointed-map)
- ( G))
- ( right-unit-law-comp-pointed-map k)))))
- where
- I : is-iso-pointed-map f
- I = is-iso-is-equiv-pointed-map f E
- g : B →∗ A
- g = pr1 (pr1 I)
- G : (f ∘∗ g) ~∗ id-pointed-map
- G = pr2 (pr1 I)
- h : B →∗ A
- h = pr1 (pr2 I)
- H : (h ∘∗ f) ~∗ id-pointed-map
- H = pr2 (pr2 I)
+ map-inv-universal-property-pointed-equiv-is-pointed-equiv :
+ (H : is-pointed-equiv f) {l : Level} (C : Pointed-Type l) →
+ (A →∗ C) → (B →∗ C)
+ map-inv-universal-property-pointed-equiv-is-pointed-equiv H C =
+ precomp-pointed-map
+ ( pointed-map-inv-is-pointed-equiv f H)
+ ( C)
+
+ is-section-map-inv-universal-property-pointed-equiv-is-pointed-equiv :
+ (H : is-pointed-equiv f) →
+ {l : Level} (C : Pointed-Type l) →
+ is-section
+ ( precomp-pointed-map f C)
+ ( map-inv-universal-property-pointed-equiv-is-pointed-equiv
+ ( H)
+ ( C))
+ is-section-map-inv-universal-property-pointed-equiv-is-pointed-equiv H C h =
+ eq-pointed-htpy
+ ( (h ∘∗ pointed-map-inv-is-pointed-equiv f H) ∘∗ f)
+ ( h)
+ ( concat-pointed-htpy
+ ( associative-comp-pointed-map h
+ ( pointed-map-inv-is-pointed-equiv f H)
+ ( f))
+ ( left-whisker-comp-pointed-htpy h
+ ( pointed-map-inv-is-pointed-equiv f H ∘∗ f)
+ ( id-pointed-map)
+ ( is-pointed-retraction-pointed-map-inv-is-pointed-equiv f H)))
+
+ section-universal-property-pointed-equiv-is-pointed-equiv :
+ (H : is-pointed-equiv f) →
+ {l : Level} (C : Pointed-Type l) →
+ section (precomp-pointed-map f C)
+ pr1 (section-universal-property-pointed-equiv-is-pointed-equiv H C) =
+ map-inv-universal-property-pointed-equiv-is-pointed-equiv H C
+ pr2 (section-universal-property-pointed-equiv-is-pointed-equiv H C) =
+ is-section-map-inv-universal-property-pointed-equiv-is-pointed-equiv
+ ( H)
+ ( C)
+
+ is-retraction-map-inv-universal-property-pointed-equiv-is-pointed-equiv :
+ (H : is-pointed-equiv f) →
+ {l : Level} (C : Pointed-Type l) →
+ is-retraction
+ ( precomp-pointed-map f C)
+ ( map-inv-universal-property-pointed-equiv-is-pointed-equiv
+ ( H)
+ ( C))
+ is-retraction-map-inv-universal-property-pointed-equiv-is-pointed-equiv
+ H C h =
+ eq-pointed-htpy
+ ( (h ∘∗ f) ∘∗ pointed-map-inv-is-pointed-equiv f H)
+ ( h)
+ ( concat-pointed-htpy
+ ( associative-comp-pointed-map h f
+ ( pointed-map-inv-is-pointed-equiv f H))
+ ( left-whisker-comp-pointed-htpy h
+ ( f ∘∗ pointed-map-inv-is-pointed-equiv f H)
+ ( id-pointed-map)
+ ( is-pointed-section-pointed-map-inv-is-pointed-equiv f H)))
+
+ retraction-universal-property-pointed-equiv-is-pointed-equiv :
+ (H : is-pointed-equiv f) →
+ {l : Level} (C : Pointed-Type l) →
+ retraction (precomp-pointed-map f C)
+ pr1 (retraction-universal-property-pointed-equiv-is-pointed-equiv H C) =
+ map-inv-universal-property-pointed-equiv-is-pointed-equiv H C
+ pr2 (retraction-universal-property-pointed-equiv-is-pointed-equiv H C) =
+ is-retraction-map-inv-universal-property-pointed-equiv-is-pointed-equiv
+ ( H)
+ ( C)
+
+ universal-property-pointed-equiv-is-pointed-equiv :
+ is-pointed-equiv f →
+ universal-property-pointed-equiv f
+ pr1 (universal-property-pointed-equiv-is-pointed-equiv H C) =
+ section-universal-property-pointed-equiv-is-pointed-equiv H C
+ pr2 (universal-property-pointed-equiv-is-pointed-equiv H C) =
+ retraction-universal-property-pointed-equiv-is-pointed-equiv H C
equiv-precomp-pointed-map :
{l1 l2 l3 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
@@ -398,123 +542,163 @@ equiv-precomp-pointed-map :
pr1 (equiv-precomp-pointed-map C f) g =
g ∘∗ (pointed-map-pointed-equiv f)
pr2 (equiv-precomp-pointed-map C f) =
- is-equiv-precomp-is-equiv-pointed-map
+ universal-property-pointed-equiv-is-pointed-equiv
( pointed-map-pointed-equiv f)
- ( is-equiv-map-equiv-pointed-equiv f)
+ ( is-equiv-map-pointed-equiv f)
( C)
```
### Postcomposing by pointed equivalences is a pointed equivalence
+#### The predicate of being an equivalence by postcomposition of pointed maps
+
```agda
module _
{l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
where
- is-equiv-is-equiv-comp-pointed-map :
- ({l : Level} (X : Pointed-Type l) → is-equiv (comp-pointed-map {A = X} f)) →
- is-equiv-pointed-map f
- is-equiv-is-equiv-comp-pointed-map H =
+ is-equiv-postcomp-pointed-map : UUω
+ is-equiv-postcomp-pointed-map =
+ {l : Level} (X : Pointed-Type l) → is-equiv (postcomp-pointed-map f X)
+```
+
+#### Any pointed map that is an equivalence by postcomposition is a pointed equivalence
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ (H : is-equiv-postcomp-pointed-map f)
+ where
+
+ pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map : B →∗ A
+ pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map =
+ map-inv-is-equiv (H B) id-pointed-map
+
+ map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map :
+ type-Pointed-Type B → type-Pointed-Type A
+ map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map =
+ map-pointed-map
+ ( pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map)
+
+ is-pointed-section-pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map :
+ is-pointed-section f
+ ( pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map)
+ is-pointed-section-pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map =
+ pointed-htpy-eq
+ ( f ∘∗ pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map)
+ ( id-pointed-map)
+ ( is-section-map-inv-is-equiv (H B) id-pointed-map)
+
+ is-section-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map :
+ is-section
+ ( map-pointed-map f)
+ ( map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map)
+ is-section-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map =
+ htpy-pointed-htpy
+ ( is-pointed-section-pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map)
+
+ is-pointed-retraction-pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map :
+ is-pointed-retraction f
+ ( pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map)
+ is-pointed-retraction-pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map =
+ pointed-htpy-eq
+ ( pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map ∘∗ f)
+ ( id-pointed-map)
+ ( is-injective-is-equiv
+ ( H A)
+ ( eq-pointed-htpy
+ ( ( f) ∘∗
+ ( ( pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map) ∘∗
+ ( f)))
+ ( f ∘∗ id-pointed-map)
+ ( concat-pointed-htpy
+ ( inv-associative-comp-pointed-map f _ f)
+ ( concat-pointed-htpy
+ ( right-whisker-comp-pointed-htpy
+ ( f ∘∗ _)
+ ( id-pointed-map)
+ ( is-pointed-section-pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map)
+ ( f))
+ ( concat-pointed-htpy
+ ( left-unit-law-comp-pointed-map f)
+ ( inv-pointed-htpy (right-unit-law-comp-pointed-map f)))))))
+
+ is-retraction-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map :
+ is-retraction
+ ( map-pointed-map f)
+ ( map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map)
+ is-retraction-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map =
+ htpy-pointed-htpy
+ ( is-pointed-retraction-pointed-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map)
+
+ is-pointed-equiv-is-equiv-postcomp-pointed-map : is-pointed-equiv f
+ is-pointed-equiv-is-equiv-postcomp-pointed-map =
is-equiv-is-invertible
- ( map-pointed-map g)
- ( pr1 G)
- ( htpy-eq
- ( ap pr1
- ( ap pr1
- ( eq-is-contr
- ( is-contr-map-is-equiv (H A) f)
- { pair
- ( g ∘∗ f)
- ( eq-htpy-pointed-map
- ( f ∘∗ (g ∘∗ f))
- ( f)
- ( concat-htpy-pointed-map
- ( f ∘∗ (g ∘∗ f))
- ( (f ∘∗ g) ∘∗ f)
- ( f)
- ( inv-associative-comp-pointed-map f g f)
- ( concat-htpy-pointed-map
- ( (f ∘∗ g) ∘∗ f)
- ( id-pointed-map ∘∗ f)
- ( f)
- ( right-whisker-comp-pointed-map
- ( f ∘∗ g)
- ( id-pointed-map)
- ( G)
- ( f))
- ( left-unit-law-comp-pointed-map f))))}
- { pair
- ( id-pointed-map)
- ( eq-htpy-pointed-map
- ( f ∘∗ id-pointed-map)
- ( f)
- ( right-unit-law-comp-pointed-map f))}))))
- where
- g : B →∗ A
- g = pr1 (center (is-contr-map-is-equiv (H B) id-pointed-map))
- G : (f ∘∗ g) ~∗ id-pointed-map
- G = map-equiv
- ( extensionality-pointed-map
- ( f ∘∗ g)
- ( id-pointed-map))
- ( pr2 (center (is-contr-map-is-equiv (H B) id-pointed-map)))
-
- is-equiv-comp-is-equiv-pointed-map :
- is-equiv-pointed-map f →
- {l : Level} (X : Pointed-Type l) → is-equiv (comp-pointed-map {A = X} f)
- is-equiv-comp-is-equiv-pointed-map E X =
- pair
- ( pair
- ( g ∘∗_)
- ( λ k →
- eq-htpy-pointed-map
- ( f ∘∗ (g ∘∗ k))
- ( k)
- ( concat-htpy-pointed-map
- ( f ∘∗ (g ∘∗ k))
- ( (f ∘∗ g) ∘∗ k)
- ( k)
- ( inv-associative-comp-pointed-map f g k)
- ( concat-htpy-pointed-map
- ( (f ∘∗ g) ∘∗ k)
- ( id-pointed-map ∘∗ k)
- ( k)
- ( right-whisker-comp-pointed-map
- ( f ∘∗ g)
- ( id-pointed-map)
- ( G)
- ( k))
- ( left-unit-law-comp-pointed-map k)))))
- ( pair
- ( h ∘∗_)
- ( λ k →
- eq-htpy-pointed-map
- ( h ∘∗ (f ∘∗ k))
- ( k)
- ( concat-htpy-pointed-map
- ( h ∘∗ (f ∘∗ k))
- ( (h ∘∗ f) ∘∗ k)
- ( k)
- ( inv-associative-comp-pointed-map h f k)
- ( concat-htpy-pointed-map
- ( (h ∘∗ f) ∘∗ k)
- ( id-pointed-map ∘∗ k)
- ( k)
- ( right-whisker-comp-pointed-map
- ( h ∘∗ f)
- ( id-pointed-map)
- ( H)
- ( k))
- ( left-unit-law-comp-pointed-map k)))))
- where
- I : is-iso-pointed-map f
- I = is-iso-is-equiv-pointed-map f E
- g : B →∗ A
- g = pr1 (pr1 I)
- G : (f ∘∗ g) ~∗ id-pointed-map
- G = pr2 (pr1 I)
- h : B →∗ A
- h = pr1 (pr2 I)
- H : (h ∘∗ f) ~∗ id-pointed-map
- H = pr2 (pr2 I)
+ ( map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map)
+ ( is-section-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map)
+ ( is-retraction-map-inv-is-pointed-equiv-is-equiv-postcomp-pointed-map)
+```
+
+#### Any pointed equivalence is an equivalence by postcomposition
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ (H : is-pointed-equiv f)
+ where
+
+ map-inv-is-equiv-postcomp-is-pointed-equiv :
+ {l : Level} (X : Pointed-Type l) → (X →∗ B) → (X →∗ A)
+ map-inv-is-equiv-postcomp-is-pointed-equiv =
+ postcomp-pointed-map (pointed-map-inv-is-pointed-equiv f H)
+
+ is-section-map-inv-is-equiv-postcomp-is-pointed-equiv :
+ {l : Level} (X : Pointed-Type l) →
+ is-section
+ ( postcomp-pointed-map f X)
+ ( map-inv-is-equiv-postcomp-is-pointed-equiv X)
+ is-section-map-inv-is-equiv-postcomp-is-pointed-equiv X h =
+ eq-pointed-htpy
+ ( f ∘∗ (pointed-map-inv-is-pointed-equiv f H ∘∗ h))
+ ( h)
+ ( concat-pointed-htpy
+ ( inv-associative-comp-pointed-map f
+ ( pointed-map-inv-is-pointed-equiv f H)
+ ( h))
+ ( concat-pointed-htpy
+ ( right-whisker-comp-pointed-htpy
+ ( f ∘∗ pointed-map-inv-is-pointed-equiv f H)
+ ( id-pointed-map)
+ ( is-pointed-section-pointed-map-inv-is-pointed-equiv f H)
+ ( h))
+ ( left-unit-law-comp-pointed-map h)))
+
+ is-retraction-map-inv-is-equiv-postcomp-is-pointed-equiv :
+ {l : Level} (X : Pointed-Type l) →
+ is-retraction
+ ( postcomp-pointed-map f X)
+ ( map-inv-is-equiv-postcomp-is-pointed-equiv X)
+ is-retraction-map-inv-is-equiv-postcomp-is-pointed-equiv X h =
+ eq-pointed-htpy
+ ( pointed-map-inv-is-pointed-equiv f H ∘∗ (f ∘∗ h))
+ ( h)
+ ( concat-pointed-htpy
+ ( inv-associative-comp-pointed-map
+ ( pointed-map-inv-is-pointed-equiv f H)
+ ( f)
+ ( h))
+ ( concat-pointed-htpy
+ ( right-whisker-comp-pointed-htpy
+ ( pointed-map-inv-is-pointed-equiv f H ∘∗ f)
+ ( id-pointed-map)
+ ( is-pointed-retraction-pointed-map-inv-is-pointed-equiv f H)
+ ( h))
+ ( left-unit-law-comp-pointed-map h)))
+
+ is-equiv-postcomp-is-pointed-equiv : is-equiv-postcomp-pointed-map f
+ is-equiv-postcomp-is-pointed-equiv X =
+ is-equiv-is-invertible
+ ( map-inv-is-equiv-postcomp-is-pointed-equiv X)
+ ( is-section-map-inv-is-equiv-postcomp-is-pointed-equiv X)
+ ( is-retraction-map-inv-is-equiv-postcomp-is-pointed-equiv X)
```
diff --git a/src/structured-types/pointed-homotopies.lagda.md b/src/structured-types/pointed-homotopies.lagda.md
index c348c745b1..381f10197d 100644
--- a/src/structured-types/pointed-homotopies.lagda.md
+++ b/src/structured-types/pointed-homotopies.lagda.md
@@ -10,14 +10,25 @@ module structured-types.pointed-homotopies where
open import foundation.action-on-higher-identifications-functions
open import foundation.action-on-identifications-binary-functions
open import foundation.action-on-identifications-functions
+open import foundation.binary-equivalences
+open import foundation.commuting-triangles-of-identifications
+open import foundation.contractible-types
open import foundation.dependent-pair-types
open import foundation.equivalences
open import foundation.function-extensionality
+open import foundation.function-types
+open import foundation.functoriality-dependent-pair-types
+open import foundation.fundamental-theorem-of-identity-types
open import foundation.homotopies
+open import foundation.homotopy-induction
open import foundation.identity-types
+open import foundation.path-algebra
open import foundation.structure-identity-principle
open import foundation.universe-levels
open import foundation.whiskering-homotopies-composition
+open import foundation.whiskering-identifications-concatenation
+
+open import foundation-core.torsorial-type-families
open import structured-types.pointed-dependent-functions
open import structured-types.pointed-families-of-types
@@ -29,248 +40,521 @@ open import structured-types.pointed-types
## Idea
-A pointed homotopy between pointed dependent functions is a pointed dependent
-function of the pointed family of pointwise identifications. Since pointed
-homotopies are defined for pointed dependent functions, a pointed homotopy
-between pointed homotopies is just an instance of a pointed homotopy.
+A {{#concept "pointed homotopy" Agda=_~∗_}} between
+[pointed dependent functions](structured-types.pointed-dependent-functions.md)
+`f := (f₀ , f₁)` and `g := (g₀ , g₁)` of type `pointed-Π A B` consists of an
+unpointed [homotopy](foundation-core.homotopies.md) `H₀ : f₀ ~ g₀` and an
+[identification](foundation-core.identity-types.md) `H₁ : f₁ = (H₀ *) ∙ g₁`
+witnessing that the triangle of identifications
+
+```text
+ H₀ *
+ f₀ * ------> g₀ *
+ \ /
+ f₁ \ / g₁
+ \ /
+ ∨ ∨
+ *
+```
+
+[commutes](foundation.commuting-triangles-of-identifications.md). This
+identification is called the
+{{#concept "base point coherence" Disambiguation="pointed homotopy" Agda=coherence-point-unpointed-htpy-pointed-Π}}
+of the pointed homotopy `H := (H₀ , H₁)`.
+
+An equivalent way of defining pointed homotopies occurs when we consider the
+type family `x ↦ f₀ x = g₀ x` over the base type `A`. This is a
+[pointed type family](structured-types.pointed-families-of-types.md), where the
+base point is the identification
+
+```text
+ f₁ ∙ inv g₁ : f₀ * = g₀ *.
+```
+
+A pointed homotopy `f ~∗ g` is therefore equivalently defined as a pointed
+dependent function of the pointed type family `x ↦ f₀ x = g₀ x`. Such a pointed
+dependent function consists of an unpointed homotopy `H₀ : f₀ ~ g₀` between the
+underlying dependent functions and an identification witnessing that the
+triangle of identifications
+
+```text
+ H₀ *
+ f₀ * ------> g₀ *
+ \ ∧
+ f₁ \ / inv g₁
+ \ /
+ ∨ /
+ *
+```
+
+[commutes](foundation.commuting-triangles-of-identifications.md). Notice that
+the identification on the right in this triangle goes up, in the inverse
+direction of the identification `g₁`.
+
+Note that in this second definition of pointed homotopies we defined pointed
+homotopies between pointed dependent functions to be certain pointed dependent
+functions. This implies that the second definition is a uniform definition that
+can easily be iterated in order to consider pointed higher homotopies. For this
+reason, we call the second definition of pointed homotopies
+[uniform pointed homotopies](structured-types.uniform-pointed-homotopies.md).
+
+Note that the difference between our main definition of pointed homotopies and
+the uniform definition of pointed homotopies is the direction of the
+identification on the right in the commuting triangle
+
+```text
+ H₀ *
+ f₀ * ------> g₀ *
+ \ /
+ f₁ \ / g₁
+ \ /
+ ∨ ∨
+ *.
+```
+
+In the definition of uniform pointed homotopies it goes in the reverse
+direction. This makes it slightly more complicated to construct an
+identification witnessing that the triangle commutes in the case of uniform
+pointed homotopies. Furthermore, this complication becomes more significant and
+bothersome when we are trying to construct a
+[pointed `2`-homotopy](structured-types.pointed-2-homotopies.md). For this
+reason, our first definition where pointed homotopies are defined to consist of
+unpointed homotopies and a base point coherence, is taken to be our main
+definition of pointed homotopy. The only disadvantage of the nonuniform
+definition of pointed homotopies is that it does not easily iterate.
-## Definition
+We will write `f ~∗ g` for the type of (nonuniform) pointed homotopies, and we
+will write `H ~²∗ K` for the nonuniform definition of pointed `2`-homotopies.
+
+## Definitions
+
+### Unpointed homotopies between pointed dependent functions
```agda
module _
{l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
- (f : pointed-Π A B)
+ (f g : pointed-Π A B)
where
- htpy-pointed-Π : pointed-Π A B → UU (l1 ⊔ l2)
- htpy-pointed-Π g =
- pointed-Π A
- ( pair
- ( λ x →
- Id
- ( function-pointed-Π f x)
- ( function-pointed-Π g x))
- ( ( preserves-point-function-pointed-Π f) ∙
- ( inv (preserves-point-function-pointed-Π g))))
-
- extensionality-pointed-Π : (g : pointed-Π A B) → Id f g ≃ htpy-pointed-Π g
- extensionality-pointed-Π =
- extensionality-Σ
- ( λ {g} q H →
- Id
- ( H (point-Pointed-Type A))
- ( preserves-point-function-pointed-Π f ∙
- inv (preserves-point-function-pointed-Π (g , q))))
- ( refl-htpy)
- ( inv (right-inv (preserves-point-function-pointed-Π f)))
- ( λ g → equiv-funext)
- ( λ p →
- ( equiv-right-transpose-eq-concat
- ( refl)
- ( p)
- ( preserves-point-function-pointed-Π f)) ∘e
- ( equiv-inv (preserves-point-function-pointed-Π f) p))
-
- eq-htpy-pointed-Π :
- (g : pointed-Π A B) → (htpy-pointed-Π g) → Id f g
- eq-htpy-pointed-Π g = map-inv-equiv (extensionality-pointed-Π g)
-
-_~∗_ :
- {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A} →
- pointed-Π A B → pointed-Π A B → UU (l1 ⊔ l2)
-_~∗_ {A = A} {B} = htpy-pointed-Π
-
-infix 6 _~∗_
-
-htpy-pointed-htpy :
- {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A} →
- (f g : pointed-Π A B) → f ~∗ g →
- function-pointed-Π f ~ function-pointed-Π g
-htpy-pointed-htpy f g = pr1
-
-triangle-pointed-htpy :
- {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A} →
- (f g : pointed-Π A B) (H : f ~∗ g) →
- ( htpy-pointed-htpy f g H (point-Pointed-Type A)) =
- ( ( preserves-point-function-pointed-Π f) ∙
- ( inv (preserves-point-function-pointed-Π g)))
-triangle-pointed-htpy f g = pr2
+ unpointed-htpy-pointed-Π : UU (l1 ⊔ l2)
+ unpointed-htpy-pointed-Π = function-pointed-Π f ~ function-pointed-Π g
```
-## Properties
+### Unpointed homotopies between pointed maps
-### Pointed homotopies are equivalent to identifications of pointed maps
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ (f g : A →∗ B)
+ where
+
+ unpointed-htpy-pointed-map : UU (l1 ⊔ l2)
+ unpointed-htpy-pointed-map = map-pointed-map f ~ map-pointed-map g
+```
+
+### The base point coherence of unpointed homotopies between pointed maps
+
+The base point coherence of pointed homotopies asserts that its underlying
+homotopy preserves the base point, in the sense that the triangle of
+identifications
+
+```text
+ H *
+ f * ------> g *
+ \ /
+ preserves-point f \ / preserves-point g
+ \ /
+ ∨ ∨
+ *
+```
+
+commutes.
```agda
module _
- {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ (f g : pointed-Π A B) (G : unpointed-htpy-pointed-Π f g)
where
- refl-htpy-pointed-map : f ~∗ f
- pr1 refl-htpy-pointed-map = refl-htpy
- pr2 refl-htpy-pointed-map = inv (right-inv (preserves-point-pointed-map f))
+ coherence-point-unpointed-htpy-pointed-Π : UU l2
+ coherence-point-unpointed-htpy-pointed-Π =
+ coherence-triangle-identifications
+ ( preserves-point-function-pointed-Π f)
+ ( preserves-point-function-pointed-Π g)
+ ( G (point-Pointed-Type A))
+```
- htpy-pointed-map : (g : A →∗ B) → UU (l1 ⊔ l2)
- htpy-pointed-map = htpy-pointed-Π f
+### Pointed homotopies
- extensionality-pointed-map : (g : A →∗ B) → Id f g ≃ (htpy-pointed-map g)
- extensionality-pointed-map = extensionality-pointed-Π f
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ (f g : pointed-Π A B)
+ where
+
+ pointed-htpy : UU (l1 ⊔ l2)
+ pointed-htpy =
+ Σ ( function-pointed-Π f ~ function-pointed-Π g)
+ ( coherence-point-unpointed-htpy-pointed-Π f g)
+
+ infix 6 _~∗_
+
+ _~∗_ : UU (l1 ⊔ l2)
+ _~∗_ = pointed-htpy
+
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B} (H : f ~∗ g)
+ where
+
+ htpy-pointed-htpy : function-pointed-Π f ~ function-pointed-Π g
+ htpy-pointed-htpy = pr1 H
- eq-htpy-pointed-map : (g : A →∗ B) → (htpy-pointed-map g) → Id f g
- eq-htpy-pointed-map g = map-inv-equiv (extensionality-pointed-map g)
+ coherence-point-pointed-htpy :
+ coherence-point-unpointed-htpy-pointed-Π f g htpy-pointed-htpy
+ coherence-point-pointed-htpy = pr2 H
```
-### The category laws for pointed maps
+### The reflexive pointed homotopy
```agda
module _
- {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ (f : pointed-Π A B)
where
- left-unit-law-comp-pointed-map :
- htpy-pointed-map (comp-pointed-map id-pointed-map f) f
- left-unit-law-comp-pointed-map =
- pair
- ( refl-htpy)
- ( ( inv (right-inv (pr2 f))) ∙
- ( ap
- ( concat'
- ( map-pointed-map f (point-Pointed-Type A))
- ( inv (pr2 f)))
- ( ( inv (ap-id (pr2 f))) ∙
- ( inv right-unit))))
-
- right-unit-law-comp-pointed-map :
- htpy-pointed-map (comp-pointed-map f id-pointed-map) f
- right-unit-law-comp-pointed-map =
- pair
- ( refl-htpy)
- ( inv (right-inv (pr2 f)))
+ refl-pointed-htpy : pointed-htpy f f
+ pr1 refl-pointed-htpy = refl-htpy
+ pr2 refl-pointed-htpy = refl
+```
+
+### Concatenation of pointed homotopies
+
+Consider three pointed dependent functions `f := (f₀ , f₁)`, `g := (g₀ , g₁)`,
+and `h := (h₀ , h₁)`, and pointed homotopies `G := (G₀ , G₁)` and
+`H := (H₀ , H₁)` between them as indicated in the diagram
+
+```text
+ G H
+ f -----> g -----> h
+```
+
+The concatenation `(G ∙h H)` of `G` and `H` has underlying unpointed homotopy
+
+```text
+ (G ∙h H)₀ := G₀ ∙h H₀.
+```
+
+The base point coherence `(G ∙h H)₁` of `(G ∙h H)` is obtained by horizontally
+pasting the commuting triangles of identifications
+
+```text
+ G₀ * H₀ *
+ f₀ * --> g₀ * ---> h₀ *
+ \ | /
+ \ | g₁ /
+ f₁ \ | / h₁
+ \ | /
+ \ | /
+ ∨ ∨ ∨
+ *.
+```
+```agda
module _
- {l1 l2 l3 l4 : Level}
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g h : pointed-Π A B} (G : f ~∗ g) (H : g ~∗ h)
where
- associative-comp-pointed-map :
- {A : Pointed-Type l1} {B : Pointed-Type l2}
- {C : Pointed-Type l3} {D : Pointed-Type l4}
- (h : C →∗ D) (g : B →∗ C) (f : A →∗ B) →
- htpy-pointed-map
- ( comp-pointed-map (comp-pointed-map h g) f)
- ( comp-pointed-map h (comp-pointed-map g f))
- associative-comp-pointed-map
- (pair h refl) (pair g refl) (pair f refl) =
- pair refl-htpy refl
+ htpy-concat-pointed-htpy : unpointed-htpy-pointed-Π f h
+ htpy-concat-pointed-htpy = htpy-pointed-htpy G ∙h htpy-pointed-htpy H
+
+ coherence-point-concat-pointed-htpy :
+ coherence-point-unpointed-htpy-pointed-Π f h htpy-concat-pointed-htpy
+ coherence-point-concat-pointed-htpy =
+ horizontal-pasting-coherence-triangle-identifications
+ ( preserves-point-function-pointed-Π f)
+ ( preserves-point-function-pointed-Π g)
+ ( preserves-point-function-pointed-Π h)
+ ( htpy-pointed-htpy G (point-Pointed-Type A))
+ ( htpy-pointed-htpy H (point-Pointed-Type A))
+ ( coherence-point-pointed-htpy G)
+ ( coherence-point-pointed-htpy H)
+
+ concat-pointed-htpy : f ~∗ h
+ pr1 concat-pointed-htpy = htpy-concat-pointed-htpy
+ pr2 concat-pointed-htpy = coherence-point-concat-pointed-htpy
+```
- inv-associative-comp-pointed-map :
- {A : Pointed-Type l1} {B : Pointed-Type l2}
- {C : Pointed-Type l3} {D : Pointed-Type l4}
- (h : C →∗ D) (g : B →∗ C) (f : A →∗ B) →
- htpy-pointed-map
- ( comp-pointed-map h (comp-pointed-map g f))
- ( comp-pointed-map (comp-pointed-map h g) f)
- inv-associative-comp-pointed-map
- (pair h refl) (pair g refl) (pair f refl) =
- pair refl-htpy refl
+### Inverses of pointed homotopies
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B} (H : f ~∗ g)
+ where
+
+ htpy-inv-pointed-htpy : unpointed-htpy-pointed-Π g f
+ htpy-inv-pointed-htpy = inv-htpy (htpy-pointed-htpy H)
+
+ coherence-point-inv-pointed-htpy :
+ coherence-point-unpointed-htpy-pointed-Π g f htpy-inv-pointed-htpy
+ coherence-point-inv-pointed-htpy =
+ transpose-top-coherence-triangle-identifications
+ ( preserves-point-function-pointed-Π g)
+ ( preserves-point-function-pointed-Π f)
+ ( htpy-pointed-htpy H (point-Pointed-Type A))
+ ( refl)
+ ( coherence-point-pointed-htpy H)
+
+ inv-pointed-htpy : g ~∗ f
+ pr1 inv-pointed-htpy = htpy-inv-pointed-htpy
+ pr2 inv-pointed-htpy = coherence-point-inv-pointed-htpy
```
-### The groupoid laws for pointed homotopies
+## Properties
+
+### Extensionality of pointed dependent function types by pointed homotopies
```agda
module _
{l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ (f : pointed-Π A B)
where
- concat-htpy-pointed-Π :
- (f g h : pointed-Π A B) →
- htpy-pointed-Π f g → htpy-pointed-Π g h → htpy-pointed-Π f h
- concat-htpy-pointed-Π f g h G H =
- pair
- ( pr1 G ∙h pr1 H)
- ( ( ap-binary (λ p q → p ∙ q) (pr2 G) (pr2 H)) ∙
- ( ( assoc (pr2 f) (inv (pr2 g)) (pr2 g ∙ inv (pr2 h))) ∙
- ( ap
- ( concat (pr2 f) (function-pointed-Π h (point-Pointed-Type A)))
- ( ( inv (assoc (inv (pr2 g)) (pr2 g) (inv (pr2 h)))) ∙
- ( ap
- ( concat' (point-Pointed-Fam A B) (inv (pr2 h)))
- ( left-inv (pr2 g)))))))
-
- inv-htpy-pointed-Π :
- (f g : pointed-Π A B) → htpy-pointed-Π f g → htpy-pointed-Π g f
- inv-htpy-pointed-Π f g H =
- pair
- ( inv-htpy (pr1 H))
- ( ( ap inv (pr2 H)) ∙
- ( ( distributive-inv-concat (pr2 f) (inv (pr2 g))) ∙
- ( ap
- ( concat'
- ( function-pointed-Π g (point-Pointed-Type A))
- ( inv (pr2 f)))
- ( inv-inv (pr2 g)))))
+ abstract
+ is-torsorial-pointed-htpy :
+ is-torsorial (pointed-htpy f)
+ is-torsorial-pointed-htpy =
+ is-torsorial-Eq-structure
+ ( is-torsorial-htpy _)
+ ( function-pointed-Π f , refl-htpy)
+ ( is-torsorial-Id _)
+
+ pointed-htpy-eq :
+ (g : pointed-Π A B) → f = g → f ~∗ g
+ pointed-htpy-eq .f refl = refl-pointed-htpy f
+
+ abstract
+ is-equiv-pointed-htpy-eq :
+ (g : pointed-Π A B) → is-equiv (pointed-htpy-eq g)
+ is-equiv-pointed-htpy-eq =
+ fundamental-theorem-id
+ ( is-torsorial-pointed-htpy)
+ ( pointed-htpy-eq)
+
+ extensionality-pointed-Π :
+ (g : pointed-Π A B) → (f = g) ≃ (f ~∗ g)
+ pr1 (extensionality-pointed-Π g) = pointed-htpy-eq g
+ pr2 (extensionality-pointed-Π g) = is-equiv-pointed-htpy-eq g
+
+ eq-pointed-htpy :
+ (g : pointed-Π A B) → f ~∗ g → f = g
+ eq-pointed-htpy g = map-inv-equiv (extensionality-pointed-Π g)
+```
+### Extensionality of pointed maps
+
+```agda
module _
- {l1 l2 l3 : Level}
- {A : Pointed-Type l1} {B : Pointed-Type l2} {C : Pointed-Type l3}
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ (f : A →∗ B)
where
- left-whisker-comp-pointed-map :
- (g : B →∗ C) (f1 f2 : A →∗ B) (H : htpy-pointed-map f1 f2) →
- htpy-pointed-map
- ( comp-pointed-map g f1)
- ( comp-pointed-map g f2)
- left-whisker-comp-pointed-map g f1 f2 H =
- pair
- ( map-pointed-map g ·l (pr1 H))
- ( ( ( ( ap² (pr1 g) (pr2 H)) ∙
- ( ap-concat (pr1 g) (pr2 f1) (inv (pr2 f2)))) ∙
- ( ap
- ( concat
- ( ap (pr1 g) (pr2 f1))
- ( map-pointed-map g
- ( map-pointed-map f2 (point-Pointed-Type A))))
- ( ( ( ( ap-inv (pr1 g) (pr2 f2)) ∙
- ( ap
- ( concat'
- ( pr1 g (point-Pointed-Fam A (constant-Pointed-Fam A B)))
- ( inv (ap (pr1 g) (pr2 f2)))))
- ( inv (right-inv (pr2 g)))) ∙
- ( assoc
- ( pr2 g)
- ( inv (pr2 g))
- ( inv (ap (pr1 g) (pr2 f2))))) ∙
- ( ap
- ( concat
- ( pr2 g)
- ( map-pointed-map g
- ( map-pointed-map f2 (point-Pointed-Type A))))
- ( inv
- ( distributive-inv-concat
- ( ap (pr1 g) (pr2 f2))
- ( pr2 g))))))) ∙
+ extensionality-pointed-map :
+ (g : A →∗ B) → (f = g) ≃ (f ~∗ g)
+ extensionality-pointed-map = extensionality-pointed-Π f
+```
+
+### Associativity of composition of pointed maps
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {C : Pointed-Type l3} {D : Pointed-Type l4}
+ (h : C →∗ D) (g : B →∗ C) (f : A →∗ B)
+ where
+
+ htpy-inv-associative-comp-pointed-map :
+ unpointed-htpy-pointed-Π (h ∘∗ (g ∘∗ f)) ((h ∘∗ g) ∘∗ f)
+ htpy-inv-associative-comp-pointed-map = refl-htpy
+
+ coherence-point-inv-associative-comp-pointed-map :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( h ∘∗ (g ∘∗ f))
+ ( (h ∘∗ g) ∘∗ f)
+ ( htpy-inv-associative-comp-pointed-map)
+ coherence-point-inv-associative-comp-pointed-map =
+ right-whisker-concat-coherence-triangle-identifications
+ ( ap
+ ( map-pointed-map h)
+ ( ( ap
+ ( map-pointed-map g)
+ ( preserves-point-pointed-map f)) ∙
+ ( preserves-point-pointed-map g)))
+ ( ap (map-pointed-map h) _)
+ ( ap (map-comp-pointed-map h g) (preserves-point-pointed-map f))
+ ( preserves-point-pointed-map h)
+ ( ( ap-concat
+ ( map-pointed-map h)
+ ( ap (map-pointed-map g) (preserves-point-pointed-map f))
+ ( preserves-point-pointed-map g)) ∙
( inv
- ( assoc
- ( ap (pr1 g) (pr2 f1))
- ( pr2 g)
- ( inv (ap (pr1 g) (pr2 f2) ∙ pr2 g)))))
+ ( right-whisker-concat
+ ( ap-comp
+ ( map-pointed-map h)
+ ( map-pointed-map g)
+ ( preserves-point-pointed-map f))
+ ( ap (map-pointed-map h) (preserves-point-pointed-map g)))))
+
+ inv-associative-comp-pointed-map :
+ h ∘∗ (g ∘∗ f) ~∗ (h ∘∗ g) ∘∗ f
+ pr1 inv-associative-comp-pointed-map =
+ htpy-inv-associative-comp-pointed-map
+ pr2 inv-associative-comp-pointed-map =
+ coherence-point-inv-associative-comp-pointed-map
+
+ htpy-associative-comp-pointed-map :
+ unpointed-htpy-pointed-Π ((h ∘∗ g) ∘∗ f) (h ∘∗ (g ∘∗ f))
+ htpy-associative-comp-pointed-map = refl-htpy
+
+ coherence-associative-comp-pointed-map :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( (h ∘∗ g) ∘∗ f)
+ ( h ∘∗ (g ∘∗ f))
+ ( htpy-associative-comp-pointed-map)
+ coherence-associative-comp-pointed-map =
+ inv coherence-point-inv-associative-comp-pointed-map
+
+ associative-comp-pointed-map :
+ (h ∘∗ g) ∘∗ f ~∗ h ∘∗ (g ∘∗ f)
+ pr1 associative-comp-pointed-map =
+ htpy-associative-comp-pointed-map
+ pr2 associative-comp-pointed-map =
+ coherence-associative-comp-pointed-map
+```
+
+### The left unit law for composition of pointed maps
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ where
+
+ htpy-left-unit-law-comp-pointed-map :
+ id ∘ map-pointed-map f ~ map-pointed-map f
+ htpy-left-unit-law-comp-pointed-map = refl-htpy
+
+ coherence-left-unit-law-comp-pointed-map :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( id-pointed-map ∘∗ f)
+ ( f)
+ ( htpy-left-unit-law-comp-pointed-map)
+ coherence-left-unit-law-comp-pointed-map =
+ right-unit ∙ ap-id (preserves-point-pointed-map f)
+
+ left-unit-law-comp-pointed-map : id-pointed-map ∘∗ f ~∗ f
+ pr1 left-unit-law-comp-pointed-map = htpy-left-unit-law-comp-pointed-map
+ pr2 left-unit-law-comp-pointed-map = coherence-left-unit-law-comp-pointed-map
+
+ htpy-inv-left-unit-law-comp-pointed-map :
+ map-pointed-map f ~ id ∘ map-pointed-map f
+ htpy-inv-left-unit-law-comp-pointed-map = refl-htpy
+
+ coherence-point-inv-left-unit-law-comp-pointed-map :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( f)
+ ( id-pointed-map ∘∗ f)
+ ( htpy-inv-left-unit-law-comp-pointed-map)
+ coherence-point-inv-left-unit-law-comp-pointed-map =
+ inv coherence-left-unit-law-comp-pointed-map
+
+ inv-left-unit-law-comp-pointed-map : f ~∗ id-pointed-map ∘∗ f
+ pr1 inv-left-unit-law-comp-pointed-map =
+ htpy-inv-left-unit-law-comp-pointed-map
+ pr2 inv-left-unit-law-comp-pointed-map =
+ coherence-point-inv-left-unit-law-comp-pointed-map
+```
+
+### The right unit law for composition of pointed maps
+```agda
module _
- {l1 l2 l3 : Level}
- {A : Pointed-Type l1} {B : Pointed-Type l2} {C : Pointed-Type l3}
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
where
- right-whisker-comp-pointed-map :
- (g1 g2 : B →∗ C) (H : htpy-pointed-map g1 g2) (f : A →∗ B) →
- htpy-pointed-map (comp-pointed-map g1 f) (comp-pointed-map g2 f)
- right-whisker-comp-pointed-map g1 g2 H (pair f refl) =
- pair (pr1 H ·r f) (pr2 H)
+ htpy-right-unit-law-comp-pointed-map :
+ map-pointed-map f ∘ id ~ map-pointed-map f
+ htpy-right-unit-law-comp-pointed-map = refl-htpy
+
+ coherence-right-unit-law-comp-pointed-map :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( f ∘∗ id-pointed-map)
+ ( f)
+ ( htpy-right-unit-law-comp-pointed-map)
+ coherence-right-unit-law-comp-pointed-map = refl
+
+ right-unit-law-comp-pointed-map : f ∘∗ id-pointed-map ~∗ f
+ pr1 right-unit-law-comp-pointed-map =
+ htpy-right-unit-law-comp-pointed-map
+ pr2 right-unit-law-comp-pointed-map =
+ coherence-right-unit-law-comp-pointed-map
+```
+### Concatenation of pointed homotopies is a binary equivalence
+
+```agda
module _
- {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g h : pointed-Π A B}
where
- concat-htpy-pointed-map :
- (f g h : A →∗ B) → htpy-pointed-map f g → htpy-pointed-map g h →
- htpy-pointed-map f h
- concat-htpy-pointed-map = concat-htpy-pointed-Π
+ abstract
+ is-equiv-concat-pointed-htpy :
+ (G : f ~∗ g) → is-equiv (λ (H : g ~∗ h) → concat-pointed-htpy G H)
+ is-equiv-concat-pointed-htpy G =
+ is-equiv-map-Σ _
+ ( is-equiv-concat-htpy (htpy-pointed-htpy G) _)
+ ( λ H →
+ is-equiv-horizontal-pasting-coherence-triangle-identifications
+ ( preserves-point-function-pointed-Π f)
+ ( preserves-point-function-pointed-Π g)
+ ( preserves-point-function-pointed-Π h)
+ ( htpy-pointed-htpy G _)
+ ( H _)
+ ( coherence-point-pointed-htpy G))
+
+ equiv-concat-pointed-htpy : f ~∗ g → (g ~∗ h) ≃ (f ~∗ h)
+ pr1 (equiv-concat-pointed-htpy G) = concat-pointed-htpy G
+ pr2 (equiv-concat-pointed-htpy G) = is-equiv-concat-pointed-htpy G
+
+ abstract
+ is-equiv-concat-pointed-htpy' :
+ (H : g ~∗ h) → is-equiv (λ (G : f ~∗ g) → concat-pointed-htpy G H)
+ is-equiv-concat-pointed-htpy' H =
+ is-equiv-map-Σ _
+ ( is-equiv-concat-htpy' _ (htpy-pointed-htpy H))
+ ( λ G →
+ is-equiv-horizontal-pasting-coherence-triangle-identifications'
+ ( preserves-point-function-pointed-Π f)
+ ( preserves-point-function-pointed-Π g)
+ ( preserves-point-function-pointed-Π h)
+ ( G _)
+ ( htpy-pointed-htpy H _)
+ ( coherence-point-pointed-htpy H))
+
+ equiv-concat-pointed-htpy' : g ~∗ h → (f ~∗ g) ≃ (f ~∗ h)
+ pr1 (equiv-concat-pointed-htpy' H) G = concat-pointed-htpy G H
+ pr2 (equiv-concat-pointed-htpy' H) = is-equiv-concat-pointed-htpy' H
+
+ is-binary-equiv-concat-pointed-htpy :
+ is-binary-equiv (λ (G : f ~∗ g) (H : g ~∗ h) → concat-pointed-htpy G H)
+ pr1 is-binary-equiv-concat-pointed-htpy = is-equiv-concat-pointed-htpy'
+ pr2 is-binary-equiv-concat-pointed-htpy = is-equiv-concat-pointed-htpy
```
+
+## See also
+
+- [Pointed 2-homotopies](structured-types.pointed-2-homotopies.md)
+- [Uniform pointed homotopies](structured-types.uniform-pointed-homotopies.md)
diff --git a/src/structured-types/pointed-isomorphisms.lagda.md b/src/structured-types/pointed-isomorphisms.lagda.md
new file mode 100644
index 0000000000..8c97d7ac80
--- /dev/null
+++ b/src/structured-types/pointed-isomorphisms.lagda.md
@@ -0,0 +1,386 @@
+# Pointed isomorphisms
+
+```agda
+module structured-types.pointed-isomorphisms where
+```
+
+Imports
+
+```agda
+open import foundation.cartesian-product-types
+open import foundation.contractible-maps
+open import foundation.contractible-types
+open import foundation.dependent-pair-types
+open import foundation.equivalences
+open import foundation.identity-types
+open import foundation.propositions
+open import foundation.retractions
+open import foundation.sections
+open import foundation.structure-identity-principle
+open import foundation.universe-levels
+
+open import structured-types.pointed-equivalences
+open import structured-types.pointed-homotopies
+open import structured-types.pointed-maps
+open import structured-types.pointed-retractions
+open import structured-types.pointed-sections
+open import structured-types.pointed-types
+```
+
+
+
+## Idea
+
+A {{#concept "pointed isomorphism" Agda=pointed-iso}} is an isomorphism in the
+wild category of pointed types, i.e., it is a
+[pointed map](structured-types.pointed-types.md) equipped with a
+[pointed section](structured-types.pointed-sections.md) and a
+[pointed retraction](structured-types.pointed-retractions.md).
+
+## Definitions
+
+### The predicate of being a pointed isomorphism
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ where
+
+ is-pointed-iso : UU (l1 ⊔ l2)
+ is-pointed-iso = pointed-section f × pointed-retraction f
+
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} {f : A →∗ B}
+ (H : is-pointed-iso f)
+ where
+
+ pointed-section-is-pointed-iso : pointed-section f
+ pointed-section-is-pointed-iso = pr1 H
+
+ pointed-retraction-is-pointed-iso : pointed-retraction f
+ pointed-retraction-is-pointed-iso = pr2 H
+
+ pointed-map-pointed-section-is-pointed-iso : B →∗ A
+ pointed-map-pointed-section-is-pointed-iso =
+ pointed-map-pointed-section f pointed-section-is-pointed-iso
+
+ is-pointed-section-pointed-section-is-pointed-iso :
+ is-pointed-section f pointed-map-pointed-section-is-pointed-iso
+ is-pointed-section-pointed-section-is-pointed-iso =
+ is-pointed-section-pointed-section f
+ pointed-section-is-pointed-iso
+
+ map-pointed-section-is-pointed-iso :
+ type-Pointed-Type B → type-Pointed-Type A
+ map-pointed-section-is-pointed-iso =
+ map-pointed-section f pointed-section-is-pointed-iso
+
+ preserves-point-pointed-map-pointed-section-is-pointed-iso :
+ map-pointed-section-is-pointed-iso (point-Pointed-Type B) =
+ point-Pointed-Type A
+ preserves-point-pointed-map-pointed-section-is-pointed-iso =
+ preserves-point-pointed-map-pointed-section f
+ pointed-section-is-pointed-iso
+
+ is-section-pointed-section-is-pointed-iso :
+ is-section (map-pointed-map f) map-pointed-section-is-pointed-iso
+ is-section-pointed-section-is-pointed-iso =
+ is-section-pointed-section f pointed-section-is-pointed-iso
+
+ section-is-pointed-iso :
+ section (map-pointed-map f)
+ section-is-pointed-iso =
+ section-pointed-section f pointed-section-is-pointed-iso
+
+ coherence-point-is-section-pointed-section-is-pointed-iso :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( f ∘∗ pointed-map-pointed-section-is-pointed-iso)
+ ( id-pointed-map)
+ ( is-section-pointed-section-is-pointed-iso)
+ coherence-point-is-section-pointed-section-is-pointed-iso =
+ coherence-point-is-section-pointed-section f
+ pointed-section-is-pointed-iso
+
+ pointed-map-pointed-retraction-is-pointed-iso : B →∗ A
+ pointed-map-pointed-retraction-is-pointed-iso =
+ pr1 pointed-retraction-is-pointed-iso
+
+ is-pointed-retraction-pointed-retraction-is-pointed-iso :
+ is-pointed-retraction f pointed-map-pointed-retraction-is-pointed-iso
+ is-pointed-retraction-pointed-retraction-is-pointed-iso =
+ pr2 pointed-retraction-is-pointed-iso
+
+ map-pointed-retraction-is-pointed-iso :
+ type-Pointed-Type B → type-Pointed-Type A
+ map-pointed-retraction-is-pointed-iso =
+ map-pointed-map pointed-map-pointed-retraction-is-pointed-iso
+
+ preserves-point-pointed-map-pointed-retraction-is-pointed-iso :
+ map-pointed-retraction-is-pointed-iso (point-Pointed-Type B) =
+ point-Pointed-Type A
+ preserves-point-pointed-map-pointed-retraction-is-pointed-iso =
+ preserves-point-pointed-map
+ pointed-map-pointed-retraction-is-pointed-iso
+
+ is-retraction-pointed-retraction-is-pointed-iso :
+ is-retraction
+ ( map-pointed-map f)
+ ( map-pointed-retraction-is-pointed-iso)
+ is-retraction-pointed-retraction-is-pointed-iso =
+ htpy-pointed-htpy
+ is-pointed-retraction-pointed-retraction-is-pointed-iso
+
+ retraction-is-pointed-iso :
+ retraction (map-pointed-map f)
+ retraction-is-pointed-iso =
+ retraction-pointed-retraction f pointed-retraction-is-pointed-iso
+
+ coherence-point-is-retraction-pointed-retraction-is-pointed-iso :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( pointed-map-pointed-retraction-is-pointed-iso ∘∗ f)
+ ( id-pointed-map)
+ ( is-retraction-pointed-retraction-is-pointed-iso)
+ coherence-point-is-retraction-pointed-retraction-is-pointed-iso =
+ coherence-point-pointed-htpy
+ is-pointed-retraction-pointed-retraction-is-pointed-iso
+```
+
+### Pointed isomorphisms
+
+```agda
+module _
+ {l1 l2 : Level} (A : Pointed-Type l1) (B : Pointed-Type l2)
+ where
+
+ pointed-iso : UU (l1 ⊔ l2)
+ pointed-iso = Σ (A →∗ B) is-pointed-iso
+
+ infix 6 _≅∗_
+
+ _≅∗_ : UU (l1 ⊔ l2)
+ _≅∗_ = pointed-iso
+
+module _
+ {l1 l2 : Level} (A : Pointed-Type l1) (B : Pointed-Type l2)
+ (f : A ≅∗ B)
+ where
+
+ pointed-map-pointed-iso : A →∗ B
+ pointed-map-pointed-iso = pr1 f
+
+ map-pointed-iso : type-Pointed-Type A → type-Pointed-Type B
+ map-pointed-iso = map-pointed-map pointed-map-pointed-iso
+
+ preserves-point-pointed-iso :
+ map-pointed-iso (point-Pointed-Type A) = point-Pointed-Type B
+ preserves-point-pointed-iso =
+ preserves-point-pointed-map pointed-map-pointed-iso
+
+ is-pointed-iso-pointed-iso : is-pointed-iso pointed-map-pointed-iso
+ is-pointed-iso-pointed-iso = pr2 f
+
+ pointed-section-pointed-iso : pointed-section pointed-map-pointed-iso
+ pointed-section-pointed-iso =
+ pointed-section-is-pointed-iso (is-pointed-iso-pointed-iso)
+
+ pointed-retraction-pointed-iso : pointed-retraction pointed-map-pointed-iso
+ pointed-retraction-pointed-iso =
+ pointed-retraction-is-pointed-iso (is-pointed-iso-pointed-iso)
+
+ pointed-map-pointed-section-pointed-iso : B →∗ A
+ pointed-map-pointed-section-pointed-iso =
+ pointed-map-pointed-section-is-pointed-iso (is-pointed-iso-pointed-iso)
+
+ is-pointed-section-pointed-section-pointed-iso :
+ is-pointed-section
+ ( pointed-map-pointed-iso)
+ ( pointed-map-pointed-section-pointed-iso)
+ is-pointed-section-pointed-section-pointed-iso =
+ is-pointed-section-pointed-section-is-pointed-iso
+ ( is-pointed-iso-pointed-iso)
+
+ map-pointed-section-pointed-iso :
+ type-Pointed-Type B → type-Pointed-Type A
+ map-pointed-section-pointed-iso =
+ map-pointed-section-is-pointed-iso
+ ( is-pointed-iso-pointed-iso)
+
+ preserves-point-pointed-map-pointed-section-pointed-iso :
+ map-pointed-section-pointed-iso (point-Pointed-Type B) =
+ point-Pointed-Type A
+ preserves-point-pointed-map-pointed-section-pointed-iso =
+ preserves-point-pointed-map-pointed-section-is-pointed-iso
+ ( is-pointed-iso-pointed-iso)
+
+ is-section-pointed-section-pointed-iso :
+ is-section (map-pointed-iso) map-pointed-section-pointed-iso
+ is-section-pointed-section-pointed-iso =
+ is-section-pointed-section-is-pointed-iso
+ ( is-pointed-iso-pointed-iso)
+
+ section-pointed-iso :
+ section (map-pointed-iso)
+ section-pointed-iso =
+ section-is-pointed-iso
+ ( is-pointed-iso-pointed-iso)
+
+ coherence-point-is-section-pointed-section-pointed-iso :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( pointed-map-pointed-iso ∘∗ pointed-map-pointed-section-pointed-iso)
+ ( id-pointed-map)
+ ( is-section-pointed-section-pointed-iso)
+ coherence-point-is-section-pointed-section-pointed-iso =
+ coherence-point-is-section-pointed-section-is-pointed-iso
+ ( is-pointed-iso-pointed-iso)
+
+ pointed-map-pointed-retraction-pointed-iso : B →∗ A
+ pointed-map-pointed-retraction-pointed-iso =
+ pointed-map-pointed-retraction-is-pointed-iso
+ ( is-pointed-iso-pointed-iso)
+
+ is-pointed-retraction-pointed-retraction-pointed-iso :
+ is-pointed-retraction
+ ( pointed-map-pointed-iso)
+ ( pointed-map-pointed-retraction-pointed-iso)
+ is-pointed-retraction-pointed-retraction-pointed-iso =
+ is-pointed-retraction-pointed-retraction-is-pointed-iso
+ ( is-pointed-iso-pointed-iso)
+
+ map-pointed-retraction-pointed-iso :
+ type-Pointed-Type B → type-Pointed-Type A
+ map-pointed-retraction-pointed-iso =
+ map-pointed-retraction-is-pointed-iso
+ ( is-pointed-iso-pointed-iso)
+
+ preserves-point-pointed-map-pointed-retraction-pointed-iso :
+ map-pointed-retraction-pointed-iso (point-Pointed-Type B) =
+ point-Pointed-Type A
+ preserves-point-pointed-map-pointed-retraction-pointed-iso =
+ preserves-point-pointed-map-pointed-retraction-is-pointed-iso
+ ( is-pointed-iso-pointed-iso)
+
+ is-retraction-pointed-retraction-pointed-iso :
+ is-retraction
+ ( map-pointed-iso)
+ ( map-pointed-retraction-pointed-iso)
+ is-retraction-pointed-retraction-pointed-iso =
+ is-retraction-pointed-retraction-is-pointed-iso
+ ( is-pointed-iso-pointed-iso)
+
+ retraction-pointed-iso :
+ retraction (map-pointed-iso)
+ retraction-pointed-iso =
+ retraction-is-pointed-iso
+ ( is-pointed-iso-pointed-iso)
+
+ coherence-point-is-retraction-pointed-retraction-pointed-iso :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( pointed-map-pointed-retraction-pointed-iso ∘∗ pointed-map-pointed-iso)
+ ( id-pointed-map)
+ ( is-retraction-pointed-retraction-pointed-iso)
+ coherence-point-is-retraction-pointed-retraction-pointed-iso =
+ coherence-point-is-retraction-pointed-retraction-is-pointed-iso
+ ( is-pointed-iso-pointed-iso)
+```
+
+## Properties
+
+### Any pointed equivalence is a pointed isomorphism
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ where
+
+ abstract
+ is-pointed-iso-is-pointed-equiv :
+ is-pointed-equiv f → is-pointed-iso f
+ pr1 (pr1 (is-pointed-iso-is-pointed-equiv H)) =
+ pointed-map-inv-is-pointed-equiv f H
+ pr2 (pr1 (is-pointed-iso-is-pointed-equiv H)) =
+ is-pointed-section-pointed-map-inv-is-pointed-equiv f H
+ pr1 (pr2 (is-pointed-iso-is-pointed-equiv H)) =
+ pointed-map-inv-is-pointed-equiv f H
+ pr2 (pr2 (is-pointed-iso-is-pointed-equiv H)) =
+ is-pointed-retraction-pointed-map-inv-is-pointed-equiv f H
+```
+
+### Any pointed isomorphism is a pointed equivalence
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ where
+
+ abstract
+ is-pointed-equiv-is-pointed-iso :
+ is-pointed-iso f → is-pointed-equiv f
+ pr1 (is-pointed-equiv-is-pointed-iso H) = section-is-pointed-iso H
+ pr2 (is-pointed-equiv-is-pointed-iso H) = retraction-is-pointed-iso H
+```
+
+### Being a pointed isomorphism is a property
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ where
+
+ abstract
+ is-contr-section-is-pointed-equiv :
+ is-pointed-equiv f → is-contr (pointed-section f)
+ is-contr-section-is-pointed-equiv H =
+ is-torsorial-Eq-structure
+ ( is-contr-section-is-equiv H)
+ ( map-inv-is-equiv H , is-section-map-inv-is-equiv H)
+ ( is-contr-map-is-equiv
+ ( is-equiv-comp _ _
+ ( is-emb-is-equiv H _ _)
+ ( is-equiv-concat' _ (preserves-point-pointed-map f)))
+ ( _))
+
+ abstract
+ is-contr-retraction-is-pointed-equiv :
+ is-pointed-equiv f → is-contr (pointed-retraction f)
+ is-contr-retraction-is-pointed-equiv H =
+ is-torsorial-Eq-structure
+ ( is-contr-retraction-is-equiv H)
+ ( map-inv-is-equiv H , is-retraction-map-inv-is-equiv H)
+ ( is-contr-map-is-equiv
+ ( is-equiv-concat _ _)
+ ( _))
+
+ abstract
+ is-contr-is-pointed-iso-is-pointed-equiv :
+ is-pointed-equiv f → is-contr (is-pointed-iso f)
+ is-contr-is-pointed-iso-is-pointed-equiv H =
+ is-contr-product
+ ( is-contr-section-is-pointed-equiv H)
+ ( is-contr-retraction-is-pointed-equiv H)
+
+ abstract
+ is-prop-is-pointed-iso : is-prop (is-pointed-iso f)
+ is-prop-is-pointed-iso =
+ is-prop-is-proof-irrelevant
+ ( λ H →
+ is-contr-is-pointed-iso-is-pointed-equiv
+ ( is-pointed-equiv-is-pointed-iso f H))
+```
+
+### Being a pointed equivalence is equivalent to being a pointed isomorphism
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ where
+
+ equiv-is-pointed-iso-is-pointed-equiv :
+ is-pointed-equiv f ≃ (is-pointed-iso f)
+ pr1 equiv-is-pointed-iso-is-pointed-equiv =
+ is-pointed-iso-is-pointed-equiv f
+ pr2 equiv-is-pointed-iso-is-pointed-equiv =
+ is-equiv-is-prop
+ ( is-property-is-equiv (map-pointed-map f))
+ ( is-prop-is-pointed-iso f)
+ ( is-pointed-equiv-is-pointed-iso f)
+```
diff --git a/src/structured-types/pointed-maps.lagda.md b/src/structured-types/pointed-maps.lagda.md
index ec9b8563f6..1a8de69535 100644
--- a/src/structured-types/pointed-maps.lagda.md
+++ b/src/structured-types/pointed-maps.lagda.md
@@ -28,6 +28,9 @@ open import structured-types.pointed-types
A pointed map from a pointed type `A` to a pointed type `B` is a base point
preserving function from `A` to `B`.
+The type `A →∗ B` of pointed maps from `A` to `B` is itself pointed by the
+[constant pointed map](structured-types.constant-pointed-maps.md).
+
## Definitions
### Pointed maps
@@ -50,16 +53,6 @@ and pointed type constructions in general, is the
the [asterisk](https://codepoints.net/U+002A) `*`.
```agda
- constant-pointed-map : (A : Pointed-Type l1) (B : Pointed-Type l2) → A →∗ B
- pr1 (constant-pointed-map A B) =
- const (type-Pointed-Type A) (type-Pointed-Type B) (point-Pointed-Type B)
- pr2 (constant-pointed-map A B) = refl
-
- pointed-map-Pointed-Type :
- Pointed-Type l1 → Pointed-Type l2 → Pointed-Type (l1 ⊔ l2)
- pr1 (pointed-map-Pointed-Type A B) = pointed-map A B
- pr2 (pointed-map-Pointed-Type A B) = constant-pointed-map A B
-
module _
{l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
where
@@ -128,18 +121,10 @@ module _
pr1 (comp-pointed-map g f) = map-comp-pointed-map g f
pr2 (comp-pointed-map g f) = preserves-point-comp-pointed-map g f
-precomp-pointed-map :
- {l1 l2 l3 : Level}
- {A : Pointed-Type l1} {B : Pointed-Type l2} (C : Pointed-Type l3) →
- A →∗ B → B →∗ C → A →∗ C
-precomp-pointed-map C f g = comp-pointed-map g f
+ infixr 15 _∘∗_
-infixr 15 _∘∗_
-_∘∗_ :
- {l1 l2 l3 : Level}
- {A : Pointed-Type l1} {B : Pointed-Type l2} {C : Pointed-Type l3} →
- B →∗ C → A →∗ B → A →∗ C
-_∘∗_ {A = A} {B} {C} = comp-pointed-map
+ _∘∗_ : B →∗ C → A →∗ B → A →∗ C
+ _∘∗_ = comp-pointed-map
```
### The identity pointed map
@@ -153,3 +138,8 @@ module _
pr1 id-pointed-map = id
pr2 id-pointed-map = refl
```
+
+## See also
+
+- [Constant pointed maps](structured-types.constant-pointed-maps.md)
+- [Precomposition of pointed maps](structured-types.precomposition-pointed-maps.md)
diff --git a/src/structured-types/pointed-retractions.lagda.md b/src/structured-types/pointed-retractions.lagda.md
new file mode 100644
index 0000000000..d4b837ac66
--- /dev/null
+++ b/src/structured-types/pointed-retractions.lagda.md
@@ -0,0 +1,178 @@
+# Pointed retractions of pointed maps
+
+```agda
+module structured-types.pointed-retractions where
+```
+
+Imports
+
+```agda
+open import foundation.action-on-identifications-functions
+open import foundation.commuting-squares-of-identifications
+open import foundation.dependent-pair-types
+open import foundation.identity-types
+open import foundation.universe-levels
+
+open import foundation-core.contractible-maps
+open import foundation-core.contractible-types
+open import foundation-core.retractions
+
+open import structured-types.pointed-homotopies
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+```
+
+
+
+## Idea
+
+A
+{{#concept "pointed retraction" Disambiguation="pointed map" Agda=pointed-retraction}}
+of a [pointed map](structured-types.pointed-maps.md) `f : A →∗ B` consists of a
+pointed map `g : B →∗ A` equipped with a
+[pointed homotopy](structured-types.pointed-homotopies.md) `H : g ∘∗ f ~∗ id`.
+
+## Definitions
+
+### The predicate of being a pointed retraction of a pointed map
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ where
+
+ is-pointed-retraction : (B →∗ A) → UU l1
+ is-pointed-retraction g = g ∘∗ f ~∗ id-pointed-map
+```
+
+### The type of pointed retractions of a pointed map
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ where
+
+ pointed-retraction : UU (l1 ⊔ l2)
+ pointed-retraction =
+ Σ (B →∗ A) (is-pointed-retraction f)
+
+ module _
+ (r : pointed-retraction)
+ where
+
+ pointed-map-pointed-retraction : B →∗ A
+ pointed-map-pointed-retraction = pr1 r
+
+ is-pointed-retraction-pointed-retraction :
+ is-pointed-retraction f pointed-map-pointed-retraction
+ is-pointed-retraction-pointed-retraction = pr2 r
+
+ map-pointed-retraction : type-Pointed-Type B → type-Pointed-Type A
+ map-pointed-retraction = map-pointed-map pointed-map-pointed-retraction
+
+ preserves-point-pointed-map-pointed-retraction :
+ map-pointed-retraction (point-Pointed-Type B) = point-Pointed-Type A
+ preserves-point-pointed-map-pointed-retraction =
+ preserves-point-pointed-map pointed-map-pointed-retraction
+
+ is-retraction-pointed-retraction :
+ is-retraction (map-pointed-map f) map-pointed-retraction
+ is-retraction-pointed-retraction =
+ htpy-pointed-htpy is-pointed-retraction-pointed-retraction
+
+ retraction-pointed-retraction : retraction (map-pointed-map f)
+ pr1 retraction-pointed-retraction = map-pointed-retraction
+ pr2 retraction-pointed-retraction = is-retraction-pointed-retraction
+
+ coherence-point-is-retraction-pointed-retraction :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( pointed-map-pointed-retraction ∘∗ f)
+ ( id-pointed-map)
+ ( is-retraction-pointed-retraction)
+ coherence-point-is-retraction-pointed-retraction =
+ coherence-point-pointed-htpy is-pointed-retraction-pointed-retraction
+```
+
+## Properties
+
+### Any retraction of a pointed map preserves the base point in a unique way making the retracting homotopy pointed
+
+Consider a [retraction](foundation-core.retractions.md) `g : B → A` of a pointed
+map `f := (f₀ , f₁) : A →∗ B`. Then `g` is base point preserving.
+
+**Proof.** Our goal is to show that `g * = *`. Since `f` is pointed, we have
+`f * = *` and hence
+
+```text
+ (ap g f₁)⁻¹ H *
+ g * -------------> g (f₀ *) -------> *.
+```
+
+In order to show that the retracting homotopy `H : g ∘ f₀ ~ id` is pointed, we
+have to show that the triangle of identifications
+
+```text
+ H *
+ g (f₀ *) -----> *
+ \ /
+ ap g f₁ ∙ ((ap g f₁)⁻¹ ∙ H *) \ / refl
+ \ /
+ ∨ ∨
+ *
+```
+
+commutes. This follows by the fact that concatenating with an inverse
+identification is inverse to concatenating with the original identification, and
+the right unit law of concatenation.
+
+Note that the pointing of `g` chosen above is the unique way making the
+retracting homotopy pointed, because the map `p ↦ ap g f₁ ∙ p` is an equivalence
+with a contractible fiber at `H * ∙ refl`.
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ (g : type-Pointed-Type B → type-Pointed-Type A)
+ (H : is-retraction (map-pointed-map f) g)
+ where
+
+ abstract
+ uniquely-preserves-point-is-retraction-pointed-map :
+ is-contr
+ ( Σ ( g (point-Pointed-Type B) = point-Pointed-Type A)
+ ( coherence-square-identifications
+ ( H (point-Pointed-Type A))
+ ( ap g (preserves-point-pointed-map f))
+ ( refl)))
+ uniquely-preserves-point-is-retraction-pointed-map =
+ is-contr-map-is-equiv
+ ( is-equiv-concat (ap g (preserves-point-pointed-map f)) _)
+ ( H (point-Pointed-Type A) ∙ refl)
+
+ preserves-point-is-retraction-pointed-map :
+ g (point-Pointed-Type B) = point-Pointed-Type A
+ preserves-point-is-retraction-pointed-map =
+ inv (ap g (preserves-point-pointed-map f)) ∙ H (point-Pointed-Type A)
+
+ pointed-map-is-retraction-pointed-map :
+ B →∗ A
+ pr1 pointed-map-is-retraction-pointed-map = g
+ pr2 pointed-map-is-retraction-pointed-map =
+ preserves-point-is-retraction-pointed-map
+
+ coherence-point-is-retraction-pointed-map :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( pointed-map-is-retraction-pointed-map ∘∗ f)
+ ( id-pointed-map)
+ ( H)
+ coherence-point-is-retraction-pointed-map =
+ ( is-section-inv-concat (ap g (preserves-point-pointed-map f)) _) ∙
+ ( inv right-unit)
+
+ is-pointed-retraction-is-retraction-pointed-map :
+ is-pointed-retraction f pointed-map-is-retraction-pointed-map
+ pr1 is-pointed-retraction-is-retraction-pointed-map =
+ H
+ pr2 is-pointed-retraction-is-retraction-pointed-map =
+ coherence-point-is-retraction-pointed-map
+```
diff --git a/src/structured-types/pointed-sections.lagda.md b/src/structured-types/pointed-sections.lagda.md
index b4e1863dcc..a3dfdb69f3 100644
--- a/src/structured-types/pointed-sections.lagda.md
+++ b/src/structured-types/pointed-sections.lagda.md
@@ -8,6 +8,8 @@ module structured-types.pointed-sections where
```agda
open import foundation.dependent-pair-types
+open import foundation.identity-types
+open import foundation.sections
open import foundation.universe-levels
open import structured-types.pointed-homotopies
@@ -19,14 +21,69 @@ open import structured-types.pointed-types
## Idea
-A **pointed section** of a pointed map `f : A →∗ B` consists of a pointed map
-`g : B →∗ A` equipped with a pointed homotopy `H : (f ∘∗ g) ~∗ id`.
+A
+{{#concept "pointed section" Disambiguation="pointed map" Agda=pointed-section}}
+of a [pointed map](structured-types.pointed-maps.md) `f : A →∗ B` consists of a
+pointed map `g : B →∗ A` equipped with a
+[pointed homotopy](structured-types.pointed-homotopies.md) `H : f ∘∗ g ~∗ id`.
+
+## Definitions
+
+### The predicate of being a pointed section of a pointed map
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ where
+
+ is-pointed-section : (B →∗ A) → UU l2
+ is-pointed-section g = f ∘∗ g ~∗ id-pointed-map
+```
+
+### The type of pointed sections of a pointed map
```agda
-pointed-section-Pointed-Type :
- {l1 l2 : Level} (A : Pointed-Type l1) (B : Pointed-Type l2) →
- (A →∗ B) → UU (l1 ⊔ l2)
-pointed-section-Pointed-Type A B f =
- Σ ( B →∗ A)
- ( λ g → (f ∘∗ g) ~∗ id-pointed-map)
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ where
+
+ pointed-section : UU (l1 ⊔ l2)
+ pointed-section =
+ Σ (B →∗ A) (is-pointed-section f)
+
+ module _
+ (s : pointed-section)
+ where
+
+ pointed-map-pointed-section : B →∗ A
+ pointed-map-pointed-section = pr1 s
+
+ is-pointed-section-pointed-section :
+ is-pointed-section f pointed-map-pointed-section
+ is-pointed-section-pointed-section = pr2 s
+
+ map-pointed-section : type-Pointed-Type B → type-Pointed-Type A
+ map-pointed-section = map-pointed-map pointed-map-pointed-section
+
+ preserves-point-pointed-map-pointed-section :
+ map-pointed-section (point-Pointed-Type B) = point-Pointed-Type A
+ preserves-point-pointed-map-pointed-section =
+ preserves-point-pointed-map pointed-map-pointed-section
+
+ is-section-pointed-section :
+ is-section (map-pointed-map f) map-pointed-section
+ is-section-pointed-section =
+ htpy-pointed-htpy is-pointed-section-pointed-section
+
+ section-pointed-section : section (map-pointed-map f)
+ pr1 section-pointed-section = map-pointed-section
+ pr2 section-pointed-section = is-section-pointed-section
+
+ coherence-point-is-section-pointed-section :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( f ∘∗ pointed-map-pointed-section)
+ ( id-pointed-map)
+ ( is-section-pointed-section)
+ coherence-point-is-section-pointed-section =
+ coherence-point-pointed-htpy is-pointed-section-pointed-section
```
diff --git a/src/structured-types/pointed-span-diagrams.lagda.md b/src/structured-types/pointed-span-diagrams.lagda.md
new file mode 100644
index 0000000000..95f98ce716
--- /dev/null
+++ b/src/structured-types/pointed-span-diagrams.lagda.md
@@ -0,0 +1,259 @@
+# Pointed span diagrams
+
+```agda
+module structured-types.pointed-span-diagrams where
+```
+
+Imports
+
+```agda
+open import foundation.dependent-pair-types
+open import foundation.identity-types
+open import foundation.morphisms-arrows
+open import foundation.universe-levels
+
+open import structured-types.morphisms-pointed-arrows
+open import structured-types.pointed-maps
+open import structured-types.pointed-spans
+open import structured-types.pointed-types
+```
+
+
+
+## Idea
+
+A {{#concept "(binary) pointed span diagram" Agda=pointed-span-diagram}} is a
+diagram of [pointed maps](structured-types.pointed-maps.md) of the form
+
+```text
+ f g
+ A <----- S -----> B.
+```
+
+In other words, a pointed span diagram consists of two
+[pointed types](structured-types.pointed-types.md) `A` and `B` and a
+[pointed span](structured-types.pointed-spans.md) from `A` to `B`.
+
+### (Binary) span diagrams of pointed types
+
+```agda
+pointed-span-diagram :
+ (l1 l2 l3 : Level) → UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3)
+pointed-span-diagram l1 l2 l3 =
+ Σ ( Pointed-Type l1)
+ ( λ A → Σ (Pointed-Type l2) (pointed-span l3 A))
+
+module _
+ {l1 l2 l3 : Level} {S : Pointed-Type l1}
+ {A : Pointed-Type l2} {B : Pointed-Type l3}
+ where
+
+ make-pointed-span-diagram :
+ (S →∗ A) → (S →∗ B) → pointed-span-diagram l2 l3 l1
+ make-pointed-span-diagram f g = (A , B , S , f , g)
+
+module _
+ {l1 l2 l3 : Level} (𝒮 : pointed-span-diagram l1 l2 l3)
+ where
+
+ pointed-domain-pointed-span-diagram : Pointed-Type l1
+ pointed-domain-pointed-span-diagram = pr1 𝒮
+
+ domain-pointed-span-diagram : UU l1
+ domain-pointed-span-diagram =
+ type-Pointed-Type pointed-domain-pointed-span-diagram
+
+ point-domain-pointed-span-diagram :
+ domain-pointed-span-diagram
+ point-domain-pointed-span-diagram =
+ point-Pointed-Type pointed-domain-pointed-span-diagram
+
+ pointed-codomain-pointed-span-diagram : Pointed-Type l2
+ pointed-codomain-pointed-span-diagram = pr1 (pr2 𝒮)
+
+ codomain-pointed-span-diagram : UU l2
+ codomain-pointed-span-diagram =
+ type-Pointed-Type pointed-codomain-pointed-span-diagram
+
+ point-codomain-pointed-span-diagram :
+ codomain-pointed-span-diagram
+ point-codomain-pointed-span-diagram =
+ point-Pointed-Type pointed-codomain-pointed-span-diagram
+
+ pointed-span-pointed-span-diagram :
+ pointed-span l3
+ ( pointed-domain-pointed-span-diagram)
+ ( pointed-codomain-pointed-span-diagram)
+ pointed-span-pointed-span-diagram = pr2 (pr2 𝒮)
+
+ spanning-pointed-type-pointed-span-diagram : Pointed-Type l3
+ spanning-pointed-type-pointed-span-diagram =
+ spanning-pointed-type-pointed-span
+ ( pointed-span-pointed-span-diagram)
+
+ spanning-type-pointed-span-diagram : UU l3
+ spanning-type-pointed-span-diagram =
+ type-Pointed-Type spanning-pointed-type-pointed-span-diagram
+
+ point-spanning-type-pointed-span-diagram :
+ spanning-type-pointed-span-diagram
+ point-spanning-type-pointed-span-diagram =
+ point-Pointed-Type spanning-pointed-type-pointed-span-diagram
+
+ left-pointed-map-pointed-span-diagram :
+ spanning-pointed-type-pointed-span-diagram →∗
+ pointed-domain-pointed-span-diagram
+ left-pointed-map-pointed-span-diagram =
+ left-pointed-map-pointed-span
+ ( pointed-span-pointed-span-diagram)
+
+ left-map-pointed-span-diagram :
+ spanning-type-pointed-span-diagram → domain-pointed-span-diagram
+ left-map-pointed-span-diagram =
+ left-map-pointed-span
+ ( pointed-span-pointed-span-diagram)
+
+ preserves-point-left-map-pointed-span-diagram :
+ left-map-pointed-span-diagram
+ ( point-spanning-type-pointed-span-diagram) =
+ point-domain-pointed-span-diagram
+ preserves-point-left-map-pointed-span-diagram =
+ preserves-point-left-map-pointed-span
+ ( pointed-span-pointed-span-diagram)
+
+ right-pointed-map-pointed-span-diagram :
+ spanning-pointed-type-pointed-span-diagram →∗
+ pointed-codomain-pointed-span-diagram
+ right-pointed-map-pointed-span-diagram =
+ right-pointed-map-pointed-span
+ ( pointed-span-pointed-span-diagram)
+
+ right-map-pointed-span-diagram :
+ spanning-type-pointed-span-diagram → codomain-pointed-span-diagram
+ right-map-pointed-span-diagram =
+ right-map-pointed-span
+ ( pointed-span-pointed-span-diagram)
+
+ preserves-point-right-map-pointed-span-diagram :
+ right-map-pointed-span-diagram
+ ( point-spanning-type-pointed-span-diagram) =
+ point-codomain-pointed-span-diagram
+ preserves-point-right-map-pointed-span-diagram =
+ preserves-point-right-map-pointed-span
+ ( pointed-span-pointed-span-diagram)
+```
+
+### The pointed span diagram obtained from a morphism of pointed arrows
+
+Given pointed maps `f : A →∗ B` and `g : X →∗ Y` and a morphism of pointed
+arrows `α : f →∗ g`, the pointed span diagram associated to `α` is the pointed
+span diagram
+
+```text
+ f α₀
+ B <----- A -----> X.
+```
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {X : Pointed-Type l3} {Y : Pointed-Type l4}
+ (f : A →∗ B) (g : X →∗ Y) (α : hom-pointed-arrow f g)
+ where
+
+ domain-span-diagram-hom-pointed-arrow : Pointed-Type l2
+ domain-span-diagram-hom-pointed-arrow = B
+
+ type-domain-span-diagram-hom-pointed-arrow : UU l2
+ type-domain-span-diagram-hom-pointed-arrow =
+ type-Pointed-Type domain-span-diagram-hom-pointed-arrow
+
+ point-domain-span-diagram-hom-pointed-arrow :
+ type-domain-span-diagram-hom-pointed-arrow
+ point-domain-span-diagram-hom-pointed-arrow =
+ point-Pointed-Type domain-span-diagram-hom-pointed-arrow
+
+ codomain-span-diagram-hom-pointed-arrow : Pointed-Type l3
+ codomain-span-diagram-hom-pointed-arrow = X
+
+ type-codomain-span-diagram-hom-pointed-arrow : UU l3
+ type-codomain-span-diagram-hom-pointed-arrow =
+ type-Pointed-Type codomain-span-diagram-hom-pointed-arrow
+
+ point-codomain-span-diagram-hom-pointed-arrow :
+ type-codomain-span-diagram-hom-pointed-arrow
+ point-codomain-span-diagram-hom-pointed-arrow =
+ point-Pointed-Type codomain-span-diagram-hom-pointed-arrow
+
+ pointed-spanning-type-hom-pointed-arrow : Pointed-Type l1
+ pointed-spanning-type-hom-pointed-arrow = A
+
+ spanning-type-hom-pointed-arrow : UU l1
+ spanning-type-hom-pointed-arrow =
+ type-Pointed-Type pointed-spanning-type-hom-pointed-arrow
+
+ point-spanning-type-hom-pointed-arrow :
+ spanning-type-hom-pointed-arrow
+ point-spanning-type-hom-pointed-arrow =
+ point-Pointed-Type pointed-spanning-type-hom-pointed-arrow
+
+ left-pointed-map-span-diagram-hom-pointed-arrow :
+ pointed-spanning-type-hom-pointed-arrow →∗
+ domain-span-diagram-hom-pointed-arrow
+ left-pointed-map-span-diagram-hom-pointed-arrow = f
+
+ left-map-span-diagram-hom-pointed-arrow :
+ spanning-type-hom-pointed-arrow → type-domain-span-diagram-hom-pointed-arrow
+ left-map-span-diagram-hom-pointed-arrow =
+ map-pointed-map left-pointed-map-span-diagram-hom-pointed-arrow
+
+ preserves-point-left-map-span-diagram-hom-pointed-arrow :
+ left-map-span-diagram-hom-pointed-arrow
+ ( point-spanning-type-hom-pointed-arrow) =
+ point-domain-span-diagram-hom-pointed-arrow
+ preserves-point-left-map-span-diagram-hom-pointed-arrow =
+ preserves-point-pointed-map
+ ( left-pointed-map-span-diagram-hom-pointed-arrow)
+
+ right-pointed-map-span-diagram-hom-pointed-arrow :
+ pointed-spanning-type-hom-pointed-arrow →∗
+ codomain-span-diagram-hom-pointed-arrow
+ right-pointed-map-span-diagram-hom-pointed-arrow =
+ pointed-map-domain-hom-pointed-arrow f g α
+
+ right-map-span-diagram-hom-pointed-arrow :
+ spanning-type-hom-pointed-arrow →
+ type-codomain-span-diagram-hom-pointed-arrow
+ right-map-span-diagram-hom-pointed-arrow =
+ map-pointed-map right-pointed-map-span-diagram-hom-pointed-arrow
+
+ preserves-point-right-map-span-diagram-hom-pointed-arrow :
+ right-map-span-diagram-hom-pointed-arrow
+ ( point-spanning-type-hom-pointed-arrow) =
+ point-codomain-span-diagram-hom-pointed-arrow
+ preserves-point-right-map-span-diagram-hom-pointed-arrow =
+ preserves-point-pointed-map
+ ( right-pointed-map-span-diagram-hom-pointed-arrow)
+
+ span-hom-pointed-arrow :
+ pointed-span l1 B X
+ pr1 span-hom-pointed-arrow =
+ A
+ pr1 (pr2 span-hom-pointed-arrow) =
+ left-pointed-map-span-diagram-hom-pointed-arrow
+ pr2 (pr2 span-hom-pointed-arrow) =
+ right-pointed-map-span-diagram-hom-pointed-arrow
+
+ span-diagram-hom-pointed-arrow : pointed-span-diagram l2 l3 l1
+ pr1 span-diagram-hom-pointed-arrow =
+ domain-span-diagram-hom-pointed-arrow
+ pr1 (pr2 span-diagram-hom-pointed-arrow) =
+ codomain-span-diagram-hom-pointed-arrow
+ pr2 (pr2 span-diagram-hom-pointed-arrow) =
+ span-hom-pointed-arrow
+```
+
+## See also
+
+- [Transposition of pointed span diagrams](structured-types.transposition-pointed-span-diagrams.md)
diff --git a/src/structured-types/pointed-spans.lagda.md b/src/structured-types/pointed-spans.lagda.md
new file mode 100644
index 0000000000..125e01bee5
--- /dev/null
+++ b/src/structured-types/pointed-spans.lagda.md
@@ -0,0 +1,115 @@
+# Pointed spans
+
+```agda
+module structured-types.pointed-spans where
+```
+
+Imports
+
+```agda
+open import foundation.cartesian-product-types
+open import foundation.dependent-pair-types
+open import foundation.identity-types
+open import foundation.spans
+open import foundation.universe-levels
+
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+```
+
+
+
+## Idea
+
+Consider two [pointed types](structured-types.pointed-types.md) `A` and `B`. A
+{{#concept "(binary) pointed span" Agda=pointed-span}} from `A` to `B` consists
+of a
+{{#concept "spanning pointed type" Disambiguation="binary pointed span" Agda=spanning-pointed-type-pointed-span}}
+`S` and a [pair](foundation.dependent-pair-types.md) of
+[pointed maps](structured-types.pointed-maps.md) `f : S →∗ A` and `g : S →∗ B`.
+The pointed types `A` and `B` in the specification of a binary span of pointed
+types are also referred to as the
+{{#concept "domain" Disambiguation="binary pointed span"}} and
+{{#concept "codomain" Disambiguation="binary pointed span"}} of the pointed
+span, respectively.
+
+## Definitions
+
+### (Binary) pointed spans
+
+```agda
+pointed-span :
+ {l1 l2 : Level} (l : Level) (A : Pointed-Type l1) (B : Pointed-Type l2) →
+ UU (l1 ⊔ l2 ⊔ lsuc l)
+pointed-span l A B = Σ (Pointed-Type l) (λ S → (S →∗ A) × (S →∗ B))
+
+module _
+ {l1 l2 l3 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ (𝒮 : pointed-span l3 A B)
+ where
+
+ spanning-pointed-type-pointed-span : Pointed-Type l3
+ spanning-pointed-type-pointed-span = pr1 𝒮
+
+ spanning-type-pointed-span : UU l3
+ spanning-type-pointed-span =
+ type-Pointed-Type spanning-pointed-type-pointed-span
+
+ point-spanning-type-pointed-span : spanning-type-pointed-span
+ point-spanning-type-pointed-span =
+ point-Pointed-Type spanning-pointed-type-pointed-span
+
+ left-pointed-map-pointed-span :
+ spanning-pointed-type-pointed-span →∗ A
+ left-pointed-map-pointed-span = pr1 (pr2 𝒮)
+
+ left-map-pointed-span :
+ spanning-type-pointed-span → type-Pointed-Type A
+ left-map-pointed-span =
+ map-pointed-map left-pointed-map-pointed-span
+
+ preserves-point-left-map-pointed-span :
+ left-map-pointed-span point-spanning-type-pointed-span =
+ point-Pointed-Type A
+ preserves-point-left-map-pointed-span =
+ preserves-point-pointed-map left-pointed-map-pointed-span
+
+ right-pointed-map-pointed-span :
+ spanning-pointed-type-pointed-span →∗ B
+ right-pointed-map-pointed-span = pr2 (pr2 𝒮)
+
+ right-map-pointed-span :
+ spanning-type-pointed-span → type-Pointed-Type B
+ right-map-pointed-span =
+ map-pointed-map right-pointed-map-pointed-span
+
+ preserves-point-right-map-pointed-span :
+ right-map-pointed-span point-spanning-type-pointed-span =
+ point-Pointed-Type B
+ preserves-point-right-map-pointed-span =
+ preserves-point-pointed-map right-pointed-map-pointed-span
+
+ span-pointed-span : span l3 (type-Pointed-Type A) (type-Pointed-Type B)
+ pr1 span-pointed-span = spanning-type-pointed-span
+ pr1 (pr2 span-pointed-span) = left-map-pointed-span
+ pr2 (pr2 span-pointed-span) = right-map-pointed-span
+```
+
+### Identity pointed spans
+
+```agda
+module _
+ {l1 : Level} {X : Pointed-Type l1}
+ where
+
+ id-pointed-span : pointed-span l1 X X
+ pr1 id-pointed-span = X
+ pr1 (pr2 id-pointed-span) = id-pointed-map
+ pr2 (pr2 id-pointed-span) = id-pointed-map
+```
+
+## See also
+
+- [Opposite pointed spans](structured-types.opposite-pointed-spans.md)
+- [Pointed span diagrams](structured-types.pointed-span-diagrams.md)
+- [Spans](foundation.spans.md)
diff --git a/src/structured-types/postcomposition-pointed-maps.lagda.md b/src/structured-types/postcomposition-pointed-maps.lagda.md
new file mode 100644
index 0000000000..03629c2259
--- /dev/null
+++ b/src/structured-types/postcomposition-pointed-maps.lagda.md
@@ -0,0 +1,40 @@
+# Postcomposition of pointed maps
+
+```agda
+module structured-types.postcomposition-pointed-maps where
+```
+
+Imports
+
+```agda
+open import foundation.universe-levels
+
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+```
+
+
+
+## Idea
+
+The
+{{#concept "postcomposition operation" Disambiguation="pointed maps" Agda=postcomp-pointed-map}}
+on [pointed maps](structured-types.pointed-maps.md) by a pointed map
+`f : A →∗ B` is a family of operations
+
+```text
+ f ∘∗ - : (X →∗ A) → (X →∗ B)
+```
+
+indexed by a [pointed type](structured-types.pointed-types.md) `X`.
+
+## Definitions
+
+### Postcomposition by pointed maps
+
+```agda
+postcomp-pointed-map :
+ {l1 l2 l3 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ (X : Pointed-Type l3) → (X →∗ A) → (X →∗ B)
+postcomp-pointed-map f X g = comp-pointed-map f g
+```
diff --git a/src/structured-types/precomposition-pointed-maps.lagda.md b/src/structured-types/precomposition-pointed-maps.lagda.md
new file mode 100644
index 0000000000..da91cbc088
--- /dev/null
+++ b/src/structured-types/precomposition-pointed-maps.lagda.md
@@ -0,0 +1,40 @@
+# Precomposition of pointed maps
+
+```agda
+module structured-types.precomposition-pointed-maps where
+```
+
+Imports
+
+```agda
+open import foundation.universe-levels
+
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+```
+
+
+
+## Idea
+
+The
+{{#concept "precomposition operation" Disambiguation="pointed maps" Agda=precomp-pointed-map}}
+on [pointed maps](structured-types.pointed-maps.md) by a pointed map
+`f : A →∗ B` is a family of operations
+
+```text
+ - ∘∗ f : (B →∗ C) → (A →∗ C)
+```
+
+indexed by a [pointed type](structured-types.pointed-types.md) `C`.
+
+## Definitions
+
+### Precomposition by pointed maps
+
+```agda
+precomp-pointed-map :
+ {l1 l2 l3 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ (C : Pointed-Type l3) → (B →∗ C) → (A →∗ C)
+precomp-pointed-map f C g = comp-pointed-map g f
+```
diff --git a/src/structured-types/transposition-pointed-span-diagrams.lagda.md b/src/structured-types/transposition-pointed-span-diagrams.lagda.md
new file mode 100644
index 0000000000..3b21a8cd57
--- /dev/null
+++ b/src/structured-types/transposition-pointed-span-diagrams.lagda.md
@@ -0,0 +1,59 @@
+# Transposition of pointed span diagrams
+
+```agda
+module structured-types.transposition-pointed-span-diagrams where
+```
+
+Imports
+
+```agda
+open import foundation.dependent-pair-types
+open import foundation.universe-levels
+
+open import structured-types.opposite-pointed-spans
+open import structured-types.pointed-span-diagrams
+```
+
+
+
+## Idea
+
+The
+{{#concept "trasposition" Disambiguation="pointed span diagram" Agda=transposition-pointed-span-diagram}}
+of a [pointed span diagram](structured-types.pointed-span-diagrams.md)
+
+```text
+ f g
+ A <----- S -----> B
+```
+
+is the pointed span diagram
+
+```text
+ g f
+ B <----- S -----> A.
+```
+
+In other words, the transposition of a pointed span diagram `(A , B , s)` is the
+pointed span diagram `(B , A , opposite-pointed-span s)` where
+`opposite-pointed-span s` is the
+[opposite](structured-types.opposite-pointed-spans.md) of the
+[pointed span](structured-types.pointed-spans.md) `s` from `A` to `B`.
+
+## Definitions
+
+### Transposition of pointed span diagrams
+
+```agda
+module _
+ {l1 l2 l3 : Level} (𝒮 : pointed-span-diagram l1 l2 l3)
+ where
+
+ transposition-pointed-span-diagram : pointed-span-diagram l2 l1 l3
+ pr1 transposition-pointed-span-diagram =
+ pointed-codomain-pointed-span-diagram 𝒮
+ pr1 (pr2 transposition-pointed-span-diagram) =
+ pointed-domain-pointed-span-diagram 𝒮
+ pr2 (pr2 transposition-pointed-span-diagram) =
+ opposite-pointed-span (pointed-span-pointed-span-diagram 𝒮)
+```
diff --git a/src/structured-types/uniform-pointed-homotopies.lagda.md b/src/structured-types/uniform-pointed-homotopies.lagda.md
new file mode 100644
index 0000000000..eb98f22918
--- /dev/null
+++ b/src/structured-types/uniform-pointed-homotopies.lagda.md
@@ -0,0 +1,317 @@
+# Uniform pointed homotopies
+
+```agda
+module structured-types.uniform-pointed-homotopies where
+```
+
+Imports
+
+```agda
+open import foundation.commuting-triangles-of-identifications
+open import foundation.dependent-pair-types
+open import foundation.equivalences
+open import foundation.function-extensionality
+open import foundation.function-types
+open import foundation.functoriality-dependent-pair-types
+open import foundation.homotopies
+open import foundation.identity-types
+open import foundation.structure-identity-principle
+open import foundation.universe-levels
+
+open import structured-types.pointed-dependent-functions
+open import structured-types.pointed-families-of-types
+open import structured-types.pointed-homotopies
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+```
+
+
+
+## Idea
+
+The concept of _uniform pointed homotopy_ is an
+[equivalent](foundation-core.equivalences.md) way of defining
+[pointed homotopies](structured-types.pointed-homotopies.md). A uniform pointed
+homotopy `H` between two
+[pointed dependent functions](structured-types.pointed-dependent-functions.md)
+`f` and `g` is defined to be a pointed dependent function of the
+[pointed type family](structured-types.pointed-families-of-types.md) of
+[identifications](foundation-core.identity-types.md) between the values of `f`
+and `g`. The main idea is that, since uniform pointed homotopies between pointed
+dependent functions are again pointed dependent functions, we can easily
+consider uniform pointed homotopies between uniform pointed homotopies and so
+on. The definition of uniform pointed homotopies is uniform in the sense that
+they can be iterated in this way. We now give a more detailed description of the
+definition.
+
+Consider two pointed dependent functions `f := (f₀ , f₁)` and `g := (g₀ , g₁)`
+in the pointed dependent function type `Π∗ A B`. Then the type family
+`x ↦ f₀ x = g₀ x` over the base type `A` is a pointed type family, where the
+base point is the identification
+
+```text
+ f₁ ∙ inv g₁ : f₀ * = g₀ *.
+```
+
+A {{#concept "uniform pointed homotopy" Agda=uniform-pointed-htpy}} from `f` to
+`g` is defined to be a
+[pointed dependent function](structured-types.pointed-dependent-functions.md) of
+the pointed type family `x ↦ f₀ x = g₀ x`. In other words, a pointed dependent
+function consists of an unpointed [homotopy](foundation-core.homotopies.md)
+`H₀ : f₀ ~ g₀` between the underlying dependent functions and an identification
+witnessing that the triangle of identifications
+
+```text
+ H₀ *
+ f₀ * ------> g₀ *
+ \ ∧
+ f₁ \ / inv g₁
+ \ /
+ ∨ /
+ *
+```
+
+[commutes](foundation.commuting-triangles-of-identifications.md).
+
+Notice that in comparison to the pointed homotopies, the identification on the
+right in this triangle goes up, in the inverse direction of the identification
+`g₁`. This makes it slightly more complicated to construct an identification
+witnessing that the triangle commutes in the case of uniform pointed homotopies.
+Furthermore, this complication becomes more significant and bothersome when we
+are trying to construct a
+[pointed `2`-homotopy](structured-types.pointed-2-homotopies.md).
+
+## Definitions
+
+### Preservation of the base point of unpointed homotopies between pointed maps
+
+The underlying homotopy of a uniform pointed homotopy preserves the base point
+in the sense that the triangle of identifications
+
+```text
+ H *
+ f * ------> g *
+ \ ∧
+ preserves-point f \ / inv (preserves-point g)
+ \ /
+ ∨ /
+ *
+```
+
+commutes.
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ (f g : pointed-Π A B) (G : unpointed-htpy-pointed-Π f g)
+ where
+
+ preserves-point-unpointed-htpy-pointed-Π : UU l2
+ preserves-point-unpointed-htpy-pointed-Π =
+ coherence-triangle-identifications
+ ( G (point-Pointed-Type A))
+ ( inv (preserves-point-function-pointed-Π g))
+ ( preserves-point-function-pointed-Π f)
+
+ compute-coherence-point-unpointed-htpy-pointed-Π :
+ coherence-point-unpointed-htpy-pointed-Π f g G ≃
+ preserves-point-unpointed-htpy-pointed-Π
+ compute-coherence-point-unpointed-htpy-pointed-Π =
+ equiv-transpose-right-coherence-triangle-identifications _ _ _
+
+ preserves-point-coherence-point-unpointed-htpy-pointed-Π :
+ coherence-point-unpointed-htpy-pointed-Π f g G →
+ preserves-point-unpointed-htpy-pointed-Π
+ preserves-point-coherence-point-unpointed-htpy-pointed-Π =
+ transpose-right-coherence-triangle-identifications _ _ _ refl
+
+ coherence-point-preserves-point-unpointed-htpy-pointed-Π :
+ preserves-point-unpointed-htpy-pointed-Π →
+ coherence-point-unpointed-htpy-pointed-Π f g G
+ coherence-point-preserves-point-unpointed-htpy-pointed-Π =
+ inv ∘ inv-right-transpose-eq-concat _ _ _
+
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B} (H : f ~∗ g)
+ where
+
+ preserves-point-pointed-htpy :
+ preserves-point-unpointed-htpy-pointed-Π f g (htpy-pointed-htpy H)
+ preserves-point-pointed-htpy =
+ preserves-point-coherence-point-unpointed-htpy-pointed-Π f g
+ ( htpy-pointed-htpy H)
+ ( coherence-point-pointed-htpy H)
+```
+
+### Uniform pointed homotopies
+
+**Note.** The operation `htpy-uniform-pointed-htpy` that converts a uniform
+pointed homotopy to an unpointed homotopy is set up with the pointed functions
+as explicit arguments, because Agda has trouble inferring them.
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ (f g : pointed-Π A B)
+ where
+
+ eq-value-Pointed-Fam : Pointed-Fam l2 A
+ pr1 eq-value-Pointed-Fam =
+ eq-value (function-pointed-Π f) (function-pointed-Π g)
+ pr2 eq-value-Pointed-Fam =
+ ( preserves-point-function-pointed-Π f) ∙
+ ( inv (preserves-point-function-pointed-Π g))
+
+ uniform-pointed-htpy : UU (l1 ⊔ l2)
+ uniform-pointed-htpy = pointed-Π A eq-value-Pointed-Fam
+
+ htpy-uniform-pointed-htpy :
+ uniform-pointed-htpy → function-pointed-Π f ~ function-pointed-Π g
+ htpy-uniform-pointed-htpy = pr1
+
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B}
+ (H : uniform-pointed-htpy f g)
+ where
+
+ preserves-point-uniform-pointed-htpy :
+ preserves-point-unpointed-htpy-pointed-Π f g
+ ( htpy-uniform-pointed-htpy f g H)
+ preserves-point-uniform-pointed-htpy = pr2 H
+
+ coherence-point-uniform-pointed-htpy :
+ coherence-point-unpointed-htpy-pointed-Π f g
+ ( htpy-uniform-pointed-htpy f g H)
+ coherence-point-uniform-pointed-htpy =
+ coherence-point-preserves-point-unpointed-htpy-pointed-Π f g
+ ( htpy-uniform-pointed-htpy f g H)
+ ( preserves-point-uniform-pointed-htpy)
+
+ pointed-htpy-uniform-pointed-htpy : f ~∗ g
+ pr1 pointed-htpy-uniform-pointed-htpy =
+ htpy-uniform-pointed-htpy f g H
+ pr2 pointed-htpy-uniform-pointed-htpy =
+ coherence-point-uniform-pointed-htpy
+
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B}
+ where
+
+ make-uniform-pointed-htpy :
+ (G : unpointed-htpy-pointed-Π f g) →
+ coherence-point-unpointed-htpy-pointed-Π f g G →
+ uniform-pointed-htpy f g
+ pr1 (make-uniform-pointed-htpy G p) = G
+ pr2 (make-uniform-pointed-htpy G p) =
+ preserves-point-coherence-point-unpointed-htpy-pointed-Π f g G p
+
+ uniform-pointed-htpy-pointed-htpy : f ~∗ g → uniform-pointed-htpy f g
+ pr1 (uniform-pointed-htpy-pointed-htpy H) = htpy-pointed-htpy H
+ pr2 (uniform-pointed-htpy-pointed-htpy H) = preserves-point-pointed-htpy H
+
+ compute-uniform-pointed-htpy : (f ~∗ g) ≃ uniform-pointed-htpy f g
+ compute-uniform-pointed-htpy =
+ equiv-tot (compute-coherence-point-unpointed-htpy-pointed-Π f g)
+```
+
+### The reflexive uniform pointed homotopy
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ (f : pointed-Π A B)
+ where
+
+ refl-uniform-pointed-htpy : uniform-pointed-htpy f f
+ pr1 refl-uniform-pointed-htpy = refl-htpy
+ pr2 refl-uniform-pointed-htpy =
+ inv (right-inv (preserves-point-function-pointed-Π f))
+```
+
+### Concatenation of uniform pointed homotopies
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g h : pointed-Π A B}
+ (G : uniform-pointed-htpy f g) (H : uniform-pointed-htpy g h)
+ where
+
+ htpy-concat-uniform-pointed-htpy : unpointed-htpy-pointed-Π f h
+ htpy-concat-uniform-pointed-htpy =
+ htpy-uniform-pointed-htpy f g G ∙h htpy-uniform-pointed-htpy g h H
+
+ coherence-point-concat-uniform-pointed-htpy :
+ coherence-point-unpointed-htpy-pointed-Π f h
+ ( htpy-concat-uniform-pointed-htpy)
+ coherence-point-concat-uniform-pointed-htpy =
+ coherence-point-concat-pointed-htpy
+ ( pointed-htpy-uniform-pointed-htpy G)
+ ( pointed-htpy-uniform-pointed-htpy H)
+
+ concat-uniform-pointed-htpy : uniform-pointed-htpy f h
+ concat-uniform-pointed-htpy =
+ make-uniform-pointed-htpy
+ ( htpy-concat-uniform-pointed-htpy)
+ ( coherence-point-concat-uniform-pointed-htpy)
+```
+
+### Inverses of uniform pointed homotopies
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ {f g : pointed-Π A B} (H : uniform-pointed-htpy f g)
+ where
+
+ htpy-inv-uniform-pointed-htpy : unpointed-htpy-pointed-Π g f
+ htpy-inv-uniform-pointed-htpy = inv-htpy (htpy-uniform-pointed-htpy f g H)
+
+ coherence-point-inv-uniform-pointed-htpy :
+ coherence-point-unpointed-htpy-pointed-Π g f htpy-inv-uniform-pointed-htpy
+ coherence-point-inv-uniform-pointed-htpy =
+ coherence-point-inv-pointed-htpy
+ ( pointed-htpy-uniform-pointed-htpy H)
+
+ inv-uniform-pointed-htpy : uniform-pointed-htpy g f
+ inv-uniform-pointed-htpy =
+ make-uniform-pointed-htpy
+ ( htpy-inv-uniform-pointed-htpy)
+ ( coherence-point-inv-uniform-pointed-htpy)
+```
+
+## Properties
+
+### Extensionality of pointed dependent function types by uniform pointed homotopies
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Fam l2 A}
+ (f : pointed-Π A B)
+ where
+
+ uniform-extensionality-pointed-Π :
+ (g : pointed-Π A B) → (f = g) ≃ uniform-pointed-htpy f g
+ uniform-extensionality-pointed-Π =
+ extensionality-Σ
+ ( λ {g} q H →
+ H (point-Pointed-Type A) =
+ preserves-point-function-pointed-Π f ∙
+ inv (preserves-point-function-pointed-Π (g , q)))
+ ( refl-htpy)
+ ( inv (right-inv (preserves-point-function-pointed-Π f)))
+ ( λ g → equiv-funext)
+ ( λ p →
+ ( equiv-right-transpose-eq-concat
+ ( refl)
+ ( p)
+ ( preserves-point-function-pointed-Π f)) ∘e
+ ( equiv-inv (preserves-point-function-pointed-Π f) p))
+
+ eq-uniform-pointed-htpy :
+ (g : pointed-Π A B) → uniform-pointed-htpy f g → f = g
+ eq-uniform-pointed-htpy g = map-inv-equiv (uniform-extensionality-pointed-Π g)
+```
diff --git a/src/structured-types/universal-property-pointed-equivalences.lagda.md b/src/structured-types/universal-property-pointed-equivalences.lagda.md
new file mode 100644
index 0000000000..1ed069812f
--- /dev/null
+++ b/src/structured-types/universal-property-pointed-equivalences.lagda.md
@@ -0,0 +1,49 @@
+# The universal property of pointed equivalences
+
+```agda
+module structured-types.universal-property-pointed-equivalences where
+```
+
+Imports
+
+```agda
+open import foundation.equivalences
+open import foundation.universe-levels
+
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+open import structured-types.precomposition-pointed-maps
+```
+
+
+
+## Idea
+
+Analogous to the
+[universal property of equivalences](foundation.universal-property-equivalences.md),
+the
+{{#concept "universal property of pointed equivalences" Agda=universal-property-pointed-equiv}}
+asserts about a [pointed map](structured-types.pointed-maps.md) `f : A →∗ B`
+that the
+[precomposition function](structured-types.precomposition-pointed-maps.md)
+
+```text
+ - ∘∗ f : (B →∗ C) → (A →∗ C)
+```
+
+is an [equivalence](foundation.equivalences.md) for every
+[pointed type](structured-types.pointed-types.md) `C`.
+
+## Definitions
+
+### The universal property of pointed equivalences
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} (f : A →∗ B)
+ where
+
+ universal-property-pointed-equiv : UUω
+ universal-property-pointed-equiv =
+ {l : Level} (C : Pointed-Type l) → is-equiv (precomp-pointed-map f C)
+```
diff --git a/src/structured-types/whiskering-pointed-2-homotopies-concatenation.lagda.md b/src/structured-types/whiskering-pointed-2-homotopies-concatenation.lagda.md
new file mode 100644
index 0000000000..7e99884c46
--- /dev/null
+++ b/src/structured-types/whiskering-pointed-2-homotopies-concatenation.lagda.md
@@ -0,0 +1,280 @@
+# Whiskering pointed homotopies with respect to concatenation
+
+```agda
+module structured-types.whiskering-pointed-2-homotopies-concatenation where
+```
+
+Imports
+
+```agda
+open import foundation.action-on-identifications-functions
+open import foundation.commuting-triangles-of-identifications
+open import foundation.dependent-pair-types
+open import foundation.identity-types
+open import foundation.path-algebra
+open import foundation.universe-levels
+open import foundation.whiskering-homotopies-concatenation
+open import foundation.whiskering-identifications-concatenation
+
+open import structured-types.pointed-2-homotopies
+open import structured-types.pointed-homotopies
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+```
+
+
+
+## Idea
+
+The [whiskering operations](foundation.whiskering-operations.md) of
+[pointed `2`-homotopies](structured-types.pointed-2-homotopies.md) with respect
+to concatenation of [pointed homotopies](structured-types.pointed-homotopies.md)
+are two operations that produce pointed 2-homotopies between concatenations of
+pointed homotopies from either a pointed 2-homotopy on the left or on the right
+of the concatenations.
+
+- The
+ {{#concept "left whiskering" Disambiguation="pointed `2`-homotopies with respect to concatenation" Agda=left-whisker-concat-pointed-2-htpy}}
+ is an operation that takes a pointed homotopy `H : f ~∗ g` and a pointed
+ `2`-homotopy `α : K ~²∗ L` between two pointed homotopies `K L : g ~∗ h` as
+ indicated in the diagram
+
+ ```text
+ K
+ H ----->
+ f -----> g -----> h,
+ L
+ ```
+
+ and returns a pointed `2`-homotopy `H ∙h K ~²∗ H ∙h L`.
+
+- The
+ {{#concept "right whiskering" Disambiguation="pointed `2`-homotopies with respect to concatenation" Agda=right-whisker-concat-pointed-2-htpy}}
+ is an operation that takes a pointed `2`-homotopy `α : H ~²∗ K` between two
+ pointed homotopies `H K : f ~∗ g` and a pointed homotopy `L : g ~∗ h` as
+ indicated in the diagram
+
+ ```text
+ H
+ ----->
+ f -----> g -----> h,
+ K L
+ ```
+
+ and returns a pointed `2`-homotopy `H ∙h L ~²∗ K ∙h L`.
+
+## Definitions
+
+### Left whiskering of pointed `2`-homotopies with respect to concatenation
+
+Consider three pointed maps `f := (f₀ , f₁)`, `g := (g₀ , g₁)`, and
+`h := (h₀ , h₁)` from `A` to `B`, a pointed homotopy `H := (H₀ , H₁) : f ~∗ g`
+and a pointed `2`-homotopy `α := (α₀ , α₁) : K ~²∗ L` between two pointed
+homotopies `K := (K₀ , K₁)` and `L := (L₀ , L₁)` from `g` to `h` as indicated in
+the diagram
+
+```text
+ K
+ H ----->
+ f -----> g -----> h.
+ L
+```
+
+The underlying homotopy of the left whiskering `H ·l∗ α : H ∙h K ~²∗ H ∙h L` is
+the homotopy
+
+```text
+ H₀ ·l α₀ : H₀ ∙h K₀ ~ H₀ ∙h L₀.
+```
+
+The base point coherence of this homotopy is an identification witnessing that
+the triangle
+
+```text
+ (H ∙h K)₁
+ f₁ --------> ((H₀ *) ∙ (K₀ *)) ∙ h₁
+ \ /
+ (H ∙h L)₁ \ / right-whisker (left-whisker (H₀ *) (α₀ *)) h₁
+ \ /
+ ∨ ∨
+ ((H₀ *) ∙ (L₀ *)) ∙ h₁
+```
+
+commutes. Here, the identifications `(H ∙h K)₁` and `(H ∙h L)₁` are the
+horizontal pastings of the
+[commuting triangles of identifications](foundation.commuting-triangles-of-identifications.md)
+
+```text
+ H₀ * K₀ * H₀ * L₀ *
+ f₀ * ---> g₀ * ----> h₀ * f₀ * ---> g₀ * ----> h₀ *
+ \ | / \ | /
+ \ H₁ | K₁ / \ H₁ | L₁ /
+ f₁ \ |g₁ / h₁ f₁ \ |g₁ / h₁
+ \ | / \ | /
+ \ | / \ | /
+ ∨ ∨ ∨ ∨ ∨ ∨
+ * *.
+```
+
+Then the triangle
+
+```text
+ horizontal-pasting H₁ K₁
+ f₁ --------> (H₀ * ∙ K₀ *) ∙ h₁
+ \ /
+ \ /
+ horizontal-pasting H₁ L₁ \ / right-whisker (left-whisker (H₀ *) (α₀ *)) h₁
+ \ /
+ ∨ ∨
+ (H₀ * ∙ K₀ *) ∙ h₁
+```
+
+commutes by left whiskering of horizontal pasting of commuting triangles of
+identifications.
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {f g h : A →∗ B} (H : f ~∗ g) (K L : g ~∗ h) (α : K ~²∗ L)
+ where
+
+ htpy-left-whisker-concat-pointed-2-htpy :
+ unpointed-htpy-pointed-htpy
+ ( concat-pointed-htpy H K)
+ ( concat-pointed-htpy H L)
+ htpy-left-whisker-concat-pointed-2-htpy =
+ left-whisker-concat-htpy (htpy-pointed-htpy H) (htpy-pointed-2-htpy α)
+
+ coherence-point-left-whisker-concat-pointed-2-htpy :
+ coherence-point-unpointed-htpy-pointed-htpy
+ ( concat-pointed-htpy H K)
+ ( concat-pointed-htpy H L)
+ ( htpy-left-whisker-concat-pointed-2-htpy)
+ coherence-point-left-whisker-concat-pointed-2-htpy =
+ left-whisker-horizontal-pasting-coherence-triangle-identifications
+ ( preserves-point-pointed-map f)
+ ( preserves-point-pointed-map g)
+ ( preserves-point-pointed-map h)
+ ( htpy-pointed-htpy H (point-Pointed-Type A))
+ ( htpy-pointed-htpy K (point-Pointed-Type A))
+ ( htpy-pointed-htpy L (point-Pointed-Type A))
+ ( coherence-point-pointed-htpy H)
+ ( coherence-point-pointed-htpy K)
+ ( coherence-point-pointed-htpy L)
+ ( htpy-pointed-2-htpy α (point-Pointed-Type A))
+ ( coherence-point-pointed-2-htpy α)
+
+ left-whisker-concat-pointed-2-htpy :
+ concat-pointed-htpy H K ~²∗ concat-pointed-htpy H L
+ pr1 left-whisker-concat-pointed-2-htpy =
+ htpy-left-whisker-concat-pointed-2-htpy
+ pr2 left-whisker-concat-pointed-2-htpy =
+ coherence-point-left-whisker-concat-pointed-2-htpy
+```
+
+### Right whiskering of pointed `2`-homotopies with respect to concatenation
+
+Consider three pointed maps `f := (f₀ , f₁)`, `g := (g₀ , g₁)`, and
+`h := (h₀ , h₁)` from `A` to `B`, a pointed `2`-homotopy
+`α := (α₀ , α₁) : H ~²∗ K` between two pointed homotopies `H := (H₀ , H₁)` and
+`K := (K₀ , K₁)` from `f` to `g` and a pointed homotopy
+`L := (L₀ , L₁) : g ~∗ h` as indicated in the diagram
+
+```text
+ H
+ ----->
+ f -----> g -----> h.
+ K L
+```
+
+The underlying homotopy of the right whiskering `α ·r∗ L : H ∙h L ~²∗ K ∙h L` is
+the homotopy
+
+```text
+ α₀ ·r L₀ : H₀ ∙h L₀ ~ K₀ ∙h L₀.
+```
+
+The base point coherence of this homotopy is an identification witnessing that
+the triangle
+
+```text
+ (H ∙h L)₁
+ f₁ --------> ((H₀ *) ∙ (L₀ *)) ∙ h₁
+ \ /
+ (K ∙h L)₁ \ / right-whisker (right-whisker (α₀ *) (L₀ *)) h₁
+ \ /
+ ∨ ∨
+ ((K₀ *) ∙ (L₀ *)) ∙ h₁
+```
+
+commutes. Here, the identifications `(H ∙h L)₁` and `(K ∙h L)₁` are the
+horizontal pastings of the
+[commuting triangles of identifications](foundation.commuting-triangles-of-identifications.md)
+
+```text
+ H₀ * L₀ * K₀ * L₀ *
+ f₀ * ---> g₀ * ----> h₀ * f₀ * ---> g₀ * ----> h₀ *
+ \ | / \ | /
+ \ H₁ | L₁ / \ K₁ | L₁ /
+ f₁ \ |g₁ / h₁ f₁ \ |g₁ / h₁
+ \ | / \ | /
+ \ | / \ | /
+ ∨ ∨ ∨ ∨ ∨ ∨
+ * *.
+```
+
+Then the triangle
+
+```text
+ horizontal-pasting H₁ L₁
+ f₁ --------> (H₀ * ∙ L₀ *) ∙ h₁
+ \ /
+ \ /
+ horizontal-pasting K₁ L₁ \ / right-whisker (right-whisker (α₀ *) (L₀ *)) h₁
+ \ /
+ ∨ ∨
+ (K₀ * ∙ L₀ *) ∙ h₁
+```
+
+commutes by right whiskering of horizontal pasting of commuting triangles of
+identifications.
+
+```agda
+module _
+ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
+ {f g h : A →∗ B} (H K : f ~∗ g) (α : H ~²∗ K) (L : g ~∗ h)
+ where
+
+ htpy-right-whisker-concat-pointed-2-htpy :
+ unpointed-htpy-pointed-htpy
+ ( concat-pointed-htpy H L)
+ ( concat-pointed-htpy K L)
+ htpy-right-whisker-concat-pointed-2-htpy =
+ right-whisker-concat-htpy (htpy-pointed-2-htpy α) (htpy-pointed-htpy L)
+
+ coherence-point-right-whisker-concat-pointed-2-htpy :
+ coherence-point-unpointed-htpy-pointed-htpy
+ ( concat-pointed-htpy H L)
+ ( concat-pointed-htpy K L)
+ ( htpy-right-whisker-concat-pointed-2-htpy)
+ coherence-point-right-whisker-concat-pointed-2-htpy =
+ right-whisker-horizontal-pasting-coherence-triangle-identifications
+ ( preserves-point-pointed-map f)
+ ( preserves-point-pointed-map g)
+ ( preserves-point-pointed-map h)
+ ( htpy-pointed-htpy H (point-Pointed-Type A))
+ ( htpy-pointed-htpy K (point-Pointed-Type A))
+ ( htpy-pointed-htpy L (point-Pointed-Type A))
+ ( coherence-point-pointed-htpy H)
+ ( coherence-point-pointed-htpy K)
+ ( coherence-point-pointed-htpy L)
+ ( htpy-pointed-2-htpy α (point-Pointed-Type A))
+ ( coherence-point-pointed-2-htpy α)
+
+ right-whisker-concat-pointed-2-htpy :
+ concat-pointed-htpy H L ~²∗ concat-pointed-htpy K L
+ pr1 right-whisker-concat-pointed-2-htpy =
+ htpy-right-whisker-concat-pointed-2-htpy
+ pr2 right-whisker-concat-pointed-2-htpy =
+ coherence-point-right-whisker-concat-pointed-2-htpy
+```
diff --git a/src/structured-types/whiskering-pointed-homotopies-composition.lagda.md b/src/structured-types/whiskering-pointed-homotopies-composition.lagda.md
new file mode 100644
index 0000000000..ce71678ce6
--- /dev/null
+++ b/src/structured-types/whiskering-pointed-homotopies-composition.lagda.md
@@ -0,0 +1,365 @@
+# Whiskering of pointed homotopies with respect to composition of pointed maps
+
+```agda
+module structured-types.whiskering-pointed-homotopies-composition where
+```
+
+Imports
+
+```agda
+open import foundation.action-on-identifications-binary-functions
+open import foundation.action-on-identifications-functions
+open import foundation.commuting-squares-of-identifications
+open import foundation.commuting-triangles-of-identifications
+open import foundation.dependent-pair-types
+open import foundation.homotopies
+open import foundation.identity-types
+open import foundation.universe-levels
+open import foundation.whiskering-homotopies-composition
+open import foundation.whiskering-identifications-concatenation
+
+open import structured-types.pointed-2-homotopies
+open import structured-types.pointed-families-of-types
+open import structured-types.pointed-homotopies
+open import structured-types.pointed-maps
+open import structured-types.pointed-types
+```
+
+
+
+## Idea
+
+The [whiskering operations](foundation.whiskering-operations.md) of
+[pointed homotopies](structured-types.pointed-homotopies.md) with respect to
+composition of [pointed maps](structured-types.pointed-maps.md) are two
+operations that produce pointed homotopies between composites of pointed maps
+from either a pointed homotopy on the left or on the right of the composition.
+
+- Consider a pointed homotopy `H : f ~∗ g` between pointed maps `f g : A →∗ B`,
+ and consider a pointed map `h : B →∗ C`, as indicated in the diagram
+
+ ```text
+ f
+ -----> h
+ A -----> B -----> C.
+ g
+ ```
+
+ The
+ {{#concept "left whiskering operation on pointed homotopies" Agda=left-whisker-comp-pointed-htpy}}
+ of `h` and `H` is a pointed homotopy
+
+ ```text
+ h ·l∗ H : h ∘∗ f ~∗ h ∘∗ g.
+ ```
+
+- Consider a pointed map `f : A →∗ B` and consider a pointed homotopy
+ `H : g ~∗ g` between tw pointed maps `g h : B →∗ C`, as indicated in the
+ diagram
+
+ ```text
+ g
+ f ----->
+ A -----> B -----> C.
+ h
+ ```
+
+ The
+ {{#concept "right whiskering operation on pointed homotopies" Agda=right-whisker-comp-pointed-htpy}}
+ of `H` and `f` is a pointed homotopy
+
+ ```text
+ H ·r∗ f : g ∘∗ f ~∗ h ∘∗ f.
+ ```
+
+## Definitions
+
+### Left whiskering of pointed homotopies
+
+Consider two pointed maps `f := (f₀ , f₁) : A →∗ B` and
+`g := (g₀ , g₁) : A →∗ B` equipped with a pointed homotopy
+`H := (H₀ , H₁) : f ~∗ g`, and consider furthermore a pointed map
+`h := (h₀ , h₁) : B →∗ C`. Then we construct a pointed homotopy
+
+```text
+ h ·l∗ H : (h ∘∗ f) ~∗ (h ∘∗ g).
+```
+
+**Construction.** The underlying homotopy of `h ·l∗ H` is the whiskered homotpy
+
+```text
+ h₀ ·l H₀.
+```
+
+For the coherence, we have to show that the triangle
+
+```text
+ ap h₀ (H₀ *)
+ h₀ (f₀ *) ------------> h₀ (g₀ *)
+ \ /
+ ap h₀ f₁ \ / ap h₀ g₁
+ ∨ ∨
+ h₀ * h₀ *
+ \ /
+ h₁ \ / h₁
+ ∨ ∨
+ ∗
+```
+
+commutes. By right whiskering of
+[commuting triangles of identifications](foundation.commuting-squares-of-identifications.md)
+with respect to concatenation it suffices to show that the triangle
+
+```text
+ ap h₀ (H₀ *)
+ h₀ (f₀ *) ---------> h₀ (g₀ *)
+ \ /
+ ap h₀ f₁ \ / ap h₀ g₁
+ \ /
+ ∨ ∨
+ h₀ *
+```
+
+commutes. By functoriality of commuting triangles of identifications, this
+follows from the fact that the triangle
+
+```text
+ H₀ *
+ f₀ * ------> g₀ *
+ \ /
+ f₁ \ / g₁
+ \ /
+ ∨ ∨
+ *
+```
+
+commutes.
+
+```agda
+module _
+ {l1 l2 l3 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Type l2} {C : Pointed-Type l3}
+ (h : B →∗ C) (f g : A →∗ B) (H : f ~∗ g)
+ where
+
+ htpy-left-whisker-comp-pointed-htpy :
+ map-comp-pointed-map h f ~ map-comp-pointed-map h g
+ htpy-left-whisker-comp-pointed-htpy =
+ map-pointed-map h ·l htpy-pointed-htpy H
+
+ coherence-point-left-whisker-comp-pointed-htpy :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( h ∘∗ f)
+ ( h ∘∗ g)
+ ( htpy-left-whisker-comp-pointed-htpy)
+ coherence-point-left-whisker-comp-pointed-htpy =
+ right-whisker-concat-coherence-triangle-identifications
+ ( ap (map-pointed-map h) (preserves-point-pointed-map f))
+ ( ap (map-pointed-map h) (preserves-point-pointed-map g))
+ ( ap
+ ( map-pointed-map h)
+ ( htpy-pointed-htpy H (point-Pointed-Type A)))
+ ( preserves-point-pointed-map h)
+ ( map-coherence-triangle-identifications
+ ( map-pointed-map h)
+ ( preserves-point-pointed-map f)
+ ( preserves-point-pointed-map g)
+ ( htpy-pointed-htpy H (point-Pointed-Type A))
+ ( coherence-point-pointed-htpy H))
+
+ left-whisker-comp-pointed-htpy : h ∘∗ f ~∗ h ∘∗ g
+ pr1 left-whisker-comp-pointed-htpy =
+ htpy-left-whisker-comp-pointed-htpy
+ pr2 left-whisker-comp-pointed-htpy =
+ coherence-point-left-whisker-comp-pointed-htpy
+```
+
+### Right whiskering of pointed homotopies
+
+Consider a pointed map `f := (f₀ , f₁) : A →∗ B` and two pointed maps
+`g := (g₀ , g₁) : B →∗ C` and `h := (h₀ , h₁) : B →∗ C` equipped with a pointed
+homotopy `H := (H₀ , H₁) : g ~∗ h`. Then we construct a pointed homotopy
+
+```text
+ H ·r∗ f : (g ∘∗ f) ~∗ (h ∘∗ f).
+```
+
+**Construction.** The underlying homotopy of `H ·r∗ f` is the homotopy
+
+```text
+ H₀ ·r f₀ : (g₀ ∘ f₀) ~ (h₀ ∘ f₀).
+```
+
+Then we have to show that the outer triangle in the diagram
+
+```text
+ H₀ (f₀ *)
+ g₀ (f₀ *) ------------> h₀ (f₀ *)
+ \ /
+ ap g₀ f₁ \ / ap h₀ f₁
+ ∨ H₀ * ∨
+ g₀ * ----> h₀ *
+ \ /
+ g₁ \ / h₁
+ ∨ ∨
+ ∗
+```
+
+commutes. This is done by vertically pasting the upper square and the lower
+triangle. The upper square commutes by inverse naturality of the homotopy `H₀`.
+The lower triangle is the base point coherence `H₁` of the pointed homotopy
+`H ≐ (H₀ , H₁)`.
+
+```agda
+module _
+ {l1 l2 l3 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Type l2} {C : Pointed-Type l3}
+ (g1 g2 : B →∗ C) (H : g1 ~∗ g2) (f : A →∗ B)
+ where
+
+ htpy-right-whisker-comp-pointed-htpy :
+ unpointed-htpy-pointed-Π (g1 ∘∗ f) (g2 ∘∗ f)
+ htpy-right-whisker-comp-pointed-htpy =
+ htpy-pointed-htpy H ·r map-pointed-map f
+
+ coherence-point-right-whisker-comp-pointed-htpy :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( g1 ∘∗ f)
+ ( g2 ∘∗ f)
+ ( htpy-right-whisker-comp-pointed-htpy)
+ coherence-point-right-whisker-comp-pointed-htpy =
+ vertical-pasting-coherence-square-coherence-triangle-identifications
+ ( htpy-pointed-htpy H _)
+ ( ap (map-pointed-map g1) _)
+ ( ap (map-pointed-map g2) _)
+ ( htpy-pointed-htpy H _)
+ ( preserves-point-pointed-map g1)
+ ( preserves-point-pointed-map g2)
+ ( inv-nat-htpy (htpy-pointed-htpy H) _)
+ ( coherence-point-pointed-htpy H)
+
+ right-whisker-comp-pointed-htpy : g1 ∘∗ f ~∗ g2 ∘∗ f
+ pr1 right-whisker-comp-pointed-htpy =
+ htpy-right-whisker-comp-pointed-htpy
+ pr2 right-whisker-comp-pointed-htpy =
+ coherence-point-right-whisker-comp-pointed-htpy
+```
+
+## Properties
+
+### Computing the left whiskering of the reflexive pointed homotopy
+
+```agda
+module _
+ {l1 l2 l3 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Type l2} {C : Pointed-Type l3}
+ (h : B →∗ C) (f : A →∗ B)
+ where
+
+ compute-refl-left-whisker-comp-pointed-htpy :
+ pointed-2-htpy
+ ( left-whisker-comp-pointed-htpy h f f (refl-pointed-htpy f))
+ ( refl-pointed-htpy (h ∘∗ f))
+ compute-refl-left-whisker-comp-pointed-htpy =
+ refl-pointed-2-htpy (refl-pointed-htpy (h ∘∗ f))
+```
+
+### Computing the right whiskering of the reflexive pointed homotopy
+
+Consider two pointed maps `f := (f₀ , f₁) : A →∗ B` and
+`g := (g₀ , g₁) : B →∗ C`. We are constructing a pointed `2`-homotopy
+
+```text
+ right-whisker-comp-pointed-htpy (refl-pointed-htpy h) f ~∗
+ refl-pointed-htpy (g ∘∗ f)
+```
+
+The underlying homotopy of this pointed `2`-homotopy is `refl-htpy`. The base
+point coherence of this homotopy is an identification witnessing that the
+triangle
+
+```text
+ H₁
+ ap g₀ f₁ ∙ g₁ ------> refl ∙ (ap g₀ f₁ ∙ g₁)
+ \ /
+ refl \ / right-whisker-concat refl (ap g₀ f₁ ∙ g₁) ≐ refl
+ \ /
+ ∨ ∨
+ refl ∙ (ap g₀ f₁ ∙ g₁)
+```
+
+commutes. Here, the identification `H₁` is the vertical pasting of the upper
+square and the lower triangle in the diagram
+
+```text
+ refl
+ g₀ (f₀ *) ------------> g₀ (f₀ *)
+ \ /
+ ap g₀ f₁ \ / ap g₀ f₁
+ ∨ refl ∨
+ g₀ * ----> g₀ *
+ \ /
+ g₁ \ / g₁
+ ∨ ∨
+ ∗.
+```
+
+The upper square in this diagram is the inverse naturality of the reflexive
+homotopy `refl-htpy` and the lower triangle in this diagram is the reflexive
+identification.
+
+Recall that the inverse naturality of the reflexive homotopy
+`inv-nat-htpy refl-htpy f₁` computes to the horizontally constant square of
+identifications. Furthermore, the vertical pasting of the horizontally constant
+square `right-unit` and any commuting triangle `refl` computes to `refl`.
+Therefore it follows that the identification `H₁` above is equal to `refl`, as
+was required to show.
+
+```agda
+module _
+ {l1 l2 l3 : Level}
+ {A : Pointed-Type l1} {B : Pointed-Type l2} {C : Pointed-Type l3}
+ (h : B →∗ C) (f : A →∗ B)
+ where
+
+ htpy-compute-refl-right-whisker-comp-pointed-htpy :
+ unpointed-htpy-pointed-htpy
+ ( right-whisker-comp-pointed-htpy h h (refl-pointed-htpy h) f)
+ ( refl-pointed-htpy (h ∘∗ f))
+ htpy-compute-refl-right-whisker-comp-pointed-htpy = refl-htpy
+
+ coherence-point-compute-refl-right-whisker-comp-pointed-htpy :
+ coherence-point-unpointed-htpy-pointed-htpy
+ ( right-whisker-comp-pointed-htpy h h (refl-pointed-htpy h) f)
+ ( refl-pointed-htpy (h ∘∗ f))
+ ( htpy-compute-refl-right-whisker-comp-pointed-htpy)
+ coherence-point-compute-refl-right-whisker-comp-pointed-htpy =
+ inv
+ ( ( right-unit) ∙
+ ( ( ap
+ ( λ t →
+ vertical-pasting-coherence-square-coherence-triangle-identifications
+ ( refl)
+ ( ap (map-pointed-map h) (preserves-point-pointed-map f))
+ ( ap (map-pointed-map h) (preserves-point-pointed-map f))
+ ( refl)
+ ( preserves-point-pointed-map h)
+ ( preserves-point-pointed-map h)
+ ( t)
+ ( refl))
+ ( inv-nat-refl-htpy
+ ( map-pointed-map h)
+ ( preserves-point-pointed-map f))) ∙
+ ( right-whisker-concat-horizontal-refl-coherence-square-identifications
+ ( ap (map-pointed-map h) (preserves-point-pointed-map f))
+ ( preserves-point-pointed-map h))))
+
+ compute-refl-right-whisker-comp-pointed-htpy :
+ pointed-2-htpy
+ ( right-whisker-comp-pointed-htpy h h (refl-pointed-htpy h) f)
+ ( refl-pointed-htpy (h ∘∗ f))
+ pr1 compute-refl-right-whisker-comp-pointed-htpy =
+ htpy-compute-refl-right-whisker-comp-pointed-htpy
+ pr2 compute-refl-right-whisker-comp-pointed-htpy =
+ coherence-point-compute-refl-right-whisker-comp-pointed-htpy
+```
diff --git a/src/synthetic-homotopy-theory.lagda.md b/src/synthetic-homotopy-theory.lagda.md
index 47dbab16c5..65723070c4 100644
--- a/src/synthetic-homotopy-theory.lagda.md
+++ b/src/synthetic-homotopy-theory.lagda.md
@@ -16,9 +16,9 @@ open import synthetic-homotopy-theory.category-of-connected-set-bundles-circle p
open import synthetic-homotopy-theory.cavallos-trick public
open import synthetic-homotopy-theory.circle public
open import synthetic-homotopy-theory.cocartesian-morphisms-arrows public
+open import synthetic-homotopy-theory.cocones-under-pointed-span-diagrams public
open import synthetic-homotopy-theory.cocones-under-sequential-diagrams public
open import synthetic-homotopy-theory.cocones-under-spans public
-open import synthetic-homotopy-theory.cocones-under-spans-of-pointed-types public
open import synthetic-homotopy-theory.codiagonals-of-maps public
open import synthetic-homotopy-theory.coequalizers public
open import synthetic-homotopy-theory.cofibers public
@@ -70,6 +70,7 @@ open import synthetic-homotopy-theory.mere-spheres public
open import synthetic-homotopy-theory.morphisms-descent-data-circle public
open import synthetic-homotopy-theory.morphisms-sequential-diagrams public
open import synthetic-homotopy-theory.multiplication-circle public
+open import synthetic-homotopy-theory.null-cocones-under-pointed-span-diagrams public
open import synthetic-homotopy-theory.plus-principle public
open import synthetic-homotopy-theory.powers-of-loops public
open import synthetic-homotopy-theory.premanifolds public
diff --git a/src/synthetic-homotopy-theory/cavallos-trick.lagda.md b/src/synthetic-homotopy-theory/cavallos-trick.lagda.md
index 86e8cc6610..54d3a7e599 100644
--- a/src/synthetic-homotopy-theory/cavallos-trick.lagda.md
+++ b/src/synthetic-homotopy-theory/cavallos-trick.lagda.md
@@ -14,6 +14,7 @@ open import foundation.homotopies
open import foundation.identity-types
open import foundation.sections
open import foundation.universe-levels
+open import foundation.whiskering-identifications-concatenation
open import structured-types.pointed-homotopies
open import structured-types.pointed-maps
@@ -39,23 +40,31 @@ module _
{l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2}
where
+ htpy-cavallos-trick :
+ (f g : A →∗ B) → section (λ (H : id ~ id) → H (point-Pointed-Type B)) →
+ (map-pointed-map f ~ map-pointed-map g) →
+ unpointed-htpy-pointed-map f g
+ htpy-cavallos-trick (f , refl) (g , q) (K , α) H a =
+ K (inv q ∙ inv (H (point-Pointed-Type A))) (f a) ∙ H a
+
+ coherence-point-cavallos-trick :
+ (f g : A →∗ B) (s : section (λ (H : id ~ id) → H (point-Pointed-Type B))) →
+ (H : map-pointed-map f ~ map-pointed-map g) →
+ coherence-point-unpointed-htpy-pointed-Π f g
+ ( htpy-cavallos-trick f g s H)
+ coherence-point-cavallos-trick (f , refl) (g , q) (K , α) H =
+ inv
+ ( ( right-whisker-concat
+ ( ( right-whisker-concat (α _) (H _)) ∙
+ ( is-section-inv-concat' (H _) (inv q)))
+ ( q)) ∙
+ ( left-inv q))
+
cavallos-trick :
(f g : A →∗ B) → section (λ (H : id ~ id) → H (point-Pointed-Type B)) →
(map-pointed-map f ~ map-pointed-map g) → f ~∗ g
- pr1 (cavallos-trick (f , refl) (g , q) (K , α) H) a =
- K (inv q ∙ inv (H (point-Pointed-Type A))) (f a) ∙ H a
- pr2 (cavallos-trick (f , refl) (g , q) (K , α) H) =
- ( ap
- ( concat' (f (point-Pointed-Type A)) (H (point-Pointed-Type A)))
- ( α (inv q ∙ inv (H (point-Pointed-Type A))))) ∙
- ( ( assoc
- ( inv q)
- ( inv (H (point-Pointed-Type A)))
- ( H (point-Pointed-Type A))) ∙
- ( ( ap
- ( concat (inv q) (g (point-Pointed-Type A)))
- ( left-inv (H (point-Pointed-Type A)))) ∙
- ( right-unit)))
+ pr1 (cavallos-trick f g s H) = htpy-cavallos-trick f g s H
+ pr2 (cavallos-trick f g s H) = coherence-point-cavallos-trick f g s H
```
## References
diff --git a/src/synthetic-homotopy-theory/cocones-under-spans-of-pointed-types.lagda.md b/src/synthetic-homotopy-theory/cocones-under-pointed-span-diagrams.lagda.md
similarity index 66%
rename from src/synthetic-homotopy-theory/cocones-under-spans-of-pointed-types.lagda.md
rename to src/synthetic-homotopy-theory/cocones-under-pointed-span-diagrams.lagda.md
index 6c606499df..cec36cef2d 100644
--- a/src/synthetic-homotopy-theory/cocones-under-spans-of-pointed-types.lagda.md
+++ b/src/synthetic-homotopy-theory/cocones-under-pointed-span-diagrams.lagda.md
@@ -1,7 +1,7 @@
-# Cocones under spans of pointed types
+# Cocones under pointed span diagrams
```agda
-module synthetic-homotopy-theory.cocones-under-spans-of-pointed-types where
+module synthetic-homotopy-theory.cocones-under-pointed-span-diagrams where
```
Imports
@@ -44,36 +44,13 @@ module _
(f : S →∗ A) (g : S →∗ B)
where
- type-cocone-Pointed-Type :
+ cocone-Pointed-Type :
{l4 : Level} → Pointed-Type l4 → UU (l1 ⊔ l2 ⊔ l3 ⊔ l4)
- type-cocone-Pointed-Type X =
+ cocone-Pointed-Type X =
Σ ( A →∗ X)
( λ i →
Σ ( B →∗ X)
( λ j → coherence-square-pointed-maps g f j i))
-
- cocone-Pointed-Type :
- {l4 : Level} → Pointed-Type l4 → Pointed-Type (l1 ⊔ l2 ⊔ l3 ⊔ l4)
- pr1 (cocone-Pointed-Type X) = type-cocone-Pointed-Type X
- pr1 (pr2 (cocone-Pointed-Type X)) = constant-pointed-map A X
- pr1 (pr2 (pr2 (cocone-Pointed-Type X))) = constant-pointed-map B X
- pr1 (pr2 (pr2 (pr2 (cocone-Pointed-Type X)))) = refl-htpy
- pr2 (pr2 (pr2 (pr2 (cocone-Pointed-Type X)))) =
- inv
- ( ( ap
- ( λ p →
- ( p ∙ refl) ∙
- ( inv
- ( preserves-point-pointed-map
- ( constant-pointed-map B X ∘∗ g))))
- ( ap-const
- ( point-Pointed-Type X)
- ( preserves-point-pointed-map f))) ∙
- ( ap
- ( λ p → inv (p ∙ refl))
- ( ap-const
- ( point-Pointed-Type X)
- ( preserves-point-pointed-map g))))
```
### Components of a cocone of pointed types
@@ -83,7 +60,7 @@ module _
{l1 l2 l3 l4 : Level}
{S : Pointed-Type l1} {A : Pointed-Type l2}
{B : Pointed-Type l3} {X : Pointed-Type l4}
- (f : S →∗ A) (g : S →∗ B) (c : type-cocone-Pointed-Type f g X)
+ (f : S →∗ A) (g : S →∗ B) (c : cocone-Pointed-Type f g X)
where
horizontal-pointed-map-cocone-Pointed-Type : A →∗ X
@@ -110,13 +87,14 @@ module _
( horizontal-pointed-map-cocone-Pointed-Type)
coherence-square-cocone-Pointed-Type = pr2 (pr2 c)
- cocone-type-cocone-Pointed-Type : cocone (pr1 f) (pr1 g) (pr1 X)
- pr1 cocone-type-cocone-Pointed-Type = horizontal-map-cocone-Pointed-Type
- pr1 (pr2 cocone-type-cocone-Pointed-Type) = vertical-map-cocone-Pointed-Type
- pr2 (pr2 cocone-type-cocone-Pointed-Type) =
+ cocone-cocone-Pointed-Type : cocone (pr1 f) (pr1 g) (pr1 X)
+ pr1 cocone-cocone-Pointed-Type = horizontal-map-cocone-Pointed-Type
+ pr1 (pr2 cocone-cocone-Pointed-Type) = vertical-map-cocone-Pointed-Type
+ pr2 (pr2 cocone-cocone-Pointed-Type) =
pr1 coherence-square-cocone-Pointed-Type
```
## See also
- [Pushouts of pointed types](synthetic-homotopy-theory.pushouts-of-pointed-types.md)
+- [Null cocones under pointed span diagrams](synthetic-homotopy-theory.null-cocones-under-pointed-span-diagrams.md)
diff --git a/src/synthetic-homotopy-theory/conjugation-loops.lagda.md b/src/synthetic-homotopy-theory/conjugation-loops.lagda.md
index f904e324d8..589c5d457d 100644
--- a/src/synthetic-homotopy-theory/conjugation-loops.lagda.md
+++ b/src/synthetic-homotopy-theory/conjugation-loops.lagda.md
@@ -80,14 +80,16 @@ module _
{x y : A} (p : x = y) → map-conjugation-Ω p ~ map-conjugation-Ω' p
htpy-compute-conjugation-Ω refl ω = right-unit
- preserves-point-compute-conjugation-Ω :
+ coherence-point-compute-conjugation-Ω :
{x y : A} (p : x = y) →
- ( htpy-compute-conjugation-Ω p refl) =
- ( preserves-point-conjugation-Ω p ∙ inv (preserves-point-conjugation-Ω' p))
- preserves-point-compute-conjugation-Ω refl = refl
+ coherence-point-unpointed-htpy-pointed-Π
+ ( conjugation-Ω p)
+ ( conjugation-Ω' p)
+ ( htpy-compute-conjugation-Ω p)
+ coherence-point-compute-conjugation-Ω refl = refl
compute-conjugation-Ω :
{x y : A} (p : x = y) → conjugation-Ω p ~∗ conjugation-Ω' p
pr1 (compute-conjugation-Ω p) = htpy-compute-conjugation-Ω p
- pr2 (compute-conjugation-Ω p) = preserves-point-compute-conjugation-Ω p
+ pr2 (compute-conjugation-Ω p) = coherence-point-compute-conjugation-Ω p
```
diff --git a/src/synthetic-homotopy-theory/eckmann-hilton-argument.lagda.md b/src/synthetic-homotopy-theory/eckmann-hilton-argument.lagda.md
index c886b19c8c..34728cf419 100644
--- a/src/synthetic-homotopy-theory/eckmann-hilton-argument.lagda.md
+++ b/src/synthetic-homotopy-theory/eckmann-hilton-argument.lagda.md
@@ -298,13 +298,13 @@ module _
3-loop-eckmann-hilton-Ω² : type-Ω³ a
3-loop-eckmann-hilton-Ω² =
- map-equiv-pointed-equiv
+ map-pointed-equiv
( pointed-equiv-2-loop-pointed-identity (Ω (A , a)) (s ∙ s))
( eckmann-hilton-Ω² s s)
inv-3-loop-eckmann-hilton-Ω² : type-Ω³ a
inv-3-loop-eckmann-hilton-Ω² =
- map-equiv-pointed-equiv
+ map-pointed-equiv
( pointed-equiv-2-loop-pointed-identity (Ω (A , a)) (s ∙ s))
( inv-eckmann-hilton-Ω² s s)
```
diff --git a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md
index 3b49db32e9..fad6799a5d 100644
--- a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md
+++ b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md
@@ -137,7 +137,7 @@ module _
equiv-map-Ω-pointed-equiv =
equiv-map-Ω-is-emb
( pointed-map-pointed-equiv e)
- ( is-emb-is-equiv (is-equiv-map-equiv-pointed-equiv e))
+ ( is-emb-is-equiv (is-equiv-map-pointed-equiv e))
```
### `pointed-map-Ω` preserves pointed equivalences
diff --git a/src/synthetic-homotopy-theory/multiplication-circle.lagda.md b/src/synthetic-homotopy-theory/multiplication-circle.lagda.md
index 99ce5de1b9..7281f2cd04 100644
--- a/src/synthetic-homotopy-theory/multiplication-circle.lagda.md
+++ b/src/synthetic-homotopy-theory/multiplication-circle.lagda.md
@@ -75,15 +75,7 @@ dependent-identification-Mul-Π-𝕊¹ :
(H : pr1 q ~ pr1 r) → Id (pr2 q ∙ p) (H base-𝕊¹ ∙ pr2 r) →
Id (tr Mul-Π-𝕊¹ p q) r
dependent-identification-Mul-Π-𝕊¹ {x} refl q r H u =
- eq-htpy-pointed-map
- ( q)
- ( r)
- ( ( H) ,
- (right-transpose-eq-concat
- ( H base-𝕊¹)
- ( pr2 r)
- ( pr2 q)
- ( inv (inv right-unit ∙ u))))
+ eq-pointed-htpy q r (H , inv right-unit ∙ u)
eq-id-id-𝕊¹-Pointed-Type :
Id (tr Mul-Π-𝕊¹ loop-𝕊¹ id-pointed-map) id-pointed-map
diff --git a/src/synthetic-homotopy-theory/null-cocones-under-pointed-span-diagrams.lagda.md b/src/synthetic-homotopy-theory/null-cocones-under-pointed-span-diagrams.lagda.md
new file mode 100644
index 0000000000..277a8788fb
--- /dev/null
+++ b/src/synthetic-homotopy-theory/null-cocones-under-pointed-span-diagrams.lagda.md
@@ -0,0 +1,167 @@
+# Null cocones under pointed span diagrams
+
+```agda
+module synthetic-homotopy-theory.null-cocones-under-pointed-span-diagrams where
+```
+
+Imports
+
+```agda
+open import foundation.action-on-identifications-functions
+open import foundation.commuting-squares-of-maps
+open import foundation.dependent-pair-types
+open import foundation.homotopies
+open import foundation.identity-types
+open import foundation.universe-levels
+open import foundation.whiskering-identifications-concatenation
+
+open import structured-types.commuting-squares-of-pointed-maps
+open import structured-types.constant-pointed-maps
+open import structured-types.pointed-homotopies
+open import structured-types.pointed-maps
+open import structured-types.pointed-span-diagrams
+open import structured-types.pointed-types
+
+open import synthetic-homotopy-theory.cocones-under-pointed-span-diagrams
+```
+
+
+
+## Idea
+
+The {{#concept "null cocone" Disambiguation="pointed span diagram"}} under a
+[pointed span diagram](structured-types.pointed-span-diagrams.md) `𝒮` given by
+
+```text
+ f g
+ A <---- S ----> B
+```
+
+with codomain `X` is the
+[cocone](synthetic-homotopy-theory.cocones-under-pointed-span-diagrams.md) under
+`𝒮` consisting of the
+[constant pointed maps](structured-types.constant-pointed-maps.md) `A →∗ X` and
+`B →∗ X` and the canonical homotopy witnessing that the square of pointed maps
+
+```text
+ g
+ S -----> B
+ | |
+ f | | const
+ ∨ ∨
+ A -----> X
+ const
+```
+
+[commutes](structured-types.commuting-squares-of-pointed-maps.md). The null
+cocone under `𝒮` provides a canonical pointing of the type
+`cocone-Pointed-Type f g`.
+
+## Definitions
+
+### Null cocones under pointed span diagrams
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} (𝒮 : pointed-span-diagram l1 l2 l3)
+ (X : Pointed-Type l4)
+ where
+
+ left-pointed-map-null-cocone-Pointed-Type :
+ pointed-domain-pointed-span-diagram 𝒮 →∗ X
+ left-pointed-map-null-cocone-Pointed-Type = constant-pointed-map _ X
+
+ left-map-null-cocone-Pointed-Type :
+ domain-pointed-span-diagram 𝒮 → type-Pointed-Type X
+ left-map-null-cocone-Pointed-Type =
+ map-pointed-map left-pointed-map-null-cocone-Pointed-Type
+
+ preserves-point-left-map-null-cocone-Pointed-Type :
+ left-map-null-cocone-Pointed-Type (point-domain-pointed-span-diagram 𝒮) =
+ point-Pointed-Type X
+ preserves-point-left-map-null-cocone-Pointed-Type =
+ preserves-point-pointed-map left-pointed-map-null-cocone-Pointed-Type
+
+ right-pointed-map-null-cocone-Pointed-Type :
+ pointed-codomain-pointed-span-diagram 𝒮 →∗ X
+ right-pointed-map-null-cocone-Pointed-Type = constant-pointed-map _ X
+
+ right-map-null-cocone-Pointed-Type :
+ codomain-pointed-span-diagram 𝒮 → type-Pointed-Type X
+ right-map-null-cocone-Pointed-Type =
+ map-pointed-map right-pointed-map-null-cocone-Pointed-Type
+
+ preserves-point-right-map-null-cocone-Pointed-Type :
+ right-map-null-cocone-Pointed-Type
+ ( point-codomain-pointed-span-diagram 𝒮) =
+ point-Pointed-Type X
+ preserves-point-right-map-null-cocone-Pointed-Type =
+ preserves-point-pointed-map right-pointed-map-null-cocone-Pointed-Type
+
+ htpy-coherence-square-null-cocone-Pointed-Type :
+ coherence-square-maps
+ ( map-pointed-map (right-pointed-map-pointed-span-diagram 𝒮))
+ ( map-pointed-map (left-pointed-map-pointed-span-diagram 𝒮))
+ ( map-constant-pointed-map (pointed-codomain-pointed-span-diagram 𝒮) X)
+ ( map-constant-pointed-map (pointed-domain-pointed-span-diagram 𝒮) X)
+ htpy-coherence-square-null-cocone-Pointed-Type = refl-htpy
+
+ coherence-point-coherence-square-null-cocone-Pointed-Type :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( constant-pointed-map _ X ∘∗ (left-pointed-map-pointed-span-diagram 𝒮))
+ ( constant-pointed-map _ X ∘∗ (right-pointed-map-pointed-span-diagram 𝒮))
+ ( htpy-coherence-square-null-cocone-Pointed-Type)
+ coherence-point-coherence-square-null-cocone-Pointed-Type =
+ right-whisker-concat
+ ( ( ap-const
+ ( point-Pointed-Type X)
+ ( preserves-point-left-map-pointed-span-diagram 𝒮)) ∙
+ ( inv
+ ( ap-const
+ ( point-Pointed-Type X)
+ ( preserves-point-right-map-pointed-span-diagram 𝒮))))
+ ( refl)
+
+ coherence-square-null-cocone-Pointed-Type :
+ coherence-square-pointed-maps
+ ( right-pointed-map-pointed-span-diagram 𝒮)
+ ( left-pointed-map-pointed-span-diagram 𝒮)
+ ( right-pointed-map-null-cocone-Pointed-Type)
+ ( left-pointed-map-null-cocone-Pointed-Type)
+ pr1 coherence-square-null-cocone-Pointed-Type =
+ htpy-coherence-square-null-cocone-Pointed-Type
+ pr2 coherence-square-null-cocone-Pointed-Type =
+ coherence-point-coherence-square-null-cocone-Pointed-Type
+
+ null-cocone-Pointed-Type :
+ cocone-Pointed-Type
+ ( left-pointed-map-pointed-span-diagram 𝒮)
+ ( right-pointed-map-pointed-span-diagram 𝒮)
+ ( X)
+ pr1 null-cocone-Pointed-Type =
+ left-pointed-map-null-cocone-Pointed-Type
+ pr1 (pr2 null-cocone-Pointed-Type) =
+ right-pointed-map-null-cocone-Pointed-Type
+ pr2 (pr2 null-cocone-Pointed-Type) =
+ coherence-square-null-cocone-Pointed-Type
+```
+
+### The pointed type of cocones under pointed span diagrams
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} (𝒮 : pointed-span-diagram l1 l2 l3)
+ (X : Pointed-Type l4)
+ where
+
+ type-cocone-pointed-type-Pointed-Type : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4)
+ type-cocone-pointed-type-Pointed-Type =
+ cocone-Pointed-Type
+ ( left-pointed-map-pointed-span-diagram 𝒮)
+ ( right-pointed-map-pointed-span-diagram 𝒮)
+ ( X)
+
+ cocone-pointed-type-Pointed-Type : Pointed-Type (l1 ⊔ l2 ⊔ l3 ⊔ l4)
+ pr1 cocone-pointed-type-Pointed-Type = type-cocone-pointed-type-Pointed-Type
+ pr2 cocone-pointed-type-Pointed-Type = null-cocone-Pointed-Type 𝒮 X
+```
diff --git a/src/synthetic-homotopy-theory/pushouts-of-pointed-types.lagda.md b/src/synthetic-homotopy-theory/pushouts-of-pointed-types.lagda.md
index cf1a1ffbe8..437ec7ff2c 100644
--- a/src/synthetic-homotopy-theory/pushouts-of-pointed-types.lagda.md
+++ b/src/synthetic-homotopy-theory/pushouts-of-pointed-types.lagda.md
@@ -15,7 +15,7 @@ open import foundation.universe-levels
open import structured-types.pointed-maps
open import structured-types.pointed-types
-open import synthetic-homotopy-theory.cocones-under-spans-of-pointed-types
+open import synthetic-homotopy-theory.cocones-under-pointed-span-diagrams
open import synthetic-homotopy-theory.pushouts
```
@@ -92,25 +92,25 @@ module _
{l4 : Level}
(f : S →∗ A) (g : S →∗ B) →
{X : Pointed-Type l4} →
- type-cocone-Pointed-Type f g X →
+ cocone-Pointed-Type f g X →
type-Pointed-Type (pushout-Pointed-Type f g) → type-Pointed-Type X
map-cogap-Pointed-Type f g c =
cogap
( map-pointed-map f)
( map-pointed-map g)
- ( cocone-type-cocone-Pointed-Type f g c)
+ ( cocone-cocone-Pointed-Type f g c)
cogap-Pointed-Type :
{l4 : Level}
(f : S →∗ A) (g : S →∗ B) →
{X : Pointed-Type l4} →
- type-cocone-Pointed-Type f g X → pushout-Pointed-Type f g →∗ X
+ cocone-Pointed-Type f g X → pushout-Pointed-Type f g →∗ X
pr1 (cogap-Pointed-Type f g c) = map-cogap-Pointed-Type f g c
pr2 (cogap-Pointed-Type f g {X} c) =
( compute-inl-cogap
( map-pointed-map f)
( map-pointed-map g)
- ( cocone-type-cocone-Pointed-Type f g c)
+ ( cocone-cocone-Pointed-Type f g c)
( point-Pointed-Type A)) ∙
( preserves-point-pointed-map
( horizontal-pointed-map-cocone-Pointed-Type f g c))
@@ -123,7 +123,7 @@ module _
{ l1 l2 l3 l4 : Level}
{S : Pointed-Type l1} {A : Pointed-Type l2} {B : Pointed-Type l3}
( f : S →∗ A) (g : S →∗ B)
- { X : Pointed-Type l4} (c : type-cocone-Pointed-Type f g X)
+ { X : Pointed-Type l4} (c : cocone-Pointed-Type f g X)
where
compute-inl-cogap-Pointed-Type :
@@ -138,7 +138,7 @@ module _
compute-inl-cogap
( map-pointed-map f)
( map-pointed-map g)
- ( cocone-type-cocone-Pointed-Type f g c)
+ ( cocone-cocone-Pointed-Type f g c)
compute-inr-cogap-Pointed-Type :
( y : type-Pointed-Type B) →
@@ -152,5 +152,5 @@ module _
compute-inr-cogap
( map-pointed-map f)
( map-pointed-map g)
- ( cocone-type-cocone-Pointed-Type f g c)
+ ( cocone-cocone-Pointed-Type f g c)
```
diff --git a/src/synthetic-homotopy-theory/smash-products-of-pointed-types.lagda.md b/src/synthetic-homotopy-theory/smash-products-of-pointed-types.lagda.md
index 209908dca1..9be1a3bf29 100644
--- a/src/synthetic-homotopy-theory/smash-products-of-pointed-types.lagda.md
+++ b/src/synthetic-homotopy-theory/smash-products-of-pointed-types.lagda.md
@@ -17,13 +17,14 @@ open import foundation.universe-levels
open import foundation.whiskering-homotopies-composition
open import foundation.whiskering-identifications-concatenation
+open import structured-types.constant-pointed-maps
open import structured-types.pointed-cartesian-product-types
open import structured-types.pointed-homotopies
open import structured-types.pointed-maps
open import structured-types.pointed-types
open import structured-types.pointed-unit-type
-open import synthetic-homotopy-theory.cocones-under-spans-of-pointed-types
+open import synthetic-homotopy-theory.cocones-under-pointed-span-diagrams
open import synthetic-homotopy-theory.pushouts
open import synthetic-homotopy-theory.pushouts-of-pointed-types
open import synthetic-homotopy-theory.wedges-of-pointed-types
@@ -74,7 +75,7 @@ and the [asterisk operator](https://codepoints.net/U+2217) `∗` (agda-input:
cogap-smash-product-Pointed-Type :
{l1 l2 l3 : Level}
{A : Pointed-Type l1} {B : Pointed-Type l2} {X : Pointed-Type l3} →
- type-cocone-Pointed-Type
+ cocone-Pointed-Type
( pointed-map-product-wedge-Pointed-Type A B)
( terminal-pointed-map (A ∨∗ B)) X →
(A ∧∗ B) →∗ X
@@ -86,7 +87,7 @@ cogap-smash-product-Pointed-Type {A = A} {B} =
map-cogap-smash-product-Pointed-Type :
{l1 l2 l3 : Level}
{A : Pointed-Type l1} {B : Pointed-Type l2} {X : Pointed-Type l3} →
- type-cocone-Pointed-Type
+ cocone-Pointed-Type
( pointed-map-product-wedge-Pointed-Type A B)
( terminal-pointed-map (A ∨∗ B))
( X) →
@@ -165,10 +166,6 @@ module _
( x)) ∙
( right-whisker-comp
( htpy-pointed-htpy
- ( inr-pushout-Pointed-Type
- ( pointed-map-product-wedge-Pointed-Type A B)
- ( terminal-pointed-map (A ∨∗ B)))
- ( inclusion-point-Pointed-Type (A ∧∗ B))
( is-initial-unit-Pointed-Type
( A ∧∗ B)
( inr-pushout-Pointed-Type
@@ -190,8 +187,7 @@ module _
coh-contraction-map-smash-product-wedge-Pointed-Type =
( map-inv-compute-dependent-identification-eq-value-function
( map-smash-product-wedge-Pointed-Type)
- ( map-pointed-map
- ( constant-pointed-map (A ∨∗ B) (A ∧∗ B)))
+ ( map-constant-pointed-map (A ∨∗ B) (A ∧∗ B))
( glue-wedge-Pointed-Type A B)
( contraction-map-smash-product-wedge-Pointed-Type
( map-inl-wedge-Pointed-Type A B (point-Pointed-Type A)))
@@ -311,7 +307,7 @@ module _
( inclusion-point-Pointed-Type A))
( map-pointed-map
( inclusion-point-Pointed-Type B))
- ( cocone-type-cocone-Pointed-Type
+ ( cocone-cocone-Pointed-Type
( inclusion-point-Pointed-Type A)
( inclusion-point-Pointed-Type B)
( cocone-product-wedge-Pointed-Type A B))
@@ -336,41 +332,83 @@ language of type theory, this means that we have a pointed equivalence:
[pointed cartesian product](structured-types.pointed-cartesian-product-types.md).
```agda
-map-universal-property-smash-product-Pointed-Type :
- {l1 l2 l3 : Level}
- (A : Pointed-Type l1) (B : Pointed-Type l2) (C : Pointed-Type l3) →
- ((A ∧∗ B) →∗ C) → (type-Pointed-Type A) → (B →∗ C)
-pr1 (map-universal-property-smash-product-Pointed-Type A B C f x) y =
- map-pointed-map f (map-smash-product-product-Pointed-Type A B (x , y))
-pr2 (map-universal-property-smash-product-Pointed-Type A B C f x) =
- ( ap
- ( map-pointed-map f)
- ( inl-glue-smash-product-Pointed-Type A B x)) ∙
- ( preserves-point-pointed-map f)
-
-universal-property-smash-product-Pointed-Type :
+module _
{l1 l2 l3 : Level}
- (A : Pointed-Type l1) (B : Pointed-Type l2) (C : Pointed-Type l3) →
- ((A ∧∗ B) →∗ C) → (A →∗ (pointed-map-Pointed-Type B C))
-pr1 (universal-property-smash-product-Pointed-Type A B C f) =
- map-universal-property-smash-product-Pointed-Type A B C f
-pr2 (universal-property-smash-product-Pointed-Type A B C f) =
- eq-htpy-pointed-Π
- ( map-universal-property-smash-product-Pointed-Type A B C
- ( f)
- ( point-Pointed-Type A))
- ( constant-pointed-map B C)
- ( ( λ y →
- ( ap
- ( map-pointed-map f)
- ( inr-glue-smash-product-Pointed-Type A B y) ∙
- ( preserves-point-pointed-map f))) ,
- ( ( right-whisker-concat
- ( ap²
- ( map-pointed-map f)
- ( inv (coh-glue-smash-product-Pointed-Type A B)))
- ( preserves-point-pointed-map f)) ∙
- ( inv right-unit)))
+ (A : Pointed-Type l1) (B : Pointed-Type l2) (C : Pointed-Type l3)
+ (f : (A ∧∗ B) →∗ C)
+ where
+
+ map-map-universal-property-smash-product-Pointed-Type :
+ type-Pointed-Type A → type-Pointed-Type B → type-Pointed-Type C
+ map-map-universal-property-smash-product-Pointed-Type x y =
+ map-pointed-map f (map-smash-product-product-Pointed-Type A B (x , y))
+
+ preserves-point-map-map-universal-property-smash-product-Pointed-Type :
+ (x : type-Pointed-Type A) →
+ map-map-universal-property-smash-product-Pointed-Type x
+ ( point-Pointed-Type B) =
+ point-Pointed-Type C
+ preserves-point-map-map-universal-property-smash-product-Pointed-Type x =
+ ( ap
+ ( map-pointed-map f)
+ ( inl-glue-smash-product-Pointed-Type A B x)) ∙
+ ( preserves-point-pointed-map f)
+
+ map-universal-property-smash-product-Pointed-Type :
+ type-Pointed-Type A → (B →∗ C)
+ pr1 (map-universal-property-smash-product-Pointed-Type x) =
+ map-map-universal-property-smash-product-Pointed-Type x
+ pr2 (map-universal-property-smash-product-Pointed-Type x) =
+ preserves-point-map-map-universal-property-smash-product-Pointed-Type x
+
+ htpy-preserves-point-map-universal-property-smash-product-Pointed-Type :
+ map-map-universal-property-smash-product-Pointed-Type
+ ( point-Pointed-Type A) ~
+ map-constant-pointed-map B C
+ htpy-preserves-point-map-universal-property-smash-product-Pointed-Type y =
+ ( ap (map-pointed-map f) (inr-glue-smash-product-Pointed-Type A B y)) ∙
+ ( preserves-point-pointed-map f)
+
+ coherence-point-preserves-point-map-universal-property-smash-product-Pointed-Type :
+ coherence-point-unpointed-htpy-pointed-Π
+ ( map-universal-property-smash-product-Pointed-Type
+ ( point-Pointed-Type A))
+ ( constant-pointed-map B C)
+ ( htpy-preserves-point-map-universal-property-smash-product-Pointed-Type)
+ coherence-point-preserves-point-map-universal-property-smash-product-Pointed-Type =
+ ( right-whisker-concat
+ ( ap²
+ ( map-pointed-map f)
+ ( coh-glue-smash-product-Pointed-Type A B))
+ ( preserves-point-pointed-map f)) ∙
+ ( inv right-unit)
+
+ pointed-htpy-preserves-point-map-universal-property-smash-product-Pointed-Type :
+ map-universal-property-smash-product-Pointed-Type (point-Pointed-Type A) ~∗
+ constant-pointed-map B C
+ pr1
+ pointed-htpy-preserves-point-map-universal-property-smash-product-Pointed-Type =
+ htpy-preserves-point-map-universal-property-smash-product-Pointed-Type
+ pr2
+ pointed-htpy-preserves-point-map-universal-property-smash-product-Pointed-Type =
+ coherence-point-preserves-point-map-universal-property-smash-product-Pointed-Type
+
+ preserves-point-map-universal-property-smash-product-Pointed-Type :
+ map-universal-property-smash-product-Pointed-Type (point-Pointed-Type A) =
+ constant-pointed-map B C
+ preserves-point-map-universal-property-smash-product-Pointed-Type =
+ eq-pointed-htpy
+ ( map-universal-property-smash-product-Pointed-Type
+ ( point-Pointed-Type A))
+ ( constant-pointed-map B C)
+ ( pointed-htpy-preserves-point-map-universal-property-smash-product-Pointed-Type)
+
+ pointed-map-universal-property-smash-product-Pointed-Type :
+ A →∗ (pointed-map-Pointed-Type B C)
+ pr1 pointed-map-universal-property-smash-product-Pointed-Type =
+ map-universal-property-smash-product-Pointed-Type
+ pr2 pointed-map-universal-property-smash-product-Pointed-Type =
+ preserves-point-map-universal-property-smash-product-Pointed-Type
```
## See also
diff --git a/src/synthetic-homotopy-theory/spectra.lagda.md b/src/synthetic-homotopy-theory/spectra.lagda.md
index 8ef4d35b61..ba8278a8ec 100644
--- a/src/synthetic-homotopy-theory/spectra.lagda.md
+++ b/src/synthetic-homotopy-theory/spectra.lagda.md
@@ -49,7 +49,7 @@ is-spectrum-Prop : {l : Level} → Prespectrum l → Prop l
is-spectrum-Prop A =
Π-Prop ℕ
( λ n →
- is-equiv-pointed-map-Prop (pointed-adjoint-structure-map-Prespectrum A n))
+ is-pointed-equiv-Prop (pointed-adjoint-structure-map-Prespectrum A n))
is-spectrum : {l : Level} → Prespectrum l → UU l
is-spectrum = type-Prop ∘ is-spectrum-Prop
@@ -98,7 +98,7 @@ module _
preserves-point-adjoint-structure-map-Prespectrum prespectrum-Spectrum
is-equiv-pointed-adjoint-structure-map-Spectrum :
- (n : ℕ) → is-equiv-pointed-map (pointed-adjoint-structure-map-Spectrum n)
+ (n : ℕ) → is-pointed-equiv (pointed-adjoint-structure-map-Spectrum n)
is-equiv-pointed-adjoint-structure-map-Spectrum = pr2 A
structure-equiv-Spectrum :
diff --git a/src/synthetic-homotopy-theory/wedges-of-pointed-types.lagda.md b/src/synthetic-homotopy-theory/wedges-of-pointed-types.lagda.md
index 86d9ae65eb..81a8394688 100644
--- a/src/synthetic-homotopy-theory/wedges-of-pointed-types.lagda.md
+++ b/src/synthetic-homotopy-theory/wedges-of-pointed-types.lagda.md
@@ -17,7 +17,7 @@ open import structured-types.pointed-maps
open import structured-types.pointed-types
open import structured-types.pointed-unit-type
-open import synthetic-homotopy-theory.cocones-under-spans-of-pointed-types
+open import synthetic-homotopy-theory.cocones-under-pointed-span-diagrams
open import synthetic-homotopy-theory.cofibers
open import synthetic-homotopy-theory.pushouts
open import synthetic-homotopy-theory.pushouts-of-pointed-types
@@ -130,7 +130,7 @@ module _
where
cocone-product-wedge-Pointed-Type :
- type-cocone-Pointed-Type
+ cocone-Pointed-Type
( inclusion-point-Pointed-Type A)
( inclusion-point-Pointed-Type B)
( A ×∗ B)