Right homotopies in model categories #
We introduce the types PrepathObject.RightHomotopy
and PathObject.RightHomotopy
of homotopies between morphisms X ⟶ Y
relative to a (pre)path object of Y
.
Given two morphisms f
and g
, we introduce the relation RightHomotopyRel f g
asserting the existence of a path object P
and
a right homotopy P.RightHomotopy f g
, and we define the quotient
type RightHomotopyClass X Y
. We show that if Y
is a fibrant
object in a model category, then RightHomotopyRel
is an equivalence
relation on X ⟶ Y
.
(This file dualizes the definitions in AlgebraicTopology.ModelCategory.LeftHomotopy
.)
References #
Given a pre-path object P
for Y
, two maps f
and g
in X ⟶ Y
are
homotopic relative to P
when there is a morphism h : X ⟶ P.P
such that h ≫ P.p₀ = f
and h ≫ P.p₁ = g
.
a morphism from the source to the pre-path object
Instances For
f : X ⟶ Y
is right homotopic to itself relative to any pre-path object.
Equations
- HomotopicalAlgebra.PrepathObject.RightHomotopy.refl P f = { h := CategoryTheory.CategoryStruct.comp f P.ι, h₀ := ⋯, h₁ := ⋯ }
Instances For
If f
and g
are homotopic relative to a pre-path object P
, then g
and f
are homotopic relative to P.symm
Instances For
If f₀
is homotopic to f₁
relative to a pre-path object P
,
and f₁
is homotopic to f₂
relative to P'
, then
f₀
is homotopic to f₂
relative to P.trans P'
.
Instances For
Right homotopies are compatible with precomposition.
Instances For
Given a path object P
for X
, two maps f
and g
in X ⟶ Y
are homotopic relative to P
when there is a morphism h : P.I ⟶ Y
such that P.i₀ ≫ h = f
and P.i₁ ≫ h = g
.
Equations
- P.RightHomotopy f g = P.RightHomotopy f g
Instances For
f : X ⟶ Y
is right homotopic to itself relative to any path object.
Equations
Instances For
If f
and g
are homotopic relative to a path object P
, then g
and f
are homotopic relative to P.symm
.
Equations
Instances For
Right homotopies are compatible with precomposition.
Equations
Instances For
If f₀ : X ⟶ Y
is homotopic to f₁
relative to a path object P
,
and f₁
is homotopic to f₂
relative to a good path object P'
,
then f₀
is homotopic to f₂
relative to the path object P.trans P'
when Y
is fibrant.
Equations
Instances For
The homotopy extension theorem: if p : A ⟶ X
is a cofibration,
l₀ : X ⟶ B
is a morphism, if there is a right homotopy h
between
the composition f₀ := i ≫ l₀
and a morphism f₁ : A ⟶ B
,
then there exists a morphism l₁ : X ⟶ B
and a right homotopy h'
from
l₀
to l₁
which is compatible with h
(in particular, i ≫ l₁ = f₁
).
The right homotopy relation on morphisms in a category with weak equivalences.
Equations
- HomotopicalAlgebra.RightHomotopyRel f g = ∃ (P : HomotopicalAlgebra.PathObject Y), Nonempty (P.RightHomotopy f g)
Instances For
In a category with weak equivalences, this is the quotient of the type
of morphisms X ⟶ Y
by the equivalence relation generated by right homotopies.
Instances For
Given f : X ⟶ Y
, this is the class of f
in the quotient RightHomotopyClass X Y
.
Instances For
The precomposition map RightHomotopyClass Y Z → (X ⟶ Y) → RightHomotopyClass X Z
.
Equations
- g.precomp f = Quot.lift (fun (g : Y ⟶ Z) => HomotopicalAlgebra.RightHomotopyClass.mk (CategoryTheory.CategoryStruct.comp f g)) ⋯ g