Documentation

Mathlib.CategoryTheory.PathCategory.Basic

The category paths on a quiver. #

When C is a quiver, paths C is the category of paths.

When the quiver is itself a category #

We provide path_composition : paths C ⥤ C.

We check that the quotient of the path category of a category by the canonical relation (paths are related if they compose to the same path) is equivalent to the original category.

def CategoryTheory.Paths (V : Type u₁) :
Type u₁

A type synonym for the category of paths in a quiver.

Equations
Instances For
    def CategoryTheory.Paths.of {V : Type u₁} [Quiver V] :

    The inclusion of a quiver V into its path category, as a prefunctor.

    Equations
    Instances For
      @[simp]
      theorem CategoryTheory.Paths.of_map {V : Type u₁} [Quiver V] {X✝ Y✝ : V} (f : X✝ Y✝) :
      of.map f = f.toPath
      @[simp]
      theorem CategoryTheory.Paths.of_obj {V : Type u₁} [Quiver V] (X : V) :
      of.obj X = X
      theorem CategoryTheory.Paths.induction_fixed_source {V : Type u₁} [Quiver V] {a : Paths V} (P : {b : Paths V} → (a b)Prop) (id : P (CategoryStruct.id a)) (comp : ∀ {u v : V} (p : a of.obj u) (q : u v), P pP (CategoryStruct.comp p (of.map q))) {b : Paths V} (f : a b) :
      P f

      To prove a property on morphisms of a path category with given source a, it suffices to prove it for the identity and prove that the property is preserved under composition on the right with length 1 paths.

      theorem CategoryTheory.Paths.induction_fixed_target {V : Type u₁} [Quiver V] {b : Paths V} (P : {a : Paths V} → (a b)Prop) (id : P (CategoryStruct.id b)) (comp : ∀ {u v : V} (p : of.obj v b) (q : u v), P pP (CategoryStruct.comp (of.map q) p)) {a : Paths V} (f : a b) :
      P f

      To prove a property on morphisms of a path category with given target b, it suffices to prove it for the identity and prove that the property is preserved under composition on the left with length 1 paths.

      theorem CategoryTheory.Paths.induction {V : Type u₁} [Quiver V] (P : {a b : Paths V} → (a b)Prop) (id : ∀ {v : V}, P (CategoryStruct.id (of.obj v))) (comp : ∀ {u v w : V} (p : of.obj u of.obj v) (q : v w), P pP (CategoryStruct.comp p (of.map q))) {a b : Paths V} (f : a b) :
      P f

      To prove a property on morphisms of a path category, it suffices to prove it for the identity and prove that the property is preserved under composition on the right with length 1 paths.

      theorem CategoryTheory.Paths.induction' {V : Type u₁} [Quiver V] (P : {a b : Paths V} → (a b)Prop) (id : ∀ {v : V}, P (CategoryStruct.id (of.obj v))) (comp : ∀ {u v w : V} (p : u v) (q : of.obj v of.obj w), P qP (CategoryStruct.comp (of.map p) q)) {a b : Paths V} (f : a b) :
      P f

      To prove a property on morphisms of a path category, it suffices to prove it for the identity and prove that the property is preserved under composition on the left with length 1 paths.

      def CategoryTheory.Paths.lift {V : Type u₁} [Quiver V] {C : Type u_1} [Category.{u_2, u_1} C] (φ : V ⥤q C) :

      Any prefunctor from V lifts to a functor from paths V

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        @[simp]
        theorem CategoryTheory.Paths.lift_nil {V : Type u₁} [Quiver V] {C : Type u_1} [Category.{u_2, u_1} C] (φ : V ⥤q C) (X : V) :
        @[simp]
        theorem CategoryTheory.Paths.lift_cons {V : Type u₁} [Quiver V] {C : Type u_1} [Category.{u_2, u_1} C] (φ : V ⥤q C) {X Y Z : V} (p : Quiver.Path X Y) (f : Y Z) :
        (lift φ).map (p.cons f) = CategoryStruct.comp ((lift φ).map p) (φ.map f)
        @[simp]
        theorem CategoryTheory.Paths.lift_toPath {V : Type u₁} [Quiver V] {C : Type u_1} [Category.{u_2, u_1} C] (φ : V ⥤q C) {X Y : V} (f : X Y) :
        (lift φ).map f.toPath = φ.map f
        theorem CategoryTheory.Paths.lift_spec {V : Type u₁} [Quiver V] {C : Type u_1} [Category.{u_2, u_1} C] (φ : V ⥤q C) :
        of ⋙q (lift φ).toPrefunctor = φ
        theorem CategoryTheory.Paths.lift_unique {V : Type u₁} [Quiver V] {C : Type u_1} [Category.{u_2, u_1} C] (φ : V ⥤q C) (Φ : Functor (Paths V) C) (hΦ : of ⋙q Φ.toPrefunctor = φ) :
        Φ = lift φ
        theorem CategoryTheory.Paths.ext_functor {V : Type u₁} [Quiver V] {C : Type u_1} [Category.{u_2, u_1} C] {F G : Functor (Paths V) C} (h_obj : F.obj = G.obj) (h : ∀ (a b : V) (e : a b), F.map e.toPath = CategoryStruct.comp (eqToHom ) (CategoryStruct.comp (G.map e.toPath) (eqToHom ))) :
        F = G

        Two functors out of a path category are equal when they agree on singleton paths.

        @[simp]
        theorem CategoryTheory.Prefunctor.mapPath_comp' (V : Type u₁) [Quiver V] (W : Type u₂) [Quiver W] (F : V ⥤q W) {X Y Z : Paths V} (f : X Y) (g : Y Z) :
        F.mapPath (CategoryStruct.comp f g) = (F.mapPath f).comp (F.mapPath g)
        @[simp]
        theorem CategoryTheory.composePath_cons {C : Type u₁} [Category.{v₁, u₁} C] {X Y Z : C} (p : Quiver.Path X Y) (e : Y Z) :
        @[simp]
        theorem CategoryTheory.composePath_toPath {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (f : X Y) :
        composePath f.toPath = f
        @[simp]

        Composition of paths as functor from the path category of a category to the category.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          @[simp]
          theorem CategoryTheory.pathComposition_map (C : Type u₁) [Category.{v₁, u₁} C] {X✝ Y✝ : Paths C} (f : X✝ Y✝) :

          The canonical relation on the path category of a category: two paths are related if they compose to the same morphism.

          Equations
          Instances For

            The functor from a category to the canonical quotient of its path category.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              theorem CategoryTheory.toQuotientPaths_map (C : Type u₁) [Category.{v₁, u₁} C] {X✝ Y✝ : C} (f : X✝ Y✝) :
              @[simp]
              theorem CategoryTheory.toQuotientPaths_obj_as (C : Type u₁) [Category.{v₁, u₁} C] (X : C) :
              ((toQuotientPaths C).obj X).as = X

              The functor from the canonical quotient of a path category of a category to the original category.

              Equations
              Instances For
                @[simp]
                theorem CategoryTheory.quotientPathsTo_map (C : Type u₁) [Category.{v₁, u₁} C] (a b : Quotient (pathsHomRel C)) (hf : a b) :
                (quotientPathsTo C).map hf = Quot.liftOn hf (fun (f : a.as b.as) => composePath f)

                The canonical quotient of the path category of a category is equivalent to the original category.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For