Documentation

Mathlib.CategoryTheory.Limits.Shapes.Equalizers

Equalizers and coequalizers #

This file defines (co)equalizers as special cases of (co)limits.

An equalizer is the categorical generalization of the subobject {a ∈ A | f(a) = g(a)} known from abelian groups or modules. It is a limit cone over the diagram formed by f and g.

A coequalizer is the dual concept.

Main definitions #

Each of these has a dual.

Main statements #

Implementation notes #

As with the other special shapes in the limits library, all the definitions here are given as abbreviations of the general statements for limits, so all the simp lemmas and theorems about general limits can be used.

References #

The type of objects for the diagram indexing a (co)equalizer.

Instances For

    Composition of morphisms in the indexing diagram for (co)equalizers.

    Equations
    Instances For

      The functor WalkingParallelPair ⥤ WalkingParallelPairᵒᵖ sending left to left and right to right.

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

        The equivalence WalkingParallelPair ⥤ WalkingParallelPairᵒᵖ sending left to left and right to right.

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

          parallelPair f g is the diagram in C consisting of the two morphisms f and g with common domain and codomain.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            def CategoryTheory.Limits.parallelPairHom {C : Type u} [Category.{v, u} C] {X Y X' Y' : C} (f g : X Y) (f' g' : X' Y') (p : X X') (q : Y Y') (wf : CategoryStruct.comp f q = CategoryStruct.comp p f') (wg : CategoryStruct.comp g q = CategoryStruct.comp p g') :

            Construct a morphism between parallel pairs.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              theorem CategoryTheory.Limits.parallelPairHom_app_zero {C : Type u} [Category.{v, u} C] {X Y X' Y' : C} (f g : X Y) (f' g' : X' Y') (p : X X') (q : Y Y') (wf : CategoryStruct.comp f q = CategoryStruct.comp p f') (wg : CategoryStruct.comp g q = CategoryStruct.comp p g') :
              (parallelPairHom f g f' g' p q wf wg).app WalkingParallelPair.zero = p
              @[simp]
              theorem CategoryTheory.Limits.parallelPairHom_app_one {C : Type u} [Category.{v, u} C] {X Y X' Y' : C} (f g : X Y) (f' g' : X' Y') (p : X X') (q : Y Y') (wf : CategoryStruct.comp f q = CategoryStruct.comp p f') (wg : CategoryStruct.comp g q = CategoryStruct.comp p g') :
              (parallelPairHom f g f' g' p q wf wg).app WalkingParallelPair.one = q

              Construct a natural isomorphism between functors out of the walking parallel pair from its components.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                def CategoryTheory.Limits.parallelPair.eqOfHomEq {C : Type u} [Category.{v, u} C] {X Y : C} {f g f' g' : X Y} (hf : f = f') (hg : g = g') :

                Construct a natural isomorphism between parallelPair f g and parallelPair f' g' given equalities f = f' and g = g'.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  @[simp]
                  theorem CategoryTheory.Limits.parallelPair.eqOfHomEq_hom_app {C : Type u} [Category.{v, u} C] {X Y : C} {f g f' g' : X Y} (hf : f = f') (hg : g = g') (X✝ : WalkingParallelPair) :
                  (eqOfHomEq hf hg).hom.app X✝ = (WalkingParallelPair.rec (Iso.refl X) (Iso.refl Y) X✝).hom
                  @[simp]
                  theorem CategoryTheory.Limits.parallelPair.eqOfHomEq_inv_app {C : Type u} [Category.{v, u} C] {X Y : C} {f g f' g' : X Y} (hf : f = f') (hg : g = g') (X✝ : WalkingParallelPair) :
                  (eqOfHomEq hf hg).inv.app X✝ = (WalkingParallelPair.rec (Iso.refl X) (Iso.refl Y) X✝).inv
                  @[reducible, inline]
                  abbrev CategoryTheory.Limits.Fork {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) :
                  Type (max u v)

                  A fork on f and g is just a Cone (parallelPair f g).

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

                    A cofork on f and g is just a Cocone (parallelPair f g).

                    Equations
                    Instances For

                      A fork t on the parallel pair f g : X ⟶ Y consists of two morphisms t.π.app zero : t.pt ⟶ X and t.π.app one : t.pt ⟶ Y. Of these, only the first one is interesting, and we give it the shorter name Fork.ι t.

                      Equations
                      Instances For
                        @[simp]
                        theorem CategoryTheory.Limits.Fork.app_zero_eq_ι {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (t : Fork f g) :

                        A cofork t on the parallelPair f g : X ⟶ Y consists of two morphisms t.ι.app zero : X ⟶ t.pt and t.ι.app one : Y ⟶ t.pt. Of these, only the second one is interesting, and we give it the shorter name Cofork.π t.

                        Equations
                        Instances For
                          @[simp]
                          theorem CategoryTheory.Limits.Cofork.app_one_eq_π {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (t : Cofork f g) :
                          @[simp]
                          def CategoryTheory.Limits.Fork.ofι {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {P : C} (ι : P X) (w : CategoryStruct.comp ι f = CategoryStruct.comp ι g) :
                          Fork f g

                          A fork on f g : X ⟶ Y is determined by the morphism ι : P ⟶ X satisfying ι ≫ f = ι ≫ g.

                          Equations
                          • One or more equations did not get rendered due to their size.
                          Instances For
                            @[simp]
                            theorem CategoryTheory.Limits.Fork.ofι_pt {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {P : C} (ι : P X) (w : CategoryStruct.comp ι f = CategoryStruct.comp ι g) :
                            (ofι ι w).pt = P
                            @[simp]
                            theorem CategoryTheory.Limits.Fork.ofι_π_app {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {P : C} (ι : P X) (w : CategoryStruct.comp ι f = CategoryStruct.comp ι g) (X✝ : WalkingParallelPair) :
                            (ofι ι w).app X✝ = WalkingParallelPair.casesOn (motive := fun (t : WalkingParallelPair) => X✝ = t(((Functor.const WalkingParallelPair).obj P).obj X✝ (parallelPair f g).obj X✝)) X✝ (fun (h : X✝ = WalkingParallelPair.zero) => ι) (fun (h : X✝ = WalkingParallelPair.one) => CategoryStruct.comp ι f)
                            def CategoryTheory.Limits.Cofork.ofπ {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {P : C} (π : Y P) (w : CategoryStruct.comp f π = CategoryStruct.comp g π) :
                            Cofork f g

                            A cofork on f g : X ⟶ Y is determined by the morphism π : Y ⟶ P satisfying f ≫ π = g ≫ π.

                            Equations
                            • One or more equations did not get rendered due to their size.
                            Instances For
                              @[simp]
                              theorem CategoryTheory.Limits.Cofork.ofπ_ι_app {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {P : C} (π : Y P) (w : CategoryStruct.comp f π = CategoryStruct.comp g π) (X✝ : WalkingParallelPair) :
                              (ofπ π w).app X✝ = WalkingParallelPair.casesOn X✝ (CategoryStruct.comp f π) π
                              @[simp]
                              theorem CategoryTheory.Limits.Cofork.ofπ_pt {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {P : C} (π : Y P) (w : CategoryStruct.comp f π = CategoryStruct.comp g π) :
                              (ofπ π w).pt = P
                              @[simp]
                              theorem CategoryTheory.Limits.Fork.ι_ofι {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {P : C} (ι : P X) (w : CategoryStruct.comp ι f = CategoryStruct.comp ι g) :
                              (ofι ι w) = ι
                              @[simp]
                              theorem CategoryTheory.Limits.Cofork.π_ofπ {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {P : C} (π : Y P) (w : CategoryStruct.comp f π = CategoryStruct.comp g π) :
                              (ofπ π w) = π
                              @[simp]
                              theorem CategoryTheory.Limits.Fork.condition {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (t : Fork f g) :
                              @[simp]
                              theorem CategoryTheory.Limits.Fork.condition_assoc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (t : Fork f g) {Z : C} (h : Y Z) :
                              @[simp]
                              theorem CategoryTheory.Limits.Cofork.condition {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (t : Cofork f g) :
                              theorem CategoryTheory.Limits.Fork.equalizer_ext {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (s : Fork f g) {W : C} {k l : W s.pt} (h : CategoryStruct.comp k s = CategoryStruct.comp l s) (j : WalkingParallelPair) :
                              CategoryStruct.comp k (s.app j) = CategoryStruct.comp l (s.app j)

                              To check whether two maps are equalized by both maps of a fork, it suffices to check it for the first map

                              theorem CategoryTheory.Limits.Cofork.coequalizer_ext {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (s : Cofork f g) {W : C} {k l : s.pt W} (h : CategoryStruct.comp s k = CategoryStruct.comp s l) (j : WalkingParallelPair) :
                              CategoryStruct.comp (s.app j) k = CategoryStruct.comp (s.app j) l

                              To check whether two maps are coequalized by both maps of a cofork, it suffices to check it for the second map

                              theorem CategoryTheory.Limits.Fork.IsLimit.hom_ext {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s : Fork f g} (hs : IsLimit s) {W : C} {k l : W s.pt} (h : CategoryStruct.comp k s = CategoryStruct.comp l s) :
                              k = l
                              theorem CategoryTheory.Limits.Cofork.IsColimit.hom_ext {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s : Cofork f g} (hs : IsColimit s) {W : C} {k l : s.pt W} (h : CategoryStruct.comp s k = CategoryStruct.comp s l) :
                              k = l
                              @[simp]
                              theorem CategoryTheory.Limits.Fork.IsLimit.lift_ι {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Fork f g} (hs : IsLimit s) :
                              CategoryStruct.comp (hs.lift t) s = t
                              @[simp]
                              theorem CategoryTheory.Limits.Fork.IsLimit.lift_ι_assoc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Fork f g} (hs : IsLimit s) {Z : C} (h : (parallelPair f g).obj WalkingParallelPair.zero Z) :
                              @[simp]
                              theorem CategoryTheory.Limits.Cofork.IsColimit.π_desc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Cofork f g} (hs : IsColimit s) :
                              CategoryStruct.comp s (hs.desc t) = t
                              @[simp]
                              theorem CategoryTheory.Limits.Cofork.IsColimit.π_desc_assoc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Cofork f g} (hs : IsColimit s) {Z : C} (h : t.pt Z) :
                              def CategoryTheory.Limits.Fork.IsLimit.lift {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s : Fork f g} (hs : IsLimit s) {W : C} (k : W X) (h : CategoryStruct.comp k f = CategoryStruct.comp k g) :
                              W s.pt

                              If s is a limit fork over f and g, then a morphism k : W ⟶ X satisfying k ≫ f = k ≫ g induces a morphism l : W ⟶ s.pt such that l ≫ fork.ι s = k.

                              Equations
                              Instances For
                                @[simp]
                                theorem CategoryTheory.Limits.Fork.IsLimit.lift_ι' {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s : Fork f g} (hs : IsLimit s) {W : C} (k : W X) (h : CategoryStruct.comp k f = CategoryStruct.comp k g) :
                                CategoryStruct.comp (lift hs k h) s = k
                                @[simp]
                                theorem CategoryTheory.Limits.Fork.IsLimit.lift_ι'_assoc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s : Fork f g} (hs : IsLimit s) {W : C} (k : W X) (h : CategoryStruct.comp k f = CategoryStruct.comp k g) {Z : C} (h✝ : (parallelPair f g).obj WalkingParallelPair.zero Z) :
                                def CategoryTheory.Limits.Fork.IsLimit.lift' {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s : Fork f g} (hs : IsLimit s) {W : C} (k : W X) (h : CategoryStruct.comp k f = CategoryStruct.comp k g) :
                                { l : W s.pt // CategoryStruct.comp l s = k }

                                If s is a limit fork over f and g, then a morphism k : W ⟶ X satisfying k ≫ f = k ≫ g induces a morphism l : W ⟶ s.pt such that l ≫ fork.ι s = k.

                                Equations
                                Instances For
                                  def CategoryTheory.Limits.Cofork.IsColimit.desc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s : Cofork f g} (hs : IsColimit s) {W : C} (k : Y W) (h : CategoryStruct.comp f k = CategoryStruct.comp g k) :
                                  s.pt W

                                  If s is a colimit cofork over f and g, then a morphism k : Y ⟶ W satisfying f ≫ k = g ≫ k induces a morphism l : s.pt ⟶ W such that cofork.π s ≫ l = k.

                                  Equations
                                  Instances For
                                    @[simp]
                                    theorem CategoryTheory.Limits.Cofork.IsColimit.π_desc' {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s : Cofork f g} (hs : IsColimit s) {W : C} (k : Y W) (h : CategoryStruct.comp f k = CategoryStruct.comp g k) :
                                    CategoryStruct.comp s (desc hs k h) = k
                                    @[simp]
                                    theorem CategoryTheory.Limits.Cofork.IsColimit.π_desc'_assoc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s : Cofork f g} (hs : IsColimit s) {W : C} (k : Y W) (h : CategoryStruct.comp f k = CategoryStruct.comp g k) {Z : C} (h✝ : W Z) :
                                    def CategoryTheory.Limits.Cofork.IsColimit.desc' {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s : Cofork f g} (hs : IsColimit s) {W : C} (k : Y W) (h : CategoryStruct.comp f k = CategoryStruct.comp g k) :
                                    { l : s.pt W // CategoryStruct.comp s l = k }

                                    If s is a colimit cofork over f and g, then a morphism k : Y ⟶ W satisfying f ≫ k = g ≫ k induces a morphism l : s.pt ⟶ W such that cofork.π s ≫ l = k.

                                    Equations
                                    Instances For
                                      theorem CategoryTheory.Limits.Fork.IsLimit.existsUnique {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s : Fork f g} (hs : IsLimit s) {W : C} (k : W X) (h : CategoryStruct.comp k f = CategoryStruct.comp k g) :
                                      ∃! l : W s.pt, CategoryStruct.comp l s = k
                                      theorem CategoryTheory.Limits.Cofork.IsColimit.existsUnique {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s : Cofork f g} (hs : IsColimit s) {W : C} (k : Y W) (h : CategoryStruct.comp f k = CategoryStruct.comp g k) :
                                      ∃! d : s.pt W, CategoryStruct.comp s d = k
                                      def CategoryTheory.Limits.Fork.IsLimit.mk {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (t : Fork f g) (lift : (s : Fork f g) → s.pt t.pt) (fac : ∀ (s : Fork f g), CategoryStruct.comp (lift s) t = s) (uniq : ∀ (s : Fork f g) (m : s.pt t.pt), CategoryStruct.comp m t = sm = lift s) :

                                      This is a slightly more convenient method to verify that a fork is a limit cone. It only asks for a proof of facts that carry any mathematical content

                                      Equations
                                      Instances For
                                        @[simp]
                                        theorem CategoryTheory.Limits.Fork.IsLimit.mk_lift {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (t : Fork f g) (lift : (s : Fork f g) → s.pt t.pt) (fac : ∀ (s : Fork f g), CategoryStruct.comp (lift s) t = s) (uniq : ∀ (s : Fork f g) (m : s.pt t.pt), CategoryStruct.comp m t = sm = lift s) (s : Fork f g) :
                                        (mk t lift fac uniq).lift s = lift s

                                        This is another convenient method to verify that a fork is a limit cone. It only asks for a proof of facts that carry any mathematical content, and allows access to the same s for all parts.

                                        Equations
                                        Instances For
                                          def CategoryTheory.Limits.Cofork.IsColimit.mk {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (t : Cofork f g) (desc : (s : Cofork f g) → t.pt s.pt) (fac : ∀ (s : Cofork f g), CategoryStruct.comp t (desc s) = s) (uniq : ∀ (s : Cofork f g) (m : t.pt s.pt), CategoryStruct.comp t m = sm = desc s) :

                                          This is a slightly more convenient method to verify that a cofork is a colimit cocone. It only asks for a proof of facts that carry any mathematical content

                                          Equations
                                          Instances For
                                            def CategoryTheory.Limits.Cofork.IsColimit.mk' {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (t : Cofork f g) (create : (s : Cofork f g) → { l : t.pt s.pt // CategoryStruct.comp t l = s ∀ {m : ((Functor.const WalkingParallelPair).obj t.pt).obj WalkingParallelPair.one ((Functor.const WalkingParallelPair).obj s.pt).obj WalkingParallelPair.one}, CategoryStruct.comp t m = sm = l }) :

                                            This is another convenient method to verify that a fork is a limit cone. It only asks for a proof of facts that carry any mathematical content, and allows access to the same s for all parts.

                                            Equations
                                            Instances For
                                              noncomputable def CategoryTheory.Limits.Fork.IsLimit.ofExistsUnique {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {t : Fork f g} (hs : ∀ (s : Fork f g), ∃! l : s.pt t.pt, CategoryStruct.comp l t = s) :

                                              Noncomputably make a limit cone from the existence of unique factorizations.

                                              Equations
                                              • One or more equations did not get rendered due to their size.
                                              Instances For
                                                noncomputable def CategoryTheory.Limits.Cofork.IsColimit.ofExistsUnique {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {t : Cofork f g} (hs : ∀ (s : Cofork f g), ∃! d : t.pt s.pt, CategoryStruct.comp t d = s) :

                                                Noncomputably make a colimit cocone from the existence of unique factorizations.

                                                Equations
                                                • One or more equations did not get rendered due to their size.
                                                Instances For
                                                  def CategoryTheory.Limits.Fork.IsLimit.homIso {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {t : Fork f g} (ht : IsLimit t) (Z : C) :
                                                  (Z t.pt) { h : Z X // CategoryStruct.comp h f = CategoryStruct.comp h g }

                                                  Given a limit cone for the pair f g : X ⟶ Y, for any Z, morphisms from Z to its point are in bijection with morphisms h : Z ⟶ X such that h ≫ f = h ≫ g. Further, this bijection is natural in Z: see Fork.IsLimit.homIso_natural. This is a special case of IsLimit.homIso', often useful to construct adjunctions.

                                                  Equations
                                                  • One or more equations did not get rendered due to their size.
                                                  Instances For
                                                    @[simp]
                                                    theorem CategoryTheory.Limits.Fork.IsLimit.homIso_symm_apply {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {t : Fork f g} (ht : IsLimit t) (Z : C) (h : { h : Z X // CategoryStruct.comp h f = CategoryStruct.comp h g }) :
                                                    (homIso ht Z).symm h = (lift' ht h )
                                                    @[simp]
                                                    theorem CategoryTheory.Limits.Fork.IsLimit.homIso_apply_coe {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {t : Fork f g} (ht : IsLimit t) (Z : C) (k : Z t.pt) :
                                                    ((homIso ht Z) k) = CategoryStruct.comp k t
                                                    theorem CategoryTheory.Limits.Fork.IsLimit.homIso_natural {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {t : Fork f g} (ht : IsLimit t) {Z Z' : C} (q : Z' Z) (k : Z t.pt) :
                                                    ((homIso ht Z') (CategoryStruct.comp q k)) = CategoryStruct.comp q ((homIso ht Z) k)

                                                    The bijection of Fork.IsLimit.homIso is natural in Z.

                                                    def CategoryTheory.Limits.Cofork.IsColimit.homIso {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {t : Cofork f g} (ht : IsColimit t) (Z : C) :
                                                    (t.pt Z) { h : Y Z // CategoryStruct.comp f h = CategoryStruct.comp g h }

                                                    Given a colimit cocone for the pair f g : X ⟶ Y, for any Z, morphisms from the cocone point to Z are in bijection with morphisms h : Y ⟶ Z such that f ≫ h = g ≫ h. Further, this bijection is natural in Z: see Cofork.IsColimit.homIso_natural. This is a special case of IsColimit.homIso', often useful to construct adjunctions.

                                                    Equations
                                                    • One or more equations did not get rendered due to their size.
                                                    Instances For
                                                      @[simp]
                                                      theorem CategoryTheory.Limits.Cofork.IsColimit.homIso_apply_coe {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {t : Cofork f g} (ht : IsColimit t) (Z : C) (k : t.pt Z) :
                                                      ((homIso ht Z) k) = CategoryStruct.comp t k
                                                      @[simp]
                                                      theorem CategoryTheory.Limits.Cofork.IsColimit.homIso_symm_apply {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {t : Cofork f g} (ht : IsColimit t) (Z : C) (h : { h : Y Z // CategoryStruct.comp f h = CategoryStruct.comp g h }) :
                                                      (homIso ht Z).symm h = (desc' ht h )
                                                      theorem CategoryTheory.Limits.Cofork.IsColimit.homIso_natural {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {t : Cofork f g} {Z Z' : C} (q : Z Z') (ht : IsColimit t) (k : t.pt Z) :
                                                      ((homIso ht Z') (CategoryStruct.comp k q)) = CategoryStruct.comp (↑((homIso ht Z) k)) q

                                                      The bijection of Cofork.IsColimit.homIso is natural in Z.

                                                      This is a helper construction that can be useful when verifying that a category has all equalizers. Given F : WalkingParallelPair ⥤ C, which is really the same as parallelPair (F.map left) (F.map right), and a fork on F.map left and F.map right, we get a cone on F.

                                                      If you're thinking about using this, have a look at hasEqualizers_of_hasLimit_parallelPair, which you may find to be an easier way of achieving your goal.

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

                                                        This is a helper construction that can be useful when verifying that a category has all coequalizers. Given F : WalkingParallelPair ⥤ C, which is really the same as parallelPair (F.map left) (F.map right), and a cofork on F.map left and F.map right, we get a cocone on F.

                                                        If you're thinking about using this, have a look at hasCoequalizers_of_hasColimit_parallelPair, which you may find to be an easier way of achieving your goal.

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

                                                          Given F : WalkingParallelPair ⥤ C, which is really the same as parallelPair (F.map left) (F.map right) and a cone on F, we get a fork on F.map left and F.map right.

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

                                                            Given F : WalkingParallelPair ⥤ C, which is really the same as parallelPair (F.map left) (F.map right) and a cocone on F, we get a cofork on F.map left and F.map right.

                                                            Equations
                                                            • One or more equations did not get rendered due to their size.
                                                            Instances For
                                                              @[simp]
                                                              theorem CategoryTheory.Limits.Fork.ι_postcompose {C : Type u} [Category.{v, u} C] {X Y : C} {f g f' g' : X Y} {α : parallelPair f g parallelPair f' g'} {c : Fork f g} :
                                                              @[simp]
                                                              theorem CategoryTheory.Limits.Cofork.π_precompose {C : Type u} [Category.{v, u} C] {X Y : C} {f g f' g' : X Y} {α : parallelPair f g parallelPair f' g'} {c : Cofork f' g'} :
                                                              def CategoryTheory.Limits.Fork.mkHom {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Fork f g} (k : s.pt t.pt) (w : CategoryStruct.comp k t = s) :
                                                              s t

                                                              Helper function for constructing morphisms between equalizer forks.

                                                              Equations
                                                              Instances For
                                                                @[simp]
                                                                theorem CategoryTheory.Limits.Fork.mkHom_hom {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Fork f g} (k : s.pt t.pt) (w : CategoryStruct.comp k t = s) :
                                                                (mkHom k w).hom = k
                                                                def CategoryTheory.Limits.Fork.ext {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Fork f g} (i : s.pt t.pt) (w : CategoryStruct.comp i.hom t = s := by aesop_cat) :
                                                                s t

                                                                To construct an isomorphism between forks, it suffices to give an isomorphism between the cone points and check that it commutes with the ι morphisms.

                                                                Equations
                                                                Instances For
                                                                  @[simp]
                                                                  theorem CategoryTheory.Limits.Fork.ext_hom {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Fork f g} (i : s.pt t.pt) (w : CategoryStruct.comp i.hom t = s := by aesop_cat) :
                                                                  (ext i w).hom = mkHom i.hom w
                                                                  @[simp]
                                                                  theorem CategoryTheory.Limits.Fork.ext_inv {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Fork f g} (i : s.pt t.pt) (w : CategoryStruct.comp i.hom t = s := by aesop_cat) :
                                                                  (ext i w).inv = mkHom i.inv
                                                                  def CategoryTheory.Limits.ForkOfι.ext {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {P : C} {ι ι' : P X} (w : CategoryStruct.comp ι f = CategoryStruct.comp ι g) (w' : CategoryStruct.comp ι' f = CategoryStruct.comp ι' g) (h : ι = ι') :
                                                                  Fork.ofι ι w Fork.ofι ι' w'

                                                                  Two forks of the form ofι are isomorphic whenever their ι's are equal.

                                                                  Equations
                                                                  Instances For
                                                                    def CategoryTheory.Limits.Fork.isoForkOfι {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (c : Fork f g) :
                                                                    c ofι c

                                                                    Every fork is isomorphic to one of the form Fork.of_ι _ _.

                                                                    Equations
                                                                    Instances For
                                                                      def CategoryTheory.Limits.Fork.isLimitOfIsos {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {X' Y' : C} (c : Fork f g) (hc : IsLimit c) {f' g' : X' Y'} (c' : Fork f' g') (e₀ : X X') (e₁ : Y Y') (e : c.pt c'.pt) (comm₁ : CategoryStruct.comp e₀.hom f' = CategoryStruct.comp f e₁.hom := by aesop_cat) (comm₂ : CategoryStruct.comp e₀.hom g' = CategoryStruct.comp g e₁.hom := by aesop_cat) (comm₃ : CategoryStruct.comp e.hom c' = CategoryStruct.comp c e₀.hom := by aesop_cat) :

                                                                      Given two forks with isomorphic components in such a way that the natural diagrams commute, then if one is a limit, then the other one is as well.

                                                                      Equations
                                                                      • One or more equations did not get rendered due to their size.
                                                                      Instances For
                                                                        def CategoryTheory.Limits.Cofork.mkHom {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Cofork f g} (k : s.pt t.pt) (w : CategoryStruct.comp s k = t) :
                                                                        s t

                                                                        Helper function for constructing morphisms between coequalizer coforks.

                                                                        Equations
                                                                        Instances For
                                                                          @[simp]
                                                                          theorem CategoryTheory.Limits.Cofork.mkHom_hom {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Cofork f g} (k : s.pt t.pt) (w : CategoryStruct.comp s k = t) :
                                                                          (mkHom k w).hom = k
                                                                          @[simp]
                                                                          theorem CategoryTheory.Limits.Fork.hom_comp_ι {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Fork f g} (f✝ : s t) :
                                                                          CategoryStruct.comp f✝.hom t = s
                                                                          @[simp]
                                                                          theorem CategoryTheory.Limits.Fork.hom_comp_ι_assoc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Fork f g} (f✝ : s t) {Z : C} (h : (parallelPair f g).obj WalkingParallelPair.zero Z) :
                                                                          @[simp]
                                                                          theorem CategoryTheory.Limits.Fork.π_comp_hom {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Cofork f g} (f✝ : s t) :
                                                                          CategoryStruct.comp s f✝.hom = t
                                                                          @[simp]
                                                                          theorem CategoryTheory.Limits.Fork.π_comp_hom_assoc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Cofork f g} (f✝ : s t) {Z : C} (h : t.pt Z) :
                                                                          def CategoryTheory.Limits.Cofork.ext {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Cofork f g} (i : s.pt t.pt) (w : CategoryStruct.comp s i.hom = t := by aesop_cat) :
                                                                          s t

                                                                          To construct an isomorphism between coforks, it suffices to give an isomorphism between the cocone points and check that it commutes with the π morphisms.

                                                                          Equations
                                                                          Instances For
                                                                            @[simp]
                                                                            theorem CategoryTheory.Limits.Cofork.ext_hom {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Cofork f g} (i : s.pt t.pt) (w : CategoryStruct.comp s i.hom = t := by aesop_cat) :
                                                                            (ext i w).hom = mkHom i.hom w
                                                                            @[simp]
                                                                            theorem CategoryTheory.Limits.Cofork.ext_inv {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {s t : Cofork f g} (i : s.pt t.pt) (w : CategoryStruct.comp s i.hom = t := by aesop_cat) :
                                                                            (ext i w).inv = mkHom i.inv
                                                                            def CategoryTheory.Limits.Cofork.isoCoforkOfπ {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (c : Cofork f g) :
                                                                            c ofπ c

                                                                            Every cofork is isomorphic to one of the form Cofork.ofπ _ _.

                                                                            Equations
                                                                            Instances For
                                                                              @[reducible, inline]
                                                                              abbrev CategoryTheory.Limits.HasEqualizer {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) :

                                                                              HasEqualizer f g represents a particular choice of limiting cone for the parallel pair of morphisms f and g.

                                                                              Equations
                                                                              Instances For
                                                                                @[reducible, inline]
                                                                                noncomputable abbrev CategoryTheory.Limits.equalizer {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) [HasEqualizer f g] :
                                                                                C

                                                                                If an equalizer of f and g exists, we can access an arbitrary choice of such by saying equalizer f g.

                                                                                Equations
                                                                                Instances For
                                                                                  @[reducible, inline]
                                                                                  noncomputable abbrev CategoryTheory.Limits.equalizer.ι {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) [HasEqualizer f g] :

                                                                                  If an equalizer of f and g exists, we can access the inclusion equalizer f g ⟶ X by saying equalizer.ι f g.

                                                                                  Equations
                                                                                  Instances For
                                                                                    @[reducible, inline]
                                                                                    noncomputable abbrev CategoryTheory.Limits.equalizer.fork {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) [HasEqualizer f g] :
                                                                                    Fork f g

                                                                                    An equalizer cone for a parallel pair f and g

                                                                                    Equations
                                                                                    Instances For
                                                                                      @[simp]
                                                                                      theorem CategoryTheory.Limits.equalizer.fork_ι {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) [HasEqualizer f g] :
                                                                                      (fork f g) = ι f g
                                                                                      @[simp]
                                                                                      noncomputable def CategoryTheory.Limits.equalizerIsEqualizer {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) [HasEqualizer f g] :

                                                                                      The equalizer built from equalizer.ι f g is limiting.

                                                                                      Equations
                                                                                      • One or more equations did not get rendered due to their size.
                                                                                      Instances For
                                                                                        @[reducible, inline]
                                                                                        noncomputable abbrev CategoryTheory.Limits.equalizer.lift {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasEqualizer f g] {W : C} (k : W X) (h : CategoryStruct.comp k f = CategoryStruct.comp k g) :

                                                                                        A morphism k : W ⟶ X satisfying k ≫ f = k ≫ g factors through the equalizer of f and g via equalizer.lift : W ⟶ equalizer f g.

                                                                                        Equations
                                                                                        Instances For
                                                                                          theorem CategoryTheory.Limits.equalizer.lift_ι {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasEqualizer f g] {W : C} (k : W X) (h : CategoryStruct.comp k f = CategoryStruct.comp k g) :
                                                                                          theorem CategoryTheory.Limits.equalizer.lift_ι_assoc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasEqualizer f g] {W : C} (k : W X) (h : CategoryStruct.comp k f = CategoryStruct.comp k g) {Z : C} (h✝ : X Z) :
                                                                                          noncomputable def CategoryTheory.Limits.equalizer.lift' {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasEqualizer f g] {W : C} (k : W X) (h : CategoryStruct.comp k f = CategoryStruct.comp k g) :
                                                                                          { l : W equalizer f g // CategoryStruct.comp l (ι f g) = k }

                                                                                          A morphism k : W ⟶ X satisfying k ≫ f = k ≫ g induces a morphism l : W ⟶ equalizer f g satisfying l ≫ equalizer.ι f g = k.

                                                                                          Equations
                                                                                          Instances For
                                                                                            theorem CategoryTheory.Limits.equalizer.hom_ext {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasEqualizer f g] {W : C} {k l : W equalizer f g} (h : CategoryStruct.comp k (ι f g) = CategoryStruct.comp l (ι f g)) :
                                                                                            k = l

                                                                                            Two maps into an equalizer are equal if they are equal when composed with the equalizer map.

                                                                                            theorem CategoryTheory.Limits.equalizer.existsUnique {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasEqualizer f g] {W : C} (k : W X) (h : CategoryStruct.comp k f = CategoryStruct.comp k g) :
                                                                                            ∃! l : W equalizer f g, CategoryStruct.comp l (ι f g) = k
                                                                                            instance CategoryTheory.Limits.equalizer.ι_mono {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasEqualizer f g] :
                                                                                            Mono (ι f g)

                                                                                            An equalizer morphism is a monomorphism

                                                                                            theorem CategoryTheory.Limits.mono_of_isLimit_fork {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {c : Fork f g} (i : IsLimit c) :
                                                                                            Mono c

                                                                                            The equalizer morphism in any limit cone is a monomorphism.

                                                                                            def CategoryTheory.Limits.idFork {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (h : f = g) :
                                                                                            Fork f g

                                                                                            The identity determines a cone on the equalizer diagram of f and g if f = g.

                                                                                            Equations
                                                                                            Instances For
                                                                                              def CategoryTheory.Limits.isLimitIdFork {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (h : f = g) :

                                                                                              The identity on X is an equalizer of (f, g), if f = g.

                                                                                              Equations
                                                                                              Instances For
                                                                                                theorem CategoryTheory.Limits.isIso_limit_cone_parallelPair_of_eq {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (h₀ : f = g) {c : Fork f g} (h : IsLimit c) :
                                                                                                IsIso c

                                                                                                Every equalizer of (f, g), where f = g, is an isomorphism.

                                                                                                theorem CategoryTheory.Limits.equalizer.ι_of_eq {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasEqualizer f g] (h : f = g) :
                                                                                                IsIso (ι f g)

                                                                                                The equalizer of (f, g), where f = g, is an isomorphism.

                                                                                                theorem CategoryTheory.Limits.isIso_limit_cone_parallelPair_of_self {C : Type u} [Category.{v, u} C] {X Y : C} {f : X Y} {c : Fork f f} (h : IsLimit c) :
                                                                                                IsIso c

                                                                                                Every equalizer of (f, f) is an isomorphism.

                                                                                                theorem CategoryTheory.Limits.isIso_limit_cone_parallelPair_of_epi {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {c : Fork f g} (h : IsLimit c) [Epi c] :
                                                                                                IsIso c

                                                                                                An equalizer that is an epimorphism is an isomorphism.

                                                                                                theorem CategoryTheory.Limits.eq_of_epi_fork_ι {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (t : Fork f g) [Epi t] :
                                                                                                f = g

                                                                                                Two morphisms are equal if there is a fork whose inclusion is epi.

                                                                                                theorem CategoryTheory.Limits.eq_of_epi_equalizer {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasEqualizer f g] [Epi (equalizer.ι f g)] :
                                                                                                f = g

                                                                                                If the equalizer of two morphisms is an epimorphism, then the two morphisms are equal.

                                                                                                instance CategoryTheory.Limits.equalizer.ι_of_self {C : Type u} [Category.{v, u} C] {X Y : C} (f : X Y) :
                                                                                                IsIso (ι f f)

                                                                                                The equalizer inclusion for (f, f) is an isomorphism.

                                                                                                noncomputable def CategoryTheory.Limits.equalizer.isoSourceOfSelf {C : Type u} [Category.{v, u} C] {X Y : C} (f : X Y) :

                                                                                                The equalizer of a morphism with itself is isomorphic to the source.

                                                                                                Equations
                                                                                                Instances For
                                                                                                  @[simp]
                                                                                                  @[reducible, inline]
                                                                                                  abbrev CategoryTheory.Limits.HasCoequalizer {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) :

                                                                                                  HasCoequalizer f g represents a particular choice of colimiting cocone for the parallel pair of morphisms f and g.

                                                                                                  Equations
                                                                                                  Instances For
                                                                                                    @[reducible, inline]
                                                                                                    noncomputable abbrev CategoryTheory.Limits.coequalizer {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) [HasCoequalizer f g] :
                                                                                                    C

                                                                                                    If a coequalizer of f and g exists, we can access an arbitrary choice of such by saying coequalizer f g.

                                                                                                    Equations
                                                                                                    Instances For
                                                                                                      @[reducible, inline]
                                                                                                      noncomputable abbrev CategoryTheory.Limits.coequalizer.π {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) [HasCoequalizer f g] :

                                                                                                      If a coequalizer of f and g exists, we can access the corresponding projection by saying coequalizer.π f g.

                                                                                                      Equations
                                                                                                      Instances For
                                                                                                        @[reducible, inline]
                                                                                                        noncomputable abbrev CategoryTheory.Limits.coequalizer.cofork {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) [HasCoequalizer f g] :
                                                                                                        Cofork f g

                                                                                                        An arbitrary choice of coequalizer cocone for a parallel pair f and g.

                                                                                                        Equations
                                                                                                        Instances For
                                                                                                          @[simp]
                                                                                                          theorem CategoryTheory.Limits.coequalizer.cofork_π {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) [HasCoequalizer f g] :
                                                                                                          (cofork f g) = π f g
                                                                                                          noncomputable def CategoryTheory.Limits.coequalizerIsCoequalizer {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) [HasCoequalizer f g] :

                                                                                                          The cofork built from coequalizer.π f g is colimiting.

                                                                                                          Equations
                                                                                                          • One or more equations did not get rendered due to their size.
                                                                                                          Instances For
                                                                                                            @[reducible, inline]
                                                                                                            noncomputable abbrev CategoryTheory.Limits.coequalizer.desc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasCoequalizer f g] {W : C} (k : Y W) (h : CategoryStruct.comp f k = CategoryStruct.comp g k) :

                                                                                                            Any morphism k : Y ⟶ W satisfying f ≫ k = g ≫ k factors through the coequalizer of f and g via coequalizer.desc : coequalizer f g ⟶ W.

                                                                                                            Equations
                                                                                                            Instances For
                                                                                                              theorem CategoryTheory.Limits.coequalizer.π_desc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasCoequalizer f g] {W : C} (k : Y W) (h : CategoryStruct.comp f k = CategoryStruct.comp g k) :
                                                                                                              theorem CategoryTheory.Limits.coequalizer.π_desc_assoc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasCoequalizer f g] {W : C} (k : Y W) (h : CategoryStruct.comp f k = CategoryStruct.comp g k) {Z : C} (h✝ : W Z) :
                                                                                                              theorem CategoryTheory.Limits.coequalizer.π_colimMap_desc {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasCoequalizer f g] {X' Y' Z : C} (f' g' : X' Y') [HasCoequalizer f' g'] (p : X X') (q : Y Y') (wf : CategoryStruct.comp f q = CategoryStruct.comp p f') (wg : CategoryStruct.comp g q = CategoryStruct.comp p g') (h : Y' Z) (wh : CategoryStruct.comp f' h = CategoryStruct.comp g' h) :
                                                                                                              noncomputable def CategoryTheory.Limits.coequalizer.desc' {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasCoequalizer f g] {W : C} (k : Y W) (h : CategoryStruct.comp f k = CategoryStruct.comp g k) :
                                                                                                              { l : coequalizer f g W // CategoryStruct.comp (π f g) l = k }

                                                                                                              Any morphism k : Y ⟶ W satisfying f ≫ k = g ≫ k induces a morphism l : coequalizer f g ⟶ W satisfying coequalizer.π ≫ g = l.

                                                                                                              Equations
                                                                                                              Instances For
                                                                                                                theorem CategoryTheory.Limits.coequalizer.hom_ext {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasCoequalizer f g] {W : C} {k l : coequalizer f g W} (h : CategoryStruct.comp (π f g) k = CategoryStruct.comp (π f g) l) :
                                                                                                                k = l

                                                                                                                Two maps from a coequalizer are equal if they are equal when composed with the coequalizer map

                                                                                                                theorem CategoryTheory.Limits.coequalizer.existsUnique {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasCoequalizer f g] {W : C} (k : Y W) (h : CategoryStruct.comp f k = CategoryStruct.comp g k) :
                                                                                                                ∃! d : coequalizer f g W, CategoryStruct.comp (π f g) d = k
                                                                                                                instance CategoryTheory.Limits.coequalizer.π_epi {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasCoequalizer f g] :
                                                                                                                Epi (π f g)

                                                                                                                A coequalizer morphism is an epimorphism

                                                                                                                theorem CategoryTheory.Limits.epi_of_isColimit_cofork {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {c : Cofork f g} (i : IsColimit c) :
                                                                                                                Epi c

                                                                                                                The coequalizer morphism in any colimit cocone is an epimorphism.

                                                                                                                def CategoryTheory.Limits.idCofork {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (h : f = g) :
                                                                                                                Cofork f g

                                                                                                                The identity determines a cocone on the coequalizer diagram of f and g, if f = g.

                                                                                                                Equations
                                                                                                                Instances For
                                                                                                                  def CategoryTheory.Limits.isColimitIdCofork {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (h : f = g) :

                                                                                                                  The identity on Y is a coequalizer of (f, g), where f = g.

                                                                                                                  Equations
                                                                                                                  Instances For
                                                                                                                    theorem CategoryTheory.Limits.isIso_colimit_cocone_parallelPair_of_eq {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (h₀ : f = g) {c : Cofork f g} (h : IsColimit c) :
                                                                                                                    IsIso c

                                                                                                                    Every coequalizer of (f, g), where f = g, is an isomorphism.

                                                                                                                    theorem CategoryTheory.Limits.coequalizer.π_of_eq {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasCoequalizer f g] (h : f = g) :
                                                                                                                    IsIso (π f g)

                                                                                                                    The coequalizer of (f, g), where f = g, is an isomorphism.

                                                                                                                    theorem CategoryTheory.Limits.isIso_colimit_cocone_parallelPair_of_self {C : Type u} [Category.{v, u} C] {X Y : C} {f : X Y} {c : Cofork f f} (h : IsColimit c) :
                                                                                                                    IsIso c

                                                                                                                    Every coequalizer of (f, f) is an isomorphism.

                                                                                                                    theorem CategoryTheory.Limits.isIso_limit_cocone_parallelPair_of_epi {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {c : Cofork f g} (h : IsColimit c) [Mono c] :
                                                                                                                    IsIso c

                                                                                                                    A coequalizer that is a monomorphism is an isomorphism.

                                                                                                                    theorem CategoryTheory.Limits.eq_of_mono_cofork_π {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (t : Cofork f g) [Mono t] :
                                                                                                                    f = g

                                                                                                                    Two morphisms are equal if there is a cofork whose projection is mono.

                                                                                                                    theorem CategoryTheory.Limits.eq_of_mono_coequalizer {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} [HasCoequalizer f g] [Mono (coequalizer.π f g)] :
                                                                                                                    f = g

                                                                                                                    If the coequalizer of two morphisms is a monomorphism, then the two morphisms are equal.

                                                                                                                    instance CategoryTheory.Limits.coequalizer.π_of_self {C : Type u} [Category.{v, u} C] {X Y : C} (f : X Y) :
                                                                                                                    IsIso (π f f)

                                                                                                                    The coequalizer projection for (f, f) is an isomorphism.

                                                                                                                    noncomputable def CategoryTheory.Limits.coequalizer.isoTargetOfSelf {C : Type u} [Category.{v, u} C] {X Y : C} (f : X Y) :

                                                                                                                    The coequalizer of a morphism with itself is isomorphic to the target.

                                                                                                                    Equations
                                                                                                                    Instances For
                                                                                                                      @[simp]
                                                                                                                      noncomputable def CategoryTheory.Limits.equalizerComparison {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) {D : Type u₂} [Category.{v₂, u₂} D] (G : Functor C D) [HasEqualizer f g] [HasEqualizer (G.map f) (G.map g)] :
                                                                                                                      G.obj (equalizer f g) equalizer (G.map f) (G.map g)

                                                                                                                      The comparison morphism for the equalizer of f,g. This is an isomorphism iff G preserves the equalizer of f,g; see CategoryTheory/Limits/Preserves/Shapes/Equalizers.lean

                                                                                                                      Equations
                                                                                                                      Instances For
                                                                                                                        @[simp]
                                                                                                                        theorem CategoryTheory.Limits.equalizerComparison_comp_π {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) {D : Type u₂} [Category.{v₂, u₂} D] (G : Functor C D) [HasEqualizer f g] [HasEqualizer (G.map f) (G.map g)] :
                                                                                                                        CategoryStruct.comp (equalizerComparison f g G) (equalizer.ι (G.map f) (G.map g)) = G.map (equalizer.ι f g)
                                                                                                                        @[simp]
                                                                                                                        theorem CategoryTheory.Limits.equalizerComparison_comp_π_assoc {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) {D : Type u₂} [Category.{v₂, u₂} D] (G : Functor C D) [HasEqualizer f g] [HasEqualizer (G.map f) (G.map g)] {Z : D} (h : G.obj X Z) :
                                                                                                                        @[simp]
                                                                                                                        theorem CategoryTheory.Limits.map_lift_equalizerComparison {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) {D : Type u₂} [Category.{v₂, u₂} D] (G : Functor C D) [HasEqualizer f g] [HasEqualizer (G.map f) (G.map g)] {Z : C} {h : Z X} (w : CategoryStruct.comp h f = CategoryStruct.comp h g) :
                                                                                                                        @[simp]
                                                                                                                        theorem CategoryTheory.Limits.map_lift_equalizerComparison_assoc {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) {D : Type u₂} [Category.{v₂, u₂} D] (G : Functor C D) [HasEqualizer f g] [HasEqualizer (G.map f) (G.map g)] {Z : C} {h : Z X} (w : CategoryStruct.comp h f = CategoryStruct.comp h g) {Z✝ : D} (h✝ : equalizer (G.map f) (G.map g) Z✝) :
                                                                                                                        noncomputable def CategoryTheory.Limits.coequalizerComparison {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) {D : Type u₂} [Category.{v₂, u₂} D] (G : Functor C D) [HasCoequalizer f g] [HasCoequalizer (G.map f) (G.map g)] :
                                                                                                                        coequalizer (G.map f) (G.map g) G.obj (coequalizer f g)

                                                                                                                        The comparison morphism for the coequalizer of f,g.

                                                                                                                        Equations
                                                                                                                        Instances For
                                                                                                                          @[simp]
                                                                                                                          theorem CategoryTheory.Limits.ι_comp_coequalizerComparison {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) {D : Type u₂} [Category.{v₂, u₂} D] (G : Functor C D) [HasCoequalizer f g] [HasCoequalizer (G.map f) (G.map g)] :
                                                                                                                          CategoryStruct.comp (coequalizer.π (G.map f) (G.map g)) (coequalizerComparison f g G) = G.map (coequalizer.π f g)
                                                                                                                          @[simp]
                                                                                                                          theorem CategoryTheory.Limits.ι_comp_coequalizerComparison_assoc {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) {D : Type u₂} [Category.{v₂, u₂} D] (G : Functor C D) [HasCoequalizer f g] [HasCoequalizer (G.map f) (G.map g)] {Z : D} (h : G.obj (coequalizer f g) Z) :
                                                                                                                          @[simp]
                                                                                                                          theorem CategoryTheory.Limits.coequalizerComparison_map_desc {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) {D : Type u₂} [Category.{v₂, u₂} D] (G : Functor C D) [HasCoequalizer f g] [HasCoequalizer (G.map f) (G.map g)] {Z : C} {h : Y Z} (w : CategoryStruct.comp f h = CategoryStruct.comp g h) :
                                                                                                                          @[simp]
                                                                                                                          theorem CategoryTheory.Limits.coequalizerComparison_map_desc_assoc {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) {D : Type u₂} [Category.{v₂, u₂} D] (G : Functor C D) [HasCoequalizer f g] [HasCoequalizer (G.map f) (G.map g)] {Z : C} {h : Y Z} (w : CategoryStruct.comp f h = CategoryStruct.comp g h) {Z✝ : D} (h✝ : G.obj Z Z✝) :
                                                                                                                          @[reducible, inline]

                                                                                                                          HasEqualizers represents a choice of equalizer for every pair of morphisms

                                                                                                                          Equations
                                                                                                                          Instances For
                                                                                                                            @[reducible, inline]

                                                                                                                            HasCoequalizers represents a choice of coequalizer for every pair of morphisms

                                                                                                                            Equations
                                                                                                                            Instances For

                                                                                                                              If C has all limits of diagrams parallelPair f g, then it has all equalizers

                                                                                                                              If C has all colimits of diagrams parallelPair f g, then it has all coequalizers

                                                                                                                              A split mono f equalizes (retraction f ≫ f) and (𝟙 Y). Here we build the cone, and show in isSplitMonoEqualizes that it is a limit cone.

                                                                                                                              Equations
                                                                                                                              Instances For
                                                                                                                                @[simp]
                                                                                                                                theorem CategoryTheory.Limits.coneOfIsSplitMono_π_app {C : Type u} [Category.{v, u} C] {X Y : C} (f : X Y) [IsSplitMono f] (X✝ : WalkingParallelPair) :
                                                                                                                                (coneOfIsSplitMono f).app X✝ = WalkingParallelPair.rec (motive := fun (t : WalkingParallelPair) => X✝ = t(X (parallelPair (CategoryStruct.id Y) (CategoryStruct.comp (retraction f) f)).obj X✝)) (fun (h : X✝ = WalkingParallelPair.zero) => f) (fun (h : X✝ = WalkingParallelPair.one) => f) X✝
                                                                                                                                @[simp]
                                                                                                                                @[simp]

                                                                                                                                A split mono f equalizes (retraction f ≫ f) and (𝟙 Y).

                                                                                                                                Equations
                                                                                                                                • One or more equations did not get rendered due to their size.
                                                                                                                                Instances For
                                                                                                                                  def CategoryTheory.Limits.splitMonoOfEqualizer (C : Type u) [Category.{v, u} C] {X Y : C} {f : X Y} {r : Y X} (hr : CategoryStruct.comp f (CategoryStruct.comp r f) = f) (h : IsLimit (Fork.ofι f )) :

                                                                                                                                  We show that the converse to isSplitMonoEqualizes is true: Whenever f equalizes (r ≫ f) and (𝟙 Y), then r is a retraction of f.

                                                                                                                                  Equations
                                                                                                                                  Instances For
                                                                                                                                    def CategoryTheory.Limits.isEqualizerCompMono {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {c : Fork f g} (i : IsLimit c) {Z : C} (h : Y Z) [hm : Mono h] :
                                                                                                                                    let_fun this := ; IsLimit (Fork.ofι c )

                                                                                                                                    The fork obtained by postcomposing an equalizer fork with a monomorphism is an equalizer.

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

                                                                                                                                      An equalizer of an idempotent morphism and the identity is split mono.

                                                                                                                                      Equations
                                                                                                                                      Instances For
                                                                                                                                        @[simp]

                                                                                                                                        The equalizer of an idempotent morphism and the identity is split mono.

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

                                                                                                                                          A split epi f coequalizes (f ≫ section_ f) and (𝟙 X). Here we build the cocone, and show in isSplitEpiCoequalizes that it is a colimit cocone.

                                                                                                                                          Equations
                                                                                                                                          Instances For
                                                                                                                                            @[simp]
                                                                                                                                            @[simp]

                                                                                                                                            A split epi f coequalizes (f ≫ section_ f) and (𝟙 X).

                                                                                                                                            Equations
                                                                                                                                            • One or more equations did not get rendered due to their size.
                                                                                                                                            Instances For
                                                                                                                                              def CategoryTheory.Limits.splitEpiOfCoequalizer (C : Type u) [Category.{v, u} C] {X Y : C} {f : X Y} {s : Y X} (hs : CategoryStruct.comp f (CategoryStruct.comp s f) = f) (h : IsColimit (Cofork.ofπ f )) :

                                                                                                                                              We show that the converse to isSplitEpiEqualizes is true: Whenever f coequalizes (f ≫ s) and (𝟙 X), then s is a section of f.

                                                                                                                                              Equations
                                                                                                                                              Instances For
                                                                                                                                                def CategoryTheory.Limits.isCoequalizerEpiComp {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} {c : Cofork f g} (i : IsColimit c) {W : C} (h : W X) [hm : Epi h] :
                                                                                                                                                let_fun this := ; IsColimit (Cofork.ofπ c this)

                                                                                                                                                The cofork obtained by precomposing a coequalizer cofork with an epimorphism is a coequalizer.

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

                                                                                                                                                  A coequalizer of an idempotent morphism and the identity is split epi.

                                                                                                                                                  Equations
                                                                                                                                                  Instances For

                                                                                                                                                    The coequalizer of an idempotent morphism and the identity is split epi.

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