Free groupoid on a category #
This file defines the free groupoid on a category, the lifting of a functor to its unique extension as a functor from the free groupoid, and proves uniqueness of this extension.
Main results #
Given a type C and a category instance on C:
CategoryTheory.FreeGroupoid C: the underlying type of the free groupoid onC.CategoryTheory.FreeGroupoid.instGroupoid: theGroupoidinstance onFreeGroupoid C.CategoryTheory.FreeGroupoid.lift: the lifting of a functorC ⥤ GwhereGis a groupoid, to a functorCategoryTheory.FreeGroupoid C ⥤ G.CategoryTheory.FreeGroupoid.lift_specandCategoryTheory.FreeGroupoid.lift_unique: the proofs that, respectively,CategoryTheory.FreeGroupoid.liftindeed is a lifting and is the unique one.CategoryTheory.Grpd.free: the free functor fromGrpdtoCatCategoryTheory.Grpd.freeForgetAdjunction: thatfreeis left adjoint toGrpd.forgetToCat.
Implementation notes #
The free groupoid on a category C is first defined by taking the free groupoid G
on the underlying quiver of C. Then the free groupoid on the category C is defined as
the quotient of G by the relation that makes the inclusion prefunctor C ⥤q G a functor.
The relation on the free groupoid on the underlying quiver of C that
promotes the prefunctor C ⥤q FreeGroupoid C into a functor
C ⥤ Quotient (FreeGroupoid.homRel C).
- map_id {C : Type u} [Category.{v, u} C] (X : C) : homRel C ((Quiver.FreeGroupoid.of C).map (CategoryStruct.id X)) (CategoryStruct.id ((Quiver.FreeGroupoid.of C).obj X))
- map_comp {C : Type u} [Category.{v, u} C] {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) : homRel C ((Quiver.FreeGroupoid.of C).map (CategoryStruct.comp f g)) (CategoryStruct.comp ((Quiver.FreeGroupoid.of C).map f) ((Quiver.FreeGroupoid.of C).map g))
Instances For
The underlying type of the free groupoid on a category,
defined by quotienting the free groupoid on the underlying quiver of C
by the relation that promotes the prefunctor C ⥤q FreeGroupoid C into a functor
C ⥤ Quotient (FreeGroupoid.homRel C).
Equations
Instances For
The localization functor from the category C to the groupoid FreeGroupoid C
Equations
- One or more equations did not get rendered due to their size.
Instances For
Construct an object in the free groupoid on C by providing an object in C.
Equations
Instances For
Construct a morphism in the free groupoid on C by providing a morphism in C.
Equations
Instances For
The lift of a functor from C to a groupoid to a functor from
FreeGroupoid C to the groupoid
Equations
Instances For
The universal property of the free groupoid.
Equations
- One or more equations did not get rendered due to their size.
Instances For
In order to define a natural isomorphism F ≅ G with F G : FreeGroupoid ⥤ D,
it suffices to do so after precomposing with FreeGroupoid.of C.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The functor between free groupoids induced by a functor between categories.
Equations
Instances For
The operation of is natural.
Equations
Instances For
The functor induced by the identity is the identity.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The functor induced by a composition is the composition of the functors they induce.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The operation lift is natural.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Functors out of the free groupoid biject with functors out of the original category.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The free groupoid construction on a category as a functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The free-forgetful adjunction between Grpd and Cat.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.