Documentation

Mathlib.CategoryTheory.Bicategory.Functor.Cat.ObjectProperty

Properties of objects in target categories of a pseudofunctor to Cat #

Given F : Pseudofunctor B Cat, we introduce a type F.ObjectProperty which consists of properties P of objects for all categories F.obj X for X : B. The typeclass P.IsClosedUnderMapObj expresses that this property is preserved by the application of the functors F.map: this allows to define a sub-pseudofunctor P.fullsubcategory : Pseudofunctor B Cat.

TODO (@joelriou) #

If F : Pseudofunctor B Cat, this is the data of a property of objects in all categories F.obj X for X : B.

  • prop (X : B) : ObjectProperty (F.obj X)

    A property of objects in the category F.obj X for all X : B.

Instances For
    @[reducible, inline]

    Given F : Pseudofunctor B Cat, P : F.ObjectProperty and X : B, this is the full subcategory of F.obj X consisting of the objects satisfying the property P.

    Equations
    Instances For

      If P is a property of objects for a pseudofunctor F to Cat, this is the condition that P is preserved by the application of the functors F.obj.

      Instances

        If P is a property of objects for a pseudofunctor F to Cat, this is the condition that all P.prop : ObjectProprety (F.obj X) for X : B are closed under isomorphisms.

        Instances

          Given a property P of objects for F : Pseudofunctor B Cat and a morphism f : X ⟶ Y in B, this is the functor P.Obj X ⥤ P.Obj Y that is induced by F.map f.

          Equations
          Instances For
            @[simp]
            theorem CategoryTheory.Pseudofunctor.ObjectProperty.map_map_hom {B : Type u} [Bicategory B] {F : Pseudofunctor B Cat} (P : F.ObjectProperty) [P.IsClosedUnderMapObj] {X Y : B} (f : X Y) {X✝ Y✝ : P.Obj X} (f✝ : X✝ Y✝) :
            ((P.map f).map f✝).hom = (F.map f).toFunctor.map f✝.hom
            @[simp]
            theorem CategoryTheory.Pseudofunctor.ObjectProperty.map_obj_obj {B : Type u} [Bicategory B] {F : Pseudofunctor B Cat} (P : F.ObjectProperty) [P.IsClosedUnderMapObj] {X Y : B} (f : X Y) (X✝ : P.Obj X) :
            ((P.map f).obj X✝).obj = (F.map f).toFunctor.obj X✝.obj
            def CategoryTheory.Pseudofunctor.ObjectProperty.map₂ {B : Type u} [Bicategory B] {F : Pseudofunctor B Cat} (P : F.ObjectProperty) [P.IsClosedUnderMapObj] {X Y : B} {f g : X Y} (α : f g) :
            P.map f P.map g

            Given a property P of objects for F : Pseudofunctor B Cat and a 2-morphism in B, this is the induced natural transformation between the induced functors on the fullsubcategories of objects satisfying P.

            Equations
            Instances For
              @[simp]
              theorem CategoryTheory.Pseudofunctor.ObjectProperty.map₂_app_hom {B : Type u} [Bicategory B] {F : Pseudofunctor B Cat} (P : F.ObjectProperty) [P.IsClosedUnderMapObj] {X Y : B} {f g : X Y} (α : f g) (X✝ : P.Obj X) :
              ((P.map₂ α).app X✝).hom = (F.map₂ α).toNatTrans.app X✝.obj

              Auxiliary definition for fullsubcategory.

              Equations
              Instances For
                @[simp]
                @[simp]

                Given a property of objects P for a pseudofunctor from B to Cat, this is the induced pseudofunctor which sends X : B to the full subcategory of F.obj B consisting of objects satisfying P.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  @[simp]
                  theorem CategoryTheory.Pseudofunctor.ObjectProperty.fullsubcategory_mapComp {B : Type u} [Bicategory B] {F : Pseudofunctor B Cat} (P : F.ObjectProperty) [P.IsClosedUnderMapObj] {a✝ b✝ c✝ : B} (f : a✝ b✝) (g : b✝ c✝) :

                  The inclusion of P.fullsubcategory in F.

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