refactor: Lint

This commit is contained in:
jstoobysmith 2024-10-09 15:20:23 +00:00
parent a39aeeed8b
commit 4054665c38
3 changed files with 109 additions and 75 deletions

View file

@ -33,8 +33,7 @@ open TensorProduct
## induction principals for pi tensor products
-/
lemma tensorProd_piTensorProd_ext
{f g : ((⨂[R] i : ι1, s1 i) ⊗[R] ⨂[R] i : ι2, s2 i) →ₗ[R] M}
lemma induction_tmul {f g : ((⨂[R] i : ι1, s1 i) ⊗[R] ⨂[R] i : ι2, s2 i) →ₗ[R] M}
(h : ∀ p q, f (PiTensorProduct.tprod R p ⊗ₜ[R] PiTensorProduct.tprod R q)
= g (PiTensorProduct.tprod R p ⊗ₜ[R] PiTensorProduct.tprod R q)) : f = g := by
apply TensorProduct.ext'
@ -56,8 +55,10 @@ lemma tensorProd_piTensorProd_ext
lemma induction_assoc
{f g : ((⨂[R] i : ι1, s1 i) ⊗[R] (⨂[R] i : ι2, s2 i) ⊗[R] ⨂[R] i : ι3, s3 i) →ₗ[R] M}
(h : ∀ p q m, f (PiTensorProduct.tprod R p ⊗ₜ[R] PiTensorProduct.tprod R q ⊗ₜ[R] PiTensorProduct.tprod R m)
= g (PiTensorProduct.tprod R p ⊗ₜ[R] PiTensorProduct.tprod R q ⊗ₜ[R] PiTensorProduct.tprod R m)) : f = g := by
(h : ∀ p q m, f (PiTensorProduct.tprod R p ⊗ₜ[R]
PiTensorProduct.tprod R q ⊗ₜ[R] PiTensorProduct.tprod R m)
= g (PiTensorProduct.tprod R p ⊗ₜ[R] PiTensorProduct.tprod R q
⊗ₜ[R] PiTensorProduct.tprod R m)) : f = g := by
apply TensorProduct.ext'
refine fun x ↦
PiTensorProduct.induction_on' x ?_ (by
@ -65,7 +66,7 @@ lemma induction_assoc
simp [map_add, add_tmul, hx, hy])
intro rx fx
intro y
simp
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod]
simp only [smul_tmul, tmul_smul, LinearMapClass.map_smul]
apply congrArg
let f' : ((⨂[R] i : ι2, s2 i) ⊗[R] ⨂[R] i : ι3, s3 i) →ₗ[R] M := {
@ -83,14 +84,15 @@ lemma induction_assoc
change f' y = g' y
apply congrFun
refine DFunLike.coe_fn_eq.mpr ?H.h.h.a
apply tensorProd_piTensorProd_ext
apply induction_tmul
intro p q
exact h fx p q
lemma induction_assoc'
{f g : (((⨂[R] i : ι1, s1 i) ⊗[R] (⨂[R] i : ι2, s2 i)) ⊗[R] ⨂[R] i : ι3, s3 i) →ₗ[R] M}
(h : ∀ p q m, f ((PiTensorProduct.tprod R p ⊗ₜ[R] PiTensorProduct.tprod R q) ⊗ₜ[R] PiTensorProduct.tprod R m)
= g ((PiTensorProduct.tprod R p ⊗ₜ[R] PiTensorProduct.tprod R q) ⊗ₜ[R] PiTensorProduct.tprod R m)) : f = g := by
(h : ∀ p q m, f ((PiTensorProduct.tprod R p ⊗ₜ[R] PiTensorProduct.tprod R q) ⊗ₜ[R]
PiTensorProduct.tprod R m) = g ((PiTensorProduct.tprod R p ⊗ₜ[R] PiTensorProduct.tprod R q)
⊗ₜ[R] PiTensorProduct.tprod R m)) : f = g := by
apply TensorProduct.ext'
intro x
refine fun y ↦
@ -98,7 +100,7 @@ lemma induction_assoc'
intro a b hy hx
simp [map_add, add_tmul, tmul_add, hy, hx])
intro ry fy
simp
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, tmul_smul, map_smul]
apply congrArg
let f' : ((⨂[R] i : ι1, s1 i) ⊗[R] ⨂[R] i : ι2, s2 i) →ₗ[R] M := {
toFun := fun y => f (y ⊗ₜ[R] PiTensorProduct.tprod R fy),
@ -115,13 +117,14 @@ lemma induction_assoc'
change f' x = g' x
apply congrFun
refine DFunLike.coe_fn_eq.mpr ?H.h.h.a
apply tensorProd_piTensorProd_ext
apply induction_tmul
intro p q
exact h p q fy
lemma induction_tmul_mod
{f g : ((⨂[R] i : ι1, s1 i) ⊗[R] N) →ₗ[R] M}
(h : ∀ p m, f (PiTensorProduct.tprod R p ⊗ₜ[R] m) = g (PiTensorProduct.tprod R p ⊗ₜ[R] m)) : f = g := by
(h : ∀ p m, f (PiTensorProduct.tprod R p ⊗ₜ[R] m) = g (PiTensorProduct.tprod R p ⊗ₜ[R] m)) :
f = g := by
apply TensorProduct.ext'
refine fun y ↦
PiTensorProduct.induction_on' y ?_ (by
@ -135,7 +138,8 @@ lemma induction_tmul_mod
lemma induction_mod_tmul
{f g : (N ⊗[R] ⨂[R] i : ι1, s1 i) →ₗ[R] M}
(h : ∀ m p, f (m ⊗ₜ[R] PiTensorProduct.tprod R p) = g (m ⊗ₜ[R] PiTensorProduct.tprod R p)) : f = g := by
(h : ∀ m p, f (m ⊗ₜ[R] PiTensorProduct.tprod R p) = g (m ⊗ₜ[R] PiTensorProduct.tprod R p)) :
f = g := by
apply TensorProduct.ext'
intro x
refine fun y ↦
@ -162,15 +166,20 @@ instance : (i : ι1 ⊕ ι2) → Module R ((fun i => Sum.elim s1 s2 i) i) := fun
| Sum.inr i => inst2' i
/-- Takes a map `(i : ι1 ⊕ ι2) → Sum.elim s1 s2 i` to the underlying map `(i : ι1) → s1 i `. -/
private def pureInl (f : (i : ι1 ⊕ ι2) → Sum.elim s1 s2 i) : (i : ι1) → s1 i :=
fun i => f (Sum.inl i)
/-- Takes a map `(i : ι1 ⊕ ι2) → Sum.elim s1 s2 i` to the underlying map `(i : ι2) → s2 i `. -/
private def pureInr (f : (i : ι1 ⊕ ι2) → Sum.elim s1 s2 i) : (i : ι2) → s2 i :=
fun i => f (Sum.inr i)
section
variable [DecidableEq (ι1 ⊕ ι2)] [DecidableEq ι1] [DecidableEq ι2]
lemma pureInl_update_left (f : (i : ι1 ⊕ ι2) → Sum.elim s1 s2 i) (x : ι1)
variable [DecidableEq (ι1 ⊕ ι2)]
omit inst1 inst2
lemma pureInl_update_left [DecidableEq ι1] (f : (i : ι1 ⊕ ι2) → Sum.elim s1 s2 i) (x : ι1)
(v1 : s1 x) : pureInl (Function.update f (Sum.inl x) v1) =
Function.update (pureInl f) x v1 := by
funext y
@ -187,7 +196,7 @@ lemma pureInr_update_left (f : (i : ι1 ⊕ ι2) → Sum.elim s1 s2 i) (x : ι1)
funext y
simp [pureInr, Function.update, Sum.inl.injEq, Sum.elim_inl]
lemma pureInr_update_right (f : (i : ι1 ⊕ ι2) → Sum.elim s1 s2 i) (x : ι2)
lemma pureInr_update_right [DecidableEq ι2] (f : (i : ι1 ⊕ ι2) → Sum.elim s1 s2 i) (x : ι2)
(v2 : s2 x) : pureInr (Function.update f (Sum.inr x) v2) =
Function.update (pureInr f) x v2 := by
funext y
@ -205,7 +214,11 @@ lemma pureInl_update_right (f : (i : ι1 ⊕ ι2) → Sum.elim s1 s2 i) (x : ι2
simp [pureInl, Function.update, Sum.inr.injEq, Sum.elim_inr]
end
def domCoprod : MultilinearMap R (Sum.elim s1 s2) ((⨂[R] i : ι1, s1 i) ⊗[R] ⨂[R] i : ι2, s2 i) where
/-- The multilinear map from `(Sum.elim s1 s2)` to `((⨂[R] i : ι1, s1 i) ⊗[R] ⨂[R] i : ι2, s2 i)`
defined by splitting elements of `(Sum.elim s1 s2)` into two parts. -/
def domCoprod :
MultilinearMap R (Sum.elim s1 s2) ((⨂[R] i : ι1, s1 i) ⊗[R] ⨂[R] i : ι2, s2 i) where
toFun f := (PiTensorProduct.tprod R (pureInl f)) ⊗ₜ
(PiTensorProduct.tprod R (pureInr f))
map_add' f xy v1 v2 := by
@ -235,10 +248,12 @@ def domCoprod : MultilinearMap R (Sum.elim s1 s2) ((⨂[R] i : ι1, s1 i) ⊗[R]
simp only [Sum.elim_inr, pureInl_update_right, pureInr_update_right, MultilinearMap.map_smul,
tmul_smul]
/-- Expand `PiTensorProduct` on sums into a `TensorProduct` of two factors. -/
def tmulSymm : (⨂[R] i : ι1 ⊕ ι2, (Sum.elim s1 s2) i) →ₗ[R]
((⨂[R] i : ι1, s1 i) ⊗[R] ⨂[R] i : ι2, s2 i) := PiTensorProduct.lift domCoprod
/-- Produces a map `(i : ι1 ⊕ ι2) → Sum.elim s1 s2 i` from a map `(i : ι1) → s1 i` and a
map `q : (i : ι2) → s2 i`. -/
def elimPureTensor (p : (i : ι1) → s1 i) (q : (i : ι2) → s2 i) : (i : ι1 ⊕ ι2) → Sum.elim s1 s2 i :=
fun x =>
match x with
@ -248,6 +263,7 @@ def elimPureTensor (p : (i : ι1) → s1 i) (q : (i : ι2) → s2 i) : (i : ι1
section
variable [DecidableEq ι1] [DecidableEq ι2]
omit inst1 inst2
lemma elimPureTensor_update_right (p : (i : ι1) → s1 i) (q : (i : ι2) → s2 i)
(y : ι2) (r : s2 y) : elimPureTensor p (Function.update q y r) =
@ -286,6 +302,8 @@ lemma elimPureTensor_update_left (p : (i : ι1) → s1 i) (q : (i : ι2) → s2
end
/-- The multilinear map valued in multilinear maps defined by combining
`(i : ι1) → s1 i` and `q : (i : ι2) → s2 i` into a PiTensorProduct. -/
def elimPureTensorMulLin : MultilinearMap R s1
(MultilinearMap R s2 (⨂[R] i : ι1 ⊕ ι2, (Sum.elim s1 s2) i)) where
toFun p := {
@ -311,6 +329,7 @@ def elimPureTensorMulLin : MultilinearMap R s1
intro y
simp
/-- Collapse a `TensorProduct` of `PiTensorProduct` into a `PiTensorProduct`. -/
def tmul : ((⨂[R] i : ι1, s1 i) ⊗[R] ⨂[R] i : ι2, s2 i) →ₗ[R]
⨂[R] i : ι1 ⊕ ι2, (Sum.elim s1 s2) i := TensorProduct.lift {
toFun := fun a ↦
@ -319,6 +338,7 @@ def tmul : ((⨂[R] i : ι1, s1 i) ⊗[R] ⨂[R] i : ι2, s2 i) →ₗ[R]
map_add' := fun a b ↦ by simp
map_smul' := fun r a ↦ by simp}
/-- THe equivalence formed by combining a `TensorProduct` into a `PiTensorProduct`. -/
def tmulEquiv : ((⨂[R] i : ι1, s1 i) ⊗[R] ⨂[R] i : ι2, s2 i) ≃ₗ[R]
⨂[R] i : ι1 ⊕ ι2, (Sum.elim s1 s2) i :=
LinearEquiv.ofLinear tmul tmulSymm
@ -336,7 +356,7 @@ def tmulEquiv : ((⨂[R] i : ι1, s1 i) ⊗[R] ⨂[R] i : ι2, s2 i) ≃ₗ[R]
| Sum.inl x => rfl
| Sum.inr x => rfl)
(by
apply tensorProd_piTensorProd_ext
apply induction_tmul
intro p q
simp only [tmulSymm, domCoprod, tmul, elimPureTensorMulLin, LinearMap.coe_comp,
Function.comp_apply, lift.tmul, LinearMap.coe_mk, AddHom.coe_mk, PiTensorProduct.lift.tprod,

View file

@ -194,35 +194,44 @@ lemma obj_ρ_empty (g : SL(2, )) : (colorFun.obj (𝟙_ (OverColor Color))).
funext i
exact Empty.elim i
/-- The unit natural transformation. -/
/-- The unit natural isomorphism. -/
def ε : 𝟙_ (Rep SL(2, )) ≅ colorFun.obj (𝟙_ (OverColor Color)) where
hom := {
hom := (PiTensorProduct.isEmptyEquiv Empty).symm.toLinearMap
comm := fun M => by
refine LinearMap.ext (fun x => ?_)
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorUnit_left,
OverColor.instMonoidalCategoryStruct_tensorUnit_hom, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', Functor.id_obj, Category.id_comp, LinearEquiv.coe_coe]
OverColor.instMonoidalCategoryStruct_tensorUnit_hom,
Action.instMonoidalCategory_tensorUnit_V, Action.tensorUnit_ρ', Functor.id_obj,
Category.id_comp, LinearEquiv.coe_coe]
erw [obj_ρ_empty M]
rfl}
inv := {
hom := (PiTensorProduct.isEmptyEquiv Empty).toLinearMap
comm := fun M => by
refine LinearMap.ext (fun x => ?_)
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorUnit_left,
OverColor.instMonoidalCategoryStruct_tensorUnit_hom, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', Functor.id_obj, Category.id_comp, LinearEquiv.coe_coe]
simp only [Action.instMonoidalCategory_tensorUnit_V, colorFun_obj_V_carrier,
OverColor.instMonoidalCategoryStruct_tensorUnit_left,
OverColor.instMonoidalCategoryStruct_tensorUnit_hom, Functor.id_obj, Action.tensorUnit_ρ']
erw [obj_ρ_empty M]
rfl}
hom_inv_id := by
ext1
simp [CategoryStruct.comp]
simp only [Action.instMonoidalCategory_tensorUnit_V, CategoryStruct.comp,
OverColor.instMonoidalCategoryStruct_tensorUnit_hom,
OverColor.instMonoidalCategoryStruct_tensorUnit_left, Functor.id_obj, Action.Hom.comp_hom,
colorFun_obj_V_carrier, LinearEquiv.comp_coe, LinearEquiv.symm_trans_self,
LinearEquiv.refl_toLinearMap, Action.id_hom]
rfl
inv_hom_id := by
ext1
simp [CategoryStruct.comp]
simp only [CategoryStruct.comp, OverColor.instMonoidalCategoryStruct_tensorUnit_hom,
OverColor.instMonoidalCategoryStruct_tensorUnit_left, Functor.id_obj, Action.Hom.comp_hom,
colorFun_obj_V_carrier, Action.instMonoidalCategory_tensorUnit_V, LinearEquiv.comp_coe,
LinearEquiv.self_trans_symm, LinearEquiv.refl_toLinearMap, Action.id_hom]
rfl
/-- An auxillary equivalence, and trivial, of modules needed to define `μModEquiv`. -/
def colorToRepSumEquiv {X Y : OverColor Color} (i : X.left ⊕ Y.left) :
Sum.elim (fun i => colorToRep (X.hom i)) (fun i => colorToRep (Y.hom i)) i ≃ₗ[]
colorToRep (Sum.elim X.hom Y.hom i) :=
@ -230,8 +239,10 @@ def colorToRepSumEquiv {X Y : OverColor Color} (i : X.left ⊕ Y.left) :
| Sum.inl _ => LinearEquiv.refl _ _
| Sum.inr _ => LinearEquiv.refl _ _
def μModEquiv (X Y : OverColor Color) : (colorFun.obj X ⊗ colorFun.obj Y).V ≃ₗ[] colorFun.obj (X ⊗ Y) :=
HepLean.PiTensorProduct.tmulEquiv ≪≫ₗ PiTensorProduct.congr colorToRepSumEquiv
/-- The equivalence of modules corresonding to the tensorate. -/
def μModEquiv (X Y : OverColor Color) :
(colorFun.obj X ⊗ colorFun.obj Y).V ≃ₗ[] colorFun.obj (X ⊗ Y) :=
HepLean.PiTensorProduct.tmulEquiv ≪≫ₗ PiTensorProduct.congr colorToRepSumEquiv
lemma μModEquiv_tmul_tprod {X Y : OverColor Color}(p : (i : X.left) → (colorToRep (X.hom i)))
(q : (i : Y.left) → (colorToRep (Y.hom i))) :
@ -245,22 +256,24 @@ lemma μModEquiv_tmul_tprod {X Y : OverColor Color}(p : (i : X.left) → (colo
Action.FunctorCategoryEquivalence.functor_obj_obj]
rw [LinearEquiv.trans_apply]
erw [HepLean.PiTensorProduct.tmulEquiv_tmul_tprod]
change (PiTensorProduct.congr colorToRepSumEquiv) ((PiTensorProduct.tprod ) (HepLean.PiTensorProduct.elimPureTensor p q)) = _
change (PiTensorProduct.congr colorToRepSumEquiv) ((PiTensorProduct.tprod )
(HepLean.PiTensorProduct.elimPureTensor p q)) = _
rw [PiTensorProduct.congr_tprod]
rfl
/-- The natural isomorphism corresponding to the tensorate. -/
def μ (X Y : OverColor Color) : colorFun.obj X ⊗ colorFun.obj Y ≅ colorFun.obj (X ⊗ Y) where
hom := {
hom := (μModEquiv X Y).toLinearMap
comm := fun M => by
refine HepLean.PiTensorProduct.tensorProd_piTensorProd_ext (fun p q => ?_)
refine HepLean.PiTensorProduct.induction_tmul (fun p q => ?_)
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Functor.id_obj, CategoryStruct.comp,
Action.instMonoidalCategory_tensorObj_V, Action.tensor_ρ', LinearMap.coe_comp,
Function.comp_apply]
change (μModEquiv X Y) (((((colorFun.obj X).ρ M) (PiTensorProduct.tprod p)) ⊗ₜ[]
(((colorFun.obj Y).ρ M) (PiTensorProduct.tprod q))))
= ((colorFun.obj (X ⊗ Y)).ρ M) ((μModEquiv X Y) ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
(((colorFun.obj Y).ρ M) (PiTensorProduct.tprod q)))) = ((colorFun.obj (X ⊗ Y)).ρ M)
((μModEquiv X Y) ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
rw [μModEquiv_tmul_tprod]
erw [obj'_ρ_tprod, obj'_ρ_tprod, obj'_ρ_tprod]
rw [μModEquiv_tmul_tprod]
@ -278,15 +291,15 @@ def μ (X Y : OverColor Color) : colorFun.obj X ⊗ colorFun.obj Y ≅ colorFun.
simp [CategoryStruct.comp]
erw [LinearEquiv.eq_comp_toLinearMap_symm,LinearMap.comp_assoc ,
LinearEquiv.toLinearMap_symm_comp_eq ]
refine HepLean.PiTensorProduct.tensorProd_piTensorProd_ext (fun p q => ?_)
refine HepLean.PiTensorProduct.induction_tmul (fun p q => ?_)
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Functor.id_obj, CategoryStruct.comp,
Action.instMonoidalCategory_tensorObj_V, Action.tensor_ρ', LinearMap.coe_comp,
Function.comp_apply]
symm
change (μModEquiv X Y) (((((colorFun.obj X).ρ M) (PiTensorProduct.tprod p)) ⊗ₜ[]
(((colorFun.obj Y).ρ M) (PiTensorProduct.tprod q))))
= ((colorFun.obj (X ⊗ Y)).ρ M) ((μModEquiv X Y) ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
(((colorFun.obj Y).ρ M) (PiTensorProduct.tprod q)))) = ((colorFun.obj (X ⊗ Y)).ρ M)
((μModEquiv X Y) ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
rw [μModEquiv_tmul_tprod]
erw [obj'_ρ_tprod, obj'_ρ_tprod, obj'_ρ_tprod]
rw [μModEquiv_tmul_tprod]
@ -312,9 +325,6 @@ def μ (X Y : OverColor Color) : colorFun.obj X ⊗ colorFun.obj Y ≅ colorFun.
LinearEquiv.symm_trans_self, LinearEquiv.refl_toLinearMap, Action.id_hom]
rfl
@[simp]
lemma μ_tmul_tprod {X Y : OverColor Color} (p : (i : X.left) → (colorToRep (X.hom i)))
(q : (i : Y.left) → (colorToRep (Y.hom i))) :
(μ X Y).hom.hom ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q) =
@ -326,7 +336,7 @@ lemma μ_natural_left {X Y : OverColor Color} (f : X ⟶ Y) (Z : OverColor Color
MonoidalCategory.whiskerRight (colorFun.map f) (colorFun.obj Z) ≫ (μ Y Z).hom =
(μ X Z).hom ≫ colorFun.map (MonoidalCategory.whiskerRight f Z) := by
ext1
refine HepLean.PiTensorProduct.tensorProd_piTensorProd_ext (fun p q => ?_)
refine HepLean.PiTensorProduct.induction_tmul (fun p q => ?_)
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Functor.id_obj, CategoryStruct.comp,
Action.Hom.comp_hom, Action.instMonoidalCategory_tensorObj_V,
@ -338,13 +348,13 @@ lemma μ_natural_left {X Y : OverColor Color} (f : X ⟶ Y) (Z : OverColor Color
((PiTensorProduct.tprod ) fun i => (colorToRepSumEquiv i)
(HepLean.PiTensorProduct.elimPureTensor p q i))
rw [colorFun.map_tprod]
have h1 : (((colorFun.map f).hom ▷ (colorFun.obj Z).V) ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
= ((colorFun.map f).hom ((PiTensorProduct.tprod ) p) ⊗ₜ[] ((PiTensorProduct.tprod ) q)) := by rfl
have h1 : (((colorFun.map f).hom ▷ (colorFun.obj Z).V) ((PiTensorProduct.tprod ) p ⊗ₜ[]
(PiTensorProduct.tprod ) q)) = ((colorFun.map f).hom
((PiTensorProduct.tprod ) p) ⊗ₜ[] ((PiTensorProduct.tprod ) q)) := by rfl
erw [h1]
rw [colorFun.map_tprod]
change (μ Y Z).hom.hom
(((PiTensorProduct.tprod ) fun i => (colorToRepCongr _) (p ((OverColor.Hom.toEquiv f).symm i))) ⊗ₜ[]
(PiTensorProduct.tprod ) q) = _
change (μ Y Z).hom.hom (((PiTensorProduct.tprod ) fun i => (colorToRepCongr _)
(p ((OverColor.Hom.toEquiv f).symm i))) ⊗ₜ[] (PiTensorProduct.tprod ) q) = _
rw [μ_tmul_tprod]
apply congrArg
funext i
@ -356,23 +366,25 @@ lemma μ_natural_right {X Y : OverColor Color} (X' : OverColor Color) (f : X ⟶
MonoidalCategory.whiskerLeft (colorFun.obj X') (colorFun.map f) ≫ (μ X' Y).hom =
(μ X' X).hom ≫ colorFun.map (MonoidalCategory.whiskerLeft X' f) := by
ext1
refine HepLean.PiTensorProduct.tensorProd_piTensorProd_ext (fun p q => ?_)
refine HepLean.PiTensorProduct.induction_tmul (fun p q => ?_)
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Functor.id_obj, CategoryStruct.comp,
Action.Hom.comp_hom, Action.instMonoidalCategory_tensorObj_V,
Action.instMonoidalCategory_whiskerLeft_hom, LinearMap.coe_comp, Function.comp_apply]
change _ = (colorFun.map (X' ◁ f)).hom ((μ X' X).hom.hom ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
change _ = (colorFun.map (X' ◁ f)).hom ((μ X' X).hom.hom
((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
rw [μ_tmul_tprod]
change _ = (colorFun.map (X' ◁ f)).hom
((PiTensorProduct.tprod ) fun i => (colorToRepSumEquiv i) (HepLean.PiTensorProduct.elimPureTensor p q i))
change _ = (colorFun.map (X' ◁ f)).hom ((PiTensorProduct.tprod ) fun i =>
(colorToRepSumEquiv i) (HepLean.PiTensorProduct.elimPureTensor p q i))
rw [map_tprod]
have h1 : (((colorFun.obj X').V ◁ (colorFun.map f).hom) ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
= ((PiTensorProduct.tprod ) p ⊗ₜ[] (colorFun.map f).hom ((PiTensorProduct.tprod ) q)) := by rfl
have h1 : (((colorFun.obj X').V ◁ (colorFun.map f).hom)
((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q))
= ((PiTensorProduct.tprod ) p ⊗ₜ[] (colorFun.map f).hom ((PiTensorProduct.tprod ) q)) := by
rfl
erw [h1]
rw [map_tprod]
change (μ X' Y).hom.hom
((PiTensorProduct.tprod ) p ⊗ₜ[]
(PiTensorProduct.tprod ) fun i => (colorToRepCongr _) (q ((OverColor.Hom.toEquiv f).symm i))) = _
change (μ X' Y).hom.hom ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) fun i =>
(colorToRepCongr _) (q ((OverColor.Hom.toEquiv f).symm i))) = _
rw [μ_tmul_tprod]
apply congrArg
funext i
@ -393,17 +405,16 @@ lemma associativity (X Y Z : OverColor Color) :
Action.instMonoidalCategory_whiskerRight_hom, LinearMap.coe_comp, Function.comp_apply,
Action.instMonoidalCategory_whiskerLeft_hom, Action.instMonoidalCategory_associator_hom_hom]
change (colorFun.map (α_ X Y Z).hom).hom ((μ (X ⊗ Y) Z).hom.hom
((((μ X Y).hom.hom ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) q)) ⊗ₜ[] (PiTensorProduct.tprod ) m))) =
(μ X (Y ⊗ Z)).hom.hom
(( ((PiTensorProduct.tprod ) p ⊗ₜ[] ((μ Y Z).hom.hom ((PiTensorProduct.tprod ) q ⊗ₜ[] (PiTensorProduct.tprod ) m)))))
((((μ X Y).hom.hom ((PiTensorProduct.tprod ) p ⊗ₜ[]
(PiTensorProduct.tprod ) q)) ⊗ₜ[] (PiTensorProduct.tprod ) m))) =
(μ X (Y ⊗ Z)).hom.hom ((((PiTensorProduct.tprod ) p ⊗ₜ[] ((μ Y Z).hom.hom
((PiTensorProduct.tprod ) q ⊗ₜ[] (PiTensorProduct.tprod ) m)))))
rw [μ_tmul_tprod, μ_tmul_tprod]
change (colorFun.map (α_ X Y Z).hom).hom
((μ (X ⊗ Y) Z).hom.hom
(((PiTensorProduct.tprod ) fun i => (colorToRepSumEquiv i) (HepLean.PiTensorProduct.elimPureTensor p q i)) ⊗ₜ[]
(PiTensorProduct.tprod ) m)) =
(μ X (Y ⊗ Z)).hom.hom
((PiTensorProduct.tprod ) p ⊗ₜ[]
(PiTensorProduct.tprod ) fun i => (colorToRepSumEquiv i) (HepLean.PiTensorProduct.elimPureTensor q m i))
change (colorFun.map (α_ X Y Z).hom).hom ((μ (X ⊗ Y) Z).hom.hom
(((PiTensorProduct.tprod ) fun i => (colorToRepSumEquiv i)
(HepLean.PiTensorProduct.elimPureTensor p q i)) ⊗ₜ[] (PiTensorProduct.tprod ) m)) =
(μ X (Y ⊗ Z)).hom.hom ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) fun i =>
(colorToRepSumEquiv i) (HepLean.PiTensorProduct.elimPureTensor q m i))
rw [μ_tmul_tprod, μ_tmul_tprod]
erw [map_tprod]
apply congrArg
@ -426,14 +437,15 @@ lemma left_unitality (X : OverColor Color) : (leftUnitor (colorFun.obj X)).hom =
OverColor.instMonoidalCategoryStruct_tensorUnit_left,
OverColor.instMonoidalCategoryStruct_tensorObj_hom,
Action.instMonoidalCategory_whiskerRight_hom, LinearMap.coe_comp, Function.comp_apply]
change TensorProduct.lid (colorFun.obj X) (x ⊗ₜ[] (PiTensorProduct.tprod ) q) = (colorFun.map (λ_ X).hom).hom
((μ (𝟙_ (OverColor Color)) X).hom.hom ((((PiTensorProduct.isEmptyEquiv Empty).symm x) ⊗ₜ[] (PiTensorProduct.tprod ) q)))
change TensorProduct.lid (colorFun.obj X) (x ⊗ₜ[] (PiTensorProduct.tprod ) q) =
(colorFun.map (λ_ X).hom).hom ((μ (𝟙_ (OverColor Color)) X).hom.hom
((((PiTensorProduct.isEmptyEquiv Empty).symm x) ⊗ₜ[] (PiTensorProduct.tprod ) q)))
simp [PiTensorProduct.isEmptyEquiv]
rw [TensorProduct.smul_tmul, TensorProduct.tmul_smul]
erw [LinearMap.map_smul, LinearMap.map_smul]
apply congrArg
change _ = (colorFun.map (λ_ X).hom).hom
((μ (𝟙_ (OverColor Color)) X).hom.hom ((PiTensorProduct.tprod ) _ ⊗ₜ[] (PiTensorProduct.tprod ) q))
change _ = (colorFun.map (λ_ X).hom).hom ((μ (𝟙_ (OverColor Color)) X).hom.hom
((PiTensorProduct.tprod ) _ ⊗ₜ[] (PiTensorProduct.tprod ) q))
rw [μ_tmul_tprod]
erw [map_tprod]
rfl
@ -452,19 +464,21 @@ lemma right_unitality (X : OverColor Color) : (MonoidalCategory.rightUnitor (col
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Action.instMonoidalCategory_whiskerLeft_hom,
LinearMap.coe_comp, Function.comp_apply]
change TensorProduct.rid (colorFun.obj X) ((PiTensorProduct.tprod ) p ⊗ₜ[] x ) =
(colorFun.map (ρ_ X).hom).hom
((μ X (𝟙_ (OverColor Color))).hom.hom ((((PiTensorProduct.tprod ) p ⊗ₜ[] ((PiTensorProduct.isEmptyEquiv Empty).symm x)))))
(colorFun.map (ρ_ X).hom).hom ((μ X (𝟙_ (OverColor Color))).hom.hom
((((PiTensorProduct.tprod ) p ⊗ₜ[] ((PiTensorProduct.isEmptyEquiv Empty).symm x)))))
simp [PiTensorProduct.isEmptyEquiv]
erw [LinearMap.map_smul, LinearMap.map_smul]
apply congrArg
change _ = (colorFun.map (ρ_ X).hom).hom
((μ X (𝟙_ (OverColor Color))).hom.hom ((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) _))
change _ = (colorFun.map (ρ_ X).hom).hom ((μ X (𝟙_ (OverColor Color))).hom.hom
((PiTensorProduct.tprod ) p ⊗ₜ[] (PiTensorProduct.tprod ) _))
rw [μ_tmul_tprod]
erw [map_tprod]
rfl
end colorFun
/-- The monoidal functor between `OverColor Color` and `Rep SL(2, )` taking a map of colors
to the corresponding tensor product representation. -/
def colorFunMon : MonoidalFunctor (OverColor Color) (Rep SL(2, )) where
toFunctor := colorFun
ε := colorFun.ε.hom
@ -475,6 +489,5 @@ def colorFunMon : MonoidalFunctor (OverColor Color) (Rep SL(2, )) where
left_unitality := colorFun.left_unitality
right_unitality := colorFun.right_unitality
end
end Fermion

View file

@ -514,9 +514,9 @@ def elimPureTensorMulLin : MultilinearMap R (fun i => 𝓣.ColorModule (cX i))
toFun p := {
toFun := fun q => PiTensorProduct.tprod R (𝓣.elimPureTensor p q)
map_add' := fun m x v1 v2 => by
simp [Sum.elim_inl, Sum.elim_inr]
simp only [elimPureTensor_update_right, MultilinearMap.map_add]
map_smul' := fun m x r v => by
simp [Sum.elim_inl, Sum.elim_inr]}
simp only [elimPureTensor_update_right, MultilinearMap.map_smul]}
map_add' p x v1 v2 := by
apply MultilinearMap.ext
intro y
@ -547,7 +547,8 @@ def domCoprod : MultilinearMap R (fun x => 𝓣.ColorModule (Sum.elim cX cY x))
MultilinearMap.map_add, ← tmul_add]
map_smul' f xy r p := by
match xy with
| Sum.inl x => simp [TensorProduct.tmul_smul, TensorProduct.smul_tmul]
| Sum.inl x => simp only [Sum.elim_inl, inlPureTensor_update_left, MultilinearMap.map_smul,
inrPureTensor_update_left, smul_tmul, tmul_smul]
| Sum.inr y => simp [TensorProduct.tmul_smul, TensorProduct.smul_tmul]
/-- The linear map combining two tensors into a single tensor