feat: Add symm of unit for complex lorentz
This commit is contained in:
parent
942ee12e60
commit
8c584431c4
3 changed files with 66 additions and 0 deletions
|
@ -193,6 +193,28 @@ lemma contr_coContrUnit (x : complexContr) :
|
||||||
simp only [Fin.isValue, one_smul]
|
simp only [Fin.isValue, one_smul]
|
||||||
repeat rw [add_assoc]
|
repeat rw [add_assoc]
|
||||||
|
|
||||||
|
/-!
|
||||||
|
|
||||||
|
## Symmetry properties of the units
|
||||||
|
|
||||||
|
-/
|
||||||
|
|
||||||
|
|
||||||
|
open CategoryTheory
|
||||||
|
|
||||||
|
lemma contrCoUnit_symm :
|
||||||
|
(contrCoUnit.hom (1 : ℂ)) = (complexContr ◁ 𝟙 _).hom ((β_ complexCo complexContr).hom.hom
|
||||||
|
(coContrUnit.hom (1 : ℂ))) := by
|
||||||
|
rw [contrCoUnit_apply_one, contrCoUnitVal_expand_tmul]
|
||||||
|
rw [coContrUnit_apply_one, coContrUnitVal_expand_tmul]
|
||||||
|
rfl
|
||||||
|
|
||||||
|
lemma coContrUnit_symm :
|
||||||
|
(coContrUnit.hom (1 : ℂ)) = (complexCo ◁ 𝟙 _).hom ((β_ complexContr complexCo).hom.hom
|
||||||
|
(contrCoUnit.hom (1 : ℂ))) := by
|
||||||
|
rw [coContrUnit_apply_one, coContrUnitVal_expand_tmul]
|
||||||
|
rw [contrCoUnit_apply_one, contrCoUnitVal_expand_tmul]
|
||||||
|
rfl
|
||||||
|
|
||||||
end Lorentz
|
end Lorentz
|
||||||
end
|
end
|
||||||
|
|
|
@ -325,5 +325,40 @@ lemma contr_rightAltRightUnit (x : altRightHanded) :
|
||||||
erw [TensorProduct.lid_tmul, TensorProduct.lid_tmul]
|
erw [TensorProduct.lid_tmul, TensorProduct.lid_tmul]
|
||||||
simp only [Fin.isValue, one_smul]
|
simp only [Fin.isValue, one_smul]
|
||||||
|
|
||||||
|
/-!
|
||||||
|
|
||||||
|
## Symmetry properties of the units
|
||||||
|
|
||||||
|
-/
|
||||||
|
open CategoryTheory
|
||||||
|
|
||||||
|
lemma altLeftLeftUnit_symm :
|
||||||
|
(altLeftLeftUnit.hom (1 : ℂ)) = (altLeftHanded ◁ 𝟙 _).hom ((β_ leftHanded altLeftHanded ).hom.hom
|
||||||
|
(leftAltLeftUnit.hom (1 : ℂ))) := by
|
||||||
|
rw [altLeftLeftUnit_apply_one, altLeftLeftUnitVal_expand_tmul]
|
||||||
|
rw [leftAltLeftUnit_apply_one, leftAltLeftUnitVal_expand_tmul]
|
||||||
|
rfl
|
||||||
|
|
||||||
|
lemma leftAltLeftUnit_symm :
|
||||||
|
(leftAltLeftUnit.hom (1 : ℂ)) = (leftHanded ◁ 𝟙 _).hom ((β_ altLeftHanded leftHanded ).hom.hom
|
||||||
|
(altLeftLeftUnit.hom (1 : ℂ))) := by
|
||||||
|
rw [altLeftLeftUnit_apply_one, altLeftLeftUnitVal_expand_tmul]
|
||||||
|
rw [leftAltLeftUnit_apply_one, leftAltLeftUnitVal_expand_tmul]
|
||||||
|
rfl
|
||||||
|
|
||||||
|
lemma altRightRightUnit_symm :
|
||||||
|
(altRightRightUnit.hom (1 : ℂ)) = (altRightHanded ◁ 𝟙 _).hom ((β_ rightHanded altRightHanded ).hom.hom
|
||||||
|
(rightAltRightUnit.hom (1 : ℂ))) := by
|
||||||
|
rw [altRightRightUnit_apply_one, altRightRightUnitVal_expand_tmul]
|
||||||
|
rw [rightAltRightUnit_apply_one, rightAltRightUnitVal_expand_tmul]
|
||||||
|
rfl
|
||||||
|
|
||||||
|
lemma rightAltRightUnit_symm :
|
||||||
|
(rightAltRightUnit.hom (1 : ℂ)) = (rightHanded ◁ 𝟙 _).hom ((β_ altRightHanded rightHanded ).hom.hom
|
||||||
|
(altRightRightUnit.hom (1 : ℂ))) := by
|
||||||
|
rw [altRightRightUnit_apply_one, altRightRightUnitVal_expand_tmul]
|
||||||
|
rw [rightAltRightUnit_apply_one, rightAltRightUnitVal_expand_tmul]
|
||||||
|
rfl
|
||||||
|
|
||||||
end
|
end
|
||||||
end Fermion
|
end Fermion
|
||||||
|
|
|
@ -175,6 +175,15 @@ def complexLorentzTensor : TensorSpecies where
|
||||||
| Color.downR => Fermion.contr_rightAltRightUnit
|
| Color.downR => Fermion.contr_rightAltRightUnit
|
||||||
| Color.up => Lorentz.contr_coContrUnit
|
| Color.up => Lorentz.contr_coContrUnit
|
||||||
| Color.down => Lorentz.contr_contrCoUnit
|
| Color.down => Lorentz.contr_contrCoUnit
|
||||||
|
unit_symm := fun c =>
|
||||||
|
match c with
|
||||||
|
| Color.upL => Fermion.altLeftLeftUnit_symm
|
||||||
|
| Color.downL => Fermion.leftAltLeftUnit_symm
|
||||||
|
| Color.upR => Fermion.altRightRightUnit_symm
|
||||||
|
| Color.downR => Fermion.rightAltRightUnit_symm
|
||||||
|
| Color.up => Lorentz.coContrUnit_symm
|
||||||
|
| Color.down => Lorentz.contrCoUnit_symm
|
||||||
|
contr_metric := by sorry
|
||||||
|
|
||||||
instance : DecidableEq complexLorentzTensor.C := Fermion.instDecidableEqColor
|
instance : DecidableEq complexLorentzTensor.C := Fermion.instDecidableEqColor
|
||||||
|
|
||||||
|
|
Loading…
Add table
Add a link
Reference in a new issue