/- 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 HepLean.Lorentz.Weyl.Unit /-! # Metrics of Weyl fermions We define the metrics for Weyl fermions, often denoted `Ξ΅` in the literature. These allow us to go from left-handed to alt-left-handed Weyl fermions and back, and from right-handed to alt-right-handed Weyl fermions and back. -/ namespace Fermion noncomputable section open Matrix open MatrixGroups open Complex open TensorProduct open CategoryTheory.MonoidalCategory /-- The raw `2x2` matrix corresponding to the metric for fermions. -/ def metricRaw : Matrix (Fin 2) (Fin 2) β„‚ := !![0, 1; -1, 0] /-- Multiplying an element of `SL(2, β„‚)` on the left with the metric `𝓔` is equivalent to multiplying the inverse-transpose of that element on the right with the metric. -/ lemma comm_metricRaw (M : SL(2,β„‚)) : M.1 * metricRaw = metricRaw * (M.1⁻¹)α΅€ := by rw [metricRaw] rw [Lorentz.SL2C.inverse_coe, eta_fin_two M.1] rw [SpecialLinearGroup.coe_inv, Matrix.adjugate_fin_two, Matrix.mul_fin_two, eta_fin_two !![M.1 1 1, -M.1 0 1; -M.1 1 0, M.1 0 0]α΅€] simp only [Fin.isValue, mul_zero, mul_neg, mul_one, zero_add, add_zero, transpose_apply, of_apply, cons_val', cons_val_zero, empty_val', cons_val_fin_one, cons_val_one, head_fin_const, head_cons, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons, zero_smul, tail_cons, one_smul, empty_vecMul, neg_smul, neg_cons, neg_neg, neg_empty, empty_mul, Equiv.symm_apply_apply] lemma metricRaw_comm (M : SL(2,β„‚)) : metricRaw * M.1 = (M.1⁻¹)α΅€ * metricRaw := by rw [metricRaw] rw [Lorentz.SL2C.inverse_coe, eta_fin_two M.1] rw [SpecialLinearGroup.coe_inv, Matrix.adjugate_fin_two, Matrix.mul_fin_two, eta_fin_two !![M.1 1 1, -M.1 0 1; -M.1 1 0, M.1 0 0]α΅€] simp only [Fin.isValue, zero_mul, one_mul, zero_add, neg_mul, add_zero, transpose_apply, of_apply, cons_val', cons_val_zero, empty_val', cons_val_fin_one, cons_val_one, head_fin_const, head_cons, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons, smul_cons, smul_eq_mul, mul_zero, mul_one, smul_empty, tail_cons, neg_smul, mul_neg, neg_cons, neg_neg, neg_zero, neg_empty, empty_vecMul, add_cons, empty_add_empty, empty_mul, Equiv.symm_apply_apply] lemma star_comm_metricRaw (M : SL(2,β„‚)) : M.1.map star * metricRaw = metricRaw * ((M.1)⁻¹)α΄΄ := by rw [metricRaw] rw [Lorentz.SL2C.inverse_coe, eta_fin_two M.1] rw [SpecialLinearGroup.coe_inv, Matrix.adjugate_fin_two, eta_fin_two !![M.1 1 1, -M.1 0 1; -M.1 1 0, M.1 0 0]α΄΄] rw [eta_fin_two (!![M.1 0 0, M.1 0 1; M.1 1 0, M.1 1 1].map star)] simp lemma metricRaw_comm_star (M : SL(2,β„‚)) : metricRaw * M.1.map star = ((M.1)⁻¹)α΄΄ * metricRaw := by rw [metricRaw] rw [Lorentz.SL2C.inverse_coe, eta_fin_two M.1] rw [SpecialLinearGroup.coe_inv, Matrix.adjugate_fin_two, eta_fin_two !![M.1 1 1, -M.1 0 1; -M.1 1 0, M.1 0 0]α΄΄] rw [eta_fin_two (!![M.1 0 0, M.1 0 1; M.1 1 0, M.1 1 1].map star)] simp /-- The metric `Ρᡃᡃ` as an element of `(leftHanded βŠ— leftHanded).V`. -/ def leftMetricVal : (leftHanded βŠ— leftHanded).V := leftLeftToMatrix.symm (- metricRaw) /-- Expansion of `leftMetricVal` into the left basis. -/ lemma leftMetricVal_expand_tmul : leftMetricVal = - leftBasis 0 βŠ—β‚œ[β„‚] leftBasis 1 + leftBasis 1 βŠ—β‚œ[β„‚] leftBasis 0 := by simp only [Action.instMonoidalCategory_tensorObj_V, leftMetricVal, Fin.isValue] erw [leftLeftToMatrix_symm_expand_tmul] simp only [metricRaw, neg_apply, of_apply, cons_val', empty_val', cons_val_fin_one, neg_smul, Finset.sum_neg_distrib, Fin.sum_univ_two, Fin.isValue, cons_val_zero, cons_val_one, head_cons, neg_add_rev, one_smul, zero_smul, neg_zero, add_zero, head_fin_const, neg_neg, zero_add] /-- The metric `Ρᡃᡃ` as a morphism `πŸ™_ (Rep β„‚ SL(2,β„‚)) ⟢ leftHanded βŠ— leftHanded`, making manifest its invariance under the action of `SL(2,β„‚)`. -/ def leftMetric : πŸ™_ (Rep β„‚ SL(2,β„‚)) ⟢ leftHanded βŠ— leftHanded where hom := ModuleCat.ofHom { toFun := fun a => let a' : β„‚ := a a' β€’ leftMetricVal, map_add' := fun x y => by simp only [add_smul] map_smul' := fun m x => by simp only [smul_smul] rfl} comm M := by refine ModuleCat.hom_ext ?_ refine LinearMap.ext fun x : β„‚ => ?_ simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V, Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.hom_comp, Function.comp_apply] change x β€’ leftMetricVal = (TensorProduct.map (leftHanded.ρ M) (leftHanded.ρ M)) (x β€’ leftMetricVal) simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul] apply congrArg simp only [Action.instMonoidalCategory_tensorObj_V, leftMetricVal, map_neg, neg_inj] erw [leftLeftToMatrix_ρ_symm] apply congrArg rw [comm_metricRaw, mul_assoc, ← @transpose_mul] simp only [SpecialLinearGroup.det_coe, isUnit_iff_ne_zero, ne_eq, one_ne_zero, not_false_eq_true, mul_nonsing_inv, transpose_one, mul_one] lemma leftMetric_apply_one : leftMetric.hom (1 : β„‚) = leftMetricVal := by change leftMetric.hom.hom.toFun (1 : β„‚) = leftMetricVal simp only [leftMetric, one_smul] /-- The metric `Ρₐₐ` as an element of `(altLeftHanded βŠ— altLeftHanded).V`. -/ def altLeftMetricVal : (altLeftHanded βŠ— altLeftHanded).V := altLeftaltLeftToMatrix.symm metricRaw /-- Expansion of `altLeftMetricVal` into the left basis. -/ lemma altLeftMetricVal_expand_tmul : altLeftMetricVal = altLeftBasis 0 βŠ—β‚œ[β„‚] altLeftBasis 1 - altLeftBasis 1 βŠ—β‚œ[β„‚] altLeftBasis 0 := by simp only [Action.instMonoidalCategory_tensorObj_V, altLeftMetricVal, Fin.isValue] erw [altLeftaltLeftToMatrix_symm_expand_tmul] simp only [metricRaw, neg_apply, of_apply, cons_val', empty_val', cons_val_fin_one, neg_smul, Finset.sum_neg_distrib, Fin.sum_univ_two, Fin.isValue, cons_val_zero, cons_val_one, head_cons, neg_add_rev, one_smul, zero_smul, neg_zero, add_zero, head_fin_const, neg_neg, zero_add] rfl /-- The metric `Ρₐₐ` as a morphism `πŸ™_ (Rep β„‚ SL(2,β„‚)) ⟢ altLeftHanded βŠ— altLeftHanded`, making manifest its invariance under the action of `SL(2,β„‚)`. -/ def altLeftMetric : πŸ™_ (Rep β„‚ SL(2,β„‚)) ⟢ altLeftHanded βŠ— altLeftHanded where hom := ModuleCat.ofHom { toFun := fun a => let a' : β„‚ := a a' β€’ altLeftMetricVal, map_add' := fun x y => by simp only [add_smul] map_smul' := fun m x => by simp only [smul_smul] rfl} comm M := by refine ModuleCat.hom_ext ?_ refine LinearMap.ext fun x : β„‚ => ?_ simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V, Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.hom_comp, Function.comp_apply] change x β€’ altLeftMetricVal = (TensorProduct.map (altLeftHanded.ρ M) (altLeftHanded.ρ M)) (x β€’ altLeftMetricVal) simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul] apply congrArg simp only [Action.instMonoidalCategory_tensorObj_V, altLeftMetricVal] erw [altLeftaltLeftToMatrix_ρ_symm] apply congrArg rw [← metricRaw_comm, mul_assoc] simp only [SpecialLinearGroup.det_coe, isUnit_iff_ne_zero, ne_eq, one_ne_zero, not_false_eq_true, mul_nonsing_inv, mul_one] lemma altLeftMetric_apply_one : altLeftMetric.hom (1 : β„‚) = altLeftMetricVal := by change altLeftMetric.hom.hom.toFun (1 : β„‚) = altLeftMetricVal simp only [altLeftMetric, one_smul] /-- The metric `Ξ΅^{dot a}^{dot a}` as an element of `(rightHanded βŠ— rightHanded).V`. -/ def rightMetricVal : (rightHanded βŠ— rightHanded).V := rightRightToMatrix.symm (- metricRaw) /-- Expansion of `rightMetricVal` into the left basis. -/ lemma rightMetricVal_expand_tmul : rightMetricVal = - rightBasis 0 βŠ—β‚œ[β„‚] rightBasis 1 + rightBasis 1 βŠ—β‚œ[β„‚] rightBasis 0 := by simp only [Action.instMonoidalCategory_tensorObj_V, rightMetricVal, Fin.isValue] erw [rightRightToMatrix_symm_expand_tmul] simp only [metricRaw, neg_apply, of_apply, cons_val', empty_val', cons_val_fin_one, neg_smul, Finset.sum_neg_distrib, Fin.sum_univ_two, Fin.isValue, cons_val_zero, cons_val_one, head_cons, neg_add_rev, one_smul, zero_smul, neg_zero, add_zero, head_fin_const, neg_neg, zero_add] /-- The metric `Ξ΅^{dot a}^{dot a}` as a morphism `πŸ™_ (Rep β„‚ SL(2,β„‚)) ⟢ rightHanded βŠ— rightHanded`, making manifest its invariance under the action of `SL(2,β„‚)`. -/ def rightMetric : πŸ™_ (Rep β„‚ SL(2,β„‚)) ⟢ rightHanded βŠ— rightHanded where hom := ModuleCat.ofHom { toFun := fun a => let a' : β„‚ := a a' β€’ rightMetricVal, map_add' := fun x y => by simp only [add_smul] map_smul' := fun m x => by simp only [smul_smul] rfl} comm M := by refine ModuleCat.hom_ext ?_ refine LinearMap.ext fun x : β„‚ => ?_ simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V, Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.hom_comp, Function.comp_apply] change x β€’ rightMetricVal = (TensorProduct.map (rightHanded.ρ M) (rightHanded.ρ M)) (x β€’ rightMetricVal) simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul] apply congrArg simp only [Action.instMonoidalCategory_tensorObj_V, rightMetricVal, map_neg, neg_inj] trans rightRightToMatrix.symm ((M.1).map star * metricRaw * ((M.1).map star)α΅€) Β· apply congrArg rw [star_comm_metricRaw, mul_assoc] have h1 : ((M.1)⁻¹ᴴ * ((M.1).map star)α΅€) = 1 := by trans (M.1)⁻¹ᴴ * ((M.1))α΄΄ Β· rfl rw [← @conjTranspose_mul] simp only [SpecialLinearGroup.det_coe, isUnit_iff_ne_zero, ne_eq, one_ne_zero, not_false_eq_true, mul_nonsing_inv, conjTranspose_one] rw [h1] simp Β· rw [← rightRightToMatrix_ρ_symm metricRaw M] rfl lemma rightMetric_apply_one : rightMetric.hom (1 : β„‚) = rightMetricVal := by change rightMetric.hom.hom.toFun (1 : β„‚) = rightMetricVal simp only [rightMetric, one_smul] /-- The metric `Ξ΅_{dot a}_{dot a}` as an element of `(altRightHanded βŠ— altRightHanded).V`. -/ def altRightMetricVal : (altRightHanded βŠ— altRightHanded).V := altRightAltRightToMatrix.symm (metricRaw) /-- Expansion of `rightMetricVal` into the left basis. -/ lemma altRightMetricVal_expand_tmul : altRightMetricVal = altRightBasis 0 βŠ—β‚œ[β„‚] altRightBasis 1 - altRightBasis 1 βŠ—β‚œ[β„‚] altRightBasis 0 := by simp only [Action.instMonoidalCategory_tensorObj_V, altRightMetricVal, Fin.isValue] erw [altRightAltRightToMatrix_symm_expand_tmul] simp only [metricRaw, neg_apply, of_apply, cons_val', empty_val', cons_val_fin_one, neg_smul, Finset.sum_neg_distrib, Fin.sum_univ_two, Fin.isValue, cons_val_zero, cons_val_one, head_cons, neg_add_rev, one_smul, zero_smul, neg_zero, add_zero, head_fin_const, neg_neg, zero_add] rfl /-- The metric `Ξ΅_{dot a}_{dot a}` as a morphism `πŸ™_ (Rep β„‚ SL(2,β„‚)) ⟢ altRightHanded βŠ— altRightHanded`, making manifest its invariance under the action of `SL(2,β„‚)`. -/ def altRightMetric : πŸ™_ (Rep β„‚ SL(2,β„‚)) ⟢ altRightHanded βŠ— altRightHanded where hom := ModuleCat.ofHom { toFun := fun a => let a' : β„‚ := a a' β€’ altRightMetricVal, map_add' := fun x y => by simp only [add_smul] map_smul' := fun m x => by simp only [smul_smul] rfl} comm M := by refine ModuleCat.hom_ext ?_ refine LinearMap.ext fun x : β„‚ => ?_ simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V, Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.hom_comp, Function.comp_apply] change x β€’ altRightMetricVal = (TensorProduct.map (altRightHanded.ρ M) (altRightHanded.ρ M)) (x β€’ altRightMetricVal) simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul] apply congrArg simp only [Action.instMonoidalCategory_tensorObj_V] trans altRightAltRightToMatrix.symm (((M.1)⁻¹).conjTranspose * metricRaw * (((M.1)⁻¹).conjTranspose)α΅€) Β· rw [altRightMetricVal] apply congrArg rw [← metricRaw_comm_star, mul_assoc] have h1 : ((M.1).map star * (M.1)⁻¹ᴴᡀ) = 1 := by refine transpose_eq_one.mp ?_ rw [@transpose_mul] simp only [transpose_transpose, RCLike.star_def] change (M.1)⁻¹ᴴ * (M.1)α΄΄ = 1 rw [← @conjTranspose_mul] simp rw [h1, mul_one] Β· rw [← altRightAltRightToMatrix_ρ_symm metricRaw M] rfl lemma altRightMetric_apply_one : altRightMetric.hom (1 : β„‚) = altRightMetricVal := by change altRightMetric.hom.hom.toFun (1 : β„‚) = altRightMetricVal simp only [altRightMetric, one_smul] /-! ## Contraction of metrics -/ lemma leftAltContraction_apply_metric : (Ξ²_ leftHanded altLeftHanded).hom.hom ((leftHanded.V ◁ (Ξ»_ altLeftHanded.V).hom) ((leftHanded.V ◁ leftAltContraction.hom β–· altLeftHanded.V) ((leftHanded.V ◁ (Ξ±_ leftHanded.V altLeftHanded.V altLeftHanded.V).inv) ((Ξ±_ leftHanded.V leftHanded.V (altLeftHanded.V βŠ— altLeftHanded.V)).hom (leftMetric.hom (1 : β„‚) βŠ—β‚œ[β„‚] altLeftMetric.hom (1 : β„‚)))))) = altLeftLeftUnit.hom (1 : β„‚) := by rw [leftMetric_apply_one, altLeftMetric_apply_one] rw [leftMetricVal_expand_tmul, altLeftMetricVal_expand_tmul] simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V, Fin.isValue, tmul_sub, add_tmul, neg_tmul, map_sub, map_add, map_neg] have h1 (x1 x2 : leftHanded) (y1 y2 :altLeftHanded) : (leftHanded.V ◁ (Ξ»_ altLeftHanded.V).hom) ((leftHanded.V ◁ leftAltContraction.hom β–· altLeftHanded.V) (((leftHanded.V ◁ (Ξ±_ leftHanded.V altLeftHanded.V altLeftHanded.V).inv) ((Ξ±_ leftHanded.V leftHanded.V (altLeftHanded.V βŠ— altLeftHanded.V)).hom ((x1 βŠ—β‚œ[β„‚] x2) βŠ—β‚œ[β„‚] y1 βŠ—β‚œ[β„‚] y2))))) = x1 βŠ—β‚œ[β„‚] ((Ξ»_ altLeftHanded.V).hom ((leftAltContraction.hom (x2 βŠ—β‚œ[β„‚] y1)) βŠ—β‚œ[β„‚] y2)) := rfl repeat rw (config := { transparency := .instances }) [h1] repeat rw [leftAltContraction_basis] simp only [Fin.isValue, Fin.val_one, Fin.val_zero, one_ne_zero, ↓reduceIte, zero_tmul, map_zero, tmul_zero, neg_zero, ModuleCat.MonoidalCategory.leftUnitor_hom_apply, one_smul, zero_add, zero_ne_one, add_zero, sub_neg_eq_add] erw [altLeftLeftUnit_apply_one, altLeftLeftUnitVal_expand_tmul] rw [add_comm] rfl lemma altLeftContraction_apply_metric : (Ξ²_ altLeftHanded leftHanded).hom.hom ((altLeftHanded.V ◁ (Ξ»_ leftHanded.V).hom) ((altLeftHanded.V ◁ altLeftContraction.hom β–· leftHanded.V) ((altLeftHanded.V ◁ (Ξ±_ altLeftHanded.V leftHanded.V leftHanded.V).inv) ((Ξ±_ altLeftHanded.V altLeftHanded.V (leftHanded.V βŠ— leftHanded.V)).hom (altLeftMetric.hom (1 : β„‚) βŠ—β‚œ[β„‚] leftMetric.hom (1 : β„‚)))))) = leftAltLeftUnit.hom (1 : β„‚) := by rw [leftMetric_apply_one, altLeftMetric_apply_one] rw [leftMetricVal_expand_tmul, altLeftMetricVal_expand_tmul] simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V, Fin.isValue, tmul_add, tmul_neg, sub_tmul, map_add, map_neg, map_sub] have h1 (x1 x2 : altLeftHanded) (y1 y2 : leftHanded) : (altLeftHanded.V ◁ (Ξ»_ leftHanded.V).hom) ((altLeftHanded.V ◁ altLeftContraction.hom β–· leftHanded.V) (((altLeftHanded.V ◁ (Ξ±_ altLeftHanded.V leftHanded.V leftHanded.V).inv) ((Ξ±_ altLeftHanded.V altLeftHanded.V (leftHanded.V βŠ— leftHanded.V)).hom ((x1 βŠ—β‚œ[β„‚] x2) βŠ—β‚œ[β„‚] y1 βŠ—β‚œ[β„‚] y2))))) = x1 βŠ—β‚œ[β„‚] ((Ξ»_ leftHanded.V).hom ((altLeftContraction.hom (x2 βŠ—β‚œ[β„‚] y1)) βŠ—β‚œ[β„‚] y2)) := rfl repeat rw (config := { transparency := .instances }) [h1] repeat rw [altLeftContraction_basis] simp only [Fin.isValue, Fin.val_one, Fin.val_zero, one_ne_zero, ↓reduceIte, zero_tmul, map_zero, tmul_zero, ModuleCat.MonoidalCategory.leftUnitor_hom_apply, one_smul, zero_sub, neg_neg, zero_ne_one, sub_zero] erw [leftAltLeftUnit_apply_one, leftAltLeftUnitVal_expand_tmul] rw [add_comm] rfl lemma rightAltContraction_apply_metric : (Ξ²_ rightHanded altRightHanded).hom.hom ((rightHanded.V ◁ (Ξ»_ altRightHanded.V).hom) ((rightHanded.V ◁ rightAltContraction.hom β–· altRightHanded.V) ((rightHanded.V ◁ (Ξ±_ rightHanded.V altRightHanded.V altRightHanded.V).inv) ((Ξ±_ rightHanded.V rightHanded.V (altRightHanded.V βŠ— altRightHanded.V)).hom (rightMetric.hom (1 : β„‚) βŠ—β‚œ[β„‚] altRightMetric.hom (1 : β„‚)))))) = altRightRightUnit.hom (1 : β„‚) := by rw [rightMetric_apply_one, altRightMetric_apply_one] rw [rightMetricVal_expand_tmul, altRightMetricVal_expand_tmul] simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V, Fin.isValue, tmul_sub, add_tmul, neg_tmul, map_sub, map_add, map_neg] have h1 (x1 x2 : rightHanded) (y1 y2 : altRightHanded) : (rightHanded.V ◁ (Ξ»_ altRightHanded.V).hom) ((rightHanded.V ◁ rightAltContraction.hom β–· altRightHanded.V) (((rightHanded.V ◁ (Ξ±_ rightHanded.V altRightHanded.V altRightHanded.V).inv) ((Ξ±_ rightHanded.V rightHanded.V (altRightHanded.V βŠ— altRightHanded.V)).hom ((x1 βŠ—β‚œ[β„‚] x2) βŠ—β‚œ[β„‚] y1 βŠ—β‚œ[β„‚] y2))))) = x1 βŠ—β‚œ[β„‚] ((Ξ»_ altRightHanded.V).hom ((rightAltContraction.hom (x2 βŠ—β‚œ[β„‚] y1)) βŠ—β‚œ[β„‚] y2)) := rfl repeat rw (config := { transparency := .instances }) [h1] repeat rw [rightAltContraction_basis] simp only [Fin.isValue, Fin.val_one, Fin.val_zero, one_ne_zero, ↓reduceIte, zero_tmul, map_zero, tmul_zero, neg_zero, ModuleCat.MonoidalCategory.leftUnitor_hom_apply, one_smul, zero_add, zero_ne_one, add_zero, sub_neg_eq_add] erw [altRightRightUnit_apply_one, altRightRightUnitVal_expand_tmul] rw [add_comm] rfl lemma altRightContraction_apply_metric : (Ξ²_ altRightHanded rightHanded).hom.hom ((altRightHanded.V ◁ (Ξ»_ rightHanded.V).hom) ((altRightHanded.V ◁ altRightContraction.hom β–· rightHanded.V) ((altRightHanded.V ◁ (Ξ±_ altRightHanded.V rightHanded.V rightHanded.V).inv) ((Ξ±_ altRightHanded.V altRightHanded.V (rightHanded.V βŠ— rightHanded.V)).hom (altRightMetric.hom (1 : β„‚) βŠ—β‚œ[β„‚] rightMetric.hom (1 : β„‚)))))) = rightAltRightUnit.hom (1 : β„‚) := by rw [rightMetric_apply_one, altRightMetric_apply_one] rw [rightMetricVal_expand_tmul, altRightMetricVal_expand_tmul] simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V, 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) ((altRightHanded.V ◁ altRightContraction.hom β–· rightHanded.V) (((altRightHanded.V ◁ (Ξ±_ altRightHanded.V rightHanded.V rightHanded.V).inv) ((Ξ±_ altRightHanded.V altRightHanded.V (rightHanded.V βŠ— rightHanded.V)).hom ((x1 βŠ—β‚œ[β„‚] x2) βŠ—β‚œ[β„‚] y1 βŠ—β‚œ[β„‚] y2))))) = x1 βŠ—β‚œ[β„‚] ((Ξ»_ rightHanded.V).hom ((altRightContraction.hom (x2 βŠ—β‚œ[β„‚] y1)) βŠ—β‚œ[β„‚] y2)) := rfl repeat rw (config := { transparency := .instances }) [h1] repeat rw [altRightContraction_basis] simp only [Fin.isValue, Fin.val_one, Fin.val_zero, one_ne_zero, ↓reduceIte, zero_tmul, map_zero, tmul_zero, ModuleCat.MonoidalCategory.leftUnitor_hom_apply, one_smul, zero_sub, neg_neg, zero_ne_one, sub_zero] erw [rightAltRightUnit_apply_one, rightAltRightUnitVal_expand_tmul] rw [add_comm] rfl end end Fermion