Skip to content

Commit db4dfba

Browse files
Amelia Livingstoneric-wieser
authored andcommitted
feat: category of coalgebras (#11972)
The category of coalgebras over a commutative ring. Mimics `Mathlib.LinearAlgebra.QuadraticForm.QuadraticModuleCat`. Co-authored-by: Eric Wieser <[email protected]>
1 parent f193746 commit db4dfba

File tree

2 files changed

+180
-0
lines changed

2 files changed

+180
-0
lines changed

Mathlib.lean

+1
Original file line numberDiff line numberDiff line change
@@ -50,6 +50,7 @@ import Mathlib.Algebra.Category.AlgebraCat.Limits
5050
import Mathlib.Algebra.Category.AlgebraCat.Monoidal
5151
import Mathlib.Algebra.Category.AlgebraCat.Symmetric
5252
import Mathlib.Algebra.Category.BoolRing
53+
import Mathlib.Algebra.Category.CoalgebraCat.Basic
5354
import Mathlib.Algebra.Category.FGModuleCat.Basic
5455
import Mathlib.Algebra.Category.FGModuleCat.Limits
5556
import Mathlib.Algebra.Category.GroupCat.Abelian
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,179 @@
1+
/-
2+
Copyright (c) 2024 Amelia Livingston. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Amelia Livingston
5+
-/
6+
import Mathlib.Algebra.Category.ModuleCat.Basic
7+
import Mathlib.RingTheory.Coalgebra.Equiv
8+
9+
/-!
10+
# The category of coalgebras
11+
12+
This file mimics `Mathlib.LinearAlgebra.QuadraticForm.QuadraticModuleCat`.
13+
14+
-/
15+
16+
open CategoryTheory
17+
18+
universe v u
19+
20+
variable (R : Type u) [CommRing R]
21+
22+
/-- The category of `R`-coalgebras. -/
23+
structure CoalgebraCat extends ModuleCat.{v} R where
24+
isCoalgebra : Coalgebra R carrier
25+
26+
attribute [instance] CoalgebraCat.isCoalgebra
27+
28+
variable {R}
29+
30+
namespace CoalgebraCat
31+
32+
open Coalgebra
33+
34+
instance : CoeSort (CoalgebraCat.{v} R) (Type v) :=
35+
⟨(·.carrier)⟩
36+
37+
@[simp] theorem moduleCat_of_toModuleCat (X : CoalgebraCat.{v} R) :
38+
ModuleCat.of R X.toModuleCat = X.toModuleCat :=
39+
rfl
40+
41+
variable (R)
42+
43+
/-- The object in the category of `R`-coalgebras associated to an `R`-coalgebra. -/
44+
@[simps]
45+
def of (X : Type v) [AddCommGroup X] [Module R X] [Coalgebra R X] :
46+
CoalgebraCat R where
47+
isCoalgebra := (inferInstance : Coalgebra R X)
48+
49+
variable {R}
50+
51+
@[simp]
52+
lemma of_comul {X : Type v} [AddCommGroup X] [Module R X] [Coalgebra R X] :
53+
Coalgebra.comul (A := of R X) = Coalgebra.comul (R := R) (A := X) := rfl
54+
55+
@[simp]
56+
lemma of_counit {X : Type v} [AddCommGroup X] [Module R X] [Coalgebra R X] :
57+
Coalgebra.counit (A := of R X) = Coalgebra.counit (R := R) (A := X) := rfl
58+
59+
/-- A type alias for `CoalgHom` to avoid confusion between the categorical and
60+
algebraic spellings of composition. -/
61+
@[ext]
62+
structure Hom (V W : CoalgebraCat.{v} R) :=
63+
/-- The underlying `CoalgHom` -/
64+
toCoalgHom : V →ₗc[R] W
65+
66+
lemma Hom.toCoalgHom_injective (V W : CoalgebraCat.{v} R) :
67+
Function.Injective (Hom.toCoalgHom : Hom V W → _) :=
68+
fun ⟨f⟩ ⟨g⟩ _ => by congr
69+
70+
instance category : Category (CoalgebraCat.{v} R) where
71+
Hom M N := Hom M N
72+
id M := ⟨CoalgHom.id R M⟩
73+
comp f g := ⟨CoalgHom.comp g.toCoalgHom f.toCoalgHom⟩
74+
id_comp g := Hom.ext _ _ <| CoalgHom.id_comp g.toCoalgHom
75+
comp_id f := Hom.ext _ _ <| CoalgHom.comp_id f.toCoalgHom
76+
assoc f g h := Hom.ext _ _ <| CoalgHom.comp_assoc h.toCoalgHom g.toCoalgHom f.toCoalgHom
77+
78+
-- TODO: if `Quiver.Hom` and the instance above were `reducible`, this wouldn't be needed.
79+
@[ext]
80+
lemma hom_ext {M N : CoalgebraCat.{v} R} (f g : M ⟶ N) (h : f.toCoalgHom = g.toCoalgHom) :
81+
f = g :=
82+
Hom.ext _ _ h
83+
84+
/-- Typecheck a `CoalgHom` as a morphism in `CoalgebraCat R`. -/
85+
abbrev ofHom {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y]
86+
[Coalgebra R X] [Coalgebra R Y] (f : X →ₗc[R] Y) :
87+
of R X ⟶ of R Y :=
88+
⟨f⟩
89+
90+
@[simp] theorem toCoalgHom_comp {M N U : CoalgebraCat.{v} R} (f : M ⟶ N) (g : N ⟶ U) :
91+
(f ≫ g).toCoalgHom = g.toCoalgHom.comp f.toCoalgHom :=
92+
rfl
93+
94+
@[simp] theorem toCoalgHom_id {M : CoalgebraCat.{v} R} :
95+
Hom.toCoalgHom (𝟙 M) = CoalgHom.id _ _ :=
96+
rfl
97+
98+
instance concreteCategory : ConcreteCategory.{v} (CoalgebraCat.{v} R) where
99+
forget :=
100+
{ obj := fun M => M
101+
map := fun f => f.toCoalgHom }
102+
forget_faithful :=
103+
{ map_injective := fun {M N} => DFunLike.coe_injective.comp <| Hom.toCoalgHom_injective _ _ }
104+
105+
instance hasForgetToModule : HasForget₂ (CoalgebraCat R) (ModuleCat R) where
106+
forget₂ :=
107+
{ obj := fun M => ModuleCat.of R M
108+
map := fun f => f.toCoalgHom.toLinearMap }
109+
110+
@[simp]
111+
theorem forget₂_obj (X : CoalgebraCat R) :
112+
(forget₂ (CoalgebraCat R) (ModuleCat R)).obj X = ModuleCat.of R X :=
113+
rfl
114+
115+
@[simp]
116+
theorem forget₂_map (X Y : CoalgebraCat R) (f : X ⟶ Y) :
117+
(forget₂ (CoalgebraCat R) (ModuleCat R)).map f = (f.toCoalgHom : X →ₗ[R] Y) :=
118+
rfl
119+
120+
end CoalgebraCat
121+
122+
namespace CoalgEquiv
123+
124+
open CoalgebraCat
125+
126+
variable {X Y Z : Type v}
127+
variable [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] [AddCommGroup Z] [Module R Z]
128+
variable [Coalgebra R X] [Coalgebra R Y] [Coalgebra R Z]
129+
130+
/-- Build an isomorphism in the category `CoalgebraCat R` from a
131+
`CoalgEquiv`. -/
132+
@[simps]
133+
def toIso (e : X ≃ₗc[R] Y) : CoalgebraCat.of R X ≅ CoalgebraCat.of R Y where
134+
hom := CoalgebraCat.ofHom e
135+
inv := CoalgebraCat.ofHom e.symm
136+
hom_inv_id := Hom.ext _ _ <| DFunLike.ext _ _ e.left_inv
137+
inv_hom_id := Hom.ext _ _ <| DFunLike.ext _ _ e.right_inv
138+
139+
@[simp] theorem toIso_refl : toIso (CoalgEquiv.refl R X) = .refl _ :=
140+
rfl
141+
142+
@[simp] theorem toIso_symm (e : X ≃ₗc[R] Y) : toIso e.symm = (toIso e).symm :=
143+
rfl
144+
145+
@[simp] theorem toIso_trans (e : X ≃ₗc[R] Y) (f : Y ≃ₗc[R] Z) :
146+
toIso (e.trans f) = toIso e ≪≫ toIso f :=
147+
rfl
148+
149+
end CoalgEquiv
150+
151+
namespace CategoryTheory.Iso
152+
153+
open Coalgebra
154+
155+
variable {X Y Z : CoalgebraCat.{v} R}
156+
157+
/-- Build a `CoalgEquiv` from an isomorphism in the category
158+
`CoalgebraCat R`. -/
159+
def toCoalgEquiv (i : X ≅ Y) : X ≃ₗc[R] Y :=
160+
{ i.hom.toCoalgHom with
161+
invFun := i.inv.toCoalgHom
162+
left_inv := fun x => CoalgHom.congr_fun (congr_arg CoalgebraCat.Hom.toCoalgHom i.3) x
163+
right_inv := fun x => CoalgHom.congr_fun (congr_arg CoalgebraCat.Hom.toCoalgHom i.4) x }
164+
165+
@[simp] theorem toCoalgEquiv_toCoalgHom (i : X ≅ Y) :
166+
i.toCoalgEquiv = i.hom.toCoalgHom := rfl
167+
168+
@[simp] theorem toCoalgEquiv_refl : toCoalgEquiv (.refl X) = .refl _ _ :=
169+
rfl
170+
171+
@[simp] theorem toCoalgEquiv_symm (e : X ≅ Y) :
172+
toCoalgEquiv e.symm = (toCoalgEquiv e).symm :=
173+
rfl
174+
175+
@[simp] theorem toCoalgEquiv_trans (e : X ≅ Y) (f : Y ≅ Z) :
176+
toCoalgEquiv (e ≪≫ f) = e.toCoalgEquiv.trans f.toCoalgEquiv :=
177+
rfl
178+
179+
end CategoryTheory.Iso

0 commit comments

Comments
 (0)