The category of quivers #
The category of (bundled) quivers, and the free/forgetful adjunction between Cat
and Quiv
.
Category of quivers.
Equations
Instances For
Equations
Equations
Category structure on Quiv
Equations
- One or more equations did not get rendered due to their size.
The forgetful functor from categories to quivers.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The identity in the category of quivers equals the identity prefunctor.
Composition in the category of quivers equals prefunctor composition.
A prefunctor V ⥤q W
induces a functor between the path categories defined by F.mapPath
.
Equations
- CategoryTheory.Cat.freeMap F = { obj := F.obj, map := fun {X Y : CategoryTheory.Paths V} => F.mapPath, map_id := ⋯, map_comp := ⋯ }
Instances For
The functor free : Quiv ⥤ Cat
preserves identities up to natural isomorphism and in fact up
to equality.
Equations
- CategoryTheory.Cat.freeMapIdIso V = CategoryTheory.NatIso.ofComponents (fun (x : CategoryTheory.Paths V) => CategoryTheory.Iso.refl ((CategoryTheory.Cat.freeMap (𝟭q V)).obj x)) ⋯
Instances For
The functor free : Quiv ⥤ Cat
preserves composition up to natural isomorphism and in fact up
to equality.
Equations
- CategoryTheory.Cat.freeMapCompIso F G = CategoryTheory.NatIso.ofComponents (fun (x : CategoryTheory.Paths V₁) => CategoryTheory.Iso.refl ((CategoryTheory.Cat.freeMap (F ⋙q G)).obj x)) ⋯
Instances For
The functor sending each quiver to its path category.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Any prefunctor into a category lifts to a functor from the path category.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Naturality of pathComposition
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Naturality of pathComposition
, which defines a natural transformation
Quiv.forget ⋙ Cat.free ⟶ 𝟭 _
.
Naturality of Paths.of
, which defines a natural transformation
𝟭 _⟶ Cat.free ⋙ Quiv.forget
.
The left triangle identity of Cat.free ⊣ Quiv.forget
as a natural isomorphism
Equations
- One or more equations did not get rendered due to their size.
Instances For
An unbundled version of the right triangle equality.
The adjunction between forming the free category on a quiver, and forgetting a category to a quiver.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The universal property of the path category of a quiver.
Equations
- One or more equations did not get rendered due to their size.