Documentation

Mathlib.CategoryTheory.Subobject.MonoOver

Monomorphisms over a fixed object #

As preparation for defining Subobject X, we set up the theory for MonoOver X := { f : Over X // Mono f.hom}.

Here MonoOver X is a thin category (a pair of objects has at most one morphism between them), so we can think of it as a preorder. However as it is not skeletal, it is not yet a partial order.

Subobject X will be defined as the skeletalization of MonoOver X.

We provide

and prove their basic properties and relationships.

Notes #

This development originally appeared in Bhavik Mehta's "Topos theory for Lean" repository, and was ported to mathlib by Kim Morrison.

@[reducible, inline]

The object property in Over X of the structure morphism being a monomorphism.

Equations
Instances For
    def CategoryTheory.MonoOver {C : Type u₁} [Category.{v₁, u₁} C] (X : C) :
    Type (max u₁ v₁)

    The category of monomorphisms into X as a full subcategory of the over category. This isn't skeletal, so it's not a partial order.

    Later we define Subobject X as the quotient of this by isomorphisms.

    Equations
    Instances For
      def CategoryTheory.MonoOver.mk' {C : Type u₁} [Category.{v₁, u₁} C] {X A : C} (f : A X) [hf : Mono f] :

      Construct a MonoOver X.

      Equations
      Instances For
        @[simp]
        theorem CategoryTheory.MonoOver.mk'_obj {C : Type u₁} [Category.{v₁, u₁} C] {X A : C} (f : A X) [hf : Mono f] :

        The inclusion from monomorphisms over X to morphisms over X.

        Equations
        Instances For
          @[simp]
          @[simp]
          theorem CategoryTheory.MonoOver.mk'_coe' {C : Type u₁} [Category.{v₁, u₁} C] {X A : C} (f : A X) [Mono f] :
          (mk' f).obj.left = A
          @[reducible, inline]
          abbrev CategoryTheory.MonoOver.arrow {C : Type u₁} [Category.{v₁, u₁} C] {X : C} (f : MonoOver X) :

          Convenience notation for the underlying arrow of a monomorphism over X.

          Equations
          Instances For
            @[simp]
            theorem CategoryTheory.MonoOver.mk'_arrow {C : Type u₁} [Category.{v₁, u₁} C] {X A : C} (f : A X) [Mono f] :
            (mk' f).arrow = f
            @[simp]

            The category of monomorphisms over X is a thin category,s which makes defining its skeleton easy.

            @[reducible, inline]
            abbrev CategoryTheory.MonoOver.homMk {C : Type u₁} [Category.{v₁, u₁} C] {X : C} {f g : MonoOver X} (h : f.obj.left g.obj.left) (w : CategoryStruct.comp h g.arrow = f.arrow := by cat_disch) :
            f g

            Convenience constructor for a morphism in monomorphisms over X.

            Equations
            Instances For
              def CategoryTheory.MonoOver.isoMk {C : Type u₁} [Category.{v₁, u₁} C] {X : C} {f g : MonoOver X} (h : f.obj.left g.obj.left) (w : CategoryStruct.comp h.hom g.arrow = f.arrow := by cat_disch) :
              f g

              Convenience constructor for an isomorphism in monomorphisms over X.

              Equations
              Instances For
                @[simp]
                theorem CategoryTheory.MonoOver.isoMk_inv {C : Type u₁} [Category.{v₁, u₁} C] {X : C} {f g : MonoOver X} (h : f.obj.left g.obj.left) (w : CategoryStruct.comp h.hom g.arrow = f.arrow := by cat_disch) :
                (isoMk h w).inv = homMk h.inv
                @[simp]
                theorem CategoryTheory.MonoOver.isoMk_hom {C : Type u₁} [Category.{v₁, u₁} C] {X : C} {f g : MonoOver X} (h : f.obj.left g.obj.left) (w : CategoryStruct.comp h.hom g.arrow = f.arrow := by cat_disch) :
                (isoMk h w).hom = homMk h.hom w

                If f : MonoOver X, then mk' f.arrow is of course just f, but not definitionally, so we package it as an isomorphism.

                Equations
                Instances For
                  def CategoryTheory.MonoOver.lift {C : Type u₁} [Category.{v₁, u₁} C] {X : C} {D : Type u₂} [Category.{v₂, u₂} D] {Y : D} (F : Functor (Over Y) (Over X)) (h : ∀ (f : MonoOver Y), Mono (F.obj ((forget Y).obj f)).hom) :

                  Lift a functor between over categories to a functor between MonoOver categories, given suitable evidence that morphisms are taken to monomorphisms.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    theorem CategoryTheory.MonoOver.lift_map {C : Type u₁} [Category.{v₁, u₁} C] {X : C} {D : Type u₂} [Category.{v₂, u₂} D] {Y : D} (F : Functor (Over Y) (Over X)) (h : ∀ (f : MonoOver Y), Mono (F.obj ((forget Y).obj f)).hom) {X✝ Y✝ : MonoOver Y} (k : X✝ Y✝) :
                    (lift F h).map k = ((forget Y).comp F).map k
                    @[simp]
                    theorem CategoryTheory.MonoOver.lift_obj_obj {C : Type u₁} [Category.{v₁, u₁} C] {X : C} {D : Type u₂} [Category.{v₂, u₂} D] {Y : D} (F : Functor (Over Y) (Over X)) (h : ∀ (f : MonoOver Y), Mono (F.obj ((forget Y).obj f)).hom) (f : MonoOver Y) :
                    ((lift F h).obj f).obj = F.obj ((forget Y).obj f)
                    def CategoryTheory.MonoOver.liftIso {C : Type u₁} [Category.{v₁, u₁} C] {X : C} {D : Type u₂} [Category.{v₂, u₂} D] {Y : D} {F₁ F₂ : Functor (Over Y) (Over X)} (h₁ : ∀ (f : MonoOver Y), Mono (F₁.obj ((forget Y).obj f)).hom) (h₂ : ∀ (f : MonoOver Y), Mono (F₂.obj ((forget Y).obj f)).hom) (i : F₁ F₂) :
                    lift F₁ h₁ lift F₂ h₂

                    Isomorphic functors Over Y ⥤ Over X lift to isomorphic functors MonoOver Y ⥤ MonoOver X.

                    Equations
                    Instances For
                      def CategoryTheory.MonoOver.liftComp {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] {X Z : C} {Y : D} (F : Functor (Over X) (Over Y)) (G : Functor (Over Y) (Over Z)) (h₁ : ∀ (f : MonoOver X), Mono (F.obj ((forget X).obj f)).hom) (h₂ : ∀ (f : MonoOver Y), Mono (G.obj ((forget Y).obj f)).hom) :
                      (lift F h₁).comp (lift G h₂) lift (F.comp G)

                      MonoOver.lift commutes with composition of functors.

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

                        MonoOver.lift preserves the identity functor.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          @[simp]
                          theorem CategoryTheory.MonoOver.lift_comm {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (F : Functor (Over Y) (Over X)) (h : ∀ (f : MonoOver Y), Mono (F.obj ((forget Y).obj f)).hom) :
                          (lift F h).comp (forget X) = (forget Y).comp F
                          @[simp]
                          theorem CategoryTheory.MonoOver.lift_obj_arrow {C : Type u₁} [Category.{v₁, u₁} C] {X : C} {D : Type u₂} [Category.{v₂, u₂} D] {Y : D} (F : Functor (Over Y) (Over X)) (h : ∀ (f : MonoOver Y), Mono (F.obj ((forget Y).obj f)).hom) (f : MonoOver Y) :
                          ((lift F h).obj f).arrow = (F.obj ((forget Y).obj f)).hom
                          def CategoryTheory.MonoOver.slice {C : Type u₁} [Category.{v₁, u₁} C] {A : C} {f : Over A} (h₁ : ∀ (g : MonoOver f), Mono (f.iteratedSliceEquiv.functor.obj ((forget f).obj g)).hom) (h₂ : ∀ (g : MonoOver f.left), Mono (f.iteratedSliceEquiv.inverse.obj ((forget f.left).obj g)).hom) :

                          Monomorphisms over an object f : Over A in an over category are equivalent to monomorphisms over the source of f.

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

                            When C has pullbacks, a morphism f : X ⟶ Y induces a functor MonoOver Y ⥤ MonoOver X, by pulling back a monomorphism along f.

                            Equations
                            Instances For

                              pullback commutes with composition (up to a natural isomorphism)

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                def CategoryTheory.MonoOver.pullbackObjIsoOfIsPullback {C : Type u₁} [Category.{v₁, u₁} C] [Limits.HasPullbacks C] {X Y : C} (f : Y X) (S : MonoOver X) (T : MonoOver Y) (f' : T.obj.left S.obj.left) (h : IsPullback f' T.arrow S.arrow f) :
                                (pullback f).obj S T

                                Given two monomorphisms S and T over X and Y and two morphisms f and f' between them forming the following pullback square:

                                (T : C) -f'-> (S : C)
                                   |             |
                                T.arrow       S.arrow
                                   |             |
                                   v             v
                                   Y -----f----> X
                                

                                we get an isomorphism between T and the pullback of S along f through the pullback functor.

                                Equations
                                Instances For
                                  def CategoryTheory.MonoOver.map {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (f : X Y) [Mono f] :

                                  We can map monomorphisms over X to monomorphisms over Y by post-composition with a monomorphism f : X ⟶ Y.

                                  Equations
                                  Instances For
                                    def CategoryTheory.MonoOver.mapComp {C : Type u₁} [Category.{v₁, u₁} C] {X Y Z : C} (f : X Y) (g : Y Z) [Mono f] [Mono g] :

                                    MonoOver.map commutes with composition (up to a natural isomorphism).

                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    Instances For
                                      @[simp]
                                      theorem CategoryTheory.MonoOver.map_obj_left {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (f : X Y) [Mono f] (g : MonoOver X) :
                                      ((map f).obj g).obj.left = g.obj.left
                                      @[simp]
                                      theorem CategoryTheory.MonoOver.map_obj_arrow {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (f : X Y) [Mono f] (g : MonoOver X) :
                                      instance CategoryTheory.MonoOver.full_map {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (f : X Y) [Mono f] :
                                      (map f).Full
                                      instance CategoryTheory.MonoOver.faithful_map {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (f : X Y) [Mono f] :

                                      Isomorphic objects have equivalent MonoOver categories.

                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For
                                        @[simp]
                                        theorem CategoryTheory.MonoOver.mapIso_inverse {C : Type u₁} [Category.{v₁, u₁} C] {A B : C} (e : A B) :
                                        @[simp]
                                        theorem CategoryTheory.MonoOver.mapIso_functor {C : Type u₁} [Category.{v₁, u₁} C] {A B : C} (e : A B) :

                                        An equivalence of categories e between C and D induces an equivalence between MonoOver X and MonoOver (e.functor.obj X) whenever X is an object of C.

                                        Equations
                                        • One or more equations did not get rendered due to their size.
                                        Instances For
                                          @[simp]
                                          theorem CategoryTheory.MonoOver.congr_functor {C : Type u₁} [Category.{v₁, u₁} C] (X : C) {D : Type u₂} [Category.{v₂, u₂} D] (e : C D) :
                                          @[simp]
                                          theorem CategoryTheory.MonoOver.congr_counitIso {C : Type u₁} [Category.{v₁, u₁} C] (X : C) {D : Type u₂} [Category.{v₂, u₂} D] (e : C D) :
                                          (congr X e).counitIso = NatIso.ofComponents (fun (Y : MonoOver (e.functor.obj X)) => isoMk (e.counitIso.app Y.obj.left) )
                                          @[simp]
                                          theorem CategoryTheory.MonoOver.congr_unitIso {C : Type u₁} [Category.{v₁, u₁} C] (X : C) {D : Type u₂} [Category.{v₂, u₂} D] (e : C D) :
                                          (congr X e).unitIso = NatIso.ofComponents (fun (Y : MonoOver X) => isoMk (e.unitIso.app Y.obj.left) )

                                          map f is left adjoint to pullback f when f is a monomorphism

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

                                            The MonoOver Y for the image inclusion for a morphism f : X ⟶ Y.

                                            Equations
                                            Instances For

                                              Taking the image of a morphism gives a functor Over X ⥤ MonoOver X.

                                              Equations
                                              • One or more equations did not get rendered due to their size.
                                              Instances For
                                                @[simp]
                                                theorem CategoryTheory.MonoOver.image_map {C : Type u₁} [Category.{v₁, u₁} C] {X : C} [Limits.HasImages C] {f g : Over X} (k : f g) :

                                                MonoOver.image : Over X ⥤ MonoOver X is left adjoint to MonoOver.forget : MonoOver X ⥤ Over X

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

                                                  Forgetting that a monomorphism over X is a monomorphism, then taking its image, is the identity functor.

                                                  Equations
                                                  Instances For

                                                    In the case where f is not a monomorphism but C has images, we can still take the "forward map" under it, which agrees with MonoOver.map f.

                                                    Equations
                                                    Instances For

                                                      When f : X ⟶ Y is a monomorphism, exists f agrees with map f.

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

                                                        exists is adjoint to pullback when images exist

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