Properties of objects in target categories of a pseudofunctor to Cat #
Given F : Pseudofunctor B Cat, we introduce a type F.ObjectProperty
which consists of properties P of objects for all categories F.obj X for X : B.
The typeclass P.IsClosedUnderMapObj expresses that this property
is preserved by the application of the functors F.map: this allows
to define a sub-pseudofunctor P.fullsubcategory : Pseudofunctor B Cat.
TODO (@joelriou) #
- Given a Grothendieck topology
Jon a categoryC, define a type classPseudofunctor.ObjectProperty.IsLocal P JextendingIsClosedUnderMapObjsaying that if an object locally satisfies the property, then it satisfies the property. Assuming this, show thatP.fullsubcategoryif a stack if the original pseudofunctor was.
If F : Pseudofunctor B Cat, this is the data of a property of
objects in all categories F.obj X for X : B.
- prop (X : B) : ObjectProperty ↑(F.obj X)
A property of objects in the category
F.obj Xfor allX : B.
Instances For
Given F : Pseudofunctor B Cat, P : F.ObjectProperty and X : B, this is
the full subcategory of F.obj X consisting of the objects satisfying the
property P.
Equations
- P.Obj X = (P.prop X).FullSubcategory
Instances For
If P is a property of objects for a pseudofunctor F to Cat,
this is the condition that P is preserved by the application of the functors F.obj.
Instances
If P is a property of objects for a pseudofunctor F to Cat, this is the
condition that all P.prop : ObjectProprety (F.obj X) for X : B are closed
under isomorphisms.
- isClosedUnderIsomorphisms (X : B) : (P.prop X).IsClosedUnderIsomorphisms
Instances
Given a property P of objects for F : Pseudofunctor B Cat and a morphism f : X ⟶ Y
in B, this is the functor P.Obj X ⥤ P.Obj Y that is induced by F.map f.
Instances For
Given a property P of objects for F : Pseudofunctor B Cat and
a 2-morphism in B, this is the induced natural transformation between
the induced functors on the fullsubcategories of objects satisfying P.
Equations
- P.map₂ α = ((P.prop Y).fullyFaithfulι.whiskeringRight (P.Obj X)).preimage ((P.prop X).ι.whiskerLeft (F.map₂ α).toNatTrans)
Instances For
Auxiliary definition for fullsubcategory.
Equations
- P.mapId X = ((P.prop X).fullyFaithfulι.whiskeringRight (P.Obj X)).preimageIso ((P.prop X).ι.isoWhiskerLeft (CategoryTheory.Cat.Hom.toNatIso (F.mapId X)))
Instances For
Auxiliary definition for fullsubcategory.
Equations
- P.mapComp f g = ((P.prop Z).fullyFaithfulι.whiskeringRight (P.Obj X)).preimageIso ((P.prop X).ι.isoWhiskerLeft (CategoryTheory.Cat.Hom.toNatIso (F.mapComp f g)))
Instances For
Given a property of objects P for a pseudofunctor from B to Cat, this is
the induced pseudofunctor which sends X : B to the full subcategory of F.obj B
consisting of objects satisfying P.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The inclusion of P.fullsubcategory in F.
Equations
- One or more equations did not get rendered due to their size.