Cardinals that are suitable for the small object argument #
In this file, given a class of morphisms I : MorphismProperty C
and
a regular cardinal κ : Cardinal.{w}
, we define a typeclass
IsCardinalForSmallObjectArgument I κ
which requires certain
smallness properties (I
is w
-small, C
is locally w
-small),
the existence of certain colimits (pushouts, coproducts of size w
,
and the condition HasIterationOfShape κ.ord.toType C
about the
existence of colimits indexed by limit ordinal smaller than or equal
to κ.ord
), and the technical assumption that if A
is the
a morphism in I
, then the functor Hom(A, _)
should commute
with the filtering colimits corresponding to relative
I
-cell complexes. (This last condition shall hold when κ
is the successor of an infinite cardinal c
such that all
these objects A
are c
-presentable, see the file Mathlib.CategoryTheory.Presentable.Basic
.)
Given I : MorphismProperty C
, we shall say that I
permits
the small object argument if there exists κ
such that
IsCardinalForSmallObjectArgument I κ
holds. See the file
Mathlib.CategoryTheory.SmallObject.Basic
for the definition of this typeclass
HasSmallObjectArgument
and an outline of the proof.
Main results #
Assuming IsCardinalForSmallObjectArgument I κ
, any morphism f : X ⟶ Y
is factored as ιObj I κ f ≫ πObj I κ f = f
. It is shown that ιObj I κ f
is a relative I
-cell complex (see SmallObject.relativeCellComplexιObj
)
and that πObj I κ f
has the right lifting property with respect to I
(see SmallObject.rlp_πObj
). This construction is obtained by
iterating to the power κ.ord.toType
the functor Arrow C ⥤ Arrow C
defined
in the file Mathlib.CategoryTheory.SmallObject.Construction
.
This factorization is functorial in f
and gives the property HasFunctorialFactorization I.rlp.llp I.rlp
.
Finally, the lemma llp_rlp_of_isCardinalForSmallObjectArgument
(and its primed version) shows that the morphisms in I.rlp.llp
are exactly
the retracts of the transfinite compositions (of shape κ.ord.toType
) of
pushouts of coproducts of morphisms in I
.
References #
- https://ncatlab.org/nlab/show/small+object+argument
Given I : MorphismProperty C
and a regular cardinal κ : Cardinal.{w}
,
this property asserts the technical conditions which allow to proceed
to the small object argument by doing a construction by transfinite
induction indexed by the well ordered type κ.ord.toType
.
- isSmall : IsSmall.{w, v, u} I
- locallySmall : LocallySmall.{w, v, u} C
- hasPushouts : Limits.HasPushouts C
- hasCoproducts : Limits.HasCoproducts C
- hasIterationOfShape : Limits.HasIterationOfShape κ.ord.toType C
- preservesColimit {A B X Y : C} (i : A ⟶ B) : I i → ∀ (f : X ⟶ Y) (hf : HomotopicalAlgebra.RelativeCellComplex (fun (x : κ.ord.toType) => I.homFamily) f), Limits.PreservesColimit hf.F (coyoneda.obj (Opposite.op A))
Instances
The successor structure on Arrow C ⥤ Arrow C
corresponding
to the iterations of the natural transformation
ε : 𝟭 (Arrow C) ⟶ SmallObject.functor I.homFamily
(see the file SmallObject.Construction
).
Equations
Instances For
For the successor structure succStruct I κ
on Arrow C ⥤ Arrow C
,
the morphism from an object to its successor induces
morphisms in C
which consists in attaching I
-cells.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The class of morphisms in Arrow C
which on the left side are
pushouts of coproducts of morphisms in I
, and which are
isomorphisms on the right side.
Equations
Instances For
The functor κ.ord.toType ⥤ Arrow C ⥤ Arrow C
corresponding to the
iterations of the successor structure succStruct I κ
.
Equations
Instances For
The colimit of iterationFunctor I κ
.
Equations
Instances For
The natural "inclusion" 𝟭 (Arrow C) ⟶ iteration I κ
.
Equations
Instances For
The morphism ιIteration I κ
is a transfinite composition of shape
κ.ord.toType
of morphisms satisfying (succStruct I κ).prop
.
Equations
Instances For
For any f : Arrow C
, the map ((ιIteration I κ).app f).right
is
a transfinite composition of isomorphisms.
Equations
- One or more equations did not get rendered due to their size.
Instances For
For any f : Arrow C
, the object ((iteration I κ).obj f).right
identifies to f.right
.
Equations
Instances For
For any f : Arrow C
and j : κ.ord.toType
, the object
(((iterationFunctor I κ).obj j).obj f).right
identifies to f.right
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
For any f : Arrow C
and j : κ.ord.toType
, the morphism
((iterationFunctor I κ).map (homOfLE (Order.le_succ j))).app f
identifies
to a morphism given by SmallObject.ε I.homFamily
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The intermediate object in the factorization given by the small object argument.
Equations
Instances For
The "inclusion" morphism in the factorization given by the the small object argument.
Equations
Instances For
The "projection" morphism in the factorization given by the the small object argument.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The morphism ιObj I κ f
is a relative I
-cell complex.
Equations
- One or more equations did not get rendered due to their size.
Instances For
When ιObj I κ f
is considered as a relative I
-cell complex,
the object at the j
th step is obtained by applying the construction
SmallObject.functorObj
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The functoriality of the intermediate object in the factorization of the small object argument.
Equations
- CategoryTheory.SmallObject.objMap I κ φ = ((CategoryTheory.SmallObject.iteration I κ).map φ).left
Instances For
The functorial factorization ιObj I κ f ≫ πObj I κ f.hom = f
with ιObj I κ f
in I.rlp.llp
and πObj I κ f.hom
in I.rlp
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
If κ
is a suitable cardinal for the small object argument for I : MorphismProperty C
,
then the class I.rlp.llp
is exactly the class of morphisms that are retracts
of transfinite compositions (of shape κ.ord.toType
) of pushouts of coproducts
of maps in I
.
If κ
is a suitable cardinal for the small object argument for I : MorphismProperty C
,
then the class I.rlp.llp
is exactly the class of morphisms that are retracts
of transfinite compositions of pushouts of coproducts of maps in I
.