Documentation

Mathlib.CategoryTheory.Subobject.Basic

Subobjects #

We define Subobject X as the quotient (by isomorphisms) of 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 a partial order.

There is a coercion from Subobject X back to the ambient category C (using choice to pick a representative), and for P : Subobject X, P.arrow : (P : C) ⟶ X is the inclusion morphism.

We provide

The subobjects of X form a preorder making them into a category. We have X ≤ Y if and only if X.arrow factors through Y.arrow: see ofLE/ofLEMk/ofMkLE/ofMkLEMk and le_of_comm. Similarly, to show that two subobjects are equal, we can supply an isomorphism between the underlying objects that commutes with the arrows (eq_of_comm).

See also

Notes #

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

Implementation note #

Currently we describe pullback, map, etc., as functors. It may be better to just say that they are monotone functions, and even avoid using categorical language entirely when describing Subobject X. (It's worth keeping this in mind in future use; it should be a relatively easy change here if it looks preferable.)

Relation to pseudoelements #

There is a separate development of pseudoelements in CategoryTheory.Abelian.Pseudoelements, as a quotient (but not by isomorphism) of Over X.

When a morphism f has an image, the image represents the same pseudoelement. In a category with images Pseudoelements X could be constructed as a quotient of MonoOver X. In fact, in an abelian category (I'm not sure in what generality beyond that), Pseudoelements X agrees with Subobject X, but we haven't developed this in mathlib yet.

We now construct the subobject lattice for X : C, as the quotient by isomorphisms of MonoOver X.

Since MonoOver X is a thin category, we use ThinSkeleton to take the quotient.

Essentially all the structure defined above on MonoOver X descends to Subobject X, with morphisms becoming inequalities, and isomorphisms becoming equations.

def CategoryTheory.Subobject {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] (X : C) :
Type (max u₁ v₁)

The category of subobjects of X : C, defined as isomorphism classes of monomorphisms into X.

Equations
Instances For
    theorem CategoryTheory.Comma.ext_iff {A : Type u₁} :
    ∀ {inst : CategoryTheory.Category.{v₁, u₁} A} {B : Type u₂} {inst_1 : CategoryTheory.Category.{v₂, u₂} B} {T : Type u₃} {inst_2 : CategoryTheory.Category.{v₃, u₃} T} {L : CategoryTheory.Functor A T} {R : CategoryTheory.Functor B T} (x y : CategoryTheory.Comma L R), x = y x.left = y.left x.right = y.right HEq x.hom y.hom
    theorem CategoryTheory.Comma.ext {A : Type u₁} :
    ∀ {inst : CategoryTheory.Category.{v₁, u₁} A} {B : Type u₂} {inst_1 : CategoryTheory.Category.{v₂, u₂} B} {T : Type u₃} {inst_2 : CategoryTheory.Category.{v₃, u₃} T} {L : CategoryTheory.Functor A T} {R : CategoryTheory.Functor B T} (x y : CategoryTheory.Comma L R), x.left = y.leftx.right = y.rightHEq x.hom y.homx = y
    def CategoryTheory.Subobject.lift {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {α : Sort u_1} {X : C} (F : A : C⦄ → (f : A X) → [inst : CategoryTheory.Mono f] → α) (h : ∀ ⦃A B : C⦄ (f : A X) (g : B X) [inst : CategoryTheory.Mono f] [inst_1 : CategoryTheory.Mono g] (i : A B), CategoryTheory.CategoryStruct.comp i.hom g = fF f = F g) :

    Declare a function on subobjects of X by specifying a function on monomorphisms with codomain X.

    Equations
    Instances For
      @[simp]
      theorem CategoryTheory.Subobject.lift_mk {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {α : Sort u_1} {X : C} (F : A : C⦄ → (f : A X) → [inst : CategoryTheory.Mono f] → α) {h : ∀ ⦃A B : C⦄ (f : A X) (g : B X) [inst : CategoryTheory.Mono f] [inst_1 : CategoryTheory.Mono g] (i : A B), CategoryTheory.CategoryStruct.comp i.hom g = fF f = F g} {A : C} (f : A X) [CategoryTheory.Mono f] :

      The category of subobjects is equivalent to the MonoOver category. It is more convenient to use the former due to the partial order instance, but oftentimes it is easier to define structures on the latter.

      Equations
      Instances For

        Use choice to pick a representative MonoOver X for each Subobject X.

        Equations
        Instances For
          noncomputable def CategoryTheory.Subobject.representativeIso {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {X : C} (A : CategoryTheory.MonoOver X) :
          CategoryTheory.Subobject.representative.obj ((CategoryTheory.toThinSkeleton (CategoryTheory.MonoOver X)).obj A) A

          Starting with A : MonoOver X, we can take its equivalence class in Subobject X then pick an arbitrary representative using representative.obj. This is isomorphic (in MonoOver X) to the original A.

          Equations
          Instances For

            Use choice to pick a representative underlying object in C for any Subobject X.

            Prefer to use the coercion P : C rather than explicitly writing underlying.obj P.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              Equations
              • CategoryTheory.Subobject.instCoeOutSubobject = { coe := fun (Y : CategoryTheory.Subobject X) => CategoryTheory.Subobject.underlying.obj Y }
              noncomputable def CategoryTheory.Subobject.underlyingIso {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {X : C} {Y : C} (f : X Y) [CategoryTheory.Mono f] :
              CategoryTheory.Subobject.underlying.obj (CategoryTheory.Subobject.mk f) X

              If we construct a Subobject Y from an explicit f : X ⟶ Y with [Mono f], then pick an arbitrary choice of underlying object (Subobject.mk f : C) back in C, it is isomorphic (in C) to the original X.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                noncomputable def CategoryTheory.Subobject.arrow {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {X : C} (Y : CategoryTheory.Subobject X) :
                CategoryTheory.Subobject.underlying.obj Y X

                The morphism in C from the arbitrarily chosen underlying object to the ambient object.

                Equations
                Instances For
                  @[simp]
                  theorem CategoryTheory.Subobject.representative_coe {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {X : C} (Y : CategoryTheory.Subobject X) :
                  (CategoryTheory.Subobject.representative.obj Y).obj.left = CategoryTheory.Subobject.underlying.obj Y
                  theorem CategoryTheory.Subobject.eq_of_comp_arrow_eq {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {X : C} {Y : C} {P : CategoryTheory.Subobject Y} {f : X CategoryTheory.Subobject.underlying.obj P} {g : X CategoryTheory.Subobject.underlying.obj P} (h : CategoryTheory.CategoryStruct.comp f (CategoryTheory.Subobject.arrow P) = CategoryTheory.CategoryStruct.comp g (CategoryTheory.Subobject.arrow P)) :
                  f = g

                  Two morphisms into a subobject are equal exactly if the morphisms into the ambient object are equal

                  theorem CategoryTheory.Subobject.mk_le_mk_of_comm {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {B : C} {A₁ : C} {A₂ : C} {f₁ : A₁ B} {f₂ : A₂ B} [CategoryTheory.Mono f₁] [CategoryTheory.Mono f₂] (g : A₁ A₂) (w : CategoryTheory.CategoryStruct.comp g f₂ = f₁) :
                  theorem CategoryTheory.Subobject.le_of_comm {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {B : C} {X : CategoryTheory.Subobject B} {Y : CategoryTheory.Subobject B} (f : CategoryTheory.Subobject.underlying.obj X CategoryTheory.Subobject.underlying.obj Y) (w : CategoryTheory.CategoryStruct.comp f (CategoryTheory.Subobject.arrow Y) = CategoryTheory.Subobject.arrow X) :
                  X Y
                  theorem CategoryTheory.Subobject.eq_of_comm {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {B : C} {X : CategoryTheory.Subobject B} {Y : CategoryTheory.Subobject B} (f : CategoryTheory.Subobject.underlying.obj X CategoryTheory.Subobject.underlying.obj Y) (w : CategoryTheory.CategoryStruct.comp f.hom (CategoryTheory.Subobject.arrow Y) = CategoryTheory.Subobject.arrow X) :
                  X = Y

                  To show that two subobjects are equal, it suffices to exhibit an isomorphism commuting with the arrows.

                  To show that two subobjects are equal, it suffices to exhibit an isomorphism commuting with the arrows.

                  To show that two subobjects are equal, it suffices to exhibit an isomorphism commuting with the arrows.

                  To show that two subobjects are equal, it suffices to exhibit an isomorphism commuting with the arrows.

                  def CategoryTheory.Subobject.ofLE {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {B : C} (X : CategoryTheory.Subobject B) (Y : CategoryTheory.Subobject B) (h : X Y) :
                  CategoryTheory.Subobject.underlying.obj X CategoryTheory.Subobject.underlying.obj Y

                  An inequality of subobjects is witnessed by some morphism between the corresponding objects.

                  Equations
                  Instances For
                    def CategoryTheory.Subobject.ofLEMk {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {B : C} {A : C} (X : CategoryTheory.Subobject B) (f : A B) [CategoryTheory.Mono f] (h : X CategoryTheory.Subobject.mk f) :
                    CategoryTheory.Subobject.underlying.obj X A

                    An inequality of subobjects is witnessed by some morphism between the corresponding objects.

                    Equations
                    Instances For
                      def CategoryTheory.Subobject.ofMkLE {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {B : C} {A : C} (f : A B) [CategoryTheory.Mono f] (X : CategoryTheory.Subobject B) (h : CategoryTheory.Subobject.mk f X) :
                      A CategoryTheory.Subobject.underlying.obj X

                      An inequality of subobjects is witnessed by some morphism between the corresponding objects.

                      Equations
                      Instances For

                        An inequality of subobjects is witnessed by some morphism between the corresponding objects.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          def CategoryTheory.Subobject.isoOfEq {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {B : C} (X : CategoryTheory.Subobject B) (Y : CategoryTheory.Subobject B) (h : X = Y) :
                          CategoryTheory.Subobject.underlying.obj X CategoryTheory.Subobject.underlying.obj Y

                          An equality of subobjects gives an isomorphism of the corresponding objects. (One could use underlying.mapIso (eqToIso h)) here, but this is more readable.)

                          Equations
                          Instances For
                            def CategoryTheory.Subobject.isoOfEqMk {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {B : C} {A : C} (X : CategoryTheory.Subobject B) (f : A B) [CategoryTheory.Mono f] (h : X = CategoryTheory.Subobject.mk f) :
                            CategoryTheory.Subobject.underlying.obj X A

                            An equality of subobjects gives an isomorphism of the corresponding objects.

                            Equations
                            Instances For
                              def CategoryTheory.Subobject.isoOfMkEq {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] {B : C} {A : C} (f : A B) [CategoryTheory.Mono f] (X : CategoryTheory.Subobject B) (h : CategoryTheory.Subobject.mk f = X) :
                              A CategoryTheory.Subobject.underlying.obj X

                              An equality of subobjects gives an isomorphism of the corresponding objects.

                              Equations
                              Instances For

                                An equality of subobjects gives an isomorphism of the corresponding objects.

                                Equations
                                Instances For

                                  Isomorphic functors become equal when lowered to Subobject. (It's not as evil as usual to talk about equality between functors because the categories are thin and skeletal.)

                                  An equivalence between MonoOver A and MonoOver B gives an equivalence between Subobject A and Subobject B.

                                  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 Subobject Y ⥤ Subobject X, by pulling back a monomorphism along f.

                                    Equations
                                    Instances For

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

                                      Equations
                                      Instances For

                                        In fact, there's a type level bijection between the subobjects of isomorphic objects, which preserves the order.

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

                                          The functor from subobjects of X to subobjects of Y given by sending the subobject S to its "image" under f, usually denoted $\exists_f$. For instance, when C is the category of types, viewing Subobject X as Set X this is just Set.image f.

                                          This functor is left adjoint to the pullback f functor (shown in existsPullbackAdj) provided both are defined, and generalises the map f functor, again provided it is defined.

                                          Equations
                                          Instances For