Category of categories #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This file contains the definition of the category Cat of all categories.
In this category objects are categories and
morphisms are functors between these categories.
Implementation notes #
Though Cat is not a concrete category, we use bundled to define
its carrier type.
Category of categories.
Equations
- category_theory.Cat.inhabited = {default := {α := Type u, str := category_theory.types}}
Construct a bundled Cat from the underlying type and the typeclass.
Equations
Bicategory structure on Cat
Equations
- category_theory.Cat.bicategory = {to_category_struct := {to_quiver := {hom := λ (C D : category_theory.Cat), ↥C ⥤ ↥D}, id := λ (C : category_theory.Cat), 𝟭 ↥C, comp := λ (C D E : category_theory.Cat) (F : C ⟶ D) (G : D ⟶ E), F ⋙ G}, hom_category := λ (C D : category_theory.Cat), category_theory.functor.category ↥C ↥D, whisker_left := λ (C D E : category_theory.Cat) (F : C ⟶ D) (G H : D ⟶ E) (η : G ⟶ H), category_theory.whisker_left F η, whisker_right := λ (C D E : category_theory.Cat) (F G : C ⟶ D) (η : F ⟶ G) (H : D ⟶ E), category_theory.whisker_right η H, associator := λ (A B C D : category_theory.Cat), category_theory.functor.associator, left_unitor := λ (A B : category_theory.Cat), category_theory.functor.left_unitor, right_unitor := λ (A B : category_theory.Cat), category_theory.functor.right_unitor, whisker_left_id' := category_theory.Cat.bicategory._proof_1, whisker_left_comp' := category_theory.Cat.bicategory._proof_2, id_whisker_left' := category_theory.Cat.bicategory._proof_3, comp_whisker_left' := category_theory.Cat.bicategory._proof_4, id_whisker_right' := category_theory.Cat.bicategory._proof_5, comp_whisker_right' := category_theory.Cat.bicategory._proof_6, whisker_right_id' := category_theory.Cat.bicategory._proof_7, whisker_right_comp' := category_theory.Cat.bicategory._proof_8, whisker_assoc' := category_theory.Cat.bicategory._proof_9, whisker_exchange' := category_theory.Cat.bicategory._proof_10, pentagon' := category_theory.Cat.bicategory._proof_11, triangle' := category_theory.Cat.bicategory._proof_12}
Cat is a strict bicategory.
Category structure on Cat
Functor that gets the set of objects of a category. It is not
called forget, because it is not a faithful functor.
Equations
- category_theory.Cat.objects = {obj := λ (C : category_theory.Cat), ↥C, map := λ (C D : category_theory.Cat) (F : C ⟶ D), F.obj, map_id' := category_theory.Cat.objects._proof_1, map_comp' := category_theory.Cat.objects._proof_2}
Instances for category_theory.Cat.objects
Any isomorphism in Cat induces an equivalence of the underlying categories.
Equations
- category_theory.Cat.equiv_of_iso γ = {functor := γ.hom, inverse := γ.inv, unit_iso := category_theory.eq_to_iso _, counit_iso := category_theory.eq_to_iso _, functor_unit_iso_comp' := _}
Embedding Type into Cat as discrete categories.
This ought to be modelled as a 2-functor!
Equations
- category_theory.Type_to_Cat = {obj := λ (X : Type u), category_theory.Cat.of (category_theory.discrete X), map := λ (X Y : Type u) (f : X ⟶ Y), category_theory.discrete.functor (category_theory.discrete.mk ∘ f), map_id' := category_theory.Type_to_Cat._proof_1, map_comp' := category_theory.Type_to_Cat._proof_2}
Instances for category_theory.Type_to_Cat
Equations
- category_theory.Type_to_Cat.full = {preimage := λ (X Y : Type u) (F : category_theory.Type_to_Cat.obj X ⟶ category_theory.Type_to_Cat.obj Y), category_theory.discrete.as ∘ F.obj ∘ category_theory.discrete.mk, witness' := category_theory.Type_to_Cat.full._proof_1}