Documentation

Mathlib.CategoryTheory.Bicategory.Opposites

Opposite bicategories #

We construct the 1-cell opposite of a bicategory B, called Bᵒᵖ. It is defined as follows

Remarks #

There are multiple notions of opposite categories for bicategories.

TODO #

Note: Cᶜᵒᵒᵖ is WIP by Christian Merten.

structure Bicategory.Opposite.Hom2 {B : Type u} [CategoryTheory.Bicategory B] {a b : Bᵒᵖ} (f g : a b) :

Type synonym for 2-morphisms in the opposite bicategory.

  • op2' :: (
    • unop2 : f.unop g.unop

      Bᵒᵖ preserves the direction of all 2-morphisms in B

  • )
Instances For
    Equations
    • One or more equations did not get rendered due to their size.
    @[simp]
    def Bicategory.Opposite.op2 {B : Type u} [CategoryTheory.Bicategory B] {a b : B} {f g : a b} (η : f g) :
    f.op g.op

    Synonym for constructor of Hom2 where the 1-morphisms f and g lie in B and not Bᵒᵖ.

    Equations
    Instances For
      @[simp]
      theorem Bicategory.Opposite.unop2_op2 {B : Type u} [CategoryTheory.Bicategory B] {a b : B} {f g : a b} (η : f g) :
      (op2 η).unop2 = η
      @[simp]
      theorem Bicategory.Opposite.op2_unop2 {B : Type u} [CategoryTheory.Bicategory B] {a b : Bᵒᵖ} {f g : a b} (η : f g) :
      op2 η.unop2 = η
      @[simp]
      theorem Bicategory.Opposite.op2_comp {B : Type u} [CategoryTheory.Bicategory B] {a b : B} {f g h : a b} (η : f g) (θ : g h) :

      The natural functor from the hom-category a ⟶ b in B to its bicategorical opposite bop b ⟶ bop a.

      Equations
      Instances For
        @[simp]
        theorem Bicategory.Opposite.opFunctor_map {B : Type u} [CategoryTheory.Bicategory B] (a b : B) {X✝ Y✝ : a b} (η : X✝ Y✝) :
        (opFunctor a b).map η = op2 η
        @[simp]
        theorem Bicategory.Opposite.opFunctor_obj {B : Type u} [CategoryTheory.Bicategory B] (a b : B) (f : a b) :
        (opFunctor a b).obj f = f.op

        The functor from the hom-category a ⟶ b in Bᵒᵖ to its bicategorical opposite unop b ⟶ unop a.

        Equations
        Instances For
          @[simp]
          theorem Bicategory.Opposite.unopFunctor_map {B : Type u} [CategoryTheory.Bicategory B] (a b : Bᵒᵖ) {X✝ Y✝ : a b} (η : X✝ Y✝) :
          (unopFunctor a b).map η = η.unop2
          @[simp]
          @[reducible, inline]
          abbrev CategoryTheory.Iso.op2 {B : Type u} [Bicategory B] {a b : B} {f g : a b} (η : f g) :
          f.op g.op

          A 2-isomorphism in B gives a 2-isomorphism in Bᵒᵖ

          Equations
          Instances For
            @[simp]
            theorem CategoryTheory.Iso.op2_inv_unop2 {B : Type u} [Bicategory B] {a b : B} {f g : a b} (η : f g) :
            η.op2.inv.unop2 = η.inv
            @[simp]
            theorem CategoryTheory.Iso.op2_hom_unop2 {B : Type u} [Bicategory B] {a b : B} {f g : a b} (η : f g) :
            η.op2.hom.unop2 = η.hom
            @[reducible, inline]
            abbrev CategoryTheory.Iso.op2_unop {B : Type u} [Bicategory B] {a b : Bᵒᵖ} {f g : a b} (η : f.unop g.unop) :
            f g

            A 2-isomorphism in B gives a 2-isomorphism in Bᵒᵖ

            Equations
            Instances For
              @[simp]
              theorem CategoryTheory.Iso.op2_unop_hom_unop2 {B : Type u} [Bicategory B] {a b : Bᵒᵖ} {f g : a b} (η : f.unop g.unop) :
              @[simp]
              theorem CategoryTheory.Iso.op2_unop_inv_unop2 {B : Type u} [Bicategory B] {a b : Bᵒᵖ} {f g : a b} (η : f.unop g.unop) :
              @[reducible, inline]
              abbrev CategoryTheory.Iso.unop2 {B : Type u} [Bicategory B] {a b : Bᵒᵖ} {f g : a b} (η : f g) :

              A 2-isomorphism in Bᵒᵖ gives a 2-isomorphism in B

              Equations
              Instances For
                @[simp]
                theorem CategoryTheory.Iso.unop2_hom {B : Type u} [Bicategory B] {a b : Bᵒᵖ} {f g : a b} (η : f g) :
                @[simp]
                theorem CategoryTheory.Iso.unop2_inv {B : Type u} [Bicategory B] {a b : Bᵒᵖ} {f g : a b} (η : f g) :
                @[reducible, inline]
                abbrev CategoryTheory.Iso.unop2_op {B : Type u} [Bicategory B] {a b : B} {f g : a b} (η : f.op g.op) :
                f g

                A 2-isomorphism in Bᵒᵖ gives a 2-isomorphism in B

                Equations
                Instances For
                  @[simp]
                  theorem CategoryTheory.Iso.unop2_op_inv {B : Type u} [Bicategory B] {a b : B} {f g : a b} (η : f.op g.op) :
                  @[simp]
                  theorem CategoryTheory.Iso.unop2_op_hom {B : Type u} [Bicategory B] {a b : B} {f g : a b} (η : f.op g.op) :
                  @[simp]
                  theorem CategoryTheory.Iso.unop2_op2 {B : Type u} [Bicategory B] {a b : Bᵒᵖ} {f g : a b} (η : f g) :
                  η.unop2.op2 = η

                  The 1-cell dual bicategory Bᵒᵖ.

                  It is defined as follows.

                  • The objects of Bᵒᵖ correspond to objects of B.
                  • The morphisms X ⟶ Y in Bᵒᵖ are the morphisms Y ⟶ X in B.
                  • The 2-morphisms f ⟶ g in Bᵒᵖ are the 2-morphisms f ⟶ g in B. In other words, the directions of the 2-morphisms are preserved.
                  Equations
                  • One or more equations did not get rendered due to their size.
                  @[simp]
                  @[simp]
                  @[simp]
                  theorem Bicategory.Opposite.bicategory_whiskerRight_unop2 {B : Type u} [CategoryTheory.Bicategory B] {a✝ b✝ c✝ : Bᵒᵖ} {f✝ g✝ : a✝ b✝} (η : f✝ g✝) (h : b✝ c✝) :