Documentation

Mathlib.CategoryTheory.Core

The core of a category #

The core of a category C is the (non-full) subcategory of C consisting of all objects, and all isomorphisms. We construct it as a CategoryTheory.Groupoid.

CategoryTheory.Core.inclusion : Core C ⥤ C gives the faithful inclusion into the original category.

Any functor F from a groupoid G into C factors through CategoryTheory.Core C, but this is not functorial with respect to F.

def CategoryTheory.Core (C : Type u₁) :
Type u₁

The core of a category C is the groupoid whose morphisms are all the isomorphisms of C.

Equations
Instances For
    Equations
    • One or more equations did not get rendered due to their size.
    @[simp]
    theorem CategoryTheory.Core.comp_hom {C : Type u₁} [Category.{v₁, u₁} C] {X Y Z : Core C} (f : X Y) (g : Y Z) :
    theorem CategoryTheory.Core.hom_ext {C : Type u₁} [Category.{v₁, u₁} C] {X Y : Core C} {f g : X Y} (h : f.hom = g.hom) :
    f = g
    theorem CategoryTheory.Core.hom_ext_iff {C : Type u₁} [Category.{v₁, u₁} C] {X Y : Core C} {f g : X Y} :
    f = g f.hom = g.hom

    The core of a category is naturally included in the category.

    Equations
    Instances For
      @[simp]
      theorem CategoryTheory.Core.inclusion_map (C : Type u₁) [Category.{v₁, u₁} C] {X✝ Y✝ : Core C} (f : X✝ Y✝) :
      (inclusion C).map f = f.hom
      @[simp]

      A functor from a groupoid to a category C factors through the core of C.

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

        We can functorially associate to any functor from a groupoid to the core of a category C, a functor from the groupoid to C, simply by composing with the embedding Core C ⥤ C.

        Equations
        Instances For

          A functor C ⥤ D induces a functor Core C ⥤ Core D.

          Equations
          Instances For
            @[simp]
            theorem CategoryTheory.Functor.core_map_hom {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (F : Functor C D) {X✝ Y✝ : Core C} (f : X✝ Y✝) :
            (F.core.map f).hom = F.map f.hom
            @[simp]
            theorem CategoryTheory.Functor.core_map_inv {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (F : Functor C D) {X✝ Y✝ : Core C} (f : X✝ Y✝) :
            @[simp]
            theorem CategoryTheory.Functor.core_obj {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (F : Functor C D) (X : Core C) :
            F.core.obj X = F.obj X

            The core of the identity functor is the identity functor on the cores.

            Equations
            Instances For

              The core of the composition of F and G is the composition of the cores.

              Equations
              Instances For
                @[simp]
                theorem CategoryTheory.Functor.coreComp_inv_app {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {E : Type u₃} [Category.{v₃, u₃} E] (F : Functor C D) (G : Functor D E) (X : Core C) :
                @[simp]
                theorem CategoryTheory.Functor.coreComp_hom_app {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {E : Type u₃} [Category.{v₃, u₃} E] (F : Functor C D) (G : Functor D E) (X : Core C) :
                def CategoryTheory.Iso.core {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F G : Functor C D} (α : F G) :

                A natural isomorphism of functors induces a natural isomorphism between their cores.

                Equations
                Instances For
                  @[simp]
                  theorem CategoryTheory.Iso.core_inv_app {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F G : Functor C D} (α : F G) (X : Core C) :
                  @[simp]
                  theorem CategoryTheory.Iso.core_hom_app {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F G : Functor C D} (α : F G) (X : Core C) :
                  α.core.hom.app X = α.app X
                  @[simp]
                  theorem CategoryTheory.Iso.coreComp {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {F G H : Functor C D} (α : F G) (β : G H) :
                  (α ≪≫ β).core = α.core ≪≫ β.core
                  @[simp]

                  Equivalent categories have equivalent cores.

                  Equations
                  Instances For
                    @[simp]
                    theorem CategoryTheory.Equivalence.core_functor_map_hom {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (E : C D) {X✝ Y✝ : Core C} (f : X✝ Y✝) :
                    @[simp]
                    theorem CategoryTheory.Equivalence.core_inverse_map_inv {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (E : C D) {X✝ Y✝ : Core D} (f : X✝ Y✝) :
                    @[simp]
                    theorem CategoryTheory.Equivalence.core_functor_map_inv {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (E : C D) {X✝ Y✝ : Core C} (f : X✝ Y✝) :
                    @[simp]
                    theorem CategoryTheory.Equivalence.core_inverse_map_hom {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] (E : C D) {X✝ Y✝ : Core D} (f : X✝ Y✝) :

                    Taking the core of a functor is functorial if we discard non-invertible natural transformations.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      @[simp]
                      theorem CategoryTheory.coreFunctor_obj_map_hom (C : Type u₁) [Category.{v₁, u₁} C] (D : Type u₂) [Category.{v₂, u₂} D] (F : Core (Functor C D)) {X✝ Y✝ : Core C} (f : X✝ Y✝) :
                      (((coreFunctor C D).obj F).map f).hom = F.map f.hom
                      @[simp]
                      theorem CategoryTheory.coreFunctor_obj_obj (C : Type u₁) [Category.{v₁, u₁} C] (D : Type u₂) [Category.{v₂, u₂} D] (F : Core (Functor C D)) (X : Core C) :
                      ((coreFunctor C D).obj F).obj X = F.obj X
                      @[simp]
                      theorem CategoryTheory.coreFunctor_map_app_hom (C : Type u₁) [Category.{v₁, u₁} C] (D : Type u₂) [Category.{v₂, u₂} D] {X✝ Y✝ : Core (Functor C D)} (η : X✝ Y✝) (X : Core C) :
                      (((coreFunctor C D).map η).app X).hom = η.hom.app X
                      @[simp]
                      theorem CategoryTheory.coreFunctor_map_app_inv (C : Type u₁) [Category.{v₁, u₁} C] (D : Type u₂) [Category.{v₂, u₂} D] {X✝ Y✝ : Core (Functor C D)} (η : X✝ Y✝) (X : Core C) :
                      (((coreFunctor C D).map η).app X).inv = η.inv.app X
                      @[simp]
                      theorem CategoryTheory.coreFunctor_obj_map_inv (C : Type u₁) [Category.{v₁, u₁} C] (D : Type u₂) [Category.{v₂, u₂} D] (F : Core (Functor C D)) {X✝ Y✝ : Core C} (f : X✝ Y✝) :
                      (((coreFunctor C D).obj F).map f).inv = F.map (inv f).hom
                      def CategoryTheory.ofEquivFunctor (m : Type u₁ → Type u₂) [EquivFunctor m] :
                      Functor (Core (Type u₁)) (Core (Type u₂))

                      ofEquivFunctor m lifts a type-level EquivFunctor to a categorical functor Core (Type u₁) ⥤ Core (Type u₂).

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