@@ -68,14 +68,26 @@ lemma coe_of {V : Type u} [add_comm_group V] [module k V] (ρ : G →* (V →ₗ
68
68
@[simp] lemma of_ρ {V : Type u} [add_comm_group V] [module k V] (ρ : G →* (V →ₗ[k] V)) :
69
69
(of ρ).ρ = ρ := rfl
70
70
71
- @[simp] lemma Action_ρ_eq_ρ {A : Rep k G} : Action.ρ A = A.ρ := rfl
71
+ lemma Action_ρ_eq_ρ {A : Rep k G} : Action.ρ A = A.ρ := rfl
72
72
73
73
/-- Allows us to apply lemmas about the underlying `ρ`, which would take an element `g : G` rather
74
74
than `g : Mon.of G` as an argument. -/
75
75
lemma of_ρ_apply {V : Type u} [add_comm_group V] [module k V]
76
76
(ρ : representation k G V) (g : Mon.of G) :
77
77
(Rep.of ρ).ρ g = ρ (g : G) := rfl
78
78
79
+ @[simp] lemma ρ_inv_self_apply {G : Type u} [group G] (A : Rep k G) (g : G) (x : A) :
80
+ A.ρ g⁻¹ (A.ρ g x) = x :=
81
+ show (A.ρ g⁻¹ * A.ρ g) x = x, by rw [←map_mul, inv_mul_self, map_one, linear_map.one_apply]
82
+
83
+ @[simp] lemma ρ_self_inv_apply {G : Type u} [group G] {A : Rep k G} (g : G) (x : A) :
84
+ A.ρ g (A.ρ g⁻¹ x) = x :=
85
+ show (A.ρ g * A.ρ g⁻¹) x = x, by rw [←map_mul, mul_inv_self, map_one, linear_map.one_apply]
86
+
87
+ lemma hom_comm_apply {A B : Rep k G} (f : A ⟶ B) (g : G) (x : A) :
88
+ f.hom (A.ρ g x) = B.ρ g (f.hom x) :=
89
+ linear_map.ext_iff.1 (f.comm g) x
90
+
79
91
variables (k G)
80
92
81
93
/-- The trivial `k`-linear `G`-representation on a `k`-module `V.` -/
@@ -234,151 +246,111 @@ end
234
246
235
247
end linearization
236
248
end
237
- section
249
+ section monoidal_closed
238
250
open Action
239
251
variables [group G] (A B C : Rep k G)
240
252
241
- noncomputable instance : monoidal_closed (Rep k G) :=
242
- monoidal_closed.of_equiv (functor_category_monoidal_equivalence _ _)
243
-
244
- /-- Explicit description of the 'internal Hom' `iHom(A, B)` of two representations `A, B`:
245
- this is `F⁻¹(iHom(F(A), F(B)))`, where `F` is an equivalence
246
- `Rep k G ≌ (single_obj G ⥤ Module k)`. Just used to prove `Rep.ihom_obj_ρ`. -/
247
- lemma ihom_obj_ρ_def :
248
- ((ihom A).obj B).ρ = (functor_category_equivalence.inverse.obj
249
- ((functor_category_equivalence.functor.obj A).closed_ihom.obj
250
- (functor_category_equivalence.functor.obj B))).ρ := rfl
251
-
252
- /-- Given `k`-linear `G`-representations `(A, ρ₁), (B, ρ₂)`, the 'internal Hom' is the
253
- representation on `Homₖ(A, B)` sending `g : G` and `f : A →ₗ[k] B` to `(ρ₂ g) ∘ₗ f ∘ₗ (ρ₁ g⁻¹)`. -/
254
- @[simp] lemma ihom_obj_ρ :
255
- ((ihom A).obj B).ρ = A.ρ.lin_hom B.ρ :=
256
- begin
257
- refine monoid_hom.ext (λ g, _),
258
- simpa only [ihom_obj_ρ_def, functor_category_equivalence.inverse_obj_ρ_apply,
259
- functor.closed_ihom_obj_map, ←functor.map_inv, single_obj.inv_as_inv],
260
- end
261
-
262
- @[simp] lemma ihom_map_hom {B C : Rep k G} (f : B ⟶ C) :
263
- ((ihom A).map f).hom = linear_map.llcomp k A B C f.hom :=
264
- rfl
265
-
266
- /-- Unfolds the unit in the adjunction `A ⊗ - ⊣ iHom(A, -)`; just used to prove
267
- `Rep.ihom_coev_app_hom`. -/
268
- lemma ihom_coev_app_def :
269
- (ihom.coev A).app B = functor_category_equivalence.unit_iso.hom.app B ≫
270
- functor_category_equivalence.inverse.map
271
- ((functor_category_equivalence.functor.obj A).closed_unit.app _ ≫
272
- (functor_category_equivalence.functor.obj A).closed_ihom.map
273
- ((functor_category_monoidal_equivalence (Module.{u} k) (Mon.of G)).μ A B)) :=
274
- rfl
275
-
276
- /-- Describes the unit in the adjunction `A ⊗ - ⊣ iHom(A, -)`; given another `k`-linear
277
- `G`-representation `B,` the `k`-linear map underlying the resulting map `B ⟶ iHom(A, A ⊗ B)` is
278
- given by flipping the arguments in the natural `k`-bilinear map `A →ₗ[k] B →ₗ[k] A ⊗ B`. -/
279
- @[simp] lemma ihom_coev_app_hom :
280
- Action.hom.hom ((ihom.coev A).app B) = (tensor_product.mk _ _ _).flip :=
281
- begin
282
- refine linear_map.ext (λ x, linear_map.ext (λ y, _)),
283
- simpa only [ihom_coev_app_def, functor.map_comp, comp_hom,
284
- functor_category_equivalence.inverse_map_hom, functor.closed_ihom_map_app,
285
- functor_category_monoidal_equivalence.μ_app],
286
- end
287
-
288
- variables {A B C}
289
-
290
- /-- Given a `k`-linear `G`-representation `A`, the adjunction `A ⊗ - ⊣ iHom(A, -)` defines a
291
- bijection `Hom(A ⊗ B, C) ≃ Hom(B, iHom(A, C))` for all `B, C`. Given `f : A ⊗ B ⟶ C`, this lemma
292
- describes the `k`-linear map underlying `f`'s image under the bijection. It is given by currying the
253
+ /-- Given a `k`-linear `G`-representation `(A, ρ₁)`, this is the 'internal Hom' functor sending
254
+ `(B, ρ₂)` to the representation `Homₖ(A, B)` that maps `g : G` and `f : A →ₗ[k] B` to
255
+ `(ρ₂ g) ∘ₗ f ∘ₗ (ρ₁ g⁻¹)`. -/
256
+ @[simps] protected def ihom (A : Rep k G) : Rep k G ⥤ Rep k G :=
257
+ { obj := λ B, Rep.of (representation.lin_hom A.ρ B.ρ),
258
+ map := λ X Y f,
259
+ { hom := Module.of_hom (linear_map.llcomp k _ _ _ f.hom),
260
+ comm' := λ g, linear_map.ext (λ x, linear_map.ext (λ y,
261
+ show f.hom (X.ρ g _) = _, by simpa only [hom_comm_apply])) },
262
+ map_id' := λ B, by ext; refl,
263
+ map_comp' := λ B C D f g, by ext; refl }
264
+
265
+ @[simp] protected lemma ihom_obj_ρ_apply {A B : Rep k G} (g : G) (x : A →ₗ[k] B) :
266
+ ((Rep.ihom A).obj B).ρ g x = (B.ρ g) ∘ₗ x ∘ₗ (A.ρ g⁻¹) := rfl
267
+
268
+ /-- Given a `k`-linear `G`-representation `A`, this is the Hom-set bijection in the adjunction
269
+ `A ⊗ - ⊣ ihom(A, -)`. It sends `f : A ⊗ B ⟶ C` to a `Rep k G` morphism defined by currying the
293
270
`k`-linear map underlying `f`, giving a map `A →ₗ[k] B →ₗ[k] C`, then flipping the arguments. -/
294
- @[simp] lemma monoidal_closed_curry_hom (f : A ⊗ B ⟶ C) :
295
- (monoidal_closed.curry f).hom = (tensor_product.curry f.hom).flip :=
296
- begin
297
- rw [monoidal_closed.curry_eq, comp_hom, ihom_coev_app_hom],
298
- refl,
299
- end
300
-
301
- /-- Given a `k`-linear `G`-representation `A`, the adjunction `A ⊗ - ⊣ iHom(A, -)` defines a
302
- bijection `Hom(A ⊗ B, C) ≃ Hom(B, iHom(A, C))` for all `B, C`. Given `f : B ⟶ iHom(A, C)`, this
303
- lemma describes the `k`-linear map underlying `f`'s image under the bijection. It is given by
304
- flipping the arguments of the `k`-linear map underlying `f`, giving a map `A →ₗ[k] B →ₗ[k] C`, then
305
- uncurrying. -/
306
- @[simp] lemma monoidal_closed_uncurry_hom (f : B ⟶ (ihom A).obj C) :
307
- (monoidal_closed.uncurry f).hom = tensor_product.uncurry _ _ _ _ f.hom.flip :=
308
- begin
309
- simp only [monoidal_closed.of_equiv_uncurry_def, comp_inv_iso_inv_app,
310
- monoidal_functor.comm_tensor_left_inv_app, comp_hom,
311
- functor_category_monoidal_equivalence.inverse_map, functor_category_equivalence.inverse_map_hom,
312
- functor_category_monoidal_equivalence.μ_iso_inv_app],
313
- ext,
314
- refl,
315
- end
316
-
317
- /-- Describes the counit in the adjunction `A ⊗ - ⊣ iHom(A, -)`; given another `k`-linear
318
- `G`-representation `B,` the `k`-linear map underlying the resulting morphism `A ⊗ iHom(A, B) ⟶ B`
319
- is given by uncurrying the map `A →ₗ[k] (A →ₗ[k] B) →ₗ[k] B` defined by flipping the arguments in
320
- the identity map on `Homₖ(A, B).` -/
321
- @[simp] lemma ihom_ev_app_hom : Action.hom.hom ((ihom.ev A).app B) =
322
- tensor_product.uncurry _ _ _ _ linear_map.id.flip :=
323
- monoidal_closed_uncurry_hom _
271
+ @[simps] protected def hom_equiv (A B C : Rep k G) : (A ⊗ B ⟶ C) ≃ (B ⟶ (Rep.ihom A).obj C) :=
272
+ { to_fun := λ f,
273
+ { hom := (tensor_product.curry f.hom).flip,
274
+ comm' := λ g,
275
+ begin
276
+ refine linear_map.ext (λ x, linear_map.ext (λ y, _)),
277
+ change f.hom (_ ⊗ₜ[k] _) = C.ρ g (f.hom (_ ⊗ₜ[k] _)),
278
+ rw [←hom_comm_apply],
279
+ change _ = f.hom ((A.ρ g * A.ρ g⁻¹) y ⊗ₜ[k] _),
280
+ simpa only [←map_mul, mul_inv_self, map_one],
281
+ end },
282
+ inv_fun := λ f,
283
+ { hom := tensor_product.uncurry k _ _ _ f.hom.flip,
284
+ comm' := λ g, tensor_product.ext' (λ x y,
285
+ begin
286
+ dsimp only [monoidal_category.tensor_left_obj, Module.comp_def, linear_map.comp_apply,
287
+ tensor_rho, Module.monoidal_category.hom_apply, tensor_product.map_tmul],
288
+ simp only [tensor_product.uncurry_apply f.hom.flip, linear_map.flip_apply,
289
+ Action_ρ_eq_ρ, hom_comm_apply f g y, Rep.ihom_obj_ρ_apply, linear_map.comp_apply,
290
+ ρ_inv_self_apply],
291
+ end ) },
292
+ left_inv := λ f, Action.hom.ext _ _ (tensor_product.ext' (λ x y, rfl)),
293
+ right_inv := λ f, by ext; refl }
294
+
295
+ instance : monoidal_closed (Rep k G) :=
296
+ { closed' := λ A,
297
+ { is_adj :=
298
+ { right := Rep.ihom A,
299
+ adj := adjunction.mk_of_hom_equiv
300
+ { hom_equiv := Rep.hom_equiv A,
301
+ hom_equiv_naturality_left_symm' := λ X Y Z f g, by ext; refl,
302
+ hom_equiv_naturality_right' := λ X Y Z f g, by ext; refl } } } }
303
+
304
+ @[simp] lemma ihom_obj_ρ_def (A B : Rep k G) : ((ihom A).obj B).ρ = ((Rep.ihom A).obj B).ρ := rfl
305
+
306
+ @[simp] lemma hom_equiv_def (A B C : Rep k G) :
307
+ (ihom.adjunction A).hom_equiv B C = Rep.hom_equiv A B C := rfl
308
+
309
+ @[simp] lemma ihom_ev_app_hom (A B : Rep k G) :
310
+ Action.hom.hom ((ihom.ev A).app B)
311
+ = tensor_product.uncurry _ _ _ _ linear_map.id.flip :=
312
+ by ext; refl
313
+
314
+ @[simp] lemma ihom_coev_app_hom (A B : Rep k G) :
315
+ Action.hom.hom ((ihom.coev A).app B) = (tensor_product.mk _ _ _).flip :=
316
+ by ext; refl
324
317
325
318
variables (A B C)
326
319
327
320
/-- There is a `k`-linear isomorphism between the sets of representation morphisms`Hom(A ⊗ B, C)`
328
321
and `Hom(B, Homₖ(A, C))`. -/
329
- noncomputable def monoidal_closed.linear_hom_equiv :
322
+ def monoidal_closed.linear_hom_equiv :
330
323
(A ⊗ B ⟶ C) ≃ₗ[k] (B ⟶ (A ⟶[Rep k G] C)) :=
331
324
{ map_add' := λ f g, rfl,
332
325
map_smul' := λ r f, rfl, ..(ihom.adjunction A).hom_equiv _ _ }
333
326
334
327
/-- There is a `k`-linear isomorphism between the sets of representation morphisms`Hom(A ⊗ B, C)`
335
328
and `Hom(A, Homₖ(B, C))`. -/
336
- noncomputable def monoidal_closed.linear_hom_equiv_comm :
329
+ def monoidal_closed.linear_hom_equiv_comm :
337
330
(A ⊗ B ⟶ C) ≃ₗ[k] (A ⟶ (B ⟶[Rep k G] C)) :=
338
331
(linear.hom_congr k (β_ A B) (iso.refl _)) ≪≫ₗ
339
332
monoidal_closed.linear_hom_equiv _ _ _
340
333
341
334
variables {A B C}
342
335
343
- lemma monoidal_closed.linear_hom_equiv_hom (f : A ⊗ B ⟶ C) :
336
+ @[simp] lemma monoidal_closed.linear_hom_equiv_hom (f : A ⊗ B ⟶ C) :
344
337
(monoidal_closed.linear_hom_equiv A B C f).hom =
345
- (tensor_product.curry f.hom).flip :=
346
- monoidal_closed_curry_hom _
338
+ (tensor_product.curry f.hom).flip := rfl
347
339
348
- lemma monoidal_closed.linear_hom_equiv_comm_hom (f : A ⊗ B ⟶ C) :
340
+ @[simp] lemma monoidal_closed.linear_hom_equiv_comm_hom (f : A ⊗ B ⟶ C) :
349
341
(monoidal_closed.linear_hom_equiv_comm A B C f).hom =
350
- tensor_product.curry f.hom :=
351
- begin
352
- dunfold monoidal_closed.linear_hom_equiv_comm,
353
- refine linear_map.ext (λ x, linear_map.ext (λ y, _)),
354
- simp only [linear_equiv.trans_apply, monoidal_closed.linear_hom_equiv_hom,
355
- linear.hom_congr_apply, iso.refl_hom, iso.symm_hom, linear_map.to_fun_eq_coe,
356
- linear_map.coe_comp, function.comp_app, linear.left_comp_apply, linear.right_comp_apply,
357
- category.comp_id, Action.comp_hom, linear_map.flip_apply, tensor_product.curry_apply,
358
- Module.coe_comp, function.comp_app, monoidal_category.braiding_inv_apply],
359
- end
342
+ tensor_product.curry f.hom := rfl
360
343
361
- lemma monoidal_closed.linear_hom_equiv_symm_hom (f : B ⟶ (A ⟶[Rep k G] C)) :
344
+ @[simp] lemma monoidal_closed.linear_hom_equiv_symm_hom (f : B ⟶ (A ⟶[Rep k G] C)) :
362
345
((monoidal_closed.linear_hom_equiv A B C).symm f).hom =
363
- tensor_product.uncurry k A B C f.hom.flip :=
364
- monoidal_closed_uncurry_hom _
346
+ tensor_product.uncurry k A B C f.hom.flip := rfl
365
347
366
- lemma monoidal_closed.linear_hom_equiv_comm_symm_hom (f : A ⟶ (B ⟶[Rep k G] C)) :
348
+ @[simp] lemma monoidal_closed.linear_hom_equiv_comm_symm_hom (f : A ⟶ (B ⟶[Rep k G] C)) :
367
349
((monoidal_closed.linear_hom_equiv_comm A B C).symm f).hom =
368
350
tensor_product.uncurry k A B C f.hom :=
369
- begin
370
- dunfold monoidal_closed.linear_hom_equiv_comm,
371
- refine tensor_product.algebra_tensor_module.curry_injective
372
- (linear_map.ext (λ x, linear_map.ext (λ y, _))),
373
- simp only [linear_equiv.trans_symm, linear_equiv.trans_apply, linear.hom_congr_symm_apply,
374
- iso.refl_inv, linear_map.coe_comp, function.comp_app, category.comp_id, Action.comp_hom,
375
- monoidal_closed.linear_hom_equiv_symm_hom, tensor_product.algebra_tensor_module.curry_apply,
376
- linear_map.coe_restrict_scalars, linear_map.to_fun_eq_coe, linear_map.flip_apply,
377
- tensor_product.curry_apply, Module.coe_comp, function.comp_app,
378
- monoidal_category.braiding_hom_apply, tensor_product.uncurry_apply],
379
- end
351
+ by ext; refl
380
352
381
- end
353
+ end monoidal_closed
382
354
end Rep
383
355
namespace representation
384
356
variables {k G : Type u} [comm_ring k] [monoid G] {V W : Type u}
0 commit comments