Merge pull request #186 from HEPLean/IndexNotation
feat: Add monoidal functor for complex lorentz tensors
This commit is contained in:
commit
1a5bff7254
7 changed files with 708 additions and 37 deletions
|
@ -63,6 +63,7 @@ import HepLean.FlavorPhysics.CKMMatrix.Rows
|
|||
import HepLean.FlavorPhysics.CKMMatrix.StandardParameterization.Basic
|
||||
import HepLean.FlavorPhysics.CKMMatrix.StandardParameterization.StandardParameters
|
||||
import HepLean.Mathematics.LinearMaps
|
||||
import HepLean.Mathematics.PiTensorProduct
|
||||
import HepLean.Mathematics.SO3.Basic
|
||||
import HepLean.Meta.AllFilePaths
|
||||
import HepLean.Meta.Informal
|
||||
|
|
372
HepLean/Mathematics/PiTensorProduct.lean
Normal file
372
HepLean/Mathematics/PiTensorProduct.lean
Normal file
|
@ -0,0 +1,372 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import Mathlib.LinearAlgebra.PiTensorProduct
|
||||
/-!
|
||||
# Pi Tensor Products
|
||||
|
||||
The purpose of this file is to define some results about Pi tensor products not currently
|
||||
in Mathlib.
|
||||
|
||||
At some point these should either be up-streamed to Mathlib or replaced with definitions already
|
||||
in Mathlib.
|
||||
|
||||
-/
|
||||
namespace HepLean.PiTensorProduct
|
||||
|
||||
noncomputable section tmulEquiv
|
||||
|
||||
variable {R ι1 ι2 ι3 M N : Type} [CommSemiring R]
|
||||
{s1 : ι1 → Type} [inst1 : (i : ι1) → AddCommMonoid (s1 i)] [inst1' : (i : ι1) → Module R (s1 i)]
|
||||
{s2 : ι2 → Type} [inst2 : (i : ι2) → AddCommMonoid (s2 i)] [inst2' : (i : ι2) → Module R (s2 i)]
|
||||
{s3 : ι3 → Type} [inst3 : (i : ι3) → AddCommMonoid (s3 i)] [inst3' : (i : ι3) → Module R (s3 i)]
|
||||
[AddCommMonoid M] [Module R M]
|
||||
[AddCommMonoid N] [Module R N]
|
||||
|
||||
open TensorProduct
|
||||
|
||||
/-!
|
||||
|
||||
## induction principals for pi tensor products
|
||||
|
||||
-/
|
||||
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'
|
||||
refine fun x ↦
|
||||
PiTensorProduct.induction_on' x ?_ (by
|
||||
intro a b hx hy y
|
||||
simp [map_add, add_tmul, hx, hy])
|
||||
intro rx fx
|
||||
refine fun y ↦
|
||||
PiTensorProduct.induction_on' y ?_ (by
|
||||
intro a b hx hy
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod] at hx hy
|
||||
simp [map_add, tmul_add, hx, hy])
|
||||
intro ry fy
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, tmul_smul, LinearMapClass.map_smul]
|
||||
apply congrArg
|
||||
simp only [smul_tmul, tmul_smul, LinearMapClass.map_smul]
|
||||
exact congrArg (HSMul.hSMul rx) (h fx fy)
|
||||
|
||||
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
|
||||
apply TensorProduct.ext'
|
||||
refine fun x ↦
|
||||
PiTensorProduct.induction_on' x ?_ (by
|
||||
intro a b hx hy y
|
||||
simp [map_add, add_tmul, hx, hy])
|
||||
intro rx fx
|
||||
intro y
|
||||
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 := {
|
||||
toFun := fun y => f (PiTensorProduct.tprod R fx ⊗ₜ[R] y),
|
||||
map_add' := fun y1 y2 => by
|
||||
simp [tmul_add]
|
||||
map_smul' := fun r y => by
|
||||
simp [tmul_smul]}
|
||||
let g' : ((⨂[R] i : ι2, s2 i) ⊗[R] ⨂[R] i : ι3, s3 i) →ₗ[R] M := {
|
||||
toFun := fun y => g (PiTensorProduct.tprod R fx ⊗ₜ[R] y),
|
||||
map_add' := fun y1 y2 => by
|
||||
simp [tmul_add]
|
||||
map_smul' := fun r y => by
|
||||
simp [tmul_smul]}
|
||||
change f' y = g' y
|
||||
apply congrFun
|
||||
refine DFunLike.coe_fn_eq.mpr ?H.h.h.a
|
||||
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
|
||||
apply TensorProduct.ext'
|
||||
intro x
|
||||
refine fun y ↦
|
||||
PiTensorProduct.induction_on' y ?_ (by
|
||||
intro a b hy hx
|
||||
simp [map_add, add_tmul, tmul_add, hy, hx])
|
||||
intro ry fy
|
||||
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),
|
||||
map_add' := fun y1 y2 => by
|
||||
simp [add_tmul]
|
||||
map_smul' := fun r y => by
|
||||
simp [smul_tmul]}
|
||||
let g' : ((⨂[R] i : ι1, s1 i) ⊗[R] ⨂[R] i : ι2, s2 i) →ₗ[R] M := {
|
||||
toFun := fun y => g (y ⊗ₜ[R] PiTensorProduct.tprod R fy),
|
||||
map_add' := fun y1 y2 => by
|
||||
simp [add_tmul]
|
||||
map_smul' := fun r y => by
|
||||
simp [smul_tmul]}
|
||||
change f' x = g' x
|
||||
apply congrFun
|
||||
refine DFunLike.coe_fn_eq.mpr ?H.h.h.a
|
||||
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
|
||||
apply TensorProduct.ext'
|
||||
refine fun y ↦
|
||||
PiTensorProduct.induction_on' y ?_ (by
|
||||
intro a b hy hx
|
||||
simp [map_add, add_tmul, tmul_add, hy, hx])
|
||||
intro ry fy
|
||||
intro x
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, smul_tmul, tmul_smul, map_smul]
|
||||
apply congrArg
|
||||
exact h fy x
|
||||
|
||||
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
|
||||
apply TensorProduct.ext'
|
||||
intro x
|
||||
refine fun y ↦
|
||||
PiTensorProduct.induction_on' y ?_ (by
|
||||
intro a b hy hx
|
||||
simp [map_add, add_tmul, tmul_add, hy, hx])
|
||||
intro ry fy
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, smul_tmul, tmul_smul, map_smul]
|
||||
apply congrArg
|
||||
exact h x fy
|
||||
/-!
|
||||
|
||||
# Dependent type version of PiTensorProduct.tmulEquiv
|
||||
-/
|
||||
|
||||
instance : (i : ι1 ⊕ ι2) → AddCommMonoid ((fun i => Sum.elim s1 s2 i) i) := fun i =>
|
||||
match i with
|
||||
| Sum.inl i => inst1 i
|
||||
| Sum.inr i => inst2 i
|
||||
|
||||
instance : (i : ι1 ⊕ ι2) → Module R ((fun i => Sum.elim s1 s2 i) i) := fun i =>
|
||||
match i with
|
||||
| Sum.inl i => inst1' i
|
||||
| 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)]
|
||||
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
|
||||
simp only [pureInl, Function.update, Sum.inl.injEq, Sum.elim_inl]
|
||||
split
|
||||
· rename_i h
|
||||
subst h
|
||||
rfl
|
||||
· rfl
|
||||
|
||||
lemma pureInr_update_left (f : (i : ι1 ⊕ ι2) → Sum.elim s1 s2 i) (x : ι1)
|
||||
(v2 : s1 x) :
|
||||
pureInr (Function.update f (Sum.inl x) v2) = (pureInr f) := by
|
||||
funext y
|
||||
simp [pureInr, Function.update, Sum.inl.injEq, Sum.elim_inl]
|
||||
|
||||
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
|
||||
simp only [pureInr, Function.update, Sum.inr.injEq, Sum.elim_inr]
|
||||
split
|
||||
· rename_i h
|
||||
subst h
|
||||
rfl
|
||||
· rfl
|
||||
|
||||
lemma pureInl_update_right (f : (i : ι1 ⊕ ι2) → Sum.elim s1 s2 i) (x : ι2)
|
||||
(v1 : s2 x) :
|
||||
pureInl (Function.update f (Sum.inr x) v1) = (pureInl f) := by
|
||||
funext y
|
||||
simp [pureInl, Function.update, Sum.inr.injEq, Sum.elim_inr]
|
||||
|
||||
end
|
||||
|
||||
/-- 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
|
||||
haveI : DecidableEq (ι1 ⊕ ι2) := inferInstance
|
||||
haveI : DecidableEq ι1 :=
|
||||
@Function.Injective.decidableEq ι1 (ι1 ⊕ ι2) Sum.inl _ Sum.inl_injective
|
||||
haveI : DecidableEq ι2 :=
|
||||
@Function.Injective.decidableEq ι2 (ι1 ⊕ ι2) Sum.inr _ Sum.inr_injective
|
||||
match xy with
|
||||
| Sum.inl xy =>
|
||||
simp only [Sum.elim_inl, pureInl_update_left, MultilinearMap.map_add, pureInr_update_left, ←
|
||||
add_tmul]
|
||||
| Sum.inr xy =>
|
||||
simp only [Sum.elim_inr, pureInr_update_right, MultilinearMap.map_add, pureInl_update_right, ←
|
||||
tmul_add]
|
||||
map_smul' f xy r p := by
|
||||
haveI : DecidableEq (ι1 ⊕ ι2) := inferInstance
|
||||
haveI : DecidableEq ι1 :=
|
||||
@Function.Injective.decidableEq ι1 (ι1 ⊕ ι2) Sum.inl _ Sum.inl_injective
|
||||
haveI : DecidableEq ι2 :=
|
||||
@Function.Injective.decidableEq ι2 (ι1 ⊕ ι2) Sum.inr _ Sum.inr_injective
|
||||
match xy with
|
||||
| Sum.inl x =>
|
||||
simp only [Sum.elim_inl, pureInl_update_left, MultilinearMap.map_smul, pureInr_update_left,
|
||||
smul_tmul, tmul_smul]
|
||||
| Sum.inr y =>
|
||||
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
|
||||
| Sum.inl x => p x
|
||||
| Sum.inr x => q x
|
||||
|
||||
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) =
|
||||
Function.update (elimPureTensor p q) (Sum.inr y) r := by
|
||||
funext x
|
||||
match x with
|
||||
| Sum.inl x =>
|
||||
simp only [Sum.elim_inl, ne_eq, reduceCtorEq, not_false_eq_true, Function.update_noteq]
|
||||
rfl
|
||||
| Sum.inr x =>
|
||||
change Function.update q y r x = _
|
||||
simp only [Function.update, Sum.inr.injEq, Sum.elim_inr]
|
||||
split_ifs
|
||||
· rename_i h
|
||||
subst h
|
||||
simp_all only
|
||||
· rfl
|
||||
|
||||
@[simp]
|
||||
lemma elimPureTensor_update_left (p : (i : ι1) → s1 i) (q : (i : ι2) → s2 i)
|
||||
(x : ι1) (r : s1 x) : elimPureTensor (Function.update p x r) q =
|
||||
Function.update (elimPureTensor p q) (Sum.inl x) r := by
|
||||
funext y
|
||||
match y with
|
||||
| Sum.inl y =>
|
||||
change (Function.update p x r) y = _
|
||||
simp only [Function.update, Sum.inl.injEq, Sum.elim_inl]
|
||||
split_ifs
|
||||
· rename_i h
|
||||
subst h
|
||||
rfl
|
||||
· rfl
|
||||
| Sum.inr y =>
|
||||
simp only [Sum.elim_inr, ne_eq, reduceCtorEq, not_false_eq_true, Function.update_noteq]
|
||||
rfl
|
||||
|
||||
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 := {
|
||||
toFun := fun q => PiTensorProduct.tprod R (elimPureTensor p q)
|
||||
map_add' := fun m x v1 v2 => by
|
||||
haveI : DecidableEq ι2 := inferInstance
|
||||
haveI := Classical.decEq ι1
|
||||
simp only [elimPureTensor_update_right, MultilinearMap.map_add]
|
||||
map_smul' := fun m x r v => by
|
||||
haveI : DecidableEq ι2 := inferInstance
|
||||
haveI := Classical.decEq ι1
|
||||
simp only [elimPureTensor_update_right, MultilinearMap.map_smul]}
|
||||
map_add' p x v1 v2 := by
|
||||
haveI : DecidableEq ι1 := inferInstance
|
||||
haveI := Classical.decEq ι2
|
||||
apply MultilinearMap.ext
|
||||
intro y
|
||||
simp
|
||||
map_smul' p x r v := by
|
||||
haveI : DecidableEq ι1 := inferInstance
|
||||
haveI := Classical.decEq ι2
|
||||
apply MultilinearMap.ext
|
||||
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 ↦
|
||||
PiTensorProduct.lift <|
|
||||
PiTensorProduct.lift (elimPureTensorMulLin) a,
|
||||
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
|
||||
(by
|
||||
apply PiTensorProduct.ext
|
||||
apply MultilinearMap.ext
|
||||
intro p
|
||||
simp only [tmul, tmulSymm, domCoprod, LinearMap.compMultilinearMap_apply,
|
||||
LinearMap.coe_comp, Function.comp_apply, PiTensorProduct.lift.tprod, MultilinearMap.coe_mk,
|
||||
lift.tmul, LinearMap.coe_mk, AddHom.coe_mk]
|
||||
simp only [elimPureTensorMulLin, MultilinearMap.coe_mk, LinearMap.id_coe, id_eq]
|
||||
apply congrArg
|
||||
funext x
|
||||
match x with
|
||||
| Sum.inl x => rfl
|
||||
| Sum.inr x => rfl)
|
||||
(by
|
||||
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,
|
||||
MultilinearMap.coe_mk, LinearMap.id_coe, id_eq]
|
||||
rfl)
|
||||
|
||||
@[simp]
|
||||
lemma tmulEquiv_tmul_tprod (p : (i : ι1) → s1 i) (q : (i : ι2) → s2 i) :
|
||||
tmulEquiv ((PiTensorProduct.tprod R) p ⊗ₜ[R] (PiTensorProduct.tprod R) q) =
|
||||
(PiTensorProduct.tprod R) (elimPureTensor p q) := by
|
||||
simp only [tmulEquiv, tmul, elimPureTensorMulLin, LinearEquiv.ofLinear_apply, lift.tmul,
|
||||
LinearMap.coe_mk, AddHom.coe_mk, PiTensorProduct.lift.tprod, MultilinearMap.coe_mk]
|
||||
|
||||
end tmulEquiv
|
||||
end HepLean.PiTensorProduct
|
|
@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
|
|||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.Tensors.ColorCat.Basic
|
||||
import HepLean.Mathematics.PiTensorProduct
|
||||
/-!
|
||||
|
||||
## Monodial functor from color cat.
|
||||
|
@ -164,6 +165,18 @@ namespace colorFun
|
|||
open CategoryTheory
|
||||
open MonoidalCategory
|
||||
|
||||
lemma map_tprod {X Y : OverColor Color} (p : (i : X.left) → (colorToRep (X.hom i)))
|
||||
(f : X ⟶ Y) : (colorFun.map f).hom (PiTensorProduct.tprod ℂ p) =
|
||||
PiTensorProduct.tprod ℂ fun (i : Y.left) => colorToRepCongr
|
||||
(OverColor.Hom.toEquiv_comp_inv_apply f i) (p ((OverColor.Hom.toEquiv f).symm i)) := by
|
||||
change (colorFun.map' f).hom ((PiTensorProduct.tprod ℂ) p) = _
|
||||
simp [colorFun.map', mapToLinearEquiv']
|
||||
erw [LinearEquiv.trans_apply]
|
||||
change (PiTensorProduct.congr fun i => colorToRepCongr _)
|
||||
((PiTensorProduct.reindex ℂ (fun x => _) (OverColor.Hom.toEquiv f))
|
||||
((PiTensorProduct.tprod ℂ) p)) = _
|
||||
rw [PiTensorProduct.reindex_tprod, PiTensorProduct.congr_tprod]
|
||||
|
||||
@[simp]
|
||||
lemma obj_ρ_empty (g : SL(2, ℂ)) : (colorFun.obj (𝟙_ (OverColor Color))).ρ g = LinearMap.id := by
|
||||
erw [colorFun.obj'_ρ]
|
||||
|
@ -181,18 +194,300 @@ lemma obj_ρ_empty (g : SL(2, ℂ)) : (colorFun.obj (𝟙_ (OverColor Color))).
|
|||
funext i
|
||||
exact Empty.elim i
|
||||
|
||||
/-- The unit natural transformation. -/
|
||||
def ε : 𝟙_ (Rep ℂ SL(2, ℂ)) ⟶ colorFun.obj (𝟙_ (OverColor Color)) where
|
||||
hom := (PiTensorProduct.isEmptyEquiv Empty).symm.toLinearMap
|
||||
comm 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]
|
||||
erw [obj_ρ_empty M]
|
||||
/-- 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]
|
||||
erw [obj_ρ_empty M]
|
||||
rfl}
|
||||
inv := {
|
||||
hom := (PiTensorProduct.isEmptyEquiv Empty).toLinearMap
|
||||
comm := fun M => by
|
||||
refine LinearMap.ext (fun x => ?_)
|
||||
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 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 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) :=
|
||||
match i with
|
||||
| Sum.inl _ => LinearEquiv.refl _ _
|
||||
| Sum.inr _ => LinearEquiv.refl _ _
|
||||
|
||||
/-- 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))) :
|
||||
(μModEquiv X Y) ((PiTensorProduct.tprod ℂ) p ⊗ₜ[ℂ] (PiTensorProduct.tprod ℂ) q) =
|
||||
(PiTensorProduct.tprod ℂ) fun i =>
|
||||
(colorToRepSumEquiv i) (HepLean.PiTensorProduct.elimPureTensor p q i) := by
|
||||
rw [μModEquiv]
|
||||
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
|
||||
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Action.instMonoidalCategory_tensorObj_V,
|
||||
Functor.id_obj, Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
||||
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)) = _
|
||||
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.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))
|
||||
rw [μModEquiv_tmul_tprod]
|
||||
erw [obj'_ρ_tprod, obj'_ρ_tprod, obj'_ρ_tprod]
|
||||
rw [μModEquiv_tmul_tprod]
|
||||
apply congrArg
|
||||
funext i
|
||||
match i with
|
||||
| Sum.inl i =>
|
||||
rfl
|
||||
| Sum.inr i =>
|
||||
rfl
|
||||
}
|
||||
inv := {
|
||||
hom := (μModEquiv X Y).symm.toLinearMap
|
||||
comm := fun M => by
|
||||
simp [CategoryStruct.comp]
|
||||
erw [LinearEquiv.eq_comp_toLinearMap_symm,LinearMap.comp_assoc,
|
||||
LinearEquiv.toLinearMap_symm_comp_eq]
|
||||
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))
|
||||
rw [μModEquiv_tmul_tprod]
|
||||
erw [obj'_ρ_tprod, obj'_ρ_tprod, obj'_ρ_tprod]
|
||||
rw [μModEquiv_tmul_tprod]
|
||||
apply congrArg
|
||||
funext i
|
||||
match i with
|
||||
| Sum.inl i =>
|
||||
rfl
|
||||
| Sum.inr i =>
|
||||
rfl}
|
||||
hom_inv_id := by
|
||||
ext1
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, CategoryStruct.comp, Action.Hom.comp_hom,
|
||||
colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
|
||||
OverColor.instMonoidalCategoryStruct_tensorObj_hom, LinearEquiv.comp_coe,
|
||||
LinearEquiv.self_trans_symm, LinearEquiv.refl_toLinearMap, Action.id_hom]
|
||||
rfl
|
||||
inv_hom_id := by
|
||||
ext1
|
||||
simp only [CategoryStruct.comp, Action.instMonoidalCategory_tensorObj_V, Action.Hom.comp_hom,
|
||||
colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
|
||||
OverColor.instMonoidalCategoryStruct_tensorObj_hom, LinearEquiv.comp_coe,
|
||||
LinearEquiv.symm_trans_self, LinearEquiv.refl_toLinearMap, Action.id_hom]
|
||||
rfl
|
||||
|
||||
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) =
|
||||
(PiTensorProduct.tprod ℂ) fun i =>
|
||||
(colorToRepSumEquiv i) (HepLean.PiTensorProduct.elimPureTensor p q i) := by
|
||||
exact μModEquiv_tmul_tprod p q
|
||||
|
||||
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.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_whiskerRight_hom, LinearMap.coe_comp, Function.comp_apply]
|
||||
change _ = (colorFun.map (MonoidalCategory.whiskerRight f Z)).hom
|
||||
((μ X Z).hom.hom ((PiTensorProduct.tprod ℂ) p ⊗ₜ[ℂ] (PiTensorProduct.tprod ℂ) q))
|
||||
rw [μ_tmul_tprod]
|
||||
change _ = (colorFun.map (f ▷ Z)).hom
|
||||
((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
|
||||
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) = _
|
||||
rw [μ_tmul_tprod]
|
||||
apply congrArg
|
||||
funext i
|
||||
match i with
|
||||
| Sum.inl i => rfl
|
||||
| Sum.inr i => rfl
|
||||
|
||||
lemma μ_natural_right {X Y : OverColor Color} (X' : OverColor Color) (f : X ⟶ Y) :
|
||||
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.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))
|
||||
rw [μ_tmul_tprod]
|
||||
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
|
||||
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))) = _
|
||||
rw [μ_tmul_tprod]
|
||||
apply congrArg
|
||||
funext i
|
||||
match i with
|
||||
| Sum.inl i => rfl
|
||||
| Sum.inr i => rfl
|
||||
|
||||
lemma associativity (X Y Z : OverColor Color) :
|
||||
whiskerRight (μ X Y).hom (colorFun.obj Z) ≫
|
||||
(μ (X ⊗ Y) Z).hom ≫ colorFun.map (associator X Y Z).hom =
|
||||
(associator (colorFun.obj X) (colorFun.obj Y) (colorFun.obj Z)).hom ≫
|
||||
whiskerLeft (colorFun.obj X) (μ Y Z).hom ≫ (μ X (Y ⊗ Z)).hom := by
|
||||
ext1
|
||||
refine HepLean.PiTensorProduct.induction_assoc' (fun p q m => ?_)
|
||||
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_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)))))
|
||||
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))
|
||||
rw [μ_tmul_tprod, μ_tmul_tprod]
|
||||
erw [map_tprod]
|
||||
apply congrArg
|
||||
funext i
|
||||
match i with
|
||||
| Sum.inl i => rfl
|
||||
| Sum.inr (Sum.inl i) => rfl
|
||||
| Sum.inr (Sum.inr i) => rfl
|
||||
|
||||
lemma left_unitality (X : OverColor Color) : (leftUnitor (colorFun.obj X)).hom =
|
||||
whiskerRight colorFun.ε.hom (colorFun.obj X) ≫
|
||||
(μ (𝟙_ (OverColor Color)) X).hom ≫ colorFun.map (leftUnitor X).hom := by
|
||||
ext1
|
||||
apply HepLean.PiTensorProduct.induction_mod_tmul (fun x q => ?_)
|
||||
simp only [colorFun_obj_V_carrier, Equivalence.symm_inverse,
|
||||
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
|
||||
Action.instMonoidalCategory_tensorUnit_V, Functor.id_obj,
|
||||
Action.instMonoidalCategory_leftUnitor_hom_hom, CategoryStruct.comp, Action.Hom.comp_hom,
|
||||
Action.instMonoidalCategory_tensorObj_V, OverColor.instMonoidalCategoryStruct_tensorObj_left,
|
||||
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)))
|
||||
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))
|
||||
rw [μ_tmul_tprod]
|
||||
erw [map_tprod]
|
||||
rfl
|
||||
|
||||
lemma right_unitality (X : OverColor Color) : (MonoidalCategory.rightUnitor (colorFun.obj X)).hom =
|
||||
whiskerLeft (colorFun.obj X) ε.hom ≫
|
||||
(μ X (𝟙_ (OverColor Color))).hom ≫ colorFun.map (MonoidalCategory.rightUnitor X).hom := by
|
||||
ext1
|
||||
apply HepLean.PiTensorProduct.induction_tmul_mod (fun p x => ?_)
|
||||
simp only [colorFun_obj_V_carrier, Functor.id_obj, Equivalence.symm_inverse,
|
||||
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
|
||||
Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_rightUnitor_hom_hom,
|
||||
CategoryStruct.comp, Action.Hom.comp_hom, Action.instMonoidalCategory_tensorObj_V,
|
||||
OverColor.instMonoidalCategoryStruct_tensorObj_left,
|
||||
OverColor.instMonoidalCategoryStruct_tensorUnit_left,
|
||||
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)))))
|
||||
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 ℂ) _))
|
||||
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
|
||||
μ X Y := (colorFun.μ X Y).hom
|
||||
μ_natural_left := colorFun.μ_natural_left
|
||||
μ_natural_right := colorFun.μ_natural_right
|
||||
associativity := colorFun.associativity
|
||||
left_unitality := colorFun.left_unitality
|
||||
right_unitality := colorFun.right_unitality
|
||||
|
||||
end
|
||||
end Fermion
|
||||
|
|
|
@ -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
|
||||
|
|
|
@ -61,6 +61,10 @@ lemma toEquiv_comp_hom (m : f ⟶ g) : g.hom ∘ (toEquiv m) = f.hom := by
|
|||
ext x
|
||||
simpa [types_comp, toEquiv] using congrFun m.hom.w x
|
||||
|
||||
lemma toEquiv_comp_inv_apply (m : f ⟶ g) (i : g.left) :
|
||||
f.hom ((OverColor.Hom.toEquiv m).symm i) = g.hom i := by
|
||||
simpa [toEquiv, types_comp] using congrFun m.inv.w i
|
||||
|
||||
end Hom
|
||||
|
||||
section monoidal
|
||||
|
@ -265,7 +269,7 @@ def map {C D : Type} (f : C → D) : MonoidalFunctor (OverColor C) (OverColor D)
|
|||
| Sum.inr x => rfl
|
||||
|
||||
/-- The tensor product on `OverColor C` as a monoidal functor. -/
|
||||
def tensor : MonoidalFunctor (OverColor C × OverColor C) (OverColor C) where
|
||||
def tensor : MonoidalFunctor (OverColor C × OverColor C) (OverColor C) where
|
||||
toFunctor := MonoidalCategory.tensor (OverColor C)
|
||||
ε := Over.isoMk (Equiv.sumEmpty Empty Empty).symm.toIso (by
|
||||
ext x
|
||||
|
@ -340,7 +344,6 @@ def finExtractOne {n : ℕ} (i : Fin n.succ) : Fin n.succ ≃ Fin 1 ⊕ Fin n :=
|
|||
(Equiv.sumAssoc (Fin 1) (Fin i) (Fin (n - i))).trans <|
|
||||
Equiv.sumCongr (Equiv.refl (Fin 1)) (finSumFinEquiv.trans (finCongr (by omega)))
|
||||
|
||||
|
||||
end OverColor
|
||||
|
||||
end IndexNotation
|
||||
|
|
|
@ -17,46 +17,46 @@ that converts a tensor tree into a dot file.
|
|||
namespace TensorTree
|
||||
|
||||
/-- Turns a nodes of tensor trees into nodes and edges of a dot file. -/
|
||||
def dotString (m : ℕ) (nt : ℕ): ∀ {n : ℕ} {c : Fin n → S.C}, TensorTree S c → String := fun
|
||||
def dotString (m : ℕ) (nt : ℕ) : ∀ {n : ℕ} {c : Fin n → S.C}, TensorTree S c → String := fun
|
||||
| tensorNode _ =>
|
||||
" node" ++ toString m ++ " [label=\"T" ++ toString nt ++ "\"];\n"
|
||||
" node" ++ toString m ++ " [label=\"T" ++ toString nt ++ "\"];\n"
|
||||
| add t1 t2 =>
|
||||
let addNode := " node" ++ toString m ++ " [label=\"+\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
let edge2 := " node" ++ toString m ++ " -> node" ++ toString (m + 2) ++ ";\n"
|
||||
let addNode := " node" ++ toString m ++ " [label=\"+\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
let edge2 := " node" ++ toString m ++ " -> node" ++ toString (m + 2) ++ ";\n"
|
||||
addNode ++ dotString (m + 1) nt t1 ++ dotString (m + 2) (nt + 1) t2 ++ edge1 ++ edge2
|
||||
| perm σ t =>
|
||||
let permNode := " node" ++ toString m ++ " [label=\"perm\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
let permNode := " node" ++ toString m ++ " [label=\"perm\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
permNode ++ dotString (m + 1) nt t ++ edge1
|
||||
| prod t1 t2 =>
|
||||
let prodNode := " node" ++ toString m ++ " [label=\"⊗\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
let edge2 := " node" ++ toString m ++ " -> node" ++ toString (2 * t1.size + m + 2) ++ ";\n"
|
||||
let prodNode := " node" ++ toString m ++ " [label=\"⊗\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
let edge2 := " node" ++ toString m ++ " -> node" ++ toString (2 * t1.size + m + 2) ++ ";\n"
|
||||
prodNode ++ dotString (m + 1) nt t1 ++ dotString (2 * t1.size + m + 2) (nt + 1) t2
|
||||
++ edge1 ++ edge2
|
||||
| smul k t =>
|
||||
let smulNode := " node" ++ toString m ++ " [label=\"smul\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
let smulNode := " node" ++ toString m ++ " [label=\"smul\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
smulNode ++ dotString (m + 1) nt t ++ edge1
|
||||
| mult _ _ t1 t2 =>
|
||||
let multNode := " node" ++ toString m ++ " [label=\"mult\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
let edge2 := " node" ++ toString m ++ " -> node" ++ toString (t1.size + m + 2) ++ ";\n"
|
||||
let multNode := " node" ++ toString m ++ " [label=\"mult\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
let edge2 := " node" ++ toString m ++ " -> node" ++ toString (t1.size + m + 2) ++ ";\n"
|
||||
multNode ++ dotString (m + 1) nt t1 ++ dotString (2 * t1.size + m + 2) (nt + 1) t2
|
||||
++ edge1 ++ edge2
|
||||
| eval _ _ t1 =>
|
||||
let evalNode := " node" ++ toString m ++ " [label=\"eval\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
let evalNode := " node" ++ toString m ++ " [label=\"eval\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
evalNode ++ dotString (m + 1) nt t1 ++ edge1
|
||||
| jiggle i t1 =>
|
||||
let jiggleNode := " node" ++ toString m ++ " [label=\"τ\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
let jiggleNode := " node" ++ toString m ++ " [label=\"τ\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
jiggleNode ++ dotString (m + 1) nt t1 ++ edge1
|
||||
| contr i j t1 =>
|
||||
let contrNode := " node" ++ toString m ++ " [label=\"contr " ++ toString i ++ " "
|
||||
let contrNode := " node" ++ toString m ++ " [label=\"contr " ++ toString i ++ " "
|
||||
++ toString j ++ "\", shape=box];\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
|
||||
contrNode ++ dotString (m + 1) nt t1 ++ edge1
|
||||
|
||||
/-- Used to form a dot graph from a tensor tree. use e.g.
|
||||
|
@ -76,7 +76,7 @@ open Lean.Elab.Command Lean Meta Lean.Elab
|
|||
syntax (name := tensorDot) "#tensor_dot " term : command
|
||||
|
||||
/-- Adapted from `Lean.Elab.Command.elabReduce` in file copyright Microsoft Corporation. -/
|
||||
unsafe def dotElab (term : Syntax) : CommandElabM Unit :=
|
||||
unsafe def dotElab (term : Syntax) : CommandElabM Unit :=
|
||||
withoutModifyingEnv <| runTermElabM fun _ => Term.withDeclName `_reduce do
|
||||
let e ← Term.elabTerm term none
|
||||
Term.synthesizeSyntheticMVarsNoPostponing
|
||||
|
|
|
@ -308,5 +308,4 @@ variable (𝓣 : TensorTree S c4)
|
|||
-/
|
||||
end ProdNode
|
||||
|
||||
|
||||
end TensorTree
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue