|
| 1 | +/- |
| 2 | +Copyright (c) 2024 Lean FRO LLC. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Kim Morrison |
| 5 | +-/ |
| 6 | +import Mathlib.CategoryTheory.Monoidal.CoherenceLemmas |
| 7 | +import Mathlib.CategoryTheory.Limits.Shapes.Terminal |
| 8 | + |
| 9 | +/-! |
| 10 | +# The category of comonoids in a monoidal category. |
| 11 | +
|
| 12 | +We define comonoids in a monoidal category `C`. |
| 13 | +
|
| 14 | +## TODO |
| 15 | +* `Comon_ C ≌ Mon_ (Cᵒᵖ)` |
| 16 | +* An oplax monoidal functor takes comonoid objects to comonoid objects. |
| 17 | + That is, a oplax monoidal functor `F : C ⥤ D` induces a functor `Comon_ C ⥤ Comon_ D`. |
| 18 | +* Comonoid objects in `C` are "just" |
| 19 | + oplax monoidal functors from the trivial monoidal category to `C`. |
| 20 | +* The category of comonoids in a braided monoidal category is monoidal. |
| 21 | + (It may suffice to transfer this across the equivalent to monoids in the opposite category.) |
| 22 | +
|
| 23 | +-/ |
| 24 | + |
| 25 | +universe v₁ v₂ u₁ u₂ u |
| 26 | + |
| 27 | +open CategoryTheory MonoidalCategory |
| 28 | + |
| 29 | +variable (C : Type u₁) [Category.{v₁} C] [MonoidalCategory.{v₁} C] |
| 30 | + |
| 31 | +/-- A comonoid object internal to a monoidal category. |
| 32 | +
|
| 33 | +When the monoidal category is preadditive, this is also sometimes called a "coalgebra object". |
| 34 | +-/ |
| 35 | +structure Comon_ where |
| 36 | + /-- The underlying object of a comonoid object. -/ |
| 37 | + X : C |
| 38 | + /-- The counit of a comonoid object. -/ |
| 39 | + counit : X ⟶ 𝟙_ C |
| 40 | + /-- The comultiplication morphism of a comonoid object. -/ |
| 41 | + comul : X ⟶ X ⊗ X |
| 42 | + counit_comul : comul ≫ (counit ▷ X) = (λ_ X).inv := by aesop_cat |
| 43 | + comul_counit : comul ≫ (X ◁ counit) = (ρ_ X).inv := by aesop_cat |
| 44 | + comul_assoc : comul ≫ (X ◁ comul) ≫ (α_ X X X).inv = comul ≫ (comul ▷ X) := by aesop_cat |
| 45 | + |
| 46 | +attribute [reassoc] Comon_.counit_comul Comon_.comul_counit |
| 47 | + |
| 48 | +attribute [reassoc (attr := simp)] Comon_.comul_assoc |
| 49 | + |
| 50 | +namespace Comon_ |
| 51 | + |
| 52 | +/-- The trivial comonoid object. We later show this is terminal in `Comon_ C`. |
| 53 | +-/ |
| 54 | +@[simps] |
| 55 | +def trivial : Comon_ C where |
| 56 | + X := 𝟙_ C |
| 57 | + counit := 𝟙 _ |
| 58 | + comul := (λ_ _).inv |
| 59 | + comul_assoc := by coherence |
| 60 | + counit_comul := by coherence |
| 61 | + comul_counit := by coherence |
| 62 | + |
| 63 | +instance : Inhabited (Comon_ C) := |
| 64 | + ⟨trivial C⟩ |
| 65 | + |
| 66 | +variable {C} |
| 67 | +variable {M : Comon_ C} |
| 68 | + |
| 69 | +@[reassoc (attr := simp)] |
| 70 | +theorem counit_comul_hom {Z : C} (f : M.X ⟶ Z) : M.comul ≫ (M.counit ⊗ f) = f ≫ (λ_ Z).inv := by |
| 71 | + rw [leftUnitor_inv_naturality, tensorHom_def, counit_comul_assoc] |
| 72 | + |
| 73 | +@[reassoc (attr := simp)] |
| 74 | +theorem comul_counit_hom {Z : C} (f : M.X ⟶ Z) : M.comul ≫ (f ⊗ M.counit) = f ≫ (ρ_ Z).inv := by |
| 75 | + rw [rightUnitor_inv_naturality, tensorHom_def', comul_counit_assoc] |
| 76 | + |
| 77 | +@[reassoc (attr := simp)] theorem comul_assoc_flip : |
| 78 | + M.comul ≫ (M.comul ▷ M.X) ≫ (α_ M.X M.X M.X).hom = M.comul ≫ (M.X ◁ M.comul) := by |
| 79 | + simp [← comul_assoc_assoc] |
| 80 | + |
| 81 | +/-- A morphism of comonoid objects. -/ |
| 82 | +@[ext] |
| 83 | +structure Hom (M N : Comon_ C) where |
| 84 | + /-- The underlying morphism of a morphism of comonoid objects. -/ |
| 85 | + hom : M.X ⟶ N.X |
| 86 | + hom_counit : hom ≫ N.counit = M.counit := by aesop_cat |
| 87 | + hom_comul : hom ≫ N.comul = M.comul ≫ (hom ⊗ hom) := by aesop_cat |
| 88 | + |
| 89 | +attribute [reassoc (attr := simp)] Hom.hom_counit Hom.hom_comul |
| 90 | + |
| 91 | +/-- The identity morphism on a comonoid object. -/ |
| 92 | +@[simps] |
| 93 | +def id (M : Comon_ C) : Hom M M where |
| 94 | + hom := 𝟙 M.X |
| 95 | + |
| 96 | +instance homInhabited (M : Comon_ C) : Inhabited (Hom M M) := |
| 97 | + ⟨id M⟩ |
| 98 | + |
| 99 | +/-- Composition of morphisms of monoid objects. -/ |
| 100 | +@[simps] |
| 101 | +def comp {M N O : Comon_ C} (f : Hom M N) (g : Hom N O) : Hom M O where |
| 102 | + hom := f.hom ≫ g.hom |
| 103 | + |
| 104 | +instance : Category (Comon_ C) where |
| 105 | + Hom M N := Hom M N |
| 106 | + id := id |
| 107 | + comp f g := comp f g |
| 108 | + |
| 109 | +@[ext] lemma ext {X Y : Comon_ C} {f g : X ⟶ Y} (w : f.hom = g.hom) : f = g := Hom.ext _ _ w |
| 110 | + |
| 111 | +@[simp] theorem id_hom' (M : Comon_ C) : (𝟙 M : Hom M M).hom = 𝟙 M.X := rfl |
| 112 | + |
| 113 | +@[simp] |
| 114 | +theorem comp_hom' {M N K : Comon_ C} (f : M ⟶ N) (g : N ⟶ K) : (f ≫ g).hom = f.hom ≫ g.hom := |
| 115 | + rfl |
| 116 | + |
| 117 | +section |
| 118 | + |
| 119 | +variable (C) |
| 120 | + |
| 121 | +/-- The forgetful functor from comonoid objects to the ambient category. -/ |
| 122 | +@[simps] |
| 123 | +def forget : Comon_ C ⥤ C where |
| 124 | + obj A := A.X |
| 125 | + map f := f.hom |
| 126 | + |
| 127 | +end |
| 128 | + |
| 129 | +instance forget_faithful : (@forget C _ _).Faithful where |
| 130 | + |
| 131 | +instance {A B : Comon_ C} (f : A ⟶ B) [e : IsIso ((forget C).map f)] : IsIso f.hom := e |
| 132 | + |
| 133 | +/-- The forgetful functor from comonoid objects to the ambient category reflects isomorphisms. -/ |
| 134 | +instance : (forget C).ReflectsIsomorphisms where |
| 135 | + reflects f e := |
| 136 | + ⟨⟨{ hom := inv f.hom }, by aesop_cat⟩⟩ |
| 137 | + |
| 138 | +/-- Construct an isomorphism of monoids by giving an isomorphism between the underlying objects |
| 139 | +and checking compatibility with unit and multiplication only in the forward direction. |
| 140 | +-/ |
| 141 | +@[simps] |
| 142 | +def mkIso {M N : Comon_ C} (f : M.X ≅ N.X) (f_counit : f.hom ≫ N.counit = M.counit) |
| 143 | + (f_comul : f.hom ≫ N.comul = M.comul ≫ (f.hom ⊗ f.hom)) : M ≅ N where |
| 144 | + hom := |
| 145 | + { hom := f.hom |
| 146 | + hom_counit := f_counit |
| 147 | + hom_comul := f_comul } |
| 148 | + inv := |
| 149 | + { hom := f.inv |
| 150 | + hom_counit := by rw [← f_counit]; simp |
| 151 | + hom_comul := by |
| 152 | + rw [← cancel_epi f.hom] |
| 153 | + slice_rhs 1 2 => rw [f_comul] |
| 154 | + simp } |
| 155 | + |
| 156 | +instance uniqueHomToTrivial (A : Comon_ C) : Unique (A ⟶ trivial C) where |
| 157 | + default := |
| 158 | + { hom := A.counit |
| 159 | + hom_counit := by dsimp; simp |
| 160 | + hom_comul := by dsimp; simp [A.comul_counit, unitors_inv_equal] } |
| 161 | + uniq f := by |
| 162 | + ext; simp |
| 163 | + rw [← Category.comp_id f.hom] |
| 164 | + erw [f.hom_counit] |
| 165 | + |
| 166 | +open CategoryTheory.Limits |
| 167 | + |
| 168 | +instance : HasTerminal (Comon_ C) := |
| 169 | + hasTerminal_of_unique (trivial C) |
| 170 | + |
| 171 | +end Comon_ |
0 commit comments