Documentation

Mathlib.CategoryTheory.ObjectProperty.FullSubcategory

The full subcategory associated to a property of objects #

Given a category C and P : ObjectProperty C, we define a category structure on the type P.FullSubcategory of objects in C satisfying P.

A subtype-like structure for full subcategories. Morphisms just ignore the property. We don't use actual subtypes since the simp-normal form ↑X of X.val does not work well for full subcategories.

  • obj : C

    The category of which this is a full subcategory

  • property : P self.obj

    The predicate satisfied by all objects in this subcategory

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

    The forgetful functor from a full subcategory into the original category ("forgetting" the condition).

    Equations
    Instances For
      @[simp]
      theorem CategoryTheory.ObjectProperty.ι_map {C : Type u} [Category.{v, u} C] (P : ObjectProperty C) {X Y : P.FullSubcategory} {f : X Y} :
      P.ι.map f = f.hom
      @[deprecated CategoryTheory.ObjectProperty.FullSubcategory.id_hom (since := "2025-12-18")]

      Alias of CategoryTheory.ObjectProperty.FullSubcategory.id_hom.

      @[deprecated CategoryTheory.ObjectProperty.FullSubcategory.comp_hom (since := "2025-12-18")]

      Alias of CategoryTheory.ObjectProperty.FullSubcategory.comp_hom.

      Constructor for morphisms in a full subcategory.

      Equations
      Instances For
        @[reducible, inline]

        The inclusion of a full subcategory is fully faithful.

        Equations
        Instances For

          Constructor for isomorphisms in P.FullSubcategory when P : ObjectProperty C.

          Equations
          Instances For

            If P and P' are properties of objects such that P ≤ P', there is an induced functor P.FullSubcategory ⥤ P'.FullSubcategory.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              theorem CategoryTheory.ObjectProperty.ιOfLE_map {C : Type u} [Category.{v, u} C] {P P' : ObjectProperty C} (h : P P') {X✝ Y✝ : P.FullSubcategory} (f : X✝ Y✝) :
              (ιOfLE h).map f = homMk f.hom
              @[simp]

              If h : P ≤ P', then ιOfLE h is fully faithful.

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

                If h : P ≤ P' is an inequality of properties of objects, this is the obvious isomorphism ιOfLE h ⋙ P'.ι ≅ P.ι.

                Equations
                Instances For
                  def CategoryTheory.ObjectProperty.lift {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (P : ObjectProperty D) (F : Functor C D) (hF : ∀ (X : C), P (F.obj X)) :

                  A functor which maps objects to objects satisfying a certain property induces a lift through the full subcategory of objects satisfying that property.

                  Equations
                  Instances For
                    @[simp]
                    theorem CategoryTheory.ObjectProperty.lift_map {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (P : ObjectProperty D) (F : Functor C D) (hF : ∀ (X : C), P (F.obj X)) {X✝ Y✝ : C} (f : X✝ Y✝) :
                    (P.lift F hF).map f = homMk (F.map f)
                    @[simp]
                    theorem CategoryTheory.ObjectProperty.lift_obj_obj {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (P : ObjectProperty D) (F : Functor C D) (hF : ∀ (X : C), P (F.obj X)) (X : C) :
                    ((P.lift F hF).obj X).obj = F.obj X
                    def CategoryTheory.ObjectProperty.liftCompιIso {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (P : ObjectProperty D) (F : Functor C D) (hF : ∀ (X : C), P (F.obj X)) :
                    (P.lift F hF).comp P.ι F

                    Composing the lift of a functor through a full subcategory with the inclusion yields the original functor. This is actually true definitionally.

                    Equations
                    Instances For
                      @[simp]
                      theorem CategoryTheory.ObjectProperty.ι_obj_lift_obj {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (P : ObjectProperty D) (F : Functor C D) (hF : ∀ (X : C), P (F.obj X)) (X : C) :
                      P.ι.obj ((P.lift F hF).obj X) = F.obj X
                      @[simp]
                      theorem CategoryTheory.ObjectProperty.ι_obj_lift_map {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (P : ObjectProperty D) (F : Functor C D) (hF : ∀ (X : C), P (F.obj X)) {X Y : C} (f : X Y) :
                      P.ι.map ((P.lift F hF).map f) = F.map f
                      instance CategoryTheory.ObjectProperty.instFaithfulFullSubcategoryLift {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (P : ObjectProperty D) (F : Functor C D) (hF : ∀ (X : C), P (F.obj X)) [F.Faithful] :
                      (P.lift F hF).Faithful
                      instance CategoryTheory.ObjectProperty.instFullFullSubcategoryLift {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (P : ObjectProperty D) (F : Functor C D) (hF : ∀ (X : C), P (F.obj X)) [F.Full] :
                      (P.lift F hF).Full
                      def CategoryTheory.ObjectProperty.liftCompιOfLEIso {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (P : ObjectProperty D) {Q : ObjectProperty D} (F : Functor C D) (hF : ∀ (X : C), P (F.obj X)) (h : P Q) :
                      (P.lift F hF).comp (ιOfLE h) Q.lift F

                      When h : P ≤ Q, this is the canonical isomorphism P.lift F hF ⋙ ιOfLE h ≅ Q.lift F _.

                      Equations
                      Instances For