Bicategories #
In this file we define typeclass for bicategories.
A bicategory B
consists of
- objects
a : B
, - 1-morphisms
f : a ⟶ b
between objectsa b : B
, and - 2-morphisms
η : f ⟶ g
between 1-morphismsf g : a ⟶ b
between objectsa b : B
.
We use u
, v
, and w
as the universe variables for objects, 1-morphisms, and 2-morphisms,
respectively.
A typeclass for bicategories extends CategoryTheory.CategoryStruct
typeclass. This means that
we have
- a composition
f ≫ g : a ⟶ c
for each 1-morphismsf : a ⟶ b
andg : b ⟶ c
, and - an identity
𝟙 a : a ⟶ a
for each objecta : B
.
For each object a b : B
, the collection of 1-morphisms a ⟶ b
has a category structure. The
2-morphisms in the bicategory are implemented as the morphisms in this family of categories.
The composition of 1-morphisms is in fact an object part of a functor
(a ⟶ b) ⥤ (b ⟶ c) ⥤ (a ⟶ c)
. The definition of bicategories in this file does not
require this functor directly. Instead, it requires the whiskering functions. For a 1-morphism
f : a ⟶ b
and a 2-morphism η : g ⟶ h
between 1-morphisms g h : b ⟶ c
, there is a
2-morphism whiskerLeft f η : f ≫ g ⟶ f ≫ h
. Similarly, for a 2-morphism η : f ⟶ g
between 1-morphisms f g : a ⟶ b
and a 1-morphism f : b ⟶ c
, there is a 2-morphism
whiskerRight η h : f ≫ h ⟶ g ≫ h
. These satisfy the exchange law
whiskerLeft f θ ≫ whiskerRight η i = whiskerRight η h ≫ whiskerLeft g θ
,
which is required as an axiom in the definition here.
In a bicategory, we can compose the 1-morphisms f : a ⟶ b
and g : b ⟶ c
to obtain
a 1-morphism f ≫ g : a ⟶ c
. This composition does not need to be strictly associative,
but there is a specified associator, α_ f g h : (f ≫ g) ≫ h ≅ f ≫ (g ≫ h)
.
There is an identity 1-morphism 𝟙 a : a ⟶ a
, with specified left and right unitor
isomorphisms λ_ f : 𝟙 a ≫ f ≅ f
and ρ_ f : f ≫ 𝟙 a ≅ f
.
These associators and unitors satisfy the pentagon and triangle equations.
See https://ncatlab.org/nlab/show/bicategory.
- Hom : B → B → Type v
- id : (X : B) → X ⟶ X
- homCategory : (a b : B) → CategoryTheory.Category.{w, v} (a ⟶ b)
- whiskerLeft : {a b c : B} → (f : a ⟶ b) → {g h : b ⟶ c} → (g ⟶ h) → (CategoryTheory.CategoryStruct.comp f g ⟶ CategoryTheory.CategoryStruct.comp f h)
- whiskerRight : {a b c : B} → {f g : a ⟶ b} → (f ⟶ g) → (h : b ⟶ c) → CategoryTheory.CategoryStruct.comp f h ⟶ CategoryTheory.CategoryStruct.comp g h
- associator : {a b c d : B} → (f : a ⟶ b) → (g : b ⟶ c) → (h : c ⟶ d) → CategoryTheory.CategoryStruct.comp (CategoryTheory.CategoryStruct.comp f g) h ≅ CategoryTheory.CategoryStruct.comp f (CategoryTheory.CategoryStruct.comp g h)
- leftUnitor : {a b : B} → (f : a ⟶ b) → CategoryTheory.CategoryStruct.comp (CategoryTheory.CategoryStruct.id a) f ≅ f
- rightUnitor : {a b : B} → (f : a ⟶ b) → CategoryTheory.CategoryStruct.comp f (CategoryTheory.CategoryStruct.id b) ≅ f
- whiskerLeft_id : ∀ {a b c : B} (f : a ⟶ b) (g : b ⟶ c), CategoryTheory.Bicategory.whiskerLeft f (CategoryTheory.CategoryStruct.id g) = CategoryTheory.CategoryStruct.id (CategoryTheory.CategoryStruct.comp f g)
- whiskerLeft_comp : ∀ {a b c : B} (f : a ⟶ b) {g h i : b ⟶ c} (η : g ⟶ h) (θ : h ⟶ i), CategoryTheory.Bicategory.whiskerLeft f (CategoryTheory.CategoryStruct.comp η θ) = CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.whiskerLeft f η) (CategoryTheory.Bicategory.whiskerLeft f θ)
- id_whiskerLeft : ∀ {a b : B} {f g : a ⟶ b} (η : f ⟶ g), CategoryTheory.Bicategory.whiskerLeft (CategoryTheory.CategoryStruct.id a) η = CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.leftUnitor f).hom (CategoryTheory.CategoryStruct.comp η (CategoryTheory.Bicategory.leftUnitor g).inv)
- comp_whiskerLeft : ∀ {a b c d : B} (f : a ⟶ b) (g : b ⟶ c) {h h' : c ⟶ d} (η : h ⟶ h'), CategoryTheory.Bicategory.whiskerLeft (CategoryTheory.CategoryStruct.comp f g) η = CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.associator f g h).hom (CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.whiskerLeft f (CategoryTheory.Bicategory.whiskerLeft g η)) (CategoryTheory.Bicategory.associator f g h').inv)
- id_whiskerRight : ∀ {a b c : B} (f : a ⟶ b) (g : b ⟶ c), CategoryTheory.Bicategory.whiskerRight (CategoryTheory.CategoryStruct.id f) g = CategoryTheory.CategoryStruct.id (CategoryTheory.CategoryStruct.comp f g)
- comp_whiskerRight : ∀ {a b c : B} {f g h : a ⟶ b} (η : f ⟶ g) (θ : g ⟶ h) (i : b ⟶ c), CategoryTheory.Bicategory.whiskerRight (CategoryTheory.CategoryStruct.comp η θ) i = CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.whiskerRight η i) (CategoryTheory.Bicategory.whiskerRight θ i)
- whiskerRight_id : ∀ {a b : B} {f g : a ⟶ b} (η : f ⟶ g), CategoryTheory.Bicategory.whiskerRight η (CategoryTheory.CategoryStruct.id b) = CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.rightUnitor f).hom (CategoryTheory.CategoryStruct.comp η (CategoryTheory.Bicategory.rightUnitor g).inv)
- whiskerRight_comp : ∀ {a b c d : B} {f f' : a ⟶ b} (η : f ⟶ f') (g : b ⟶ c) (h : c ⟶ d), CategoryTheory.Bicategory.whiskerRight η (CategoryTheory.CategoryStruct.comp g h) = CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.associator f g h).inv (CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.whiskerRight (CategoryTheory.Bicategory.whiskerRight η g) h) (CategoryTheory.Bicategory.associator f' g h).hom)
- whisker_assoc : ∀ {a b c d : B} (f : a ⟶ b) {g g' : b ⟶ c} (η : g ⟶ g') (h : c ⟶ d), CategoryTheory.Bicategory.whiskerRight (CategoryTheory.Bicategory.whiskerLeft f η) h = CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.associator f g h).hom (CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.whiskerLeft f (CategoryTheory.Bicategory.whiskerRight η h)) (CategoryTheory.Bicategory.associator f g' h).inv)
- whisker_exchange : ∀ {a b c : B} {f g : a ⟶ b} {h i : b ⟶ c} (η : f ⟶ g) (θ : h ⟶ i), CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.whiskerLeft f θ) (CategoryTheory.Bicategory.whiskerRight η i) = CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.whiskerRight η h) (CategoryTheory.Bicategory.whiskerLeft g θ)
- pentagon : ∀ {a b c d e : B} (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) (i : d ⟶ e), CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.whiskerRight (CategoryTheory.Bicategory.associator f g h).hom i) (CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.associator f (CategoryTheory.CategoryStruct.comp g h) i).hom (CategoryTheory.Bicategory.whiskerLeft f (CategoryTheory.Bicategory.associator g h i).hom)) = CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.associator (CategoryTheory.CategoryStruct.comp f g) h i).hom (CategoryTheory.Bicategory.associator f g (CategoryTheory.CategoryStruct.comp h i)).hom
- triangle : ∀ {a b c : B} (f : a ⟶ b) (g : b ⟶ c), CategoryTheory.CategoryStruct.comp (CategoryTheory.Bicategory.associator f (CategoryTheory.CategoryStruct.id b) g).hom (CategoryTheory.Bicategory.whiskerLeft f (CategoryTheory.Bicategory.leftUnitor g).hom) = CategoryTheory.Bicategory.whiskerRight (CategoryTheory.Bicategory.rightUnitor f).hom g
Instances
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- CategoryTheory.Bicategory.termα_ = Lean.ParserDescr.node `CategoryTheory.Bicategory.termα_ 1024 (Lean.ParserDescr.symbol "α_")
Instances For
Equations
- CategoryTheory.Bicategory.«termλ_» = Lean.ParserDescr.node `CategoryTheory.Bicategory.termλ_ 1024 (Lean.ParserDescr.symbol "λ_")
Instances For
Equations
- CategoryTheory.Bicategory.termρ_ = Lean.ParserDescr.node `CategoryTheory.Bicategory.termρ_ 1024 (Lean.ParserDescr.symbol "ρ_")
Instances For
Simp-normal form for 2-morphisms #
Rewriting involving associators and unitors could be very complicated. We try to ease this
complexity by putting carefully chosen simp lemmas that rewrite any 2-morphisms into simp-normal
form defined below. Rewriting into simp-normal form is also useful when applying (forthcoming)
coherence
tactic.
The simp-normal form of 2-morphisms is defined to be an expression that has the minimal number of parentheses. More precisely,
- it is a composition of 2-morphisms like
η₁ ≫ η₂ ≫ η₃ ≫ η₄ ≫ η₅
such that eachηᵢ
is either a structural 2-morphisms (2-morphisms made up only of identities, associators, unitors) or non-structural 2-morphisms, and - each non-structural 2-morphism in the composition is of the form
f₁ ◁ f₂ ◁ f₃ ◁ η ▷ f₄ ▷ f₅
, where eachfᵢ
is a 1-morphism that is not the identity or a composite andη
is a non-structural 2-morphisms that is also not the identity or a composite.
Note that f₁ ◁ f₂ ◁ f₃ ◁ η ▷ f₄ ▷ f₅
is actually f₁ ◁ (f₂ ◁ (f₃ ◁ ((η ▷ f₄) ▷ f₅)))
.
The left whiskering of a 2-isomorphism is a 2-isomorphism.
Equations
- CategoryTheory.Bicategory.whiskerLeftIso f η = { hom := CategoryTheory.Bicategory.whiskerLeft f η.hom, inv := CategoryTheory.Bicategory.whiskerLeft f η.inv, hom_inv_id := ⋯, inv_hom_id := ⋯ }
Instances For
Equations
- ⋯ = ⋯
The right whiskering of a 2-isomorphism is a 2-isomorphism.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- ⋯ = ⋯
We state it as a simp lemma, which is regarded as an involved version of
id_whiskerRight f g : 𝟙 f ▷ g = 𝟙 (f ≫ g)
.
Precomposition of a 1-morphism as a functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Precomposition of a 1-morphism as a functor from the category of 1-morphisms a ⟶ b
into the
category of functors (b ⟶ c) ⥤ (a ⟶ c)
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Postcomposition of a 1-morphism as a functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Postcomposition of a 1-morphism as a functor from the category of 1-morphisms b ⟶ c
into the
category of functors (a ⟶ b) ⥤ (a ⟶ c)
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Left component of the associator as a natural isomorphism.
Equations
- CategoryTheory.Bicategory.associatorNatIsoLeft a g h = CategoryTheory.NatIso.ofComponents (fun (x : a ⟶ b) => CategoryTheory.Bicategory.associator x g h) ⋯
Instances For
Middle component of the associator as a natural isomorphism.
Equations
- CategoryTheory.Bicategory.associatorNatIsoMiddle f h = CategoryTheory.NatIso.ofComponents (fun (x : b ⟶ c) => CategoryTheory.Bicategory.associator f x h) ⋯
Instances For
Right component of the associator as a natural isomorphism.
Equations
- CategoryTheory.Bicategory.associatorNatIsoRight f g d = CategoryTheory.NatIso.ofComponents (fun (x : c ⟶ d) => CategoryTheory.Bicategory.associator f g x) ⋯
Instances For
Left unitor as a natural isomorphism.
Equations
- CategoryTheory.Bicategory.leftUnitorNatIso a b = CategoryTheory.NatIso.ofComponents (fun (x : a ⟶ b) => CategoryTheory.Bicategory.leftUnitor x) ⋯
Instances For
Right unitor as a natural isomorphism.
Equations
- CategoryTheory.Bicategory.rightUnitorNatIso a b = CategoryTheory.NatIso.ofComponents (fun (x : a ⟶ b) => CategoryTheory.Bicategory.rightUnitor x) ⋯