Documentation

Mathlib.CategoryTheory.Limits.Shapes.Opposites.Equalizers

Equalizers and coequalizers in C and Cᵒᵖ #

We construct equalizers and coequalizers in the opposite categories.

The canonical isomorphism relating parallelPair f.op g.op and (parallelPair f g).op

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

    The canonical isomorphism relating (parallelPair f g).op and parallelPair f.op g.op

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      def CategoryTheory.Limits.Cofork.unop {C : Type u₁} [Category.{v₁, u₁} C] {X Y : Cᵒᵖ} {f g : X Y} (c : Cofork f g) :

      The obvious map Cofork f g → Fork f.unop g.unop

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        theorem CategoryTheory.Limits.Cofork.unop_ι {C : Type u₁} [Category.{v₁, u₁} C] {X Y : Cᵒᵖ} {f g : X Y} (c : Cofork f g) :
        def CategoryTheory.Limits.Cofork.op {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} {f g : X Y} (c : Cofork f g) :
        Fork f.op g.op

        The obvious map Cofork f g → Fork f.op g.op

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          theorem CategoryTheory.Limits.Cofork.op_ι {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} {f g : X Y} (c : Cofork f g) :
          c.op.ι = c.π.op
          def CategoryTheory.Limits.Fork.unop {C : Type u₁} [Category.{v₁, u₁} C] {X Y : Cᵒᵖ} {f g : X Y} (c : Fork f g) :

          The obvious map Fork f g → Cofork f.unop g.unop

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            theorem CategoryTheory.Limits.Fork.unop_π {C : Type u₁} [Category.{v₁, u₁} C] {X Y : Cᵒᵖ} {f g : X Y} (c : Fork f g) :
            def CategoryTheory.Limits.Fork.op {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} {f g : X Y} (c : Fork f g) :

            The obvious map Fork f g → Cofork f.op g.op

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              theorem CategoryTheory.Limits.Fork.op_pt {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} {f g : X Y} (c : Fork f g) :
              @[simp]
              theorem CategoryTheory.Limits.Fork.op_π {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} {f g : X Y} (c : Fork f g) :
              c.op.π = c.ι.op
              theorem CategoryTheory.Limits.Cofork.op_unop_π {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} {f g : X Y} (c : Cofork f g) :
              c.op.unop.π = c.π
              theorem CategoryTheory.Limits.Cofork.unop_op_π {C : Type u₁} [Category.{v₁, u₁} C] {X Y : Cᵒᵖ} {f g : X Y} (c : Cofork f g) :
              c.unop.op.π = c.π
              def CategoryTheory.Limits.Cofork.opUnopIso {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} {f g : X Y} (c : Cofork f g) :
              c.op.unop c

              If c is a cofork, then c.op.unop is isomorphic to c.

              Equations
              Instances For
                def CategoryTheory.Limits.Cofork.unopOpIso {C : Type u₁} [Category.{v₁, u₁} C] {X Y : Cᵒᵖ} {f g : X Y} (c : Cofork f g) :
                c.unop.op c

                If c is a cofork in Cᵒᵖ, then c.unop.op is isomorphic to c.

                Equations
                Instances For
                  theorem CategoryTheory.Limits.Fork.op_unop_ι {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} {f g : X Y} (c : Fork f g) :
                  c.op.unop.ι = c.ι
                  theorem CategoryTheory.Limits.Fork.unop_op_ι {C : Type u₁} [Category.{v₁, u₁} C] {X Y : Cᵒᵖ} {f g : X Y} (c : Fork f g) :
                  c.unop.op.ι = c.ι
                  def CategoryTheory.Limits.Fork.opUnopIso {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} {f g : X Y} (c : Fork f g) :
                  c.op.unop c

                  If c is a fork, then c.op.unop is isomorphic to c.

                  Equations
                  Instances For
                    def CategoryTheory.Limits.Fork.unopOpIso {C : Type u₁} [Category.{v₁, u₁} C] {X Y : Cᵒᵖ} {f g : X Y} (c : Fork f g) :
                    c.unop.op c

                    If c is a fork in Cᵒᵖ, then c.unop.op is isomorphic to c.

                    Equations
                    Instances For
                      noncomputable def CategoryTheory.Limits.Cofork.isColimitEquivIsLimitOp {C : Type u₁} [Category.{v₁, u₁} C] {X Y : C} {f g : X Y} (c : Cofork f g) :

                      A cofork is a colimit cocone if and only if the corresponding fork in the opposite category is a limit cone.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        noncomputable def CategoryTheory.Limits.Cofork.isColimitEquivIsLimitUnop {C : Type u₁} [Category.{v₁, u₁} C] {X Y : Cᵒᵖ} {f g : X Y} (c : Cofork f g) :

                        A cofork is a colimit cocone in Cᵒᵖ if and only if the corresponding fork in C is a limit cone.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          def CategoryTheory.Limits.Cofork.ofπOpIsoOfι {C : Type u₁} [Category.{v₁, u₁} C] {X Y P : C} {f g : X Y} (π π' : Y P) (w : CategoryStruct.comp f π = CategoryStruct.comp g π) (w' : CategoryStruct.comp π'.op f.op = CategoryStruct.comp π'.op g.op) (h : π = π') :
                          (ofπ π w).op Fork.ofι π'.op w'

                          The canonical isomorphism between (Cofork.ofπ π w).op and Fork.ofι π.op w'.

                          Equations
                          Instances For
                            def CategoryTheory.Limits.Cofork.ofπUnopIsoOfι {C : Type u₁} [Category.{v₁, u₁} C] {X Y P : Cᵒᵖ} {f g : X Y} (π π' : Y P) (w : CategoryStruct.comp f π = CategoryStruct.comp g π) (w' : CategoryStruct.comp π'.unop f.unop = CategoryStruct.comp π'.unop g.unop) (h : π = π') :
                            (ofπ π w).unop Fork.ofι π'.unop w'

                            The canonical isomorphism between (Cofork.ofπ π w).unop and Fork.ofι π.unop w'.

                            Equations
                            Instances For
                              def CategoryTheory.Limits.Cofork.isColimitOfπEquivIsLimitOp {C : Type u₁} [Category.{v₁, u₁} C] {X Y P : C} {f g : X Y} (π π' : Y P) (w : CategoryStruct.comp f π = CategoryStruct.comp g π) (w' : CategoryStruct.comp π'.op f.op = CategoryStruct.comp π'.op g.op) (h : π = π') :

                              Cofork.ofπ π w is a colimit cocone if and only if Fork.ofι π.op w' in the opposite category is a limit cone.

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

                                Cofork.ofπ π w is a colimit cocone in Cᵒᵖ if and only if Fork.ofι π'.unop w' in C is a limit cone.

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

                                  A fork is a limit cone if and only if the corresponding cofork in the opposite category is a colimit cocone.

                                  Equations
                                  Instances For

                                    A fork is a limit cone in Cᵒᵖ if and only if the corresponding cofork in C is a colimit cocone.

                                    Equations
                                    Instances For
                                      def CategoryTheory.Limits.Fork.ofιOpIsoOfπ {C : Type u₁} [Category.{v₁, u₁} C] {X Y P : C} {f g : X Y} (ι ι' : P X) (w : CategoryStruct.comp ι f = CategoryStruct.comp ι g) (w' : CategoryStruct.comp f.op ι'.op = CategoryStruct.comp g.op ι'.op) (h : ι = ι') :
                                      (ofι ι w).op Cofork.ofπ ι'.op w'

                                      The canonical isomorphism between (Fork.ofι ι w).op and Cofork.ofπ ι.op w'.

                                      Equations
                                      Instances For
                                        def CategoryTheory.Limits.Fork.ofιUnopIsoOfπ {C : Type u₁} [Category.{v₁, u₁} C] {X Y P : Cᵒᵖ} {f g : X Y} (ι ι' : P X) (w : CategoryStruct.comp ι f = CategoryStruct.comp ι g) (w' : CategoryStruct.comp f.unop ι'.unop = CategoryStruct.comp g.unop ι'.unop) (h : ι = ι') :
                                        (ofι ι w).unop Cofork.ofπ ι'.unop w'

                                        The canonical isomorphism between (Fork.ofι ι w).unop and Cofork.ofπ ι.unop w.unop.

                                        Equations
                                        Instances For
                                          def CategoryTheory.Limits.Fork.isLimitOfιEquivIsColimitOp {C : Type u₁} [Category.{v₁, u₁} C] {X Y P : C} {f g : X Y} (ι ι' : P X) (w : CategoryStruct.comp ι f = CategoryStruct.comp ι g) (w' : CategoryStruct.comp f.op ι'.op = CategoryStruct.comp g.op ι'.op) (h : ι = ι') :

                                          Fork.ofι ι w is a limit cone if and only if Cofork.ofπ ι'.op w' in the opposite category is a colimit cocone.

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

                                            Fork.ofι ι w is a limit cone in Cᵒᵖ if and only if Cofork.ofπ ι.unop w.unop in C is a colimit cocone.

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

                                              Cofork.ofπ f pullback.condition is a colimit cocone if and only if Fork.ofι f.op pushout.condition in the opposite category is a limit cone.

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

                                                Cofork.ofπ f pullback.condition is a colimit cocone in Cᵒᵖ if and only if Fork.ofι f.unop pushout.condition in C is a limit cone.

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

                                                  Fork.ofι f pushout.condition is a limit cone if and only if Cofork.ofπ f.op pullback.condition in the opposite category is a colimit cocone.

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

                                                    Fork.ofι f pushout.condition is a limit cone in Cᵒᵖ if and only if Cofork.ofπ f.op pullback.condition in C is a colimit cocone.

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