Path objects #
We introduce a notion of path object for an object A : C
in a model category.
It consists of an object P
, a weak equivalence ι : A ⟶ P
equipped with two retractions
p₀
and p₁
. This notion shall be important in the definition of "right homotopies"
in model categories.
This file dualizes the definitions in the file AlgebraicTopology.ModelCategory.Cylinder
.
Implementation notes #
The most important definition in this file is PathObject A
. This structure
extends another structure PrepathObject A
(which does not assume that C
has a notion of weak equivalences, which can be interesting in situations
where we have not yet obtained the model category axioms).
The good properties of path objects are stated as typeclasses PathObject.IsGood
and PathObject.IsVeryGood
.
The existence of very good path objects in model categories is stated
in the lemma PathObject.exists_very_good
.
References #
- Daniel G. Quillen, Homotopical algebra
- https://ncatlab.org/nlab/show/path+space+object
A pre-path object for A : C
is the data of a morphism
ι : A ⟶ P
equipped with two retractions.
- P : C
the underlying object of a (pre)path object
the first "projection" from the (pre)path object
the second "projection" from the (pre)path object
the diagonal of the (pre)path object
Instances For
The pre-path object obtained by switching the two projections.
Instances For
The gluing of two pre-path objects.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The map from P.P
to the product of two copies of A
, when P
is
a pre-path object for A
. P
shall be a good path object
when this morphism is a fibration.
Equations
- P.p = CategoryTheory.Limits.prod.lift P.p₀ P.p₁
Instances For
In a category with weak equivalences, a path object is the
data of a weak equivalence ι : A ⟶ P
equipped with two retractions.
- P : C
- weakEquivalence_ι : WeakEquivalence self.ι
Instances For
The path object obtained by switching the two projections.
Instances For
A path object P
is good if the morphism
P.p : P.P ⟶ A ⨯ A
is a fibration.
Instances
A good path object P
is very good if P.ι
is a (trivial) cofibration.
- fibration_p : Fibration P.p
- cofibration_ι : Cofibration P.ι
Instances
A path object for A
can be obtained from a factorization of the obvious
map A ⟶ A ⨯ A
as a trivial cofibration followed by a fibration.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The gluing of two good path objects.