Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit ee89acd

Browse files
committed
feat(representation_theory): fdRep k G is k-linear with finite dimensional hom spaces (#13789)
Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
1 parent f020f41 commit ee89acd

File tree

10 files changed

+156
-24
lines changed

10 files changed

+156
-24
lines changed

src/algebra/category/FinVect.lean

Lines changed: 30 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -13,9 +13,17 @@ import algebra.category.Module.monoidal
1313
# The category of finite dimensional vector spaces
1414
1515
This introduces `FinVect K`, the category of finite dimensional vector spaces over a field `K`.
16-
It is implemented as a full subcategory on a subtype of `Module K`, which inherits monoidal and
17-
symmetric structure as `finite_dimensional K` is a monoidal predicate.
18-
We also provide a right rigid monoidal category instance.
16+
It is implemented as a full subcategory on a subtype of `Module K`.
17+
18+
We first create the instance as a `K`-linear category,
19+
then as a `K`-linear monoidal category and then as a right-rigid monoidal category.
20+
21+
## Future work
22+
23+
* Show that `FinVect K` is a symmetric monoidal category (it is already monoidal).
24+
* Show that `FinVect K` is abelian.
25+
* Show that `FinVect K` is rigid (it is already right rigid).
26+
1927
-/
2028
noncomputable theory
2129

@@ -36,8 +44,9 @@ instance closed_predicate_finite_dimensional :
3644
{ prop_ihom' := λ X Y hX hY, by exactI @linear_map.finite_dimensional K _ X _ _ hX Y _ _ hY }
3745

3846
/-- Define `FinVect` as the subtype of `Module.{u} K` of finite dimensional vector spaces. -/
39-
@[derive [large_category, concrete_category, monoidal_category, symmetric_category,
40-
monoidal_closed]]
47+
@[derive [large_category, concrete_category, preadditive, linear K,
48+
monoidal_category, symmetric_category, monoidal_preadditive, monoidal_linear K,
49+
monoidal_closed]]
4150
def FinVect := full_subcategory (λ (V : Module.{u} K), finite_dimensional K V)
4251

4352
namespace FinVect
@@ -50,12 +59,28 @@ instance : inhabited (FinVect K) := ⟨⟨Module.of K K, finite_dimensional.fini
5059
def of (V : Type u) [add_comm_group V] [module K V] [finite_dimensional K V] : FinVect K :=
5160
⟨Module.of K V, by { change finite_dimensional K V, apply_instance }⟩
5261

62+
instance (V W : FinVect K) : finite_dimensional K (V ⟶ W) :=
63+
(by apply_instance : finite_dimensional K (V.obj →ₗ[K] W.obj))
64+
5365
instance : has_forget₂ (FinVect.{u} K) (Module.{u} K) :=
5466
by { dsimp [FinVect], apply_instance, }
5567

5668
instance : full (forget₂ (FinVect K) (Module.{u} K)) :=
5769
{ preimage := λ X Y f, f, }
5870

71+
/-- The forgetful functor `FinVect K ⥤ Module K` as a monoidal functor. -/
72+
def forget₂_monoidal : monoidal_functor (FinVect K) (Module.{u} K) :=
73+
monoidal_category.full_monoidal_subcategory_inclusion _
74+
75+
instance forget₂_monoidal_faithful : faithful (forget₂_monoidal K).to_functor :=
76+
by { dsimp [forget₂_monoidal], apply_instance, }
77+
78+
instance forget₂_monoidal_additive : (forget₂_monoidal K).to_functor.additive :=
79+
by { dsimp [forget₂_monoidal], apply_instance, }
80+
81+
instance forget₂_monoidal_linear : (forget₂_monoidal K).to_functor.linear K :=
82+
by { dsimp [forget₂_monoidal], apply_instance, }
83+
5984
variables (V W : FinVect K)
6085

6186
@[simp] lemma ihom_obj : (ihom V).obj W = FinVect.of K (V.obj →ₗ[K] W.obj) := rfl

src/category_theory/linear/default.lean

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -88,13 +88,18 @@ section induced_category
8888
universes u'
8989
variables {C} {D : Type u'} (F : D → C)
9090

91-
instance induced_category.category : linear.{w v} R (induced_category C F) :=
91+
instance induced_category : linear.{w v} R (induced_category C F) :=
9292
{ hom_module := λ X Y, @linear.hom_module R _ C _ _ _ (F X) (F Y),
9393
smul_comp' := λ P Q R f f' g, smul_comp' _ _ _ _ _ _,
9494
comp_smul' := λ P Q R f g g', comp_smul' _ _ _ _ _ _, }
9595

9696
end induced_category
9797

98+
instance full_subcategory (Z : C → Prop) : linear.{w v} R (full_subcategory Z) :=
99+
{ hom_module := λ X Y, @linear.hom_module R _ C _ _ _ X.obj Y.obj,
100+
smul_comp' := λ P Q R f f' g, smul_comp' _ _ _ _ _ _,
101+
comp_smul' := λ P Q R f g g', comp_smul' _ _ _ _ _ _, }
102+
98103
variables (R)
99104

100105
/-- Composition by a fixed left argument as an `R`-linear map. -/

src/category_theory/linear/linear_functor.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -71,6 +71,10 @@ instance induced_functor_linear : functor.linear R (induced_functor F) := {}
7171

7272
end induced_category
7373

74+
instance full_subcategory_inclusion_linear
75+
{C : Type*} [category C] [preadditive C] [category_theory.linear R C] (Z : C → Prop) :
76+
(full_subcategory_inclusion Z).linear R := {}
77+
7478
section
7579

7680
variables {R} {C D : Type*} [category C] [category D]

src/category_theory/monoidal/linear.lean

Lines changed: 22 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -35,11 +35,32 @@ restate_axiom monoidal_linear.tensor_smul'
3535
restate_axiom monoidal_linear.smul_tensor'
3636
attribute [simp] monoidal_linear.tensor_smul monoidal_linear.smul_tensor
3737

38-
variables [monoidal_linear R C]
38+
variables {C} [monoidal_linear R C]
3939

4040
instance tensor_left_linear (X : C) : (tensor_left X).linear R := {}
4141
instance tensor_right_linear (X : C) : (tensor_right X).linear R := {}
4242
instance tensoring_left_linear (X : C) : ((tensoring_left C).obj X).linear R := {}
4343
instance tensoring_right_linear (X : C) : ((tensoring_right C).obj X).linear R := {}
4444

45+
/-- A faithful linear monoidal functor to a linear monoidal category
46+
ensures that the domain is linear monoidal. -/
47+
def monoidal_linear_of_faithful
48+
{D : Type*} [category D] [preadditive D] [linear R D]
49+
[monoidal_category D] [monoidal_preadditive D]
50+
(F : monoidal_functor D C) [faithful F.to_functor]
51+
[F.to_functor.additive] [F.to_functor.linear R] :
52+
monoidal_linear R D :=
53+
{ tensor_smul' := begin
54+
intros,
55+
apply F.to_functor.map_injective,
56+
simp only [F.to_functor.map_smul r (f ⊗ g), F.to_functor.map_smul r g, F.map_tensor,
57+
monoidal_linear.tensor_smul, linear.smul_comp, linear.comp_smul],
58+
end,
59+
smul_tensor' := begin
60+
intros,
61+
apply F.to_functor.map_injective,
62+
simp only [F.to_functor.map_smul r (f ⊗ g), F.to_functor.map_smul r f, F.map_tensor,
63+
monoidal_linear.smul_tensor, linear.smul_comp, linear.comp_smul],
64+
end, }
65+
4566
end category_theory

src/category_theory/monoidal/preadditive.lean

Lines changed: 22 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Scott Morrison
55
-/
66
import category_theory.preadditive.additive_functor
7-
import category_theory.monoidal.category
7+
import category_theory.monoidal.functor
88

99
/-!
1010
# Preadditive monoidal categories
@@ -41,7 +41,7 @@ restate_axiom monoidal_preadditive.tensor_add'
4141
restate_axiom monoidal_preadditive.add_tensor'
4242
attribute [simp] monoidal_preadditive.tensor_zero monoidal_preadditive.zero_tensor
4343

44-
variables [monoidal_preadditive C]
44+
variables {C} [monoidal_preadditive C]
4545

4646
local attribute [simp] monoidal_preadditive.tensor_add monoidal_preadditive.add_tensor
4747

@@ -50,6 +50,26 @@ instance tensor_right_additive (X : C) : (tensor_right X).additive := {}
5050
instance tensoring_left_additive (X : C) : ((tensoring_left C).obj X).additive := {}
5151
instance tensoring_right_additive (X : C) : ((tensoring_right C).obj X).additive := {}
5252

53+
/-- A faithful additive monoidal functor to a monoidal preadditive category
54+
ensures that the domain is monoidal preadditive. -/
55+
def monoidal_preadditive_of_faithful {D : Type*} [category D] [preadditive D] [monoidal_category D]
56+
(F : monoidal_functor D C) [faithful F.to_functor] [F.to_functor.additive] :
57+
monoidal_preadditive D :=
58+
{ tensor_zero' := by { intros, apply F.to_functor.map_injective, simp [F.map_tensor], },
59+
zero_tensor' := by { intros, apply F.to_functor.map_injective, simp [F.map_tensor], },
60+
tensor_add' := begin
61+
intros,
62+
apply F.to_functor.map_injective,
63+
simp only [F.map_tensor, F.to_functor.map_add, preadditive.comp_add, preadditive.add_comp,
64+
monoidal_preadditive.tensor_add],
65+
end,
66+
add_tensor' := begin
67+
intros,
68+
apply F.to_functor.map_injective,
69+
simp only [F.map_tensor, F.to_functor.map_add, preadditive.comp_add, preadditive.add_comp,
70+
monoidal_preadditive.add_tensor],
71+
end, }
72+
5373
open_locale big_operators
5474

5575
lemma tensor_sum {P Q R S : C} {J : Type*} (s : finset J) (f : P ⟶ Q) (g : J → (R ⟶ S)) :

src/category_theory/monoidal/subcategory.lean

Lines changed: 25 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Antoine Labelle
55
-/
66
import category_theory.monoidal.braided
7+
import category_theory.monoidal.linear
78
import category_theory.concrete_category.basic
9+
import category_theory.preadditive.additive_functor
10+
import category_theory.linear.linear_functor
811
import category_theory.closed.monoidal
912

1013
/-!
@@ -82,6 +85,28 @@ instance full_monoidal_subcategory.full :
8285
instance full_monoidal_subcategory.faithful :
8386
faithful (full_monoidal_subcategory_inclusion P).to_functor := full_subcategory.faithful P
8487

88+
section
89+
90+
variables [preadditive C]
91+
92+
instance full_monoidal_subcategory_inclusion_additive :
93+
(full_monoidal_subcategory_inclusion P).to_functor.additive :=
94+
functor.full_subcategory_inclusion_additive _
95+
96+
instance [monoidal_preadditive C] : monoidal_preadditive (full_subcategory P) :=
97+
monoidal_preadditive_of_faithful (full_monoidal_subcategory_inclusion P)
98+
99+
variables (R : Type*) [ring R] [linear R C]
100+
101+
instance full_monoidal_subcategory_inclusion_linear :
102+
(full_monoidal_subcategory_inclusion P).to_functor.linear R :=
103+
functor.full_subcategory_inclusion_linear R _
104+
105+
instance [monoidal_preadditive C] [monoidal_linear R C] : monoidal_linear R (full_subcategory P) :=
106+
monoidal_linear_of_faithful R (full_monoidal_subcategory_inclusion P)
107+
108+
end
109+
85110
variables {P} {P' : C → Prop} [monoidal_predicate P']
86111

87112
/-- An implication of predicates `P → P'` induces a monoidal functor between full monoidal

src/category_theory/preadditive/additive_functor.lean

Lines changed: 5 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -98,6 +98,10 @@ instance induced_functor_additive : functor.additive (induced_functor F) := {}
9898

9999
end induced_category
100100

101+
instance full_subcategory_inclusion_additive
102+
{C : Type*} [category C] [preadditive C] (Z : C → Prop) :
103+
(full_subcategory_inclusion Z).additive := {}
104+
101105
section
102106
-- To talk about preservation of biproducts we need to specify universes explicitly.
103107

@@ -154,7 +158,7 @@ full_subcategory (λ (F : C ⥤ D), F.additive)
154158
infixr ` ⥤+ `:26 := AdditiveFunctor
155159

156160
instance : preadditive (C ⥤+ D) :=
157-
preadditive.induced_category.category _
161+
preadditive.induced_category _
158162

159163
/-- An additive functor is in particular a functor. -/
160164
@[derive full, derive faithful]

src/category_theory/preadditive/default.lean

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -83,13 +83,18 @@ section induced_category
8383
universes u'
8484
variables {C} {D : Type u'} (F : D → C)
8585

86-
instance induced_category.category : preadditive.{v} (induced_category C F) :=
86+
instance induced_category : preadditive.{v} (induced_category C F) :=
8787
{ hom_group := λ P Q, @preadditive.hom_group C _ _ (F P) (F Q),
8888
add_comp' := λ P Q R f f' g, add_comp' _ _ _ _ _ _,
8989
comp_add' := λ P Q R f g g', comp_add' _ _ _ _ _ _, }
9090

9191
end induced_category
9292

93+
instance full_subcategory (Z : C → Prop) : preadditive.{v} (full_subcategory Z) :=
94+
{ hom_group := λ P Q, @preadditive.hom_group C _ _ P.obj Q.obj,
95+
add_comp' := λ P Q R f f' g, add_comp' _ _ _ _ _ _,
96+
comp_add' := λ P Q R f g g', comp_add' _ _ _ _ _ _, }
97+
9398
instance (X : C) : add_comm_group (End X) := by { dsimp [End], apply_instance, }
9499

95100
instance (X : C) : ring (End X) :=

src/representation_theory/Action.lean

Lines changed: 22 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -267,9 +267,15 @@ section has_zero_morphisms
267267
variables [has_zero_morphisms V]
268268

269269
instance : has_zero_morphisms (Action V G) :=
270-
{ has_zero := λ X Y, ⟨⟨0, by tidy⟩⟩, }
270+
{ has_zero := λ X Y, ⟨⟨0, by { intro g, simp }⟩⟩,
271+
comp_zero' := λ P Q f R, by { ext1, simp },
272+
zero_comp' := λ P Q R f, by { ext1, simp }, }
271273

272-
instance : functor.preserves_zero_morphisms (functor_category_equivalence V G).functor := {}
274+
instance forget_preserves_zero_morphisms : functor.preserves_zero_morphisms (forget V G) := {}
275+
instance forget₂_preserves_zero_morphisms [concrete_category V] :
276+
functor.preserves_zero_morphisms (forget₂ (Action V G) V) := {}
277+
instance functor_category_equivalence_preserves_zero_morphisms :
278+
functor.preserves_zero_morphisms (functor_category_equivalence V G).functor := {}
273279

274280
end has_zero_morphisms
275281

@@ -289,8 +295,12 @@ instance : preadditive (Action V G) :=
289295
add_comp' := by { intros, ext, exact preadditive.add_comp _ _ _ _ _ _, },
290296
comp_add' := by { intros, ext, exact preadditive.comp_add _ _ _ _ _ _, }, }
291297

292-
instance : functor.additive (functor_category_equivalence V G).functor := {}
293-
instance forget_additive : functor.additive (forget V G) := {}
298+
instance forget_additive :
299+
functor.additive (forget V G) := {}
300+
instance forget₂_additive [concrete_category V] :
301+
functor.additive (forget₂ (Action V G) V) := {}
302+
instance functor_category_equivalence_additive :
303+
functor.additive (functor_category_equivalence V G).functor := {}
294304

295305
@[simp] lemma zero_hom {X Y : Action V G} : (0 : X ⟶ Y).hom = 0 := rfl
296306
@[simp] lemma neg_hom {X Y : Action V G} (f : X ⟶ Y) : (-f).hom = -f.hom := rfl
@@ -315,7 +325,12 @@ instance : linear R (Action V G) :=
315325
smul_comp' := by { intros, ext, exact linear.smul_comp _ _ _ _ _ _, },
316326
comp_smul' := by { intros, ext, exact linear.comp_smul _ _ _ _ _ _, }, }
317327

318-
instance : functor.linear R (functor_category_equivalence V G).functor := {}
328+
instance forget_linear :
329+
functor.linear R (forget V G) := {}
330+
instance forget₂_linear [concrete_category V] :
331+
functor.linear R (forget₂ (Action V G) V) := {}
332+
instance functor_category_equivalence_linear :
333+
functor.linear R (functor_category_equivalence V G).functor := {}
319334

320335
@[simp] lemma smul_hom {X Y : Action V G} (r : R) (f : X ⟶ Y) : (r • f).hom = r • f.hom := rfl
321336

@@ -412,10 +427,10 @@ instance [symmetric_category V] : symmetric_category (Action V G) :=
412427
symmetric_category_of_faithful (forget_braided V G)
413428

414429
section
415-
local attribute [simp] monoidal_preadditive.tensor_add monoidal_preadditive.add_tensor
416-
417430
variables [preadditive V] [monoidal_preadditive V]
418431

432+
local attribute [simp] monoidal_preadditive.tensor_add monoidal_preadditive.add_tensor
433+
419434
instance : monoidal_preadditive (Action V G) := {}
420435

421436
variables {R : Type*} [semiring R] [linear R V] [monoidal_linear R V]

src/representation_theory/fdRep.lean

Lines changed: 14 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -17,13 +17,15 @@ Also `V.ρ` gives the homomorphism `G →* (V →ₗ[k] V)`.
1717
Conversely, given a homomorphism `ρ : G →* (V →ₗ[k] V)`,
1818
you can construct the bundled representation as `Rep.of ρ`.
1919
20-
We verify that `fdRep k G` is a rigid monoidal category.
20+
We verify that `fdRep k G` is a `k`-linear monoidal category, and right rigid when `G` is a group.
2121
2222
## TODO
23+
* `fdRep k G ≌ full_subcategory (finite_dimensional k)`
24+
* Upgrade the right rigid structure to a rigid structure (this just needs to be done for `FinVect`).
2325
* `fdRep k G` has all finite (co)limits.
2426
* `fdRep k G` is abelian.
2527
* `fdRep k G ≌ FinVect (monoid_algebra k G)` (this will require generalising `FinVect` first).
26-
* Upgrade the right rigid structure to a rigid structure.
28+
2729
-/
2830

2931
universes u
@@ -32,14 +34,16 @@ open category_theory
3234
open category_theory.limits
3335

3436
/-- The category of finite dimensional `k`-linear representations of a monoid `G`. -/
35-
@[derive [large_category, concrete_category/-, has_limits, has_colimits-/]]
37+
@[derive [large_category, concrete_category, preadditive]]
3638
abbreviation fdRep (k G : Type u) [field k] [monoid G] :=
3739
Action (FinVect.{u} k) (Mon.of G)
3840

3941
namespace fdRep
4042

4143
variables {k G : Type u} [field k] [monoid G]
4244

45+
instance : linear k (fdRep k G) := by apply_instance
46+
4347
instance : has_coe_to_sort (fdRep k G) (Type u) := concrete_category.has_coe_to_sort _
4448

4549
instance (V : fdRep k G) : add_comm_group V :=
@@ -51,6 +55,11 @@ by { change module k ((forget₂ (fdRep k G) (FinVect k)).obj V).obj, apply_inst
5155
instance (V : fdRep k G) : finite_dimensional k V :=
5256
by { change finite_dimensional k ((forget₂ (fdRep k G) (FinVect k)).obj V).obj, apply_instance, }
5357

58+
/-- All hom spaces are finite dimensional. -/
59+
instance (V W : fdRep k G) : finite_dimensional k (V ⟶ W) :=
60+
finite_dimensional.of_injective
61+
((forget₂ (fdRep k G) (FinVect k)).map_linear_map k) (functor.map_injective _)
62+
5463
/-- The monoid homomorphism corresponding to the action of `G` onto `V : fdRep k G`. -/
5564
def ρ (V : fdRep k G) : G →* (V →ₗ[k] V) := V.ρ
5665

@@ -66,9 +75,6 @@ begin
6675
exact (i.hom.comm g).symm,
6776
end
6877

69-
-- This works well with the new design for representations:
70-
example (V : fdRep k G) : G →* (V →ₗ[k] V) := V.ρ
71-
7278
/-- Lift an unbundled representation to `fdRep`. -/
7379
@[simps ρ]
7480
def of {V : Type u} [add_comm_group V] [module k V] [finite_dimensional k V]
@@ -80,6 +86,8 @@ instance : has_forget₂ (fdRep k G) (Rep k G) :=
8086

8187
-- Verify that the monoidal structure is available.
8288
example : monoidal_category (fdRep k G) := by apply_instance
89+
example : monoidal_preadditive (fdRep k G) := by apply_instance
90+
example : monoidal_linear k (fdRep k G) := by apply_instance
8391

8492
end fdRep
8593

0 commit comments

Comments
 (0)