refactor: Lint
This commit is contained in:
parent
833a570ce8
commit
7d983f5b4b
5 changed files with 23 additions and 20 deletions
|
@ -130,7 +130,6 @@ lemma coMetric_apply_one : coMetric.hom (1 : ℂ) = coMetricVal := by
|
||||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||||
coMetric, AddHom.toFun_eq_coe, AddHom.coe_mk, one_smul]
|
coMetric, AddHom.toFun_eq_coe, AddHom.coe_mk, one_smul]
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## Contraction of metrics
|
## Contraction of metrics
|
||||||
|
|
|
@ -145,10 +145,11 @@ lemma contr_contrCoUnit (x : complexCo) :
|
||||||
Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
|
Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
|
||||||
Finset.sum_singleton, Fin.sum_univ_three, tmul_add, add_tmul, smul_tmul, tmul_smul, map_add,
|
Finset.sum_singleton, Fin.sum_univ_three, tmul_add, add_tmul, smul_tmul, tmul_smul, map_add,
|
||||||
_root_.map_smul]
|
_root_.map_smul]
|
||||||
have h1' (x y : CoeSort.coe complexCo) (z : CoeSort.coe complexContr) :
|
have h1' (x y : CoeSort.coe complexCo) (z : CoeSort.coe complexContr) :
|
||||||
(α_ complexCo.V complexContr.V complexCo.V).inv (x ⊗ₜ[ℂ] z ⊗ₜ[ℂ] y) = (x ⊗ₜ[ℂ] z) ⊗ₜ[ℂ] y := rfl
|
(α_ complexCo.V complexContr.V complexCo.V).inv (x ⊗ₜ[ℂ] z ⊗ₜ[ℂ] y) = (x ⊗ₜ[ℂ] z) ⊗ₜ[ℂ] y := rfl
|
||||||
repeat rw [h1']
|
repeat rw [h1']
|
||||||
have h1'' ( y : CoeSort.coe complexCo) (z : CoeSort.coe complexCo ⊗[ℂ] CoeSort.coe complexContr) :
|
have h1'' (y : CoeSort.coe complexCo)
|
||||||
|
(z : CoeSort.coe complexCo ⊗[ℂ] CoeSort.coe complexContr) :
|
||||||
(coContrContraction.hom ▷ complexCo.V) (z ⊗ₜ[ℂ] y) = (coContrContraction.hom z) ⊗ₜ[ℂ] y := rfl
|
(coContrContraction.hom ▷ complexCo.V) (z ⊗ₜ[ℂ] y) = (coContrContraction.hom z) ⊗ₜ[ℂ] y := rfl
|
||||||
repeat rw (config := { transparency := .instances }) [h1'']
|
repeat rw (config := { transparency := .instances }) [h1'']
|
||||||
repeat rw [coContrContraction_basis']
|
repeat rw [coContrContraction_basis']
|
||||||
|
@ -178,11 +179,14 @@ lemma contr_coContrUnit (x : complexContr) :
|
||||||
Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
|
Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
|
||||||
Finset.sum_singleton, Fin.sum_univ_three, tmul_add, add_tmul, smul_tmul, tmul_smul, map_add,
|
Finset.sum_singleton, Fin.sum_univ_three, tmul_add, add_tmul, smul_tmul, tmul_smul, map_add,
|
||||||
_root_.map_smul]
|
_root_.map_smul]
|
||||||
have h1' (x y : CoeSort.coe complexContr) (z : CoeSort.coe complexCo) :
|
have h1' (x y : CoeSort.coe complexContr) (z : CoeSort.coe complexCo) :
|
||||||
(α_ complexContr.V complexCo.V complexContr.V).inv (x ⊗ₜ[ℂ] z ⊗ₜ[ℂ] y) = (x ⊗ₜ[ℂ] z) ⊗ₜ[ℂ] y := rfl
|
(α_ complexContr.V complexCo.V complexContr.V).inv
|
||||||
|
(x ⊗ₜ[ℂ] z ⊗ₜ[ℂ] y) = (x ⊗ₜ[ℂ] z) ⊗ₜ[ℂ] y := rfl
|
||||||
repeat rw [h1']
|
repeat rw [h1']
|
||||||
have h1'' ( y : CoeSort.coe complexContr) (z : CoeSort.coe complexContr ⊗[ℂ] CoeSort.coe complexCo) :
|
have h1'' (y : CoeSort.coe complexContr)
|
||||||
(contrCoContraction.hom ▷ complexContr.V) (z ⊗ₜ[ℂ] y) = (contrCoContraction.hom z) ⊗ₜ[ℂ] y := rfl
|
(z : CoeSort.coe complexContr ⊗[ℂ] CoeSort.coe complexCo) :
|
||||||
|
(contrCoContraction.hom ▷ complexContr.V) (z ⊗ₜ[ℂ] y) =
|
||||||
|
(contrCoContraction.hom z) ⊗ₜ[ℂ] y := rfl
|
||||||
repeat rw (config := { transparency := .instances }) [h1'']
|
repeat rw (config := { transparency := .instances }) [h1'']
|
||||||
repeat rw [contrCoContraction_basis']
|
repeat rw [contrCoContraction_basis']
|
||||||
simp only [Fin.isValue, Action.instMonoidalCategory_tensorUnit_V, ↓reduceIte, reduceCtorEq,
|
simp only [Fin.isValue, Action.instMonoidalCategory_tensorUnit_V, ↓reduceIte, reduceCtorEq,
|
||||||
|
@ -199,7 +203,6 @@ lemma contr_coContrUnit (x : complexContr) :
|
||||||
|
|
||||||
-/
|
-/
|
||||||
|
|
||||||
|
|
||||||
open CategoryTheory
|
open CategoryTheory
|
||||||
|
|
||||||
lemma contrCoUnit_symm :
|
lemma contrCoUnit_symm :
|
||||||
|
|
|
@ -352,8 +352,8 @@ lemma rightAltContraction_apply_metric : (β_ rightHanded altRightHanded).hom.ho
|
||||||
((rightHanded.V ◁ rightAltContraction.hom ▷ altRightHanded.V) (((rightHanded.V ◁
|
((rightHanded.V ◁ rightAltContraction.hom ▷ altRightHanded.V) (((rightHanded.V ◁
|
||||||
(α_ rightHanded.V altRightHanded.V altRightHanded.V).inv)
|
(α_ rightHanded.V altRightHanded.V altRightHanded.V).inv)
|
||||||
((α_ rightHanded.V rightHanded.V (altRightHanded.V ⊗ altRightHanded.V)).hom
|
((α_ rightHanded.V rightHanded.V (altRightHanded.V ⊗ altRightHanded.V)).hom
|
||||||
((x1 ⊗ₜ[ℂ] x2) ⊗ₜ[ℂ] y1 ⊗ₜ[ℂ] y2)))))
|
((x1 ⊗ₜ[ℂ] x2) ⊗ₜ[ℂ] y1 ⊗ₜ[ℂ] y2))))) = x1 ⊗ₜ[ℂ] ((λ_ altRightHanded.V).hom
|
||||||
= x1 ⊗ₜ[ℂ] ((λ_ altRightHanded.V).hom ((rightAltContraction.hom (x2 ⊗ₜ[ℂ] y1)) ⊗ₜ[ℂ] y2)) := rfl
|
((rightAltContraction.hom (x2 ⊗ₜ[ℂ] y1)) ⊗ₜ[ℂ] y2)) := rfl
|
||||||
repeat rw (config := { transparency := .instances }) [h1]
|
repeat rw (config := { transparency := .instances }) [h1]
|
||||||
repeat rw [rightAltContraction_basis]
|
repeat rw [rightAltContraction_basis]
|
||||||
simp only [Fin.isValue, Fin.val_one, Fin.val_zero, one_ne_zero, ↓reduceIte, zero_tmul, map_zero,
|
simp only [Fin.isValue, Fin.val_one, Fin.val_zero, one_ne_zero, ↓reduceIte, zero_tmul, map_zero,
|
||||||
|
@ -374,7 +374,8 @@ lemma altRightContraction_apply_metric : (β_ altRightHanded rightHanded).hom.ho
|
||||||
rw [rightMetricVal_expand_tmul, altRightMetricVal_expand_tmul]
|
rw [rightMetricVal_expand_tmul, altRightMetricVal_expand_tmul]
|
||||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||||
Fin.isValue, tmul_add, tmul_neg, sub_tmul, map_add, map_neg, map_sub]
|
Fin.isValue, tmul_add, tmul_neg, sub_tmul, map_add, map_neg, map_sub]
|
||||||
have h1 (x1 x2 : altRightHanded) (y1 y2 : rightHanded) : (altRightHanded.V ◁ (λ_ rightHanded.V).hom)
|
have h1 (x1 x2 : altRightHanded) (y1 y2 : rightHanded) :
|
||||||
|
(altRightHanded.V ◁ (λ_ rightHanded.V).hom)
|
||||||
((altRightHanded.V ◁ altRightContraction.hom ▷ rightHanded.V) (((altRightHanded.V ◁
|
((altRightHanded.V ◁ altRightContraction.hom ▷ rightHanded.V) (((altRightHanded.V ◁
|
||||||
(α_ altRightHanded.V rightHanded.V rightHanded.V).inv)
|
(α_ altRightHanded.V rightHanded.V rightHanded.V).inv)
|
||||||
((α_ altRightHanded.V altRightHanded.V (rightHanded.V ⊗ rightHanded.V)).hom
|
((α_ altRightHanded.V altRightHanded.V (rightHanded.V ⊗ rightHanded.V)).hom
|
||||||
|
|
|
@ -251,7 +251,7 @@ lemma contr_altLeftLeftUnit (x : leftHanded) :
|
||||||
simp only [Fin.isValue, one_smul]
|
simp only [Fin.isValue, one_smul]
|
||||||
|
|
||||||
/-- Contraction on the right with `leftAltLeftUnit` does nothing. -/
|
/-- Contraction on the right with `leftAltLeftUnit` does nothing. -/
|
||||||
lemma contr_leftAltLeftUnit (x : altLeftHanded) :
|
lemma contr_leftAltLeftUnit (x : altLeftHanded) :
|
||||||
(λ_ altLeftHanded).hom.hom
|
(λ_ altLeftHanded).hom.hom
|
||||||
(((altLeftContraction) ▷ altLeftHanded).hom
|
(((altLeftContraction) ▷ altLeftHanded).hom
|
||||||
((α_ _ _ altLeftHanded).inv.hom
|
((α_ _ _ altLeftHanded).inv.hom
|
||||||
|
@ -333,29 +333,29 @@ lemma contr_rightAltRightUnit (x : altRightHanded) :
|
||||||
open CategoryTheory
|
open CategoryTheory
|
||||||
|
|
||||||
lemma altLeftLeftUnit_symm :
|
lemma altLeftLeftUnit_symm :
|
||||||
(altLeftLeftUnit.hom (1 : ℂ)) = (altLeftHanded ◁ 𝟙 _).hom ((β_ leftHanded altLeftHanded ).hom.hom
|
(altLeftLeftUnit.hom (1 : ℂ)) = (altLeftHanded ◁ 𝟙 _).hom
|
||||||
(leftAltLeftUnit.hom (1 : ℂ))) := by
|
((β_ leftHanded altLeftHanded).hom.hom (leftAltLeftUnit.hom (1 : ℂ))) := by
|
||||||
rw [altLeftLeftUnit_apply_one, altLeftLeftUnitVal_expand_tmul]
|
rw [altLeftLeftUnit_apply_one, altLeftLeftUnitVal_expand_tmul]
|
||||||
rw [leftAltLeftUnit_apply_one, leftAltLeftUnitVal_expand_tmul]
|
rw [leftAltLeftUnit_apply_one, leftAltLeftUnitVal_expand_tmul]
|
||||||
rfl
|
rfl
|
||||||
|
|
||||||
lemma leftAltLeftUnit_symm :
|
lemma leftAltLeftUnit_symm :
|
||||||
(leftAltLeftUnit.hom (1 : ℂ)) = (leftHanded ◁ 𝟙 _).hom ((β_ altLeftHanded leftHanded ).hom.hom
|
(leftAltLeftUnit.hom (1 : ℂ)) = (leftHanded ◁ 𝟙 _).hom ((β_ altLeftHanded leftHanded).hom.hom
|
||||||
(altLeftLeftUnit.hom (1 : ℂ))) := by
|
(altLeftLeftUnit.hom (1 : ℂ))) := by
|
||||||
rw [altLeftLeftUnit_apply_one, altLeftLeftUnitVal_expand_tmul]
|
rw [altLeftLeftUnit_apply_one, altLeftLeftUnitVal_expand_tmul]
|
||||||
rw [leftAltLeftUnit_apply_one, leftAltLeftUnitVal_expand_tmul]
|
rw [leftAltLeftUnit_apply_one, leftAltLeftUnitVal_expand_tmul]
|
||||||
rfl
|
rfl
|
||||||
|
|
||||||
lemma altRightRightUnit_symm :
|
lemma altRightRightUnit_symm :
|
||||||
(altRightRightUnit.hom (1 : ℂ)) = (altRightHanded ◁ 𝟙 _).hom ((β_ rightHanded altRightHanded ).hom.hom
|
(altRightRightUnit.hom (1 : ℂ)) = (altRightHanded ◁ 𝟙 _).hom
|
||||||
(rightAltRightUnit.hom (1 : ℂ))) := by
|
((β_ rightHanded altRightHanded).hom.hom (rightAltRightUnit.hom (1 : ℂ))) := by
|
||||||
rw [altRightRightUnit_apply_one, altRightRightUnitVal_expand_tmul]
|
rw [altRightRightUnit_apply_one, altRightRightUnitVal_expand_tmul]
|
||||||
rw [rightAltRightUnit_apply_one, rightAltRightUnitVal_expand_tmul]
|
rw [rightAltRightUnit_apply_one, rightAltRightUnitVal_expand_tmul]
|
||||||
rfl
|
rfl
|
||||||
|
|
||||||
lemma rightAltRightUnit_symm :
|
lemma rightAltRightUnit_symm :
|
||||||
(rightAltRightUnit.hom (1 : ℂ)) = (rightHanded ◁ 𝟙 _).hom ((β_ altRightHanded rightHanded ).hom.hom
|
(rightAltRightUnit.hom (1 : ℂ)) = (rightHanded ◁ 𝟙 _).hom
|
||||||
(altRightRightUnit.hom (1 : ℂ))) := by
|
((β_ altRightHanded rightHanded).hom.hom (altRightRightUnit.hom (1 : ℂ))) := by
|
||||||
rw [altRightRightUnit_apply_one, altRightRightUnitVal_expand_tmul]
|
rw [altRightRightUnit_apply_one, altRightRightUnitVal_expand_tmul]
|
||||||
rw [rightAltRightUnit_apply_one, rightAltRightUnitVal_expand_tmul]
|
rw [rightAltRightUnit_apply_one, rightAltRightUnitVal_expand_tmul]
|
||||||
rfl
|
rfl
|
||||||
|
|
|
@ -183,7 +183,7 @@ def complexLorentzTensor : TensorSpecies where
|
||||||
| Color.downR => Fermion.rightAltRightUnit_symm
|
| Color.downR => Fermion.rightAltRightUnit_symm
|
||||||
| Color.up => Lorentz.coContrUnit_symm
|
| Color.up => Lorentz.coContrUnit_symm
|
||||||
| Color.down => Lorentz.contrCoUnit_symm
|
| Color.down => Lorentz.contrCoUnit_symm
|
||||||
contr_metric := fun c =>
|
contr_metric := fun c =>
|
||||||
match c with
|
match c with
|
||||||
| Color.upL => by simpa using Fermion.leftAltContraction_apply_metric
|
| Color.upL => by simpa using Fermion.leftAltContraction_apply_metric
|
||||||
| Color.downL => by simpa using Fermion.altLeftContraction_apply_metric
|
| Color.downL => by simpa using Fermion.altLeftContraction_apply_metric
|
||||||
|
|
Loading…
Add table
Add a link
Reference in a new issue