Timings for NullHomotopy.v
Require Import HoTT.Basics.
Require Import Types.Sigma.
Local Open Scope path_scope.
(** * Null homotopies of maps *)
(** Geometrically, a nullhomotopy of a map [f : X -> Y] is an extension of [f] to a map [Cone X -> Y]. One might more simply call it e.g. [Constant f], but that is a little ambiguous: it could also reasonably mean e.g. a factorisation of [f] through [ Trunc -1 X ]. (Should the unique map [0 -> Y] be constant in one way, or in [Y]-many ways?) *)
Definition NullHomotopy {X Y : Type} (f : X -> Y)
:= {y : Y & forall x:X, f x = y}.
Lemma istrunc_nullhomotopy {n : trunc_index}
{X Y : Type} (f : X -> Y) `{IsTrunc n Y}
: IsTrunc n (NullHomotopy f).
apply @istrunc_sigma; auto.
apply (@istrunc_forall _).
Definition nullhomotopy_homotopic {X Y : Type} {f g : X -> Y} (p : f == g)
: NullHomotopy f -> NullHomotopy g.
intros x; exact ((p x)^ @ e x).
Definition nullhomotopy_composeR {X Y Z : Type} (f : X -> Y) (g : Y -> Z)
: NullHomotopy g -> NullHomotopy (g o f).
intros x; exact (e (f x)).
Definition nullhomotopy_composeL {X Y Z : Type} (f : X -> Y) (g : Y -> Z)
: NullHomotopy f -> NullHomotopy (g o f).
intros x; exact (ap g (e x)).
Definition cancelL_nullhomotopy_equiv
{X Y Z : Type} (f : X -> Y) (g : Y -> Z) `{IsEquiv _ _ g}
: NullHomotopy (g o f) -> NullHomotopy f.
intros x; apply moveL_equiv_V, e.
Definition cancelR_nullhomotopy_equiv
{X Y Z : Type} (f : X -> Y) (g : Y -> Z) `{IsEquiv _ _ f}
: NullHomotopy (g o f) -> NullHomotopy g.
intros y; transitivity (g (f (f^-1 y))).
symmetry; apply ap, eisretr.
Definition nullhomotopy_ap {X Y : Type} (f : X -> Y) (x1 x2 : X)
: NullHomotopy f -> NullHomotopy (@ap _ _ f x1 x2).
refine (concat_Ap e p @ _).
refine (_ @ concat_p1 _); apply ap.