Documentation

Mathlib.CategoryTheory.Groupoid.FreeGroupoidOfCategory

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:

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).

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
        @[reducible, inline]

        Construct an object in the free groupoid on C by providing an object in C.

        Equations
        Instances For
          @[reducible, inline]
          abbrev CategoryTheory.FreeGroupoid.homMk {C : Type u} [Category.{v, u} C] {X Y : C} (f : X Y) :
          mk X mk Y

          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
              theorem CategoryTheory.FreeGroupoid.lift_spec {C : Type u} [Category.{v, u} C] {G : Type u₁} [Groupoid G] (φ : Functor C G) :
              (of C).comp (lift φ) = φ
              @[simp]
              theorem CategoryTheory.FreeGroupoid.lift_obj_mk {C : Type u} [Category.{v, u} C] {E : Type u₂} [Groupoid E] (φ : Functor C E) (X : C) :
              (lift φ).obj (mk X) = φ.obj X
              @[simp]
              theorem CategoryTheory.FreeGroupoid.lift_map_homMk {C : Type u} [Category.{v, u} C] {E : Type u₂} [Groupoid E] (φ : Functor C E) {X Y : C} (f : X Y) :
              (lift φ).map (homMk f) = φ.map f
              theorem CategoryTheory.FreeGroupoid.lift_unique {C : Type u} [Category.{v, u} C] {G : Type u₁} [Groupoid G] (φ : Functor C G) (Φ : Functor (FreeGroupoid C) G) ( : (of C).comp Φ = φ) :
              Φ = lift φ
              theorem CategoryTheory.FreeGroupoid.lift_comp {C : Type u} [Category.{v, u} C] {G : Type u₁} [Groupoid G] {H : Type u₂} [Groupoid H] (φ : Functor C G) (ψ : Functor G H) :
              lift (φ.comp ψ) = (lift φ).comp ψ

              The universal property of the free groupoid.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                def CategoryTheory.FreeGroupoid.liftNatIso {C : Type u} [Category.{v, u} C] {G : Type u₁} [Groupoid G] (F₁ F₂ : Functor (FreeGroupoid C) G) (τ : (of C).comp F₁ (of C).comp F₂) :
                F₁ F₂

                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
                  @[simp]
                  theorem CategoryTheory.FreeGroupoid.liftNatIso_hom_app {C : Type u} [Category.{v, u} C] {G : Type u₁} [Groupoid G] (F₁ F₂ : Functor (FreeGroupoid C) G) (τ : (of C).comp F₁ (of C).comp F₂) (X : C) :
                  (liftNatIso F₁ F₂ τ).hom.app (mk X) = τ.hom.app X
                  @[simp]
                  theorem CategoryTheory.FreeGroupoid.liftNatIso_inv_app {C : Type u} [Category.{v, u} C] {G : Type u₁} [Groupoid G] (F₁ F₂ : Functor (FreeGroupoid C) G) (τ : (of C).comp F₁ (of C).comp F₂) (X : C) :
                  (liftNatIso F₁ F₂ τ).inv.app (mk X) = τ.inv.app X

                  The functor between free groupoids induced by a functor between categories.

                  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
                      def CategoryTheory.FreeGroupoid.mapComp {C : Type u} [Category.{v, u} C] {D : Type u₁} [Category.{v₁, u₁} D] {E : Type u₂} [Category.{v₂, u₂} E] (φ : Functor C D) (φ' : Functor D E) :
                      map (φ.comp φ') (map φ).comp (map φ')

                      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
                        @[simp]
                        theorem CategoryTheory.FreeGroupoid.mapComp_hom_app {C : Type u} [Category.{v, u} C] {D : Type u₁} [Category.{v₁, u₁} D] {E : Type u₂} [Category.{v₂, u₂} E] (φ : Functor C D) (φ' : Functor D E) (X : FreeGroupoid C) :
                        (mapComp φ φ').hom.app X = CategoryStruct.id ((map (φ.comp φ')).obj X)
                        @[simp]
                        theorem CategoryTheory.FreeGroupoid.mapComp_inv_app {C : Type u} [Category.{v, u} C] {D : Type u₁} [Category.{v₁, u₁} D] {E : Type u₂} [Category.{v₂, u₂} E] (φ : Functor C D) (φ' : Functor D E) (X : FreeGroupoid C) :
                        (mapComp φ φ').inv.app X = CategoryStruct.id (((map φ).comp (map φ')).obj X)
                        theorem CategoryTheory.FreeGroupoid.map_comp {C : Type u} [Category.{v, u} C] {D : Type u₁} [Category.{v₁, u₁} D] {E : Type u₂} [Category.{v₂, u₂} E] (φ : Functor C D) (φ' : Functor D E) :
                        map (φ.comp φ') = (map φ).comp (map φ')
                        @[simp]
                        theorem CategoryTheory.FreeGroupoid.map_obj_mk {C : Type u} [Category.{v, u} C] {D : Type u₁} [Category.{v₁, u₁} D] (φ : Functor C D) (X : C) :
                        (map φ).obj (mk X) = mk (φ.obj X)
                        @[simp]
                        theorem CategoryTheory.FreeGroupoid.map_map_homMk {C : Type u} [Category.{v, u} C] {D : Type u₁} [Category.{v₁, u₁} D] (φ : Functor C D) {X Y : C} (f : X Y) :
                        (map φ).map (homMk f) = homMk (φ.map f)
                        theorem CategoryTheory.FreeGroupoid.map_comp_lift {C : Type u} [Category.{v, u} C] {D : Type u₁} [Category.{v₁, u₁} D] {E : Type u₂} [Groupoid E] (F : Functor C D) (G : Functor D E) :
                        (map F).comp (lift G) = lift (F.comp G)
                        def CategoryTheory.FreeGroupoid.mapCompLift {C : Type u} [Category.{v, u} C] {D : Type u₁} [Category.{v₁, u₁} D] {E : Type u₂} [Groupoid E] (F : Functor C D) (G : Functor D E) :
                        (map F).comp (lift G) lift (F.comp G)

                        The operation lift is natural.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          @[simp]
                          theorem CategoryTheory.FreeGroupoid.mapCompLift_hom_app {C : Type u} [Category.{v, u} C] {D : Type u₁} [Category.{v₁, u₁} D] {E : Type u₂} [Groupoid E] (F : Functor C D) (G : Functor D E) (X : FreeGroupoid C) :
                          @[simp]

                          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
                              @[simp]
                              @[simp]

                              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.