Documentation

Mathlib.CategoryTheory.Limits.Shapes.IsTerminal

Initial and terminal objects in a category. #

In this file we define the predicates IsTerminal and IsInitial as well as the class InitialMonoClass.

The classes HasTerminal and HasInitial and the associated notations for terminal and initial objects are defined in Terminal.lean.

References #

Construct a cone for the empty diagram given an object.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    @[simp]
    theorem CategoryTheory.Limits.asEmptyCone_π_app {C : Type u₁} [Category.{v₁, u₁} C] (X : C) (X_1 : Discrete PEmpty.{1}) :
    (asEmptyCone X).app X_1 = id (Discrete.casesOn X_1 fun (as : PEmpty.{1}) => .elim)
    @[simp]

    Construct a cocone for the empty diagram given an object.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      theorem CategoryTheory.Limits.asEmptyCocone_ι_app {C : Type u₁} [Category.{v₁, u₁} C] (X : C) (X_1 : Discrete PEmpty.{1}) :
      (asEmptyCocone X).app X_1 = id (Discrete.casesOn X_1 fun (as : PEmpty.{1}) => .elim)
      @[reducible, inline]
      abbrev CategoryTheory.Limits.IsTerminal {C : Type u₁} [Category.{v₁, u₁} C] (X : C) :
      Type (max u₁ v₁)

      X is terminal if the cone it induces on the empty diagram is limiting.

      Equations
      Instances For
        @[reducible, inline]
        abbrev CategoryTheory.Limits.IsInitial {C : Type u₁} [Category.{v₁, u₁} C] (X : C) :
        Type (max u₁ v₁)

        X is initial if the cocone it induces on the empty diagram is colimiting.

        Equations
        Instances For
          def CategoryTheory.Limits.isTerminalEquivUnique {C : Type u₁} [Category.{v₁, u₁} C] (F : Functor (Discrete PEmpty.{1}) C) (Y : C) :
          IsLimit { pt := Y, π := { app := fun (X : Discrete PEmpty.{1}) => id (Discrete.casesOn X fun (as : PEmpty.{1}) => .elim), naturality := } } ((X : C) → Unique (X Y))

          An object Y is terminal iff for every X there is a unique morphism X ⟶ Y.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            def CategoryTheory.Limits.IsTerminal.ofUnique {C : Type u₁} [Category.{v₁, u₁} C] (Y : C) [h : (X : C) → Unique (X Y)] :

            An object Y is terminal if for every X there is a unique morphism X ⟶ Y (as an instance).

            Equations
            Instances For
              def CategoryTheory.Limits.IsTerminal.ofUniqueHom {C : Type u₁} [Category.{v₁, u₁} C] {Y : C} (h : (X : C) → X Y) (uniq : ∀ (X : C) (m : X Y), m = h X) :

              An object Y is terminal if for every X there is a unique morphism X ⟶ Y (as explicit arguments).

              Equations
              Instances For

                If α is a preorder with top, then is a terminal object.

                Equations
                Instances For

                  Transport a term of type IsTerminal across an isomorphism.

                  Equations
                  Instances For

                    If X and Y are isomorphic, then X is terminal iff Y is.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      def CategoryTheory.Limits.isInitialEquivUnique {C : Type u₁} [Category.{v₁, u₁} C] (F : Functor (Discrete PEmpty.{1}) C) (X : C) :
                      IsColimit { pt := X, ι := { app := fun (X_1 : Discrete PEmpty.{1}) => id (Discrete.casesOn X_1 fun (as : PEmpty.{1}) => .elim), naturality := } } ((Y : C) → Unique (X Y))

                      An object X is initial iff for every Y there is a unique morphism X ⟶ Y.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        def CategoryTheory.Limits.IsInitial.ofUnique {C : Type u₁} [Category.{v₁, u₁} C] (X : C) [h : (Y : C) → Unique (X Y)] :

                        An object X is initial if for every Y there is a unique morphism X ⟶ Y (as an instance).

                        Equations
                        Instances For
                          def CategoryTheory.Limits.IsInitial.ofUniqueHom {C : Type u₁} [Category.{v₁, u₁} C] {X : C} (h : (Y : C) → X Y) (uniq : ∀ (Y : C) (m : X Y), m = h Y) :

                          An object X is initial if for every Y there is a unique morphism X ⟶ Y (as explicit arguments).

                          Equations
                          Instances For

                            If α is a preorder with bot, then is an initial object.

                            Equations
                            Instances For
                              def CategoryTheory.Limits.IsInitial.ofIso {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (hX : IsInitial X) (i : X Y) :

                              Transport a term of type is_initial across an isomorphism.

                              Equations
                              Instances For

                                If X and Y are isomorphic, then X is initial iff Y is.

                                Equations
                                • One or more equations did not get rendered due to their size.
                                Instances For
                                  def CategoryTheory.Limits.IsTerminal.from {C : Type u₁} [Category.{v₁, u₁} C] {X : C} (t : IsTerminal X) (Y : C) :
                                  Y X

                                  Give the morphism to a terminal object from any other.

                                  Equations
                                  Instances For
                                    theorem CategoryTheory.Limits.IsTerminal.hom_ext {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (t : IsTerminal X) (f g : Y X) :
                                    f = g

                                    Any two morphisms to a terminal object are equal.

                                    @[simp]
                                    theorem CategoryTheory.Limits.IsTerminal.comp_from {C : Type u₁} [Category.{v₁, u₁} C] {Z : C} (t : IsTerminal Z) {X Y : C} (f : X Y) :
                                    CategoryStruct.comp f (t.from Y) = t.from X
                                    def CategoryTheory.Limits.IsInitial.to {C : Type u₁} [Category.{v₁, u₁} C] {X : C} (t : IsInitial X) (Y : C) :
                                    X Y

                                    Give the morphism from an initial object to any other.

                                    Equations
                                    Instances For
                                      theorem CategoryTheory.Limits.IsInitial.hom_ext {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (t : IsInitial X) (f g : X Y) :
                                      f = g

                                      Any two morphisms from an initial object are equal.

                                      @[simp]
                                      theorem CategoryTheory.Limits.IsInitial.to_comp {C : Type u₁} [Category.{v₁, u₁} C] {X : C} (t : IsInitial X) {Y Z : C} (f : Y Z) :
                                      CategoryStruct.comp (t.to Y) f = t.to Z

                                      Any morphism from a terminal object is split mono.

                                      Any morphism to an initial object is split epi.

                                      theorem CategoryTheory.Limits.IsTerminal.mono_from {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (t : IsTerminal X) (f : X Y) :

                                      Any morphism from a terminal object is mono.

                                      theorem CategoryTheory.Limits.IsInitial.epi_to {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (t : IsInitial X) (f : Y X) :
                                      Epi f

                                      Any morphism to an initial object is epi.

                                      If T and T' are terminal, they are isomorphic.

                                      Equations
                                      • hT.uniqueUpToIso hT' = { hom := hT'.from T, inv := hT.from T', hom_inv_id := , inv_hom_id := }
                                      Instances For
                                        @[simp]
                                        theorem CategoryTheory.Limits.IsTerminal.uniqueUpToIso_hom {C : Type u₁} [Category.{v₁, u₁} C] {T T' : C} (hT : IsTerminal T) (hT' : IsTerminal T') :
                                        (hT.uniqueUpToIso hT').hom = hT'.from T
                                        @[simp]
                                        theorem CategoryTheory.Limits.IsTerminal.uniqueUpToIso_inv {C : Type u₁} [Category.{v₁, u₁} C] {T T' : C} (hT : IsTerminal T) (hT' : IsTerminal T') :
                                        (hT.uniqueUpToIso hT').inv = hT.from T'
                                        def CategoryTheory.Limits.IsInitial.uniqueUpToIso {C : Type u₁} [Category.{v₁, u₁} C] {I I' : C} (hI : IsInitial I) (hI' : IsInitial I') :
                                        I I'

                                        If I and I' are initial, they are isomorphic.

                                        Equations
                                        • hI.uniqueUpToIso hI' = { hom := hI.to I', inv := hI'.to I, hom_inv_id := , inv_hom_id := }
                                        Instances For
                                          @[simp]
                                          theorem CategoryTheory.Limits.IsInitial.uniqueUpToIso_inv {C : Type u₁} [Category.{v₁, u₁} C] {I I' : C} (hI : IsInitial I) (hI' : IsInitial I') :
                                          (hI.uniqueUpToIso hI').inv = hI'.to I
                                          @[simp]
                                          theorem CategoryTheory.Limits.IsInitial.uniqueUpToIso_hom {C : Type u₁} [Category.{v₁, u₁} C] {I I' : C} (hI : IsInitial I) (hI' : IsInitial I') :
                                          (hI.uniqueUpToIso hI').hom = hI.to I'
                                          def CategoryTheory.Limits.isLimitChangeEmptyCone (C : Type u₁) [Category.{v₁, u₁} C] {F₁ : Functor (Discrete PEmpty.{w + 1}) C} {F₂ : Functor (Discrete PEmpty.{w' + 1}) C} {c₁ : Cone F₁} (hl : IsLimit c₁) (c₂ : Cone F₂) (hi : c₁.pt c₂.pt) :
                                          IsLimit c₂

                                          Being terminal is independent of the empty diagram, its universe, and the cone over it, as long as the cone points are isomorphic.

                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For
                                            def CategoryTheory.Limits.isLimitEmptyConeEquiv (C : Type u₁) [Category.{v₁, u₁} C] {F₁ : Functor (Discrete PEmpty.{w + 1}) C} {F₂ : Functor (Discrete PEmpty.{w' + 1}) C} (c₁ : Cone F₁) (c₂ : Cone F₂) (h : c₁.pt c₂.pt) :
                                            IsLimit c₁ IsLimit c₂

                                            Replacing an empty cone in IsLimit by another with the same cone point is an equivalence.

                                            Equations
                                            • One or more equations did not get rendered due to their size.
                                            Instances For
                                              def CategoryTheory.Limits.isColimitChangeEmptyCocone (C : Type u₁) [Category.{v₁, u₁} C] {F₁ : Functor (Discrete PEmpty.{w + 1}) C} {F₂ : Functor (Discrete PEmpty.{w' + 1}) C} {c₁ : Cocone F₁} (hl : IsColimit c₁) (c₂ : Cocone F₂) (hi : c₁.pt c₂.pt) :

                                              Being initial is independent of the empty diagram, its universe, and the cocone over it, as long as the cocone points are isomorphic.

                                              Equations
                                              • One or more equations did not get rendered due to their size.
                                              Instances For
                                                def CategoryTheory.Limits.isColimitEmptyCoconeEquiv (C : Type u₁) [Category.{v₁, u₁} C] {F₁ : Functor (Discrete PEmpty.{w + 1}) C} {F₂ : Functor (Discrete PEmpty.{w' + 1}) C} (c₁ : Cocone F₁) (c₂ : Cocone F₂) (h : c₁.pt c₂.pt) :

                                                Replacing an empty cocone in IsColimit by another with the same cocone point is an equivalence.

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

                                                  An initial object is terminal in the opposite category.

                                                  Equations
                                                  Instances For

                                                    An initial object in the opposite category is terminal in the original category.

                                                    Equations
                                                    Instances For

                                                      A terminal object is initial in the opposite category.

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

                                                        A terminal object in the opposite category is initial in the original category.

                                                        Equations
                                                        Instances For

                                                          A category is an InitialMonoClass if the canonical morphism of an initial object is a monomorphism. In practice, this is most useful when given an arbitrary morphism out of the chosen initial object, see initial.mono_from. Given a terminal object, this is equivalent to the assumption that the unique morphism from initial to terminal is a monomorphism, which is the second of Freyd's axioms for an AT category.

                                                          TODO: This is a condition satisfied by categories with zero objects and morphisms.

                                                          • isInitial_mono_from {I : C} (X : C) (hI : IsInitial I) : Mono (hI.to X)

                                                            The map from the (any as stated) initial object to any other object is a monomorphism

                                                          Instances
                                                            theorem CategoryTheory.Limits.InitialMonoClass.of_isInitial {C : Type u₁} [Category.{v₁, u₁} C] {I : C} (hI : IsInitial I) (h : ∀ (X : C), Mono (hI.to X)) :

                                                            To show a category is an InitialMonoClass it suffices to give an initial object such that every morphism out of it is a monomorphism.

                                                            To show a category is an InitialMonoClass it suffices to show the unique morphism from an initial object to a terminal object is a monomorphism.

                                                            From a functor F : J ⥤ C, given an initial object of J, construct a cone for J. In limitOfDiagramInitial we show it is a limit cone.

                                                            Equations
                                                            Instances For
                                                              @[simp]
                                                              theorem CategoryTheory.Limits.coneOfDiagramInitial_π_app {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {X : J} (tX : IsInitial X) (F : Functor J C) (j : J) :
                                                              (coneOfDiagramInitial tX F).app j = F.map (tX.to j)
                                                              @[simp]
                                                              theorem CategoryTheory.Limits.coneOfDiagramInitial_pt {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {X : J} (tX : IsInitial X) (F : Functor J C) :
                                                              (coneOfDiagramInitial tX F).pt = F.obj X

                                                              From a functor F : J ⥤ C, given an initial object of J, show the cone coneOfDiagramInitial is a limit.

                                                              Equations
                                                              Instances For
                                                                def CategoryTheory.Limits.coneOfDiagramTerminal {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {X : J} (hX : IsTerminal X) (F : Functor J C) [∀ (i j : J) (f : i j), IsIso (F.map f)] :

                                                                From a functor F : J ⥤ C, given a terminal object of J, construct a cone for J, provided that the morphisms in the diagram are isomorphisms. In limitOfDiagramTerminal we show it is a limit cone.

                                                                Equations
                                                                Instances For
                                                                  @[simp]
                                                                  theorem CategoryTheory.Limits.coneOfDiagramTerminal_pt {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {X : J} (hX : IsTerminal X) (F : Functor J C) [∀ (i j : J) (f : i j), IsIso (F.map f)] :
                                                                  (coneOfDiagramTerminal hX F).pt = F.obj X
                                                                  @[simp]
                                                                  theorem CategoryTheory.Limits.coneOfDiagramTerminal_π_app {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {X : J} (hX : IsTerminal X) (F : Functor J C) [∀ (i j : J) (f : i j), IsIso (F.map f)] (x✝ : J) :
                                                                  (coneOfDiagramTerminal hX F).app x✝ = inv (F.map (hX.from x✝))
                                                                  def CategoryTheory.Limits.limitOfDiagramTerminal {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {X : J} (hX : IsTerminal X) (F : Functor J C) [∀ (i j : J) (f : i j), IsIso (F.map f)] :

                                                                  From a functor F : J ⥤ C, given a terminal object of J and that the morphisms in the diagram are isomorphisms, show the cone coneOfDiagramTerminal is a limit.

                                                                  Equations
                                                                  Instances For

                                                                    From a functor F : J ⥤ C, given a terminal object of J, construct a cocone for J. In colimitOfDiagramTerminal we show it is a colimit cocone.

                                                                    Equations
                                                                    Instances For
                                                                      @[simp]
                                                                      theorem CategoryTheory.Limits.coconeOfDiagramTerminal_pt {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {X : J} (tX : IsTerminal X) (F : Functor J C) :
                                                                      (coconeOfDiagramTerminal tX F).pt = F.obj X
                                                                      @[simp]
                                                                      theorem CategoryTheory.Limits.coconeOfDiagramTerminal_ι_app {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {X : J} (tX : IsTerminal X) (F : Functor J C) (j : J) :
                                                                      (coconeOfDiagramTerminal tX F).app j = F.map (tX.from j)

                                                                      From a functor F : J ⥤ C, given a terminal object of J, show the cocone coconeOfDiagramTerminal is a colimit.

                                                                      Equations
                                                                      Instances For
                                                                        theorem CategoryTheory.Limits.IsColimit.isIso_ι_app_of_isTerminal {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {F : Functor J C} {c : Cocone F} (hc : IsColimit c) (X : J) (hX : IsTerminal X) :
                                                                        IsIso (c.app X)
                                                                        def CategoryTheory.Limits.coconeOfDiagramInitial {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {X : J} (hX : IsInitial X) (F : Functor J C) [∀ (i j : J) (f : i j), IsIso (F.map f)] :

                                                                        From a functor F : J ⥤ C, given an initial object of J, construct a cocone for J, provided that the morphisms in the diagram are isomorphisms. In colimitOfDiagramInitial we show it is a colimit cocone.

                                                                        Equations
                                                                        Instances For
                                                                          @[simp]
                                                                          theorem CategoryTheory.Limits.coconeOfDiagramInitial_ι_app {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {X : J} (hX : IsInitial X) (F : Functor J C) [∀ (i j : J) (f : i j), IsIso (F.map f)] (x✝ : J) :
                                                                          (coconeOfDiagramInitial hX F).app x✝ = inv (F.map (hX.to x✝))
                                                                          @[simp]
                                                                          theorem CategoryTheory.Limits.coconeOfDiagramInitial_pt {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {X : J} (hX : IsInitial X) (F : Functor J C) [∀ (i j : J) (f : i j), IsIso (F.map f)] :
                                                                          (coconeOfDiagramInitial hX F).pt = F.obj X
                                                                          def CategoryTheory.Limits.colimitOfDiagramInitial {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {X : J} (hX : IsInitial X) (F : Functor J C) [∀ (i j : J) (f : i j), IsIso (F.map f)] :

                                                                          From a functor F : J ⥤ C, given an initial object of J and that the morphisms in the diagram are isomorphisms, show the cone coconeOfDiagramInitial is a colimit.

                                                                          Equations
                                                                          Instances For
                                                                            theorem CategoryTheory.Limits.IsLimit.isIso_π_app_of_isInitial {C : Type u₁} [Category.{v₁, u₁} C] {J : Type u} [Category.{v, u} J] {F : Functor J C} {c : Cone F} (hc : IsLimit c) (X : J) (hX : IsInitial X) :
                                                                            IsIso (c.app X)
                                                                            theorem CategoryTheory.Limits.isIso_of_isTerminal {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (hX : IsTerminal X) (hY : IsTerminal Y) (f : X Y) :

                                                                            Any morphism between terminal objects is an isomorphism.

                                                                            theorem CategoryTheory.Limits.isIso_of_isInitial {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} (hX : IsInitial X) (hY : IsInitial Y) (f : X Y) :

                                                                            Any morphism between initial objects is an isomorphism.