feat: Add Pauli-matrices as tensor.
This commit is contained in:
parent
a60ade65f0
commit
691b7e112e
13 changed files with 989 additions and 154 deletions
|
@ -24,11 +24,11 @@ open Complex
|
|||
open TensorProduct
|
||||
open CategoryTheory.MonoidalCategory
|
||||
|
||||
/-- The left-alt-left unit `δₐᵃ` as an element of `(leftHanded ⊗ altLeftHanded).V`. -/
|
||||
/-- The left-alt-left unit `δᵃₐ` as an element of `(leftHanded ⊗ altLeftHanded).V`. -/
|
||||
def leftAltLeftUnitVal : (leftHanded ⊗ altLeftHanded).V :=
|
||||
leftAltLeftToMatrix.symm 1
|
||||
|
||||
/-- The left-alt-left unit `δₐᵃ` as a morphism `𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ leftHanded ⊗ altLeftHanded `,
|
||||
/-- The left-alt-left unit `δᵃₐ` as a morphism `𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ leftHanded ⊗ altLeftHanded `,
|
||||
manifesting the invariance under the `SL(2,ℂ)` action. -/
|
||||
def leftAltLeftUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ leftHanded ⊗ altLeftHanded where
|
||||
hom := {
|
||||
|
@ -55,11 +55,11 @@ def leftAltLeftUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ leftHanded ⊗ altLeftHanded
|
|||
apply congrArg
|
||||
simp
|
||||
|
||||
/-- The alt-left-left unit `δᵃₐ` as an element of `(altLeftHanded ⊗ leftHanded).V`. -/
|
||||
/-- The alt-left-left unit `δₐᵃ` as an element of `(altLeftHanded ⊗ leftHanded).V`. -/
|
||||
def altLeftLeftUnitVal : (altLeftHanded ⊗ leftHanded).V :=
|
||||
altLeftLeftToMatrix.symm 1
|
||||
|
||||
/-- The alt-left-left unit `δᵃₐ` as a morphism `𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ altLeftHanded ⊗ leftHanded `,
|
||||
/-- The alt-left-left unit `δₐᵃ` as a morphism `𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ altLeftHanded ⊗ leftHanded `,
|
||||
manifesting the invariance under the `SL(2,ℂ)` action. -/
|
||||
def altLeftLeftUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ altLeftHanded ⊗ leftHanded where
|
||||
hom := {
|
||||
|
@ -87,12 +87,12 @@ def altLeftLeftUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ altLeftHanded ⊗ leftHanded
|
|||
simp only [mul_one, ← transpose_mul, SpecialLinearGroup.det_coe, isUnit_iff_ne_zero, ne_eq,
|
||||
one_ne_zero, not_false_eq_true, mul_nonsing_inv, transpose_one]
|
||||
|
||||
/-- The right-alt-right unit `δ_{dot a}^{dot a}` as an element of
|
||||
/-- The right-alt-right unit `δ^{dot a}_{dot a}` as an element of
|
||||
`(rightHanded ⊗ altRightHanded).V`. -/
|
||||
def rightAltRightUnitVal : (rightHanded ⊗ altRightHanded).V :=
|
||||
rightAltRightToMatrix.symm 1
|
||||
|
||||
/-- The right-alt-right unit `δ_{dot a}^{dot a}` as a morphism
|
||||
/-- The right-alt-right unit `δ^{dot a}_{dot a}` as a morphism
|
||||
`𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ rightHanded ⊗ altRightHanded`, manifesting
|
||||
the invariance under the `SL(2,ℂ)` action. -/
|
||||
def rightAltRightUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ rightHanded ⊗ altRightHanded where
|
||||
|
@ -126,12 +126,12 @@ def rightAltRightUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ rightHanded ⊗ altRightHa
|
|||
rw [@conjTranspose_nonsing_inv]
|
||||
simp
|
||||
|
||||
/-- The alt-right-right unit `δ^{dot a}_{dot a}` as an element of
|
||||
/-- The alt-right-right unit `δ_{dot a}^{dot a}` as an element of
|
||||
`(rightHanded ⊗ altRightHanded).V`. -/
|
||||
def altRightRightUnitVal : (altRightHanded ⊗ rightHanded).V :=
|
||||
altRightRightToMatrix.symm 1
|
||||
|
||||
/-- The alt-right-right unit `δ^{dot a}_{dot a}` as a morphism
|
||||
/-- The alt-right-right unit `δ_{dot a}^{dot a}` as a morphism
|
||||
`𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ altRightHanded ⊗ rightHanded`, manifesting
|
||||
the invariance under the `SL(2,ℂ)` action. -/
|
||||
def altRightRightUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ altRightHanded ⊗ rightHanded where
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue