Documentation

Mathlib.CategoryTheory.Category.Quiv

The category of quivers #

The category of (bundled) quivers, and the free/forgetful adjunction between Cat and Quiv.

def CategoryTheory.Quiv :
Type (max (u + 1) u (v + 1))

Category of quivers.

Equations
Instances For
    instance CategoryTheory.Quiv.str' (C : Quiv) :
    Quiver C
    Equations

    Construct a bundled Quiv from the underlying type and the typeclass.

    Equations
    Instances For

      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
        @[simp]
        @[simp]
        theorem CategoryTheory.Quiv.forget_map {X✝ Y✝ : Cat} (F : X✝ Y✝) :

        The identity in the category of quivers equals the identity prefunctor.

        theorem CategoryTheory.Quiv.comp_eq_comp {X Y Z : Quiv} (F : X Y) (G : Y Z) :

        Composition in the category of quivers equals prefunctor composition.

        Prefunctors between quivers define arrows in Quiv.

        Equations
        Instances For
          def CategoryTheory.Prefunctor.ofQuivHom {C D : Quiv} (F : C D) :
          C ⥤q D

          Arrows in Quiv define prefunctors.

          Equations
          Instances For
            @[simp]
            def CategoryTheory.Cat.freeMap {V : Type u_1} {W : Type u_2} [Quiver V] [Quiver W] (F : V ⥤q W) :

            A prefunctor V ⥤q W induces a functor between the path categories defined by F.mapPath.

            Equations
            Instances For
              @[simp]
              theorem CategoryTheory.Cat.freeMap_obj {V : Type u_1} {W : Type u_2} [Quiver V] [Quiver W] (F : V ⥤q W) (a✝ : V) :
              (freeMap F).obj a✝ = F.obj a✝
              @[simp]
              theorem CategoryTheory.Cat.freeMap_map {V : Type u_1} {W : Type u_2} [Quiver V] [Quiver W] (F : V ⥤q W) {X✝ Y✝ : Paths V} (a✝ : Quiver.Path X✝ Y✝) :
              (freeMap F).map a✝ = F.mapPath a✝

              The functor free : Quiv ⥤ Cat preserves identities up to natural isomorphism and in fact up to equality.

              Equations
              Instances For
                def CategoryTheory.Cat.freeMapCompIso {V₁ : Type u₁} {V₂ : Type u₂} {V₃ : Type u₃} [Quiver V₁] [Quiver V₂] [Quiver V₃] (F : V₁ ⥤q V₂) (G : V₂ ⥤q V₃) :

                The functor free : Quiv ⥤ Cat preserves composition up to natural isomorphism and in fact up to equality.

                Equations
                Instances For
                  @[simp]
                  theorem CategoryTheory.Cat.freeMapCompIso_hom_app {V₁ : Type u₁} {V₂ : Type u₂} {V₃ : Type u₃} [Quiver V₁] [Quiver V₂] [Quiver V₃] (F : V₁ ⥤q V₂) (G : V₂ ⥤q V₃) (X : Paths V₁) :
                  @[simp]
                  theorem CategoryTheory.Cat.freeMapCompIso_inv_app {V₁ : Type u₁} {V₂ : Type u₂} {V₃ : Type u₃} [Quiver V₁] [Quiver V₂] [Quiver V₃] (F : V₁ ⥤q V₂) (G : V₂ ⥤q V₃) (X : Paths V₁) :
                  theorem CategoryTheory.Cat.freeMap_comp {V₁ : Type u₁} {V₂ : Type u₂} {V₃ : Type u₃} [Quiver V₁] [Quiver V₂] [Quiver V₃] (F : V₁ ⥤q V₂) (G : V₂ ⥤q V₃) :

                  The functor sending each quiver to its path category.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    theorem CategoryTheory.Cat.free_map {X✝ Y✝ : Quiv} (F : X✝ Y✝) :
                    @[simp]
                    def CategoryTheory.Quiv.equivOfIso {V W : Quiv} (e : V W) :
                    V W

                    An isomorphism of quivers defines an equivalence on carrier types.

                    Equations
                    Instances For
                      @[simp]
                      theorem CategoryTheory.Quiv.equivOfIso_symm_apply {V W : Quiv} (e : V W) (a✝ : W) :
                      (equivOfIso e).symm a✝ = e.inv.obj a✝
                      @[simp]
                      theorem CategoryTheory.Quiv.equivOfIso_apply {V W : Quiv} (e : V W) (a✝ : V) :
                      (equivOfIso e) a✝ = e.hom.obj a✝
                      @[simp]
                      theorem CategoryTheory.Quiv.inv_obj_hom_obj_of_iso {V W : Quiv} (e : V W) (X : V) :
                      e.inv.obj (e.hom.obj X) = X
                      @[simp]
                      theorem CategoryTheory.Quiv.hom_obj_inv_obj_of_iso {V W : Quiv} (e : V W) (Y : W) :
                      e.hom.obj (e.inv.obj Y) = Y
                      theorem CategoryTheory.Quiv.hom_map_inv_map_of_iso {V W : Quiv} (e : V W) {X Y : W} (f : X Y) :
                      e.hom.map (e.inv.map f) = Quiver.homOfEq f
                      theorem CategoryTheory.Quiv.inv_map_hom_map_of_iso {V W : Quiv} (e : V W) {X Y : V} (f : X Y) :
                      e.inv.map (e.hom.map f) = Quiver.homOfEq f
                      def CategoryTheory.Quiv.homEquivOfIso {V W : Quiv} (e : V W) {X Y : V} :
                      (X Y) (e.hom.obj X e.hom.obj Y)

                      An isomorphism of quivers defines an equivalence on hom types.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        @[simp]
                        theorem CategoryTheory.Quiv.homEquivOfIso_apply {V W : Quiv} (e : V W) {X Y : V} (f : X Y) :
                        @[simp]
                        theorem CategoryTheory.Quiv.homEquivOfIso_symm_apply {V W : Quiv} (e : V W) {X Y : V} (g : e.hom.obj X e.hom.obj Y) :
                        @[simp]
                        theorem CategoryTheory.Quiv.homOfEq_map_homOfEq {V W : Type u} [Quiver V] [Quiver W] (e : V W) (he : (X Y : V) → (X Y) (e X e Y)) {X Y : V} (f : X Y) {X' Y' : V} (hX : X = X') (hY : Y = Y') {X'' Y'' : W} (hX' : e X' = X'') (hY' : e Y' = Y'') :
                        Quiver.homOfEq ((he X' Y') (Quiver.homOfEq f hX hY)) hX' hY' = Quiver.homOfEq ((he X Y) f)
                        def CategoryTheory.Quiv.isoOfEquiv {V W : Type u} [Quiver V] [Quiver W] (e : V W) (he : (X Y : V) → (X Y) (e X e Y)) :
                        of V of W

                        Compatible equivalences of types and hom-types induce an isomorphism of quivers.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          def CategoryTheory.Quiv.lift {V : Type u} [Quiver V] {C : Type u₁} [Category.{v₁, u₁} C] (F : V ⥤q C) :

                          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
                            @[simp]
                            theorem CategoryTheory.Quiv.lift_map {V : Type u} [Quiver V] {C : Type u₁} [Category.{v₁, u₁} C] (F : V ⥤q C) {X✝ Y✝ : Paths V} (f : X✝ Y✝) :
                            @[simp]
                            theorem CategoryTheory.Quiv.lift_obj {V : Type u} [Quiver V] {C : Type u₁} [Category.{v₁, u₁} C] (F : V ⥤q C) (X : Paths V) :
                            (lift F).obj X = F.obj X

                            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.forgetCat.free ⟶ 𝟭 _.

                              Naturality of Paths.of, which defines a natural transformation 𝟭 _⟶ Cat.freeQuiv.forget.

                              The left triangle identity of Cat.freeQuiv.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.
                                  Instances For