PhysLean/HepLean/Lorentz/Weyl/Unit.lean

358 lines
17 KiB
Text
Raw Normal View History

/-
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
-/
2024-11-09 17:35:02 +00:00
import HepLean.Lorentz.Weyl.Two
/-!
# Units of Weyl fermions
We define the units for Weyl fermions, often denoted `δ` in the literature.
-/
namespace Fermion
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open CategoryTheory.MonoidalCategory
2024-10-16 10:39:11 +00:00
/-- The left-alt-left unit `δᵃₐ` as an element of `(leftHanded ⊗ altLeftHanded).V`. -/
def leftAltLeftUnitVal : (leftHanded ⊗ altLeftHanded).V :=
leftAltLeftToMatrix.symm 1
2024-10-24 15:04:37 +00:00
/-- Expansion of `leftAltLeftUnitVal` into the basis. -/
lemma leftAltLeftUnitVal_expand_tmul : leftAltLeftUnitVal =
leftBasis 0 ⊗ₜ[] altLeftBasis 0 + leftBasis 1 ⊗ₜ[] altLeftBasis 1 := by
simp only [Action.instMonoidalCategory_tensorObj_V, leftAltLeftUnitVal, Fin.isValue]
erw [leftAltLeftToMatrix_symm_expand_tmul]
simp only [Fin.sum_univ_two, Fin.isValue, one_apply_eq, one_smul, ne_eq, zero_ne_one,
not_false_eq_true, one_apply_ne, zero_smul, add_zero, one_ne_zero, zero_add]
2024-10-16 10:39:11 +00:00
/-- 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 := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • leftAltLeftUnitVal,
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 : => ?_
2024-10-15 11:39:40 +00:00
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.hom_comp,
2024-10-15 11:39:40 +00:00
Function.comp_apply]
change x • leftAltLeftUnitVal =
(TensorProduct.map (leftHanded.ρ M) (altLeftHanded.ρ M)) (x • leftAltLeftUnitVal)
2024-10-15 11:39:40 +00:00
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
2024-10-15 11:39:40 +00:00
simp only [Action.instMonoidalCategory_tensorObj_V, leftAltLeftUnitVal]
erw [leftAltLeftToMatrix_ρ_symm]
apply congrArg
simp
2024-10-24 15:04:37 +00:00
lemma leftAltLeftUnit_apply_one : leftAltLeftUnit.hom (1 : ) = leftAltLeftUnitVal := by
change leftAltLeftUnit.hom.hom.toFun (1 : ) = leftAltLeftUnitVal
simp only [leftAltLeftUnit, one_smul]
2024-10-24 15:04:37 +00:00
2024-10-16 10:39:11 +00:00
/-- The alt-left-left unit `δₐᵃ` as an element of `(altLeftHanded ⊗ leftHanded).V`. -/
def altLeftLeftUnitVal : (altLeftHanded ⊗ leftHanded).V :=
altLeftLeftToMatrix.symm 1
2024-10-24 15:04:37 +00:00
/-- Expansion of `altLeftLeftUnitVal` into the basis. -/
lemma altLeftLeftUnitVal_expand_tmul : altLeftLeftUnitVal =
altLeftBasis 0 ⊗ₜ[] leftBasis 0 + altLeftBasis 1 ⊗ₜ[] leftBasis 1 := by
simp only [Action.instMonoidalCategory_tensorObj_V, altLeftLeftUnitVal, Fin.isValue]
erw [altLeftLeftToMatrix_symm_expand_tmul]
simp only [Fin.sum_univ_two, Fin.isValue, one_apply_eq, one_smul, ne_eq, zero_ne_one,
not_false_eq_true, one_apply_ne, zero_smul, add_zero, one_ne_zero, zero_add]
2024-10-16 10:39:11 +00:00
/-- 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 := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • altLeftLeftUnitVal,
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 : => ?_
2024-10-15 11:39:40 +00:00
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.hom_comp,
2024-10-15 11:39:40 +00:00
Function.comp_apply]
change x • altLeftLeftUnitVal =
(TensorProduct.map (altLeftHanded.ρ M) (leftHanded.ρ M)) (x • altLeftLeftUnitVal)
2024-10-15 11:39:40 +00:00
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
2024-10-15 11:53:24 +00:00
simp only [Action.instMonoidalCategory_tensorObj_V, altLeftLeftUnitVal]
erw [altLeftLeftToMatrix_ρ_symm]
apply congrArg
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]
2024-11-12 06:33:58 +00:00
/-- Applying the morphism `altLeftLeftUnit` to `1` returns `altLeftLeftUnitVal`. -/
2024-10-24 15:04:37 +00:00
lemma altLeftLeftUnit_apply_one : altLeftLeftUnit.hom (1 : ) = altLeftLeftUnitVal := by
change altLeftLeftUnit.hom.hom.toFun (1 : ) = altLeftLeftUnitVal
simp only [altLeftLeftUnit, one_smul]
2024-10-24 15:04:37 +00:00
2024-10-16 10:39:11 +00:00
/-- The right-alt-right unit `δ^{dot a}_{dot a}` as an element of
`(rightHanded ⊗ altRightHanded).V`. -/
def rightAltRightUnitVal : (rightHanded ⊗ altRightHanded).V :=
rightAltRightToMatrix.symm 1
2024-10-24 15:04:37 +00:00
/-- Expansion of `rightAltRightUnitVal` into the basis. -/
lemma rightAltRightUnitVal_expand_tmul : rightAltRightUnitVal =
rightBasis 0 ⊗ₜ[] altRightBasis 0 + rightBasis 1 ⊗ₜ[] altRightBasis 1 := by
simp only [Action.instMonoidalCategory_tensorObj_V, rightAltRightUnitVal, Fin.isValue]
erw [rightAltRightToMatrix_symm_expand_tmul]
simp only [Fin.sum_univ_two, Fin.isValue, one_apply_eq, one_smul, ne_eq, zero_ne_one,
not_false_eq_true, one_apply_ne, zero_smul, add_zero, one_ne_zero, zero_add]
2024-10-16 10:39:11 +00:00
/-- 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
hom := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • rightAltRightUnitVal,
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 : => ?_
2024-10-15 11:39:40 +00:00
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.hom_comp,
2024-10-15 11:39:40 +00:00
Function.comp_apply]
change x • rightAltRightUnitVal =
(TensorProduct.map (rightHanded.ρ M) (altRightHanded.ρ M)) (x • rightAltRightUnitVal)
2024-10-15 11:39:40 +00:00
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
2024-10-15 11:39:40 +00:00
simp only [Action.instMonoidalCategory_tensorObj_V, rightAltRightUnitVal]
erw [rightAltRightToMatrix_ρ_symm]
apply congrArg
2024-10-15 11:39:40 +00:00
simp only [RCLike.star_def, mul_one]
symm
refine transpose_eq_one.mp ?h.h.h.a
2024-10-15 11:39:40 +00:00
simp only [transpose_mul, transpose_transpose]
change (M.1)⁻¹ᴴ * (M.1)ᴴ = 1
rw [@conjTranspose_nonsing_inv]
simp
2024-10-24 15:04:37 +00:00
lemma rightAltRightUnit_apply_one : rightAltRightUnit.hom (1 : ) = rightAltRightUnitVal := by
change rightAltRightUnit.hom.hom.toFun (1 : ) = rightAltRightUnitVal
simp only [rightAltRightUnit, one_smul]
2024-10-24 15:04:37 +00:00
2024-10-16 10:39:11 +00:00
/-- The alt-right-right unit `δ_{dot a}^{dot a}` as an element of
`(rightHanded ⊗ altRightHanded).V`. -/
def altRightRightUnitVal : (altRightHanded ⊗ rightHanded).V :=
altRightRightToMatrix.symm 1
2024-10-24 15:04:37 +00:00
/-- Expansion of `altRightRightUnitVal` into the basis. -/
lemma altRightRightUnitVal_expand_tmul : altRightRightUnitVal =
altRightBasis 0 ⊗ₜ[] rightBasis 0 + altRightBasis 1 ⊗ₜ[] rightBasis 1 := by
simp only [Action.instMonoidalCategory_tensorObj_V, altRightRightUnitVal, Fin.isValue]
erw [altRightRightToMatrix_symm_expand_tmul]
simp only [Fin.sum_univ_two, Fin.isValue, one_apply_eq, one_smul, ne_eq, zero_ne_one,
not_false_eq_true, one_apply_ne, zero_smul, add_zero, one_ne_zero, zero_add]
2024-10-16 10:39:11 +00:00
/-- 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
hom := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • altRightRightUnitVal,
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 : => ?_
2024-10-15 11:39:40 +00:00
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.hom_comp,
2024-10-15 11:39:40 +00:00
Function.comp_apply]
change x • altRightRightUnitVal =
(TensorProduct.map (altRightHanded.ρ M) (rightHanded.ρ M)) (x • altRightRightUnitVal)
2024-10-15 11:39:40 +00:00
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
2024-10-15 11:53:24 +00:00
simp only [Action.instMonoidalCategory_tensorObj_V, altRightRightUnitVal]
erw [altRightRightToMatrix_ρ_symm]
apply congrArg
2024-10-15 11:39:40 +00:00
simp only [mul_one, RCLike.star_def]
symm
2024-10-15 11:39:40 +00:00
change (M.1)⁻¹ᴴ * (M.1)ᴴ = 1
rw [@conjTranspose_nonsing_inv]
simp
2024-10-24 15:04:37 +00:00
lemma altRightRightUnit_apply_one : altRightRightUnit.hom (1 : ) = altRightRightUnitVal := by
change altRightRightUnit.hom.hom.toFun (1 : ) = altRightRightUnitVal
simp only [altRightRightUnit, one_smul]
2024-10-24 15:04:37 +00:00
/-!
## Contraction of the units
-/
/-- Contraction on the right with `altLeftLeftUnit` does nothing. -/
lemma contr_altLeftLeftUnit (x : leftHanded) :
2024-10-24 15:04:37 +00:00
(λ_ leftHanded).hom.hom
(((leftAltContraction) ▷ leftHanded).hom
((α_ _ _ leftHanded).inv.hom
(x ⊗ₜ[] altLeftLeftUnit.hom (1 : )))) = x := by
obtain ⟨c, hc⟩ := (mem_span_range_iff_exists_fun ).mp (Basis.mem_span leftBasis x)
subst hc
rw [altLeftLeftUnit_apply_one, altLeftLeftUnitVal_expand_tmul]
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.instMonoidalCategory_leftUnitor_hom_hom, Action.instMonoidalCategory_whiskerRight_hom,
Action.instMonoidalCategory_associator_inv_hom, CategoryTheory.Equivalence.symm_inverse,
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
Fin.sum_univ_two, Fin.isValue, tmul_add, add_tmul, smul_tmul, tmul_smul, map_add,
_root_.map_smul]
have h1 (x y : leftHanded) (z : altLeftHanded) : (leftAltContraction.hom ▷ leftHanded.V)
((α_ leftHanded.V altLeftHanded.V leftHanded.V).inv (x ⊗ₜ[] z ⊗ₜ[] y)) =
(leftAltContraction.hom (x ⊗ₜ[] z)) ⊗ₜ[] y := rfl
erw [h1, h1, h1, h1]
repeat rw [leftAltContraction_basis]
simp only [Fin.isValue, leftUnitor, ModuleCat.MonoidalCategory.leftUnitor, ModuleCat.of_coe,
CategoryTheory.Iso.trans_hom, LinearEquiv.toModuleIso_hom_hom, ModuleCat.ofSelfIso_hom,
CategoryTheory.Category.comp_id, Action.instMonoidalCategory_tensorUnit_V, Fin.val_zero,
↓reduceIte, Fin.val_one, one_ne_zero, zero_tmul, map_zero, smul_zero, add_zero, zero_ne_one,
zero_add]
erw [TensorProduct.lid_tmul, TensorProduct.lid_tmul]
simp only [Fin.isValue, one_smul]
/-- Contraction on the right with `leftAltLeftUnit` does nothing. -/
2024-10-24 16:42:25 +00:00
lemma contr_leftAltLeftUnit (x : altLeftHanded) :
(λ_ altLeftHanded).hom.hom
(((altLeftContraction) ▷ altLeftHanded).hom
((α_ _ _ altLeftHanded).inv.hom
(x ⊗ₜ[] leftAltLeftUnit.hom (1 : )))) = x := by
obtain ⟨c, hc⟩ := (mem_span_range_iff_exists_fun ).mp (Basis.mem_span altLeftBasis x)
subst hc
rw [leftAltLeftUnit_apply_one, leftAltLeftUnitVal_expand_tmul]
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.instMonoidalCategory_leftUnitor_hom_hom, Action.instMonoidalCategory_whiskerRight_hom,
Action.instMonoidalCategory_associator_inv_hom, CategoryTheory.Equivalence.symm_inverse,
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
Fin.sum_univ_two, Fin.isValue, tmul_add, add_tmul, smul_tmul, tmul_smul, map_add,
_root_.map_smul]
have h1 (x y : altLeftHanded) (z : leftHanded) : (altLeftContraction.hom ▷ altLeftHanded.V)
((α_ altLeftHanded.V leftHanded.V altLeftHanded.V).inv (x ⊗ₜ[] z ⊗ₜ[] y)) =
(altLeftContraction.hom (x ⊗ₜ[] z)) ⊗ₜ[] y := rfl
erw [h1, h1, h1, h1]
repeat rw [altLeftContraction_basis]
simp only [Fin.isValue, Action.instMonoidalCategory_tensorUnit_V, Fin.val_zero, ↓reduceIte,
Fin.val_one, one_ne_zero, zero_tmul, map_zero, smul_zero, add_zero, zero_ne_one, zero_add]
erw [TensorProduct.lid_tmul, TensorProduct.lid_tmul]
simp only [Fin.isValue, one_smul]
/-- Contraction on the right with `altRightRightUnit` does nothing. -/
lemma contr_altRightRightUnit (x : rightHanded) :
(λ_ rightHanded).hom.hom
(((rightAltContraction) ▷ rightHanded).hom
((α_ _ _ rightHanded).inv.hom
(x ⊗ₜ[] altRightRightUnit.hom (1 : )))) = x := by
obtain ⟨c, hc⟩ := (mem_span_range_iff_exists_fun ).mp (Basis.mem_span rightBasis x)
subst hc
rw [altRightRightUnit_apply_one, altRightRightUnitVal_expand_tmul]
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.instMonoidalCategory_leftUnitor_hom_hom, Action.instMonoidalCategory_whiskerRight_hom,
Action.instMonoidalCategory_associator_inv_hom, CategoryTheory.Equivalence.symm_inverse,
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
Fin.sum_univ_two, Fin.isValue, tmul_add, add_tmul, smul_tmul, tmul_smul, map_add,
_root_.map_smul]
have h1 (x y : rightHanded) (z : altRightHanded) : (rightAltContraction.hom ▷ rightHanded.V)
((α_ rightHanded.V altRightHanded.V rightHanded.V).inv (x ⊗ₜ[] z ⊗ₜ[] y)) =
(rightAltContraction.hom (x ⊗ₜ[] z)) ⊗ₜ[] y := rfl
erw [h1, h1, h1, h1]
repeat rw [rightAltContraction_basis]
simp only [Fin.isValue, Action.instMonoidalCategory_tensorUnit_V, Fin.val_zero, ↓reduceIte,
Fin.val_one, one_ne_zero, zero_tmul, map_zero, smul_zero, add_zero, zero_ne_one, zero_add]
erw [TensorProduct.lid_tmul, TensorProduct.lid_tmul]
simp only [Fin.isValue, one_smul]
/-- Contraction on the right with `rightAltRightUnit` does nothing. -/
lemma contr_rightAltRightUnit (x : altRightHanded) :
(λ_ altRightHanded).hom.hom
(((altRightContraction) ▷ altRightHanded).hom
((α_ _ _ altRightHanded).inv.hom
(x ⊗ₜ[] rightAltRightUnit.hom (1 : )))) = x := by
obtain ⟨c, hc⟩ := (mem_span_range_iff_exists_fun ).mp (Basis.mem_span altRightBasis x)
subst hc
rw [rightAltRightUnit_apply_one, rightAltRightUnitVal_expand_tmul]
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.instMonoidalCategory_leftUnitor_hom_hom, Action.instMonoidalCategory_whiskerRight_hom,
Action.instMonoidalCategory_associator_inv_hom, CategoryTheory.Equivalence.symm_inverse,
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
Fin.sum_univ_two, Fin.isValue, tmul_add, add_tmul, smul_tmul, tmul_smul, map_add,
_root_.map_smul]
have h1 (x y : altRightHanded) (z : rightHanded) : (altRightContraction.hom ▷ altRightHanded.V)
((α_ altRightHanded.V rightHanded.V altRightHanded.V).inv (x ⊗ₜ[] z ⊗ₜ[] y)) =
(altRightContraction.hom (x ⊗ₜ[] z)) ⊗ₜ[] y := rfl
erw [h1, h1, h1, h1]
repeat rw [altRightContraction_basis]
simp only [Fin.isValue, Action.instMonoidalCategory_tensorUnit_V, Fin.val_zero, ↓reduceIte,
Fin.val_one, one_ne_zero, zero_tmul, map_zero, smul_zero, add_zero, zero_ne_one, zero_add]
erw [TensorProduct.lid_tmul, TensorProduct.lid_tmul]
simp only [Fin.isValue, one_smul]
/-!
## Symmetry properties of the units
-/
open CategoryTheory
lemma altLeftLeftUnit_symm :
2024-10-24 16:42:25 +00:00
(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 :
2024-10-24 16:42:25 +00:00
(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 :
2024-10-24 16:42:25 +00:00
(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 :
2024-10-24 16:42:25 +00:00
(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 Fermion