Left homotopies in model categories #
We introduce the types Precylinder.LeftHomotopy
and Cylinder.LeftHomotopy
of homotopies between morphisms X ⟶ Y
relative to a (pre)cylinder of X
.
Given two morphisms f
and g
, we introduce the relation LeftHomotopyRel f g
asserting the existence of a cylinder object P
and
a left homotopy P.LeftHomotopy f g
, and we define the quotient
type LeftHomotopyClass X Y
. We show that if X
is a cofibrant
object in a model category, then LeftHomotopyRel
is an equivalence
relation on X ⟶ Y
.
References #
Given a precylinder 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
.
a morphism from the (pre)cylinder object to the target
Instances For
f : X ⟶ Y
is left homotopic to itself relative to any precylinder.
Equations
- HomotopicalAlgebra.Precylinder.LeftHomotopy.refl P f = { h := CategoryTheory.CategoryStruct.comp P.π f, h₀ := ⋯, h₁ := ⋯ }
Instances For
If f
and g
are homotopic relative to a precylinder P
, then g
and f
are homotopic relative to P.symm
Instances For
If f₀
is homotopic to f₁
relative to a precylinder P
,
and f₁
is homotopic to f₂
relative to P'
, then
f₀
is homotopic to f₂
relative to P.trans P'
.
Instances For
Left homotopies are compatible with postcomposition.
Instances For
Given a cylinder 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.LeftHomotopy f g = P.LeftHomotopy f g
Instances For
f : X ⟶ Y
is left homotopic to itself relative to any cylinder.
Equations
Instances For
If f
and g
are homotopic relative to a cylinder P
, then g
and f
are homotopic relative to P.symm
.
Equations
Instances For
Left homotopies are compatible with postcomposition.
Equations
Instances For
If f₀ : X ⟶ Y
is homotopic to f₁
relative to a cylinder P
,
and f₁
is homotopic to f₂
relative to a good cylinder P'
,
then f₀
is homotopic to f₂
relative to the cylinder P.trans P'
when X
is cofibrant.
Equations
Instances For
The covering homotopy theorem: if p : E ⟶ B
is a fibration,
l₀ : A ⟶ E
is a morphism, if there is a left homotopy h
between
the composition f₀ := l₀ ≫ p
and a morphism f₁ : A ⟶ B
,
then there exists a morphism l₁ : A ⟶ E
and a left homotopy h'
from
l₀
to l₁
which is compatible with h
(in particular, l₁ ≫ p = f₁
).
The left homotopy relation on morphisms in a category with weak equivalences.
Equations
- HomotopicalAlgebra.LeftHomotopyRel f g = ∃ (P : HomotopicalAlgebra.Cylinder X), Nonempty (P.LeftHomotopy 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 left homotopies.
Instances For
Given f : X ⟶ Y
, this is the class of f
in the quotient LeftHomotopyClass X Y
.
Instances For
The postcomposition map LeftHomotopyClass X Y → (Y ⟶ Z) → LeftHomotopyClass X Z
.
Equations
- f.postcomp g = Quot.lift (fun (f : X ⟶ Y) => HomotopicalAlgebra.LeftHomotopyClass.mk (CategoryTheory.CategoryStruct.comp f g)) ⋯ f