refactor: Move WeylFermion
This commit is contained in:
parent
2fad30f920
commit
58ea861113
10 changed files with 22 additions and 22 deletions
297
HepLean/Lorentz/Weyl/Basic.lean
Normal file
297
HepLean/Lorentz/Weyl/Basic.lean
Normal file
|
@ -0,0 +1,297 @@
|
|||
/-
|
||||
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.Meta.Informal
|
||||
import HepLean.SpaceTime.SL2C.Basic
|
||||
import Mathlib.RepresentationTheory.Rep
|
||||
import HepLean.Lorentz.Weyl.Modules
|
||||
import Mathlib.Logic.Equiv.TransferInstance
|
||||
/-!
|
||||
|
||||
# Weyl fermions
|
||||
|
||||
A good reference for the material in this file is:
|
||||
https://particle.physics.ucdavis.edu/modernsusy/slides/slideimages/spinorfeynrules.pdf
|
||||
|
||||
-/
|
||||
|
||||
namespace Fermion
|
||||
noncomputable section
|
||||
|
||||
open Matrix
|
||||
open MatrixGroups
|
||||
open Complex
|
||||
open TensorProduct
|
||||
|
||||
/-- The vector space ℂ^2 carrying the fundamental representation of SL(2,C).
|
||||
In index notation corresponds to a Weyl fermion with indices ψ^a. -/
|
||||
def leftHanded : Rep ℂ SL(2,ℂ) := Rep.of {
|
||||
toFun := fun M => {
|
||||
toFun := fun (ψ : LeftHandedModule) =>
|
||||
LeftHandedModule.toFin2ℂEquiv.symm (M.1 *ᵥ ψ.toFin2ℂ),
|
||||
map_add' := by
|
||||
intro ψ ψ'
|
||||
simp [mulVec_add]
|
||||
map_smul' := by
|
||||
intro r ψ
|
||||
simp [mulVec_smul]}
|
||||
map_one' := by
|
||||
ext i
|
||||
simp
|
||||
map_mul' := fun M N => by
|
||||
simp only [SpecialLinearGroup.coe_mul]
|
||||
ext1 x
|
||||
simp only [LinearMap.coe_mk, AddHom.coe_mk, LinearMap.mul_apply, LinearEquiv.apply_symm_apply,
|
||||
mulVec_mulVec]}
|
||||
|
||||
/-- The standard basis on left-handed Weyl fermions. -/
|
||||
def leftBasis : Basis (Fin 2) ℂ leftHanded := Basis.ofEquivFun
|
||||
(Equiv.linearEquiv ℂ LeftHandedModule.toFin2ℂFun)
|
||||
|
||||
@[simp]
|
||||
lemma leftBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 2) :
|
||||
(LinearMap.toMatrix leftBasis leftBasis) (leftHanded.ρ M) i j = M.1 i j := by
|
||||
rw [LinearMap.toMatrix_apply]
|
||||
simp only [leftBasis, Basis.coe_ofEquivFun, Basis.ofEquivFun_repr_apply]
|
||||
change (M.1 *ᵥ (Pi.single j 1)) i = _
|
||||
simp only [mulVec_single, mul_one]
|
||||
|
||||
@[simp]
|
||||
lemma leftBasis_toFin2ℂ (i : Fin 2) : (leftBasis i).toFin2ℂ = Pi.single i 1 := by
|
||||
simp only [leftBasis, Basis.coe_ofEquivFun]
|
||||
rfl
|
||||
|
||||
/-- The vector space ℂ^2 carrying the representation of SL(2,C) given by
|
||||
M → (M⁻¹)ᵀ. In index notation corresponds to a Weyl fermion with indices ψ_a. -/
|
||||
def altLeftHanded : Rep ℂ SL(2,ℂ) := Rep.of {
|
||||
toFun := fun M => {
|
||||
toFun := fun (ψ : AltLeftHandedModule) =>
|
||||
AltLeftHandedModule.toFin2ℂEquiv.symm ((M.1⁻¹)ᵀ *ᵥ ψ.toFin2ℂ),
|
||||
map_add' := by
|
||||
intro ψ ψ'
|
||||
simp [mulVec_add]
|
||||
map_smul' := by
|
||||
intro r ψ
|
||||
simp [mulVec_smul]}
|
||||
map_one' := by
|
||||
ext i
|
||||
simp
|
||||
map_mul' := fun M N => by
|
||||
ext1 x
|
||||
simp only [SpecialLinearGroup.coe_mul, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.mul_apply,
|
||||
LinearEquiv.apply_symm_apply, mulVec_mulVec, EmbeddingLike.apply_eq_iff_eq]
|
||||
refine (congrFun (congrArg _ ?_) _)
|
||||
rw [Matrix.mul_inv_rev]
|
||||
exact transpose_mul _ _}
|
||||
|
||||
/-- The standard basis on alt-left-handed Weyl fermions. -/
|
||||
def altLeftBasis : Basis (Fin 2) ℂ altLeftHanded := Basis.ofEquivFun
|
||||
(Equiv.linearEquiv ℂ AltLeftHandedModule.toFin2ℂFun)
|
||||
|
||||
@[simp]
|
||||
lemma altLeftBasis_toFin2ℂ (i : Fin 2) : (altLeftBasis i).toFin2ℂ = Pi.single i 1 := by
|
||||
simp only [altLeftBasis, Basis.coe_ofEquivFun]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma altLeftBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 2) :
|
||||
(LinearMap.toMatrix altLeftBasis altLeftBasis) (altLeftHanded.ρ M) i j = (M.1⁻¹)ᵀ i j := by
|
||||
rw [LinearMap.toMatrix_apply]
|
||||
simp only [altLeftBasis, Basis.coe_ofEquivFun, Basis.ofEquivFun_repr_apply, transpose_apply]
|
||||
change ((M.1⁻¹)ᵀ *ᵥ (Pi.single j 1)) i = _
|
||||
simp only [mulVec_single, transpose_apply, mul_one]
|
||||
|
||||
/-- The vector space ℂ^2 carrying the conjugate representation of SL(2,C).
|
||||
In index notation corresponds to a Weyl fermion with indices ψ^{dot a}. -/
|
||||
def rightHanded : Rep ℂ SL(2,ℂ) := Rep.of {
|
||||
toFun := fun M => {
|
||||
toFun := fun (ψ : RightHandedModule) =>
|
||||
RightHandedModule.toFin2ℂEquiv.symm (M.1.map star *ᵥ ψ.toFin2ℂ),
|
||||
map_add' := by
|
||||
intro ψ ψ'
|
||||
simp [mulVec_add]
|
||||
map_smul' := by
|
||||
intro r ψ
|
||||
simp [mulVec_smul]}
|
||||
map_one' := by
|
||||
ext i
|
||||
simp
|
||||
map_mul' := fun M N => by
|
||||
ext1 x
|
||||
simp only [SpecialLinearGroup.coe_mul, RCLike.star_def, Matrix.map_mul, LinearMap.coe_mk,
|
||||
AddHom.coe_mk, LinearMap.mul_apply, LinearEquiv.apply_symm_apply, mulVec_mulVec]}
|
||||
|
||||
/-- The standard basis on right-handed Weyl fermions. -/
|
||||
def rightBasis : Basis (Fin 2) ℂ rightHanded := Basis.ofEquivFun
|
||||
(Equiv.linearEquiv ℂ RightHandedModule.toFin2ℂFun)
|
||||
|
||||
@[simp]
|
||||
lemma rightBasis_toFin2ℂ (i : Fin 2) : (rightBasis i).toFin2ℂ = Pi.single i 1 := by
|
||||
simp only [rightBasis, Basis.coe_ofEquivFun]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma rightBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 2) :
|
||||
(LinearMap.toMatrix rightBasis rightBasis) (rightHanded.ρ M) i j = (M.1.map star) i j := by
|
||||
rw [LinearMap.toMatrix_apply]
|
||||
simp only [rightBasis, Basis.coe_ofEquivFun, Basis.ofEquivFun_repr_apply, transpose_apply]
|
||||
change (M.1.map star *ᵥ (Pi.single j 1)) i = _
|
||||
simp only [mulVec_single, transpose_apply, mul_one]
|
||||
|
||||
/-- The vector space ℂ^2 carrying the representation of SL(2,C) given by
|
||||
M → (M⁻¹)^†.
|
||||
In index notation this corresponds to a Weyl fermion with index `ψ_{dot a}`. -/
|
||||
def altRightHanded : Rep ℂ SL(2,ℂ) := Rep.of {
|
||||
toFun := fun M => {
|
||||
toFun := fun (ψ : AltRightHandedModule) =>
|
||||
AltRightHandedModule.toFin2ℂEquiv.symm ((M.1⁻¹).conjTranspose *ᵥ ψ.toFin2ℂ),
|
||||
map_add' := by
|
||||
intro ψ ψ'
|
||||
simp [mulVec_add]
|
||||
map_smul' := by
|
||||
intro r ψ
|
||||
simp [mulVec_smul]}
|
||||
map_one' := by
|
||||
ext i
|
||||
simp
|
||||
map_mul' := fun M N => by
|
||||
ext1 x
|
||||
simp only [SpecialLinearGroup.coe_mul, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.mul_apply,
|
||||
LinearEquiv.apply_symm_apply, mulVec_mulVec, EmbeddingLike.apply_eq_iff_eq]
|
||||
refine (congrFun (congrArg _ ?_) _)
|
||||
rw [Matrix.mul_inv_rev]
|
||||
exact conjTranspose_mul _ _}
|
||||
|
||||
/-- The standard basis on alt-right-handed Weyl fermions. -/
|
||||
def altRightBasis : Basis (Fin 2) ℂ altRightHanded := Basis.ofEquivFun
|
||||
(Equiv.linearEquiv ℂ AltRightHandedModule.toFin2ℂFun)
|
||||
|
||||
@[simp]
|
||||
lemma altRightBasis_toFin2ℂ (i : Fin 2) : (altRightBasis i).toFin2ℂ = Pi.single i 1 := by
|
||||
simp only [altRightBasis, Basis.coe_ofEquivFun]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma altRightBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 2) :
|
||||
(LinearMap.toMatrix altRightBasis altRightBasis) (altRightHanded.ρ M) i j =
|
||||
((M.1⁻¹).conjTranspose) i j := by
|
||||
rw [LinearMap.toMatrix_apply]
|
||||
simp only [altRightBasis, Basis.coe_ofEquivFun, Basis.ofEquivFun_repr_apply, transpose_apply]
|
||||
change ((M.1⁻¹).conjTranspose *ᵥ (Pi.single j 1)) i = _
|
||||
simp only [mulVec_single, transpose_apply, mul_one]
|
||||
|
||||
/-!
|
||||
|
||||
## Equivalences between Weyl fermion vector spaces.
|
||||
|
||||
-/
|
||||
|
||||
/-- The morphism between the representation `leftHanded` and the representation
|
||||
`altLeftHanded` defined by multiplying an element of
|
||||
`leftHanded` by the matrix `εᵃ⁰ᵃ¹ = !![0, 1; -1, 0]]`. -/
|
||||
def leftHandedToAlt : leftHanded ⟶ altLeftHanded where
|
||||
hom := {
|
||||
toFun := fun ψ => AltLeftHandedModule.toFin2ℂEquiv.symm (!![0, 1; -1, 0] *ᵥ ψ.toFin2ℂ),
|
||||
map_add' := by
|
||||
intro ψ ψ'
|
||||
simp only [mulVec_add, LinearEquiv.map_add]
|
||||
map_smul' := by
|
||||
intro a ψ
|
||||
simp only [mulVec_smul, LinearEquiv.map_smul]
|
||||
rfl}
|
||||
comm := by
|
||||
intro M
|
||||
refine LinearMap.ext (fun ψ => ?_)
|
||||
change AltLeftHandedModule.toFin2ℂEquiv.symm (!![0, 1; -1, 0] *ᵥ M.1 *ᵥ ψ.val) =
|
||||
AltLeftHandedModule.toFin2ℂEquiv.symm ((M.1⁻¹)ᵀ *ᵥ !![0, 1; -1, 0] *ᵥ ψ.val)
|
||||
apply congrArg
|
||||
rw [mulVec_mulVec, mulVec_mulVec, SpaceTime.SL2C.inverse_coe, eta_fin_two M.1]
|
||||
refine congrFun (congrArg _ ?_) _
|
||||
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
|
||||
|
||||
lemma leftHandedToAlt_hom_apply (ψ : leftHanded) :
|
||||
leftHandedToAlt.hom ψ =
|
||||
AltLeftHandedModule.toFin2ℂEquiv.symm (!![0, 1; -1, 0] *ᵥ ψ.toFin2ℂ) := rfl
|
||||
|
||||
/-- The morphism from `altLeftHanded` to
|
||||
`leftHanded` defined by multiplying an element of
|
||||
altLeftHandedWeyl by the matrix `εₐ₁ₐ₂ = !![0, -1; 1, 0]`. -/
|
||||
def leftHandedAltTo : altLeftHanded ⟶ leftHanded where
|
||||
hom := {
|
||||
toFun := fun ψ =>
|
||||
LeftHandedModule.toFin2ℂEquiv.symm (!![0, -1; 1, 0] *ᵥ ψ.toFin2ℂ),
|
||||
map_add' := by
|
||||
intro ψ ψ'
|
||||
simp only [map_add]
|
||||
rw [mulVec_add, LinearEquiv.map_add]
|
||||
map_smul' := by
|
||||
intro a ψ
|
||||
simp only [LinearEquiv.map_smul]
|
||||
rw [mulVec_smul, LinearEquiv.map_smul]
|
||||
rfl}
|
||||
comm := by
|
||||
intro M
|
||||
refine LinearMap.ext (fun ψ => ?_)
|
||||
change LeftHandedModule.toFin2ℂEquiv.symm (!![0, -1; 1, 0] *ᵥ (M.1⁻¹)ᵀ *ᵥ ψ.val) =
|
||||
LeftHandedModule.toFin2ℂEquiv.symm (M.1 *ᵥ !![0, -1; 1, 0] *ᵥ ψ.val)
|
||||
rw [EquivLike.apply_eq_iff_eq, mulVec_mulVec, mulVec_mulVec, SpaceTime.SL2C.inverse_coe,
|
||||
eta_fin_two M.1]
|
||||
refine congrFun (congrArg _ ?_) _
|
||||
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
|
||||
|
||||
lemma leftHandedAltTo_hom_apply (ψ : altLeftHanded) :
|
||||
leftHandedAltTo.hom ψ =
|
||||
LeftHandedModule.toFin2ℂEquiv.symm (!![0, -1; 1, 0] *ᵥ ψ.toFin2ℂ) := rfl
|
||||
|
||||
/-- The equivalence between the representation `leftHanded` and the representation
|
||||
`altLeftHanded` defined by multiplying an element of
|
||||
`leftHanded` by the matrix `εᵃ⁰ᵃ¹ = !![0, 1; -1, 0]]`. -/
|
||||
def leftHandedAltEquiv : leftHanded ≅ altLeftHanded where
|
||||
hom := leftHandedToAlt
|
||||
inv := leftHandedAltTo
|
||||
hom_inv_id := by
|
||||
ext ψ
|
||||
simp only [Action.comp_hom, ModuleCat.coe_comp, Function.comp_apply, Action.id_hom,
|
||||
ModuleCat.id_apply]
|
||||
rw [leftHandedAltTo_hom_apply, leftHandedToAlt_hom_apply]
|
||||
rw [AltLeftHandedModule.toFin2ℂ, LinearEquiv.apply_symm_apply, mulVec_mulVec]
|
||||
rw [show (!![0, -1; (1 : ℂ), 0] * !![0, 1; -1, 0]) = 1 by simpa using Eq.symm one_fin_two]
|
||||
rw [one_mulVec]
|
||||
rfl
|
||||
inv_hom_id := by
|
||||
ext ψ
|
||||
simp only [Action.comp_hom, ModuleCat.coe_comp, Function.comp_apply, Action.id_hom,
|
||||
ModuleCat.id_apply]
|
||||
rw [leftHandedAltTo_hom_apply, leftHandedToAlt_hom_apply, LeftHandedModule.toFin2ℂ,
|
||||
LinearEquiv.apply_symm_apply, mulVec_mulVec]
|
||||
rw [show (!![0, (1 : ℂ); -1, 0] * !![0, -1; 1, 0]) = 1 by simpa using Eq.symm one_fin_two]
|
||||
rw [one_mulVec]
|
||||
rfl
|
||||
|
||||
lemma leftHandedAltEquiv_hom_hom_apply (ψ : leftHanded) :
|
||||
leftHandedAltEquiv.hom.hom ψ =
|
||||
AltLeftHandedModule.toFin2ℂEquiv.symm (!![0, 1; -1, 0] *ᵥ ψ.toFin2ℂ) := rfl
|
||||
|
||||
lemma leftHandedAltEquiv_inv_hom_apply (ψ : altLeftHanded) :
|
||||
leftHandedAltEquiv.inv.hom ψ =
|
||||
LeftHandedModule.toFin2ℂEquiv.symm (!![0, -1; 1, 0] *ᵥ ψ.toFin2ℂ) := rfl
|
||||
|
||||
informal_definition rightHandedWeylAltEquiv where
|
||||
math :≈ "The linear equiv between rightHandedWeyl and altRightHandedWeyl given
|
||||
by multiplying an element of rightHandedWeyl by the matrix `εᵃ⁰ᵃ¹ = !![0, 1; -1, 0]]`"
|
||||
deps :≈ [``rightHanded, ``altRightHanded]
|
||||
|
||||
informal_lemma rightHandedWeylAltEquiv_equivariant where
|
||||
math :≈ "The linear equiv rightHandedWeylAltEquiv is equivariant with respect to the
|
||||
action of SL(2,C) on rightHandedWeyl and altRightHandedWeyl."
|
||||
deps :≈ [``rightHandedWeylAltEquiv]
|
||||
|
||||
end
|
||||
|
||||
end Fermion
|
279
HepLean/Lorentz/Weyl/Contraction.lean
Normal file
279
HepLean/Lorentz/Weyl/Contraction.lean
Normal file
|
@ -0,0 +1,279 @@
|
|||
/-
|
||||
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.Basic
|
||||
/-!
|
||||
|
||||
# Contraction of Weyl fermions
|
||||
|
||||
We define the contraction of Weyl fermions.
|
||||
|
||||
-/
|
||||
|
||||
namespace Fermion
|
||||
noncomputable section
|
||||
|
||||
open Matrix
|
||||
open MatrixGroups
|
||||
open Complex
|
||||
open TensorProduct
|
||||
|
||||
/-!
|
||||
|
||||
## Contraction of Weyl fermions.
|
||||
|
||||
-/
|
||||
open CategoryTheory.MonoidalCategory
|
||||
|
||||
/-- The bi-linear map corresponding to contraction of a left-handed Weyl fermion with a
|
||||
alt-left-handed Weyl fermion. -/
|
||||
def leftAltBi : leftHanded →ₗ[ℂ] altLeftHanded →ₗ[ℂ] ℂ where
|
||||
toFun ψ := {
|
||||
toFun := fun φ => ψ.toFin2ℂ ⬝ᵥ φ.toFin2ℂ,
|
||||
map_add' := by
|
||||
intro φ φ'
|
||||
simp only [map_add]
|
||||
rw [dotProduct_add]
|
||||
map_smul' := by
|
||||
intro r φ
|
||||
simp only [LinearEquiv.map_smul]
|
||||
rw [dotProduct_smul]
|
||||
rfl}
|
||||
map_add' ψ ψ':= by
|
||||
refine LinearMap.ext (fun φ => ?_)
|
||||
simp only [map_add, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.add_apply]
|
||||
rw [add_dotProduct]
|
||||
map_smul' r ψ := by
|
||||
refine LinearMap.ext (fun φ => ?_)
|
||||
simp only [LinearEquiv.map_smul, LinearMap.coe_mk, AddHom.coe_mk]
|
||||
rw [smul_dotProduct]
|
||||
rfl
|
||||
|
||||
/-- The bi-linear map corresponding to contraction of a alt-left-handed Weyl fermion with a
|
||||
left-handed Weyl fermion. -/
|
||||
def altLeftBi : altLeftHanded →ₗ[ℂ] leftHanded →ₗ[ℂ] ℂ where
|
||||
toFun ψ := {
|
||||
toFun := fun φ => ψ.toFin2ℂ ⬝ᵥ φ.toFin2ℂ,
|
||||
map_add' := by
|
||||
intro φ φ'
|
||||
simp only [map_add]
|
||||
rw [dotProduct_add]
|
||||
map_smul' := by
|
||||
intro r φ
|
||||
simp only [LinearEquiv.map_smul]
|
||||
rw [dotProduct_smul]
|
||||
rfl}
|
||||
map_add' ψ ψ':= by
|
||||
refine LinearMap.ext (fun φ => ?_)
|
||||
simp only [map_add, add_dotProduct, vec2_dotProduct, Fin.isValue, LinearMap.coe_mk,
|
||||
AddHom.coe_mk, LinearMap.add_apply]
|
||||
map_smul' ψ ψ' := by
|
||||
refine LinearMap.ext (fun φ => ?_)
|
||||
simp only [_root_.map_smul, smul_dotProduct, vec2_dotProduct, Fin.isValue, smul_eq_mul,
|
||||
LinearMap.coe_mk, AddHom.coe_mk, RingHom.id_apply, LinearMap.smul_apply]
|
||||
|
||||
/-- The bi-linear map corresponding to contraction of a right-handed Weyl fermion with a
|
||||
alt-right-handed Weyl fermion. -/
|
||||
def rightAltBi : rightHanded →ₗ[ℂ] altRightHanded →ₗ[ℂ] ℂ where
|
||||
toFun ψ := {
|
||||
toFun := fun φ => ψ.toFin2ℂ ⬝ᵥ φ.toFin2ℂ,
|
||||
map_add' := by
|
||||
intro φ φ'
|
||||
simp only [map_add]
|
||||
rw [dotProduct_add]
|
||||
map_smul' := by
|
||||
intro r φ
|
||||
simp only [LinearEquiv.map_smul]
|
||||
rw [dotProduct_smul]
|
||||
rfl}
|
||||
map_add' ψ ψ':= by
|
||||
refine LinearMap.ext (fun φ => ?_)
|
||||
simp only [map_add, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.add_apply]
|
||||
rw [add_dotProduct]
|
||||
map_smul' r ψ := by
|
||||
refine LinearMap.ext (fun φ => ?_)
|
||||
simp only [LinearEquiv.map_smul, LinearMap.coe_mk, AddHom.coe_mk]
|
||||
rw [smul_dotProduct]
|
||||
rfl
|
||||
|
||||
/-- The bi-linear map corresponding to contraction of a alt-right-handed Weyl fermion with a
|
||||
right-handed Weyl fermion. -/
|
||||
def altRightBi : altRightHanded →ₗ[ℂ] rightHanded →ₗ[ℂ] ℂ where
|
||||
toFun ψ := {
|
||||
toFun := fun φ => ψ.toFin2ℂ ⬝ᵥ φ.toFin2ℂ,
|
||||
map_add' := by
|
||||
intro φ φ'
|
||||
simp only [map_add]
|
||||
rw [dotProduct_add]
|
||||
map_smul' := by
|
||||
intro r φ
|
||||
simp only [LinearEquiv.map_smul]
|
||||
rw [dotProduct_smul]
|
||||
rfl}
|
||||
map_add' ψ ψ':= by
|
||||
refine LinearMap.ext (fun φ => ?_)
|
||||
simp only [map_add, add_dotProduct, vec2_dotProduct, Fin.isValue, LinearMap.coe_mk,
|
||||
AddHom.coe_mk, LinearMap.add_apply]
|
||||
map_smul' ψ ψ' := by
|
||||
refine LinearMap.ext (fun φ => ?_)
|
||||
simp only [_root_.map_smul, smul_dotProduct, vec2_dotProduct, Fin.isValue, smul_eq_mul,
|
||||
LinearMap.coe_mk, AddHom.coe_mk, RingHom.id_apply, LinearMap.smul_apply]
|
||||
|
||||
/-- The linear map from leftHandedWeyl ⊗ altLeftHandedWeyl to ℂ given by
|
||||
summing over components of leftHandedWeyl and altLeftHandedWeyl in the
|
||||
standard basis (i.e. the dot product).
|
||||
Physically, the contraction of a left-handed Weyl fermion with a alt-left-handed Weyl fermion.
|
||||
In index notation this is ψ^a φ_a. -/
|
||||
def leftAltContraction : leftHanded ⊗ altLeftHanded ⟶ 𝟙_ (Rep ℂ SL(2,ℂ)) where
|
||||
hom := TensorProduct.lift leftAltBi
|
||||
comm M := TensorProduct.ext' fun ψ φ => by
|
||||
change (M.1 *ᵥ ψ.toFin2ℂ) ⬝ᵥ (M.1⁻¹ᵀ *ᵥ φ.toFin2ℂ) = ψ.toFin2ℂ ⬝ᵥ φ.toFin2ℂ
|
||||
rw [dotProduct_mulVec, vecMul_transpose, mulVec_mulVec]
|
||||
simp
|
||||
|
||||
lemma leftAltContraction_hom_tmul (ψ : leftHanded) (φ : altLeftHanded) :
|
||||
leftAltContraction.hom (ψ ⊗ₜ φ) = ψ.toFin2ℂ ⬝ᵥ φ.toFin2ℂ := by
|
||||
rfl
|
||||
|
||||
lemma leftAltContraction_basis (i j : Fin 2) :
|
||||
leftAltContraction.hom (leftBasis i ⊗ₜ altLeftBasis j) = if i.1 = j.1 then (1 : ℂ) else 0 := by
|
||||
rw [leftAltContraction_hom_tmul]
|
||||
simp only [Action.instMonoidalCategory_tensorUnit_V, leftBasis_toFin2ℂ, altLeftBasis_toFin2ℂ,
|
||||
dotProduct_single, mul_one]
|
||||
rw [Pi.single_apply]
|
||||
simp only [Fin.ext_iff]
|
||||
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
|
||||
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
|
||||
|
||||
/-- The linear map from altLeftHandedWeyl ⊗ leftHandedWeyl to ℂ given by
|
||||
summing over components of altLeftHandedWeyl and leftHandedWeyl in the
|
||||
standard basis (i.e. the dot product).
|
||||
Physically, the contraction of a alt-left-handed Weyl fermion with a left-handed Weyl fermion.
|
||||
In index notation this is φ_a ψ^a. -/
|
||||
def altLeftContraction : altLeftHanded ⊗ leftHanded ⟶ 𝟙_ (Rep ℂ SL(2,ℂ)) where
|
||||
hom := TensorProduct.lift altLeftBi
|
||||
comm M := TensorProduct.ext' fun φ ψ => by
|
||||
change (M.1⁻¹ᵀ *ᵥ φ.toFin2ℂ) ⬝ᵥ (M.1 *ᵥ ψ.toFin2ℂ) = φ.toFin2ℂ ⬝ᵥ ψ.toFin2ℂ
|
||||
rw [dotProduct_mulVec, mulVec_transpose, vecMul_vecMul]
|
||||
simp
|
||||
|
||||
lemma altLeftContraction_hom_tmul (φ : altLeftHanded) (ψ : leftHanded) :
|
||||
altLeftContraction.hom (φ ⊗ₜ ψ) = φ.toFin2ℂ ⬝ᵥ ψ.toFin2ℂ := by
|
||||
rfl
|
||||
|
||||
lemma altLeftContraction_basis (i j : Fin 2) :
|
||||
altLeftContraction.hom (altLeftBasis i ⊗ₜ leftBasis j) = if i.1 = j.1 then (1 : ℂ) else 0 := by
|
||||
rw [altLeftContraction_hom_tmul]
|
||||
simp only [Action.instMonoidalCategory_tensorUnit_V, leftBasis_toFin2ℂ, altLeftBasis_toFin2ℂ,
|
||||
dotProduct_single, mul_one]
|
||||
rw [Pi.single_apply]
|
||||
simp only [Fin.ext_iff]
|
||||
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
|
||||
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
|
||||
|
||||
/--
|
||||
The linear map from rightHandedWeyl ⊗ altRightHandedWeyl to ℂ given by
|
||||
summing over components of rightHandedWeyl and altRightHandedWeyl in the
|
||||
standard basis (i.e. the dot product).
|
||||
The contraction of a right-handed Weyl fermion with a left-handed Weyl fermion.
|
||||
In index notation this is ψ^{dot a} φ_{dot a}.
|
||||
-/
|
||||
def rightAltContraction : rightHanded ⊗ altRightHanded ⟶ 𝟙_ (Rep ℂ SL(2,ℂ)) where
|
||||
hom := TensorProduct.lift rightAltBi
|
||||
comm M := TensorProduct.ext' fun ψ φ => by
|
||||
change (M.1.map star *ᵥ ψ.toFin2ℂ) ⬝ᵥ (M.1⁻¹.conjTranspose *ᵥ φ.toFin2ℂ) =
|
||||
ψ.toFin2ℂ ⬝ᵥ φ.toFin2ℂ
|
||||
have h1 : (M.1)⁻¹ᴴ = ((M.1)⁻¹.map star)ᵀ := by rfl
|
||||
rw [dotProduct_mulVec, h1, vecMul_transpose, mulVec_mulVec]
|
||||
have h2 : ((M.1)⁻¹.map star * (M.1).map star) = 1 := by
|
||||
refine transpose_inj.mp ?_
|
||||
rw [transpose_mul]
|
||||
change M.1.conjTranspose * (M.1)⁻¹.conjTranspose = 1ᵀ
|
||||
rw [← @conjTranspose_mul]
|
||||
simp only [SpecialLinearGroup.det_coe, isUnit_iff_ne_zero, ne_eq, one_ne_zero,
|
||||
not_false_eq_true, nonsing_inv_mul, conjTranspose_one, transpose_one]
|
||||
rw [h2]
|
||||
simp only [one_mulVec, vec2_dotProduct, Fin.isValue, RightHandedModule.toFin2ℂEquiv_apply,
|
||||
AltRightHandedModule.toFin2ℂEquiv_apply]
|
||||
|
||||
lemma rightAltContraction_hom_tmul (ψ : rightHanded) (φ : altRightHanded) :
|
||||
rightAltContraction.hom (ψ ⊗ₜ φ) = ψ.toFin2ℂ ⬝ᵥ φ.toFin2ℂ := by
|
||||
rfl
|
||||
|
||||
lemma rightAltContraction_basis (i j : Fin 2) :
|
||||
rightAltContraction.hom (rightBasis i ⊗ₜ altRightBasis j) =
|
||||
if i.1 = j.1 then (1 : ℂ) else 0 := by
|
||||
rw [rightAltContraction_hom_tmul]
|
||||
simp only [Action.instMonoidalCategory_tensorUnit_V, rightBasis_toFin2ℂ, altRightBasis_toFin2ℂ,
|
||||
dotProduct_single, mul_one]
|
||||
rw [Pi.single_apply]
|
||||
simp only [Fin.ext_iff]
|
||||
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
|
||||
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
|
||||
|
||||
/--
|
||||
The linear map from altRightHandedWeyl ⊗ rightHandedWeyl to ℂ given by
|
||||
summing over components of altRightHandedWeyl and rightHandedWeyl in the
|
||||
standard basis (i.e. the dot product).
|
||||
The contraction of a right-handed Weyl fermion with a left-handed Weyl fermion.
|
||||
In index notation this is φ_{dot a} ψ^{dot a}.
|
||||
-/
|
||||
def altRightContraction : altRightHanded ⊗ rightHanded ⟶ 𝟙_ (Rep ℂ SL(2,ℂ)) where
|
||||
hom := TensorProduct.lift altRightBi
|
||||
comm M := TensorProduct.ext' fun φ ψ => by
|
||||
change (M.1⁻¹.conjTranspose *ᵥ φ.toFin2ℂ) ⬝ᵥ (M.1.map star *ᵥ ψ.toFin2ℂ) =
|
||||
φ.toFin2ℂ ⬝ᵥ ψ.toFin2ℂ
|
||||
have h1 : (M.1)⁻¹ᴴ = ((M.1)⁻¹.map star)ᵀ := by rfl
|
||||
rw [dotProduct_mulVec, h1, mulVec_transpose, vecMul_vecMul]
|
||||
have h2 : ((M.1)⁻¹.map star * (M.1).map star) = 1 := by
|
||||
refine transpose_inj.mp ?_
|
||||
rw [transpose_mul]
|
||||
change M.1.conjTranspose * (M.1)⁻¹.conjTranspose = 1ᵀ
|
||||
rw [← @conjTranspose_mul]
|
||||
simp only [SpecialLinearGroup.det_coe, isUnit_iff_ne_zero, ne_eq, one_ne_zero,
|
||||
not_false_eq_true, nonsing_inv_mul, conjTranspose_one, transpose_one]
|
||||
rw [h2]
|
||||
simp only [vecMul_one, vec2_dotProduct, Fin.isValue, AltRightHandedModule.toFin2ℂEquiv_apply,
|
||||
RightHandedModule.toFin2ℂEquiv_apply]
|
||||
|
||||
lemma altRightContraction_hom_tmul (φ : altRightHanded) (ψ : rightHanded) :
|
||||
altRightContraction.hom (φ ⊗ₜ ψ) = φ.toFin2ℂ ⬝ᵥ ψ.toFin2ℂ := by
|
||||
rfl
|
||||
|
||||
lemma altRightContraction_basis (i j : Fin 2) :
|
||||
altRightContraction.hom (altRightBasis i ⊗ₜ rightBasis j) =
|
||||
if i.1 = j.1 then (1 : ℂ) else 0 := by
|
||||
rw [altRightContraction_hom_tmul]
|
||||
simp only [Action.instMonoidalCategory_tensorUnit_V, rightBasis_toFin2ℂ, altRightBasis_toFin2ℂ,
|
||||
dotProduct_single, mul_one]
|
||||
rw [Pi.single_apply]
|
||||
simp only [Fin.ext_iff]
|
||||
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
|
||||
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
|
||||
|
||||
/-!
|
||||
|
||||
## Symmetry properties
|
||||
|
||||
-/
|
||||
|
||||
lemma leftAltContraction_tmul_symm (ψ : leftHanded) (φ : altLeftHanded) :
|
||||
leftAltContraction.hom (ψ ⊗ₜ[ℂ] φ) = altLeftContraction.hom (φ ⊗ₜ[ℂ] ψ) := by
|
||||
rw [leftAltContraction_hom_tmul, altLeftContraction_hom_tmul, dotProduct_comm]
|
||||
|
||||
lemma altLeftContraction_tmul_symm (φ : altLeftHanded) (ψ : leftHanded) :
|
||||
altLeftContraction.hom (φ ⊗ₜ[ℂ] ψ) = leftAltContraction.hom (ψ ⊗ₜ[ℂ] φ) := by
|
||||
rw [leftAltContraction_tmul_symm]
|
||||
|
||||
lemma rightAltContraction_tmul_symm (ψ : rightHanded) (φ : altRightHanded) :
|
||||
rightAltContraction.hom (ψ ⊗ₜ[ℂ] φ) = altRightContraction.hom (φ ⊗ₜ[ℂ] ψ) := by
|
||||
rw [rightAltContraction_hom_tmul, altRightContraction_hom_tmul, dotProduct_comm]
|
||||
|
||||
lemma altRightContraction_tmul_symm (φ : altRightHanded) (ψ : rightHanded) :
|
||||
altRightContraction.hom (φ ⊗ₜ[ℂ] ψ) = rightAltContraction.hom (ψ ⊗ₜ[ℂ] φ) := by
|
||||
rw [rightAltContraction_tmul_symm]
|
||||
|
||||
end
|
||||
end Fermion
|
394
HepLean/Lorentz/Weyl/Metric.lean
Normal file
394
HepLean/Lorentz/Weyl/Metric.lean
Normal file
|
@ -0,0 +1,394 @@
|
|||
/-
|
||||
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.Basic
|
||||
import HepLean.Lorentz.Weyl.Contraction
|
||||
import Mathlib.LinearAlgebra.TensorProduct.Matrix
|
||||
import HepLean.Lorentz.Weyl.Two
|
||||
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]
|
||||
|
||||
lemma comm_metricRaw (M : SL(2,ℂ)) : M.1 * metricRaw = metricRaw * (M.1⁻¹)ᵀ := by
|
||||
rw [metricRaw]
|
||||
rw [SpaceTime.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 [SpaceTime.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 [SpaceTime.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 [SpaceTime.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 := {
|
||||
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
|
||||
ext x : 2
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
|
||||
Function.comp_apply]
|
||||
let x' : ℂ := x
|
||||
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.toFun (1 : ℂ) = leftMetricVal
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
leftMetric, AddHom.toFun_eq_coe, AddHom.coe_mk, 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 := {
|
||||
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
|
||||
ext x : 2
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
|
||||
Function.comp_apply]
|
||||
let x' : ℂ := x
|
||||
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.toFun (1 : ℂ) = altLeftMetricVal
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
altLeftMetric, AddHom.toFun_eq_coe, AddHom.coe_mk, 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 := {
|
||||
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
|
||||
ext x : 2
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
|
||||
Function.comp_apply]
|
||||
let x' : ℂ := x
|
||||
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.toFun (1 : ℂ) = rightMetricVal
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
rightMetric, AddHom.toFun_eq_coe, AddHom.coe_mk, 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 := {
|
||||
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
|
||||
ext x : 2
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
|
||||
Function.comp_apply]
|
||||
let x' : ℂ := x
|
||||
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.toFun (1 : ℂ) = altRightMetricVal
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
altRightMetric, AddHom.toFun_eq_coe, AddHom.coe_mk, 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
|
211
HepLean/Lorentz/Weyl/Modules.lean
Normal file
211
HepLean/Lorentz/Weyl/Modules.lean
Normal file
|
@ -0,0 +1,211 @@
|
|||
/-
|
||||
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.Meta.Informal
|
||||
import HepLean.SpaceTime.SL2C.Basic
|
||||
import Mathlib.RepresentationTheory.Rep
|
||||
import Mathlib.Logic.Equiv.TransferInstance
|
||||
/-!
|
||||
|
||||
## Modules associated with Fermions
|
||||
|
||||
Weyl fermions live in the vector space `ℂ^2`, defined here as `Fin 2 → ℂ`.
|
||||
However if we simply define the Module of Weyl fermions as `Fin 2 → ℂ` we get casting problems,
|
||||
where e.g. left-handed fermions can be cast to right-handed fermions etc.
|
||||
To overcome this, for each type of Weyl fermion we define a structure that wraps `Fin 2 → ℂ`,
|
||||
and these structures we define the instance of a module. This prevents casting between different
|
||||
types of fermions.
|
||||
|
||||
-/
|
||||
|
||||
namespace Fermion
|
||||
noncomputable section
|
||||
|
||||
section LeftHanded
|
||||
|
||||
/-- The module in which left handed fermions live. This is equivalent to `Fin 2 → ℂ`. -/
|
||||
structure LeftHandedModule where
|
||||
/-- The underlying value in `Fin 2 → ℂ`. -/
|
||||
val : Fin 2 → ℂ
|
||||
|
||||
namespace LeftHandedModule
|
||||
|
||||
/-- The equivalence between `LeftHandedModule` and `Fin 2 → ℂ`. -/
|
||||
def toFin2ℂFun : LeftHandedModule ≃ (Fin 2 → ℂ) where
|
||||
toFun v := v.val
|
||||
invFun f := ⟨f⟩
|
||||
left_inv _ := rfl
|
||||
right_inv _ := rfl
|
||||
|
||||
/-- The instance of `AddCommMonoid` on `LeftHandedModule` defined via its equivalence
|
||||
with `Fin 2 → ℂ`. -/
|
||||
instance : AddCommMonoid LeftHandedModule := Equiv.addCommMonoid toFin2ℂFun
|
||||
|
||||
/-- The instance of `AddCommGroup` on `LeftHandedModule` defined via its equivalence
|
||||
with `Fin 2 → ℂ`. -/
|
||||
instance : AddCommGroup LeftHandedModule := Equiv.addCommGroup toFin2ℂFun
|
||||
|
||||
/-- The instance of `Module` on `LeftHandedModule` defined via its equivalence
|
||||
with `Fin 2 → ℂ`. -/
|
||||
instance : Module ℂ LeftHandedModule := Equiv.module ℂ toFin2ℂFun
|
||||
|
||||
/-- The linear equivalence between `LeftHandedModule` and `(Fin 2 → ℂ)`. -/
|
||||
@[simps!]
|
||||
def toFin2ℂEquiv : LeftHandedModule ≃ₗ[ℂ] (Fin 2 → ℂ) where
|
||||
toFun := toFin2ℂFun
|
||||
map_add' := fun _ _ => rfl
|
||||
map_smul' := fun _ _ => rfl
|
||||
invFun := toFin2ℂFun.symm
|
||||
left_inv := fun _ => rfl
|
||||
right_inv := fun _ => rfl
|
||||
|
||||
/-- The underlying element of `Fin 2 → ℂ` of a element in `LeftHandedModule` defined
|
||||
through the linear equivalence `toFin2ℂEquiv`. -/
|
||||
abbrev toFin2ℂ (ψ : LeftHandedModule) := toFin2ℂEquiv ψ
|
||||
|
||||
end LeftHandedModule
|
||||
|
||||
end LeftHanded
|
||||
|
||||
section AltLeftHanded
|
||||
|
||||
/-- The module in which alt-left handed fermions live. This is equivalent to `Fin 2 → ℂ`. -/
|
||||
structure AltLeftHandedModule where
|
||||
/-- The underlying value in `Fin 2 → ℂ`. -/
|
||||
val : Fin 2 → ℂ
|
||||
|
||||
namespace AltLeftHandedModule
|
||||
|
||||
/-- The equivalence between `AltLeftHandedModule` and `Fin 2 → ℂ`. -/
|
||||
def toFin2ℂFun : AltLeftHandedModule ≃ (Fin 2 → ℂ) where
|
||||
toFun v := v.val
|
||||
invFun f := ⟨f⟩
|
||||
left_inv _ := rfl
|
||||
right_inv _ := rfl
|
||||
|
||||
/-- The instance of `AddCommMonoid` on `AltLeftHandedModule` defined via its equivalence
|
||||
with `Fin 2 → ℂ`. -/
|
||||
instance : AddCommMonoid AltLeftHandedModule := Equiv.addCommMonoid toFin2ℂFun
|
||||
|
||||
/-- The instance of `AddCommGroup` on `AltLeftHandedModule` defined via its equivalence
|
||||
with `Fin 2 → ℂ`. -/
|
||||
instance : AddCommGroup AltLeftHandedModule := Equiv.addCommGroup toFin2ℂFun
|
||||
|
||||
/-- The instance of `Module` on `AltLeftHandedModule` defined via its equivalence
|
||||
with `Fin 2 → ℂ`. -/
|
||||
instance : Module ℂ AltLeftHandedModule := Equiv.module ℂ toFin2ℂFun
|
||||
|
||||
/-- The linear equivalence between `AltLeftHandedModule` and `(Fin 2 → ℂ)`. -/
|
||||
@[simps!]
|
||||
def toFin2ℂEquiv : AltLeftHandedModule ≃ₗ[ℂ] (Fin 2 → ℂ) where
|
||||
toFun := toFin2ℂFun
|
||||
map_add' := fun _ _ => rfl
|
||||
map_smul' := fun _ _ => rfl
|
||||
invFun := toFin2ℂFun.symm
|
||||
left_inv := fun _ => rfl
|
||||
right_inv := fun _ => rfl
|
||||
|
||||
/-- The underlying element of `Fin 2 → ℂ` of a element in `AltLeftHandedModule` defined
|
||||
through the linear equivalence `toFin2ℂEquiv`. -/
|
||||
abbrev toFin2ℂ (ψ : AltLeftHandedModule) := toFin2ℂEquiv ψ
|
||||
|
||||
end AltLeftHandedModule
|
||||
|
||||
end AltLeftHanded
|
||||
|
||||
section RightHanded
|
||||
|
||||
/-- The module in which right handed fermions live. This is equivalent to `Fin 2 → ℂ`. -/
|
||||
structure RightHandedModule where
|
||||
/-- The underlying value in `Fin 2 → ℂ`. -/
|
||||
val : Fin 2 → ℂ
|
||||
|
||||
namespace RightHandedModule
|
||||
|
||||
/-- The equivalence between `RightHandedModule` and `Fin 2 → ℂ`. -/
|
||||
def toFin2ℂFun : RightHandedModule ≃ (Fin 2 → ℂ) where
|
||||
toFun v := v.val
|
||||
invFun f := ⟨f⟩
|
||||
left_inv _ := rfl
|
||||
right_inv _ := rfl
|
||||
|
||||
/-- The instance of `AddCommMonoid` on `RightHandedModule` defined via its equivalence
|
||||
with `Fin 2 → ℂ`. -/
|
||||
instance : AddCommMonoid RightHandedModule := Equiv.addCommMonoid toFin2ℂFun
|
||||
|
||||
/-- The instance of `AddCommGroup` on `RightHandedModule` defined via its equivalence
|
||||
with `Fin 2 → ℂ`. -/
|
||||
instance : AddCommGroup RightHandedModule := Equiv.addCommGroup toFin2ℂFun
|
||||
|
||||
/-- The instance of `Module` on `RightHandedModule` defined via its equivalence
|
||||
with `Fin 2 → ℂ`. -/
|
||||
instance : Module ℂ RightHandedModule := Equiv.module ℂ toFin2ℂFun
|
||||
|
||||
/-- The linear equivalence between `RightHandedModule` and `(Fin 2 → ℂ)`. -/
|
||||
@[simps!]
|
||||
def toFin2ℂEquiv : RightHandedModule ≃ₗ[ℂ] (Fin 2 → ℂ) where
|
||||
toFun := toFin2ℂFun
|
||||
map_add' := fun _ _ => rfl
|
||||
map_smul' := fun _ _ => rfl
|
||||
invFun := toFin2ℂFun.symm
|
||||
left_inv := fun _ => rfl
|
||||
right_inv := fun _ => rfl
|
||||
|
||||
/-- The underlying element of `Fin 2 → ℂ` of a element in `RightHandedModule` defined
|
||||
through the linear equivalence `toFin2ℂEquiv`. -/
|
||||
abbrev toFin2ℂ (ψ : RightHandedModule) := toFin2ℂEquiv ψ
|
||||
|
||||
end RightHandedModule
|
||||
|
||||
end RightHanded
|
||||
|
||||
section AltRightHanded
|
||||
|
||||
/-- The module in which alt-right handed fermions live. This is equivalent to `Fin 2 → ℂ`. -/
|
||||
structure AltRightHandedModule where
|
||||
/-- The underlying value in `Fin 2 → ℂ`. -/
|
||||
val : Fin 2 → ℂ
|
||||
|
||||
namespace AltRightHandedModule
|
||||
|
||||
/-- The equivalence between `AltRightHandedModule` and `Fin 2 → ℂ`. -/
|
||||
def toFin2ℂFun : AltRightHandedModule ≃ (Fin 2 → ℂ) where
|
||||
toFun v := v.val
|
||||
invFun f := ⟨f⟩
|
||||
left_inv _ := rfl
|
||||
right_inv _ := rfl
|
||||
|
||||
/-- The instance of `AddCommMonoid` on `AltRightHandedModule` defined via its equivalence
|
||||
with `Fin 2 → ℂ`. -/
|
||||
instance : AddCommMonoid AltRightHandedModule := Equiv.addCommMonoid toFin2ℂFun
|
||||
|
||||
/-- The instance of `AddCommGroup` on `AltRightHandedModule` defined via its equivalence
|
||||
with `Fin 2 → ℂ`. -/
|
||||
instance : AddCommGroup AltRightHandedModule := Equiv.addCommGroup toFin2ℂFun
|
||||
|
||||
/-- The instance of `Module` on `AltRightHandedModule` defined via its equivalence
|
||||
with `Fin 2 → ℂ`. -/
|
||||
instance : Module ℂ AltRightHandedModule := Equiv.module ℂ toFin2ℂFun
|
||||
|
||||
/-- The linear equivalence between `AltRightHandedModule` and `(Fin 2 → ℂ)`. -/
|
||||
@[simps!]
|
||||
def toFin2ℂEquiv : AltRightHandedModule ≃ₗ[ℂ] (Fin 2 → ℂ) where
|
||||
toFun := toFin2ℂFun
|
||||
map_add' := fun _ _ => rfl
|
||||
map_smul' := fun _ _ => rfl
|
||||
invFun := toFin2ℂFun.symm
|
||||
left_inv := fun _ => rfl
|
||||
right_inv := fun _ => rfl
|
||||
|
||||
/-- The underlying element of `Fin 2 → ℂ` of a element in `AltRightHandedModule` defined
|
||||
through the linear equivalence `toFin2ℂEquiv`. -/
|
||||
abbrev toFin2ℂ (ψ : AltRightHandedModule) := toFin2ℂEquiv ψ
|
||||
|
||||
end AltRightHandedModule
|
||||
|
||||
end AltRightHanded
|
||||
|
||||
end
|
||||
end Fermion
|
746
HepLean/Lorentz/Weyl/Two.lean
Normal file
746
HepLean/Lorentz/Weyl/Two.lean
Normal file
|
@ -0,0 +1,746 @@
|
|||
/-
|
||||
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.Basic
|
||||
import HepLean.Lorentz.Weyl.Contraction
|
||||
import Mathlib.LinearAlgebra.TensorProduct.Matrix
|
||||
/-!
|
||||
|
||||
# Tensor product of two Weyl fermion
|
||||
|
||||
-/
|
||||
|
||||
namespace Fermion
|
||||
noncomputable section
|
||||
|
||||
open Matrix
|
||||
open MatrixGroups
|
||||
open Complex
|
||||
open TensorProduct
|
||||
open CategoryTheory.MonoidalCategory
|
||||
|
||||
/-!
|
||||
|
||||
## Equivalences to matrices.
|
||||
|
||||
-/
|
||||
|
||||
/-- Equivalence of `leftHanded ⊗ leftHanded` to `2 x 2` complex matrices. -/
|
||||
def leftLeftToMatrix : (leftHanded ⊗ leftHanded).V ≃ₗ[ℂ] Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
(Basis.tensorProduct leftBasis leftBasis).repr ≪≫ₗ
|
||||
Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2) ≪≫ₗ
|
||||
LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)
|
||||
|
||||
/-- Expanding `leftLeftToMatrix` in terms of the standard basis. -/
|
||||
lemma leftLeftToMatrix_symm_expand_tmul (M : Matrix (Fin 2) (Fin 2) ℂ) :
|
||||
leftLeftToMatrix.symm M = ∑ i, ∑ j, M i j • (leftBasis i ⊗ₜ[ℂ] leftBasis j) := by
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, leftLeftToMatrix,
|
||||
LinearEquiv.trans_symm, LinearEquiv.trans_apply, Basis.repr_symm_apply]
|
||||
rw [Finsupp.linearCombination_apply_of_mem_supported ℂ (s := Finset.univ)]
|
||||
· erw [Finset.sum_product]
|
||||
refine Finset.sum_congr rfl (fun i _ => Finset.sum_congr rfl (fun j _ => ?_))
|
||||
erw [Basis.tensorProduct_apply leftBasis leftBasis i j]
|
||||
rfl
|
||||
· simp
|
||||
|
||||
/-- Equivalence of `altLeftHanded ⊗ altLeftHanded` to `2 x 2` complex matrices. -/
|
||||
def altLeftaltLeftToMatrix : (altLeftHanded ⊗ altLeftHanded).V ≃ₗ[ℂ] Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
(Basis.tensorProduct altLeftBasis altLeftBasis).repr ≪≫ₗ
|
||||
Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2) ≪≫ₗ
|
||||
LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)
|
||||
|
||||
/-- Expanding `altLeftaltLeftToMatrix` in terms of the standard basis. -/
|
||||
lemma altLeftaltLeftToMatrix_symm_expand_tmul (M : Matrix (Fin 2) (Fin 2) ℂ) :
|
||||
altLeftaltLeftToMatrix.symm M = ∑ i, ∑ j, M i j • (altLeftBasis i ⊗ₜ[ℂ] altLeftBasis j) := by
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, altLeftaltLeftToMatrix,
|
||||
LinearEquiv.trans_symm, LinearEquiv.trans_apply, Basis.repr_symm_apply]
|
||||
rw [Finsupp.linearCombination_apply_of_mem_supported ℂ (s := Finset.univ)]
|
||||
· erw [Finset.sum_product]
|
||||
refine Finset.sum_congr rfl (fun i _ => Finset.sum_congr rfl (fun j _ => ?_))
|
||||
erw [Basis.tensorProduct_apply altLeftBasis altLeftBasis i j]
|
||||
rfl
|
||||
· simp
|
||||
|
||||
/-- Equivalence of `leftHanded ⊗ altLeftHanded` to `2 x 2` complex matrices. -/
|
||||
def leftAltLeftToMatrix : (leftHanded ⊗ altLeftHanded).V ≃ₗ[ℂ] Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
(Basis.tensorProduct leftBasis altLeftBasis).repr ≪≫ₗ
|
||||
Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2) ≪≫ₗ
|
||||
LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)
|
||||
|
||||
/-- Expanding `leftAltLeftToMatrix` in terms of the standard basis. -/
|
||||
lemma leftAltLeftToMatrix_symm_expand_tmul (M : Matrix (Fin 2) (Fin 2) ℂ) :
|
||||
leftAltLeftToMatrix.symm M = ∑ i, ∑ j, M i j • (leftBasis i ⊗ₜ[ℂ] altLeftBasis j) := by
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, leftAltLeftToMatrix,
|
||||
LinearEquiv.trans_symm, LinearEquiv.trans_apply, Basis.repr_symm_apply]
|
||||
rw [Finsupp.linearCombination_apply_of_mem_supported ℂ (s := Finset.univ)]
|
||||
· erw [Finset.sum_product]
|
||||
refine Finset.sum_congr rfl (fun i _ => Finset.sum_congr rfl (fun j _ => ?_))
|
||||
erw [Basis.tensorProduct_apply leftBasis altLeftBasis i j]
|
||||
rfl
|
||||
· simp
|
||||
|
||||
/-- Equivalence of `altLeftHanded ⊗ leftHanded` to `2 x 2` complex matrices. -/
|
||||
def altLeftLeftToMatrix : (altLeftHanded ⊗ leftHanded).V ≃ₗ[ℂ] Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
(Basis.tensorProduct altLeftBasis leftBasis).repr ≪≫ₗ
|
||||
Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2) ≪≫ₗ
|
||||
LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)
|
||||
|
||||
/-- Expanding `altLeftLeftToMatrix` in terms of the standard basis. -/
|
||||
lemma altLeftLeftToMatrix_symm_expand_tmul (M : Matrix (Fin 2) (Fin 2) ℂ) :
|
||||
altLeftLeftToMatrix.symm M = ∑ i, ∑ j, M i j • (altLeftBasis i ⊗ₜ[ℂ] leftBasis j) := by
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, altLeftLeftToMatrix,
|
||||
LinearEquiv.trans_symm, LinearEquiv.trans_apply, Basis.repr_symm_apply]
|
||||
rw [Finsupp.linearCombination_apply_of_mem_supported ℂ (s := Finset.univ)]
|
||||
· erw [Finset.sum_product]
|
||||
refine Finset.sum_congr rfl (fun i _ => Finset.sum_congr rfl (fun j _ => ?_))
|
||||
erw [Basis.tensorProduct_apply altLeftBasis leftBasis i j]
|
||||
rfl
|
||||
· simp
|
||||
|
||||
/-- Equivalence of `rightHanded ⊗ rightHanded` to `2 x 2` complex matrices. -/
|
||||
def rightRightToMatrix : (rightHanded ⊗ rightHanded).V ≃ₗ[ℂ] Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
(Basis.tensorProduct rightBasis rightBasis).repr ≪≫ₗ
|
||||
Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2) ≪≫ₗ
|
||||
LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)
|
||||
|
||||
/-- Expanding `rightRightToMatrix` in terms of the standard basis. -/
|
||||
lemma rightRightToMatrix_symm_expand_tmul (M : Matrix (Fin 2) (Fin 2) ℂ) :
|
||||
rightRightToMatrix.symm M = ∑ i, ∑ j, M i j • (rightBasis i ⊗ₜ[ℂ] rightBasis j) := by
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, rightRightToMatrix,
|
||||
LinearEquiv.trans_symm, LinearEquiv.trans_apply, Basis.repr_symm_apply]
|
||||
rw [Finsupp.linearCombination_apply_of_mem_supported ℂ (s := Finset.univ)]
|
||||
· erw [Finset.sum_product]
|
||||
refine Finset.sum_congr rfl (fun i _ => Finset.sum_congr rfl (fun j _ => ?_))
|
||||
erw [Basis.tensorProduct_apply rightBasis rightBasis i j]
|
||||
rfl
|
||||
· simp
|
||||
|
||||
/-- Equivalence of `altRightHanded ⊗ altRightHanded` to `2 x 2` complex matrices. -/
|
||||
def altRightAltRightToMatrix : (altRightHanded ⊗ altRightHanded).V ≃ₗ[ℂ] Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
(Basis.tensorProduct altRightBasis altRightBasis).repr ≪≫ₗ
|
||||
Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2) ≪≫ₗ
|
||||
LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)
|
||||
|
||||
/-- Expanding `altRightAltRightToMatrix` in terms of the standard basis. -/
|
||||
lemma altRightAltRightToMatrix_symm_expand_tmul (M : Matrix (Fin 2) (Fin 2) ℂ) :
|
||||
altRightAltRightToMatrix.symm M =
|
||||
∑ i, ∑ j, M i j • (altRightBasis i ⊗ₜ[ℂ] altRightBasis j) := by
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, altRightAltRightToMatrix,
|
||||
LinearEquiv.trans_symm, LinearEquiv.trans_apply, Basis.repr_symm_apply]
|
||||
rw [Finsupp.linearCombination_apply_of_mem_supported ℂ (s := Finset.univ)]
|
||||
· erw [Finset.sum_product]
|
||||
refine Finset.sum_congr rfl (fun i _ => Finset.sum_congr rfl (fun j _ => ?_))
|
||||
erw [Basis.tensorProduct_apply altRightBasis altRightBasis i j]
|
||||
rfl
|
||||
· simp
|
||||
|
||||
/-- Equivalence of `rightHanded ⊗ altRightHanded` to `2 x 2` complex matrices. -/
|
||||
def rightAltRightToMatrix : (rightHanded ⊗ altRightHanded).V ≃ₗ[ℂ] Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
(Basis.tensorProduct rightBasis altRightBasis).repr ≪≫ₗ
|
||||
Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2) ≪≫ₗ
|
||||
LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)
|
||||
|
||||
/-- Expanding `rightAltRightToMatrix` in terms of the standard basis. -/
|
||||
lemma rightAltRightToMatrix_symm_expand_tmul (M : Matrix (Fin 2) (Fin 2) ℂ) :
|
||||
rightAltRightToMatrix.symm M = ∑ i, ∑ j, M i j • (rightBasis i ⊗ₜ[ℂ] altRightBasis j) := by
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, rightAltRightToMatrix, LinearEquiv.trans_symm,
|
||||
LinearEquiv.trans_apply, Basis.repr_symm_apply]
|
||||
rw [Finsupp.linearCombination_apply_of_mem_supported ℂ (s := Finset.univ)]
|
||||
· erw [Finset.sum_product]
|
||||
refine Finset.sum_congr rfl (fun i _ => Finset.sum_congr rfl (fun j _ => ?_))
|
||||
erw [Basis.tensorProduct_apply rightBasis altRightBasis i j]
|
||||
rfl
|
||||
· simp
|
||||
|
||||
/-- Equivalence of `altRightHanded ⊗ rightHanded` to `2 x 2` complex matrices. -/
|
||||
def altRightRightToMatrix : (altRightHanded ⊗ rightHanded).V ≃ₗ[ℂ] Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
(Basis.tensorProduct altRightBasis rightBasis).repr ≪≫ₗ
|
||||
Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2) ≪≫ₗ
|
||||
LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)
|
||||
|
||||
/-- Expanding `altRightRightToMatrix` in terms of the standard basis. -/
|
||||
lemma altRightRightToMatrix_symm_expand_tmul (M : Matrix (Fin 2) (Fin 2) ℂ) :
|
||||
altRightRightToMatrix.symm M = ∑ i, ∑ j, M i j • (altRightBasis i ⊗ₜ[ℂ] rightBasis j) := by
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, altRightRightToMatrix, LinearEquiv.trans_symm,
|
||||
LinearEquiv.trans_apply, Basis.repr_symm_apply]
|
||||
rw [Finsupp.linearCombination_apply_of_mem_supported ℂ (s := Finset.univ)]
|
||||
· erw [Finset.sum_product]
|
||||
refine Finset.sum_congr rfl (fun i _ => Finset.sum_congr rfl (fun j _ => ?_))
|
||||
erw [Basis.tensorProduct_apply altRightBasis rightBasis i j]
|
||||
rfl
|
||||
· simp
|
||||
|
||||
/-- Equivalence of `altLeftHanded ⊗ altRightHanded` to `2 x 2` complex matrices. -/
|
||||
def altLeftAltRightToMatrix : (altLeftHanded ⊗ altRightHanded).V ≃ₗ[ℂ] Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
(Basis.tensorProduct altLeftBasis altRightBasis).repr ≪≫ₗ
|
||||
Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2) ≪≫ₗ
|
||||
LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)
|
||||
|
||||
/-- Expanding `altLeftAltRightToMatrix` in terms of the standard basis. -/
|
||||
lemma altLeftAltRightToMatrix_symm_expand_tmul (M : Matrix (Fin 2) (Fin 2) ℂ) :
|
||||
altLeftAltRightToMatrix.symm M = ∑ i, ∑ j, M i j • (altLeftBasis i ⊗ₜ[ℂ] altRightBasis j) := by
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, altLeftAltRightToMatrix,
|
||||
LinearEquiv.trans_symm, LinearEquiv.trans_apply, Basis.repr_symm_apply]
|
||||
rw [Finsupp.linearCombination_apply_of_mem_supported ℂ (s := Finset.univ)]
|
||||
· erw [Finset.sum_product]
|
||||
refine Finset.sum_congr rfl (fun i _ => Finset.sum_congr rfl (fun j _ => ?_))
|
||||
erw [Basis.tensorProduct_apply altLeftBasis altRightBasis i j]
|
||||
rfl
|
||||
· simp
|
||||
|
||||
/-- Equivalence of `leftHanded ⊗ rightHanded` to `2 x 2` complex matrices. -/
|
||||
def leftRightToMatrix : (leftHanded ⊗ rightHanded).V ≃ₗ[ℂ] Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
(Basis.tensorProduct leftBasis rightBasis).repr ≪≫ₗ
|
||||
Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2) ≪≫ₗ
|
||||
LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)
|
||||
|
||||
/-- Expanding `leftRightToMatrix` in terms of the standard basis. -/
|
||||
lemma leftRightToMatrix_symm_expand_tmul (M : Matrix (Fin 2) (Fin 2) ℂ) :
|
||||
leftRightToMatrix.symm M = ∑ i, ∑ j, M i j • (leftBasis i ⊗ₜ[ℂ] rightBasis j) := by
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, leftRightToMatrix, LinearEquiv.trans_symm,
|
||||
LinearEquiv.trans_apply, Basis.repr_symm_apply]
|
||||
rw [Finsupp.linearCombination_apply_of_mem_supported ℂ (s := Finset.univ)]
|
||||
· erw [Finset.sum_product]
|
||||
refine Finset.sum_congr rfl (fun i _ => Finset.sum_congr rfl (fun j _ => ?_))
|
||||
erw [Basis.tensorProduct_apply leftBasis rightBasis i j]
|
||||
rfl
|
||||
· simp
|
||||
|
||||
/-!
|
||||
|
||||
## Group actions
|
||||
|
||||
-/
|
||||
|
||||
/-- The group action of `SL(2,ℂ)` on `leftHanded ⊗ leftHanded` is equivalent to
|
||||
`M.1 * leftLeftToMatrix v * (M.1)ᵀ`. -/
|
||||
lemma leftLeftToMatrix_ρ (v : (leftHanded ⊗ leftHanded).V) (M : SL(2,ℂ)) :
|
||||
leftLeftToMatrix (TensorProduct.map (leftHanded.ρ M) (leftHanded.ρ M) v) =
|
||||
M.1 * leftLeftToMatrix v * (M.1)ᵀ := by
|
||||
nth_rewrite 1 [leftLeftToMatrix]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
|
||||
trans (LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)) ((LinearMap.toMatrix
|
||||
(leftBasis.tensorProduct leftBasis) (leftBasis.tensorProduct leftBasis)
|
||||
(TensorProduct.map (leftHanded.ρ M) (leftHanded.ρ M)))
|
||||
*ᵥ ((Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2))
|
||||
((leftBasis.tensorProduct leftBasis).repr (v))))
|
||||
· apply congrArg
|
||||
have h1 := (LinearMap.toMatrix_mulVec_repr (leftBasis.tensorProduct leftBasis)
|
||||
(leftBasis.tensorProduct leftBasis) (TensorProduct.map (leftHanded.ρ M) (leftHanded.ρ M)) v)
|
||||
erw [h1]
|
||||
rfl
|
||||
rw [TensorProduct.toMatrix_map]
|
||||
funext i j
|
||||
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
|
||||
((LinearMap.toMatrix leftBasis leftBasis) (leftHanded.ρ M))
|
||||
((LinearMap.toMatrix leftBasis leftBasis) (leftHanded.ρ M)) (i, j) k)
|
||||
* leftLeftToMatrix v k.1 k.2) = _
|
||||
erw [Finset.sum_product]
|
||||
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
|
||||
have h1 : ∑ x : Fin 2, (∑ j : Fin 2, M.1 i j * leftLeftToMatrix v j x) * M.1 j x
|
||||
= ∑ x : Fin 2, ∑ x1 : Fin 2, (M.1 i x1 * leftLeftToMatrix v x1 x) * M.1 j x := by
|
||||
congr
|
||||
funext x
|
||||
rw [Finset.sum_mul]
|
||||
erw [h1]
|
||||
rw [Finset.sum_comm]
|
||||
congr
|
||||
funext x
|
||||
congr
|
||||
funext x1
|
||||
simp only [leftBasis_ρ_apply, Finsupp.linearEquivFunOnFinite_apply,
|
||||
Action.instMonoidalCategory_tensorObj_V]
|
||||
rw [mul_assoc]
|
||||
nth_rewrite 2 [mul_comm]
|
||||
rw [← mul_assoc]
|
||||
|
||||
/-- The group action of `SL(2,ℂ)` on `altLeftHanded ⊗ altLeftHanded` is equivalent to
|
||||
`(M.1⁻¹)ᵀ * leftLeftToMatrix v * (M.1⁻¹)`. -/
|
||||
lemma altLeftaltLeftToMatrix_ρ (v : (altLeftHanded ⊗ altLeftHanded).V) (M : SL(2,ℂ)) :
|
||||
altLeftaltLeftToMatrix (TensorProduct.map (altLeftHanded.ρ M) (altLeftHanded.ρ M) v) =
|
||||
(M.1⁻¹)ᵀ * altLeftaltLeftToMatrix v * (M.1⁻¹) := by
|
||||
nth_rewrite 1 [altLeftaltLeftToMatrix]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
|
||||
trans (LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)) ((LinearMap.toMatrix
|
||||
(altLeftBasis.tensorProduct altLeftBasis) (altLeftBasis.tensorProduct altLeftBasis)
|
||||
(TensorProduct.map (altLeftHanded.ρ M) (altLeftHanded.ρ M)))
|
||||
*ᵥ ((Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2))
|
||||
((altLeftBasis.tensorProduct altLeftBasis).repr v)))
|
||||
· apply congrArg
|
||||
have h1 := (LinearMap.toMatrix_mulVec_repr (altLeftBasis.tensorProduct altLeftBasis)
|
||||
(altLeftBasis.tensorProduct altLeftBasis)
|
||||
(TensorProduct.map (altLeftHanded.ρ M) (altLeftHanded.ρ M)) v)
|
||||
erw [h1]
|
||||
rfl
|
||||
rw [TensorProduct.toMatrix_map]
|
||||
funext i j
|
||||
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
|
||||
((LinearMap.toMatrix altLeftBasis altLeftBasis) (altLeftHanded.ρ M))
|
||||
((LinearMap.toMatrix altLeftBasis altLeftBasis) (altLeftHanded.ρ M)) (i, j) k)
|
||||
* altLeftaltLeftToMatrix v k.1 k.2) = _
|
||||
erw [Finset.sum_product]
|
||||
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
|
||||
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (M.1)⁻¹ x1 i * altLeftaltLeftToMatrix v x1 x) * (M.1)⁻¹ x j
|
||||
= ∑ x : Fin 2, ∑ x1 : Fin 2, ((M.1)⁻¹ x1 i * altLeftaltLeftToMatrix v x1 x) * (M.1)⁻¹ x j := by
|
||||
congr
|
||||
funext x
|
||||
rw [Finset.sum_mul]
|
||||
erw [h1]
|
||||
rw [Finset.sum_comm]
|
||||
congr
|
||||
funext x
|
||||
congr
|
||||
funext x1
|
||||
simp only [altLeftBasis_ρ_apply, transpose_apply, Action.instMonoidalCategory_tensorObj_V]
|
||||
ring
|
||||
|
||||
/-- The group action of `SL(2,ℂ)` on `leftHanded ⊗ altLeftHanded` is equivalent to
|
||||
`M.1 * leftAltLeftToMatrix v * (M.1⁻¹)`. -/
|
||||
lemma leftAltLeftToMatrix_ρ (v : (leftHanded ⊗ altLeftHanded).V) (M : SL(2,ℂ)) :
|
||||
leftAltLeftToMatrix (TensorProduct.map (leftHanded.ρ M) (altLeftHanded.ρ M) v) =
|
||||
M.1 * leftAltLeftToMatrix v * (M.1⁻¹) := by
|
||||
nth_rewrite 1 [leftAltLeftToMatrix]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
|
||||
trans (LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)) ((LinearMap.toMatrix
|
||||
(leftBasis.tensorProduct altLeftBasis) (leftBasis.tensorProduct altLeftBasis)
|
||||
(TensorProduct.map (leftHanded.ρ M) (altLeftHanded.ρ M)))
|
||||
*ᵥ ((Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2))
|
||||
((leftBasis.tensorProduct altLeftBasis).repr (v))))
|
||||
· apply congrArg
|
||||
have h1 := (LinearMap.toMatrix_mulVec_repr (leftBasis.tensorProduct altLeftBasis)
|
||||
(leftBasis.tensorProduct altLeftBasis)
|
||||
(TensorProduct.map (leftHanded.ρ M) (altLeftHanded.ρ M)) v)
|
||||
erw [h1]
|
||||
rfl
|
||||
rw [TensorProduct.toMatrix_map]
|
||||
funext i j
|
||||
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
|
||||
((LinearMap.toMatrix leftBasis leftBasis) (leftHanded.ρ M))
|
||||
((LinearMap.toMatrix altLeftBasis altLeftBasis) (altLeftHanded.ρ M)) (i, j) k)
|
||||
* leftAltLeftToMatrix v k.1 k.2) = _
|
||||
erw [Finset.sum_product]
|
||||
simp_rw [kroneckerMap_apply, Matrix.mul_apply]
|
||||
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, M.1 i x1 * leftAltLeftToMatrix v x1 x) * (M.1⁻¹) x j
|
||||
= ∑ x : Fin 2, ∑ x1 : Fin 2, (M.1 i x1 * leftAltLeftToMatrix v x1 x) * (M.1⁻¹) x j := by
|
||||
congr
|
||||
funext x
|
||||
rw [Finset.sum_mul]
|
||||
erw [h1]
|
||||
rw [Finset.sum_comm]
|
||||
congr
|
||||
funext x
|
||||
congr
|
||||
funext x1
|
||||
simp only [leftBasis_ρ_apply, altLeftBasis_ρ_apply, transpose_apply,
|
||||
Action.instMonoidalCategory_tensorObj_V]
|
||||
ring
|
||||
|
||||
/-- The group action of `SL(2,ℂ)` on `altLeftHanded ⊗ leftHanded` is equivalent to
|
||||
`(M.1⁻¹)ᵀ * leftAltLeftToMatrix v * (M.1)ᵀ`. -/
|
||||
lemma altLeftLeftToMatrix_ρ (v : (altLeftHanded ⊗ leftHanded).V) (M : SL(2,ℂ)) :
|
||||
altLeftLeftToMatrix (TensorProduct.map (altLeftHanded.ρ M) (leftHanded.ρ M) v) =
|
||||
(M.1⁻¹)ᵀ * altLeftLeftToMatrix v * (M.1)ᵀ := by
|
||||
nth_rewrite 1 [altLeftLeftToMatrix]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
|
||||
trans (LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)) ((LinearMap.toMatrix
|
||||
(altLeftBasis.tensorProduct leftBasis) (altLeftBasis.tensorProduct leftBasis)
|
||||
(TensorProduct.map (altLeftHanded.ρ M) (leftHanded.ρ M)))
|
||||
*ᵥ ((Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2))
|
||||
((altLeftBasis.tensorProduct leftBasis).repr (v))))
|
||||
· apply congrArg
|
||||
have h1 := (LinearMap.toMatrix_mulVec_repr (altLeftBasis.tensorProduct leftBasis)
|
||||
(altLeftBasis.tensorProduct leftBasis)
|
||||
(TensorProduct.map (altLeftHanded.ρ M) (leftHanded.ρ M)) v)
|
||||
erw [h1]
|
||||
rfl
|
||||
rw [TensorProduct.toMatrix_map]
|
||||
funext i j
|
||||
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
|
||||
((LinearMap.toMatrix altLeftBasis altLeftBasis) (altLeftHanded.ρ M))
|
||||
((LinearMap.toMatrix leftBasis leftBasis) (leftHanded.ρ M)) (i, j) k)
|
||||
* altLeftLeftToMatrix v k.1 k.2) = _
|
||||
erw [Finset.sum_product]
|
||||
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
|
||||
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (M.1)⁻¹ x1 i * altLeftLeftToMatrix v x1 x) * M.1 j x
|
||||
= ∑ x : Fin 2, ∑ x1 : Fin 2, ((M.1)⁻¹ x1 i * altLeftLeftToMatrix v x1 x) * M.1 j x:= by
|
||||
congr
|
||||
funext x
|
||||
rw [Finset.sum_mul]
|
||||
erw [h1]
|
||||
rw [Finset.sum_comm]
|
||||
congr
|
||||
funext x
|
||||
congr
|
||||
funext x1
|
||||
simp only [altLeftBasis_ρ_apply, leftBasis_ρ_apply, transpose_apply,
|
||||
Action.instMonoidalCategory_tensorObj_V]
|
||||
ring
|
||||
|
||||
/-- The group action of `SL(2,ℂ)` on `rightHanded ⊗ rightHanded` is equivalent to
|
||||
`(M.1.map star) * rightRightToMatrix v * ((M.1.map star))ᵀ`. -/
|
||||
lemma rightRightToMatrix_ρ (v : (rightHanded ⊗ rightHanded).V) (M : SL(2,ℂ)) :
|
||||
rightRightToMatrix (TensorProduct.map (rightHanded.ρ M) (rightHanded.ρ M) v) =
|
||||
(M.1.map star) * rightRightToMatrix v * ((M.1.map star))ᵀ := by
|
||||
nth_rewrite 1 [rightRightToMatrix]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
|
||||
trans (LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)) ((LinearMap.toMatrix
|
||||
(rightBasis.tensorProduct rightBasis) (rightBasis.tensorProduct rightBasis)
|
||||
(TensorProduct.map (rightHanded.ρ M) (rightHanded.ρ M)))
|
||||
*ᵥ ((Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2))
|
||||
((rightBasis.tensorProduct rightBasis).repr (v))))
|
||||
· apply congrArg
|
||||
have h1 := (LinearMap.toMatrix_mulVec_repr (rightBasis.tensorProduct rightBasis)
|
||||
(rightBasis.tensorProduct rightBasis)
|
||||
(TensorProduct.map (rightHanded.ρ M) (rightHanded.ρ M)) v)
|
||||
erw [h1]
|
||||
rfl
|
||||
rw [TensorProduct.toMatrix_map]
|
||||
funext i j
|
||||
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
|
||||
((LinearMap.toMatrix rightBasis rightBasis) (rightHanded.ρ M))
|
||||
((LinearMap.toMatrix rightBasis rightBasis) (rightHanded.ρ M)) (i, j) k)
|
||||
* rightRightToMatrix v k.1 k.2) = _
|
||||
erw [Finset.sum_product]
|
||||
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
|
||||
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (M.1.map star) i x1 * rightRightToMatrix v x1 x) *
|
||||
(M.1.map star) j x = ∑ x : Fin 2, ∑ x1 : Fin 2,
|
||||
((M.1.map star) i x1 * rightRightToMatrix v x1 x) * (M.1.map star) j x:= by
|
||||
congr
|
||||
funext x
|
||||
rw [Finset.sum_mul]
|
||||
erw [h1]
|
||||
rw [Finset.sum_comm]
|
||||
congr
|
||||
funext x
|
||||
congr
|
||||
funext x1
|
||||
simp only [rightBasis_ρ_apply, Finsupp.linearEquivFunOnFinite_apply,
|
||||
Action.instMonoidalCategory_tensorObj_V]
|
||||
ring
|
||||
|
||||
/-- The group action of `SL(2,ℂ)` on `altRightHanded ⊗ altRightHanded` is equivalent to
|
||||
`((M.1⁻¹).conjTranspose * rightRightToMatrix v * (((M.1⁻¹).conjTranspose)ᵀ`. -/
|
||||
lemma altRightAltRightToMatrix_ρ (v : (altRightHanded ⊗ altRightHanded).V) (M : SL(2,ℂ)) :
|
||||
altRightAltRightToMatrix (TensorProduct.map (altRightHanded.ρ M) (altRightHanded.ρ M) v) =
|
||||
((M.1⁻¹).conjTranspose) * altRightAltRightToMatrix v * (((M.1⁻¹).conjTranspose)ᵀ) := by
|
||||
nth_rewrite 1 [altRightAltRightToMatrix]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
|
||||
trans (LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)) ((LinearMap.toMatrix
|
||||
(altRightBasis.tensorProduct altRightBasis) (altRightBasis.tensorProduct altRightBasis)
|
||||
(TensorProduct.map (altRightHanded.ρ M) (altRightHanded.ρ M)))
|
||||
*ᵥ ((Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2))
|
||||
((altRightBasis.tensorProduct altRightBasis).repr (v))))
|
||||
· apply congrArg
|
||||
have h1 := (LinearMap.toMatrix_mulVec_repr (altRightBasis.tensorProduct altRightBasis)
|
||||
(altRightBasis.tensorProduct altRightBasis)
|
||||
(TensorProduct.map (altRightHanded.ρ M) (altRightHanded.ρ M)) v)
|
||||
erw [h1]
|
||||
rfl
|
||||
rw [TensorProduct.toMatrix_map]
|
||||
funext i j
|
||||
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
|
||||
((LinearMap.toMatrix altRightBasis altRightBasis) (altRightHanded.ρ M))
|
||||
((LinearMap.toMatrix altRightBasis altRightBasis) (altRightHanded.ρ M)) (i, j) k)
|
||||
* altRightAltRightToMatrix v k.1 k.2) = _
|
||||
erw [Finset.sum_product]
|
||||
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
|
||||
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (↑M)⁻¹ᴴ i x1 * altRightAltRightToMatrix v x1 x) *
|
||||
(↑M)⁻¹ᴴ j x = ∑ x : Fin 2, ∑ x1 : Fin 2,
|
||||
((↑M)⁻¹ᴴ i x1 * altRightAltRightToMatrix v x1 x) * (↑M)⁻¹ᴴ j x := by
|
||||
congr
|
||||
funext x
|
||||
rw [Finset.sum_mul]
|
||||
erw [h1]
|
||||
rw [Finset.sum_comm]
|
||||
congr
|
||||
funext x
|
||||
congr
|
||||
funext x1
|
||||
simp only [altRightBasis_ρ_apply, transpose_apply, Action.instMonoidalCategory_tensorObj_V]
|
||||
ring
|
||||
|
||||
/-- The group action of `SL(2,ℂ)` on `rightHanded ⊗ altRightHanded` is equivalent to
|
||||
`(M.1.map star) * rightAltRightToMatrix v * (((M.1⁻¹).conjTranspose)ᵀ`. -/
|
||||
lemma rightAltRightToMatrix_ρ (v : (rightHanded ⊗ altRightHanded).V) (M : SL(2,ℂ)) :
|
||||
rightAltRightToMatrix (TensorProduct.map (rightHanded.ρ M) (altRightHanded.ρ M) v) =
|
||||
(M.1.map star) * rightAltRightToMatrix v * (((M.1⁻¹).conjTranspose)ᵀ) := by
|
||||
nth_rewrite 1 [rightAltRightToMatrix]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
|
||||
trans (LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)) ((LinearMap.toMatrix
|
||||
(rightBasis.tensorProduct altRightBasis) (rightBasis.tensorProduct altRightBasis)
|
||||
(TensorProduct.map (rightHanded.ρ M) (altRightHanded.ρ M)))
|
||||
*ᵥ ((Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2))
|
||||
((rightBasis.tensorProduct altRightBasis).repr (v))))
|
||||
· apply congrArg
|
||||
have h1 := (LinearMap.toMatrix_mulVec_repr (rightBasis.tensorProduct altRightBasis)
|
||||
(rightBasis.tensorProduct altRightBasis)
|
||||
(TensorProduct.map (rightHanded.ρ M) (altRightHanded.ρ M)) v)
|
||||
erw [h1]
|
||||
rfl
|
||||
rw [TensorProduct.toMatrix_map]
|
||||
funext i j
|
||||
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
|
||||
((LinearMap.toMatrix rightBasis rightBasis) (rightHanded.ρ M))
|
||||
((LinearMap.toMatrix altRightBasis altRightBasis) (altRightHanded.ρ M)) (i, j) k)
|
||||
* rightAltRightToMatrix v k.1 k.2) = _
|
||||
erw [Finset.sum_product]
|
||||
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
|
||||
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (M.1.map star) i x1 * rightAltRightToMatrix v x1 x)
|
||||
* (↑M)⁻¹ᴴ j x = ∑ x : Fin 2, ∑ x1 : Fin 2,
|
||||
((M.1.map star) i x1 * rightAltRightToMatrix v x1 x) * (↑M)⁻¹ᴴ j x := by
|
||||
congr
|
||||
funext x
|
||||
rw [Finset.sum_mul]
|
||||
erw [h1]
|
||||
rw [Finset.sum_comm]
|
||||
congr
|
||||
funext x
|
||||
congr
|
||||
funext x1
|
||||
simp only [rightBasis_ρ_apply, altRightBasis_ρ_apply, transpose_apply,
|
||||
Action.instMonoidalCategory_tensorObj_V]
|
||||
ring
|
||||
|
||||
/-- The group action of `SL(2,ℂ)` on `altRightHanded ⊗ rightHanded` is equivalent to
|
||||
`((M.1⁻¹).conjTranspose * rightAltRightToMatrix v * ((M.1.map star)).ᵀ`. -/
|
||||
lemma altRightRightToMatrix_ρ (v : (altRightHanded ⊗ rightHanded).V) (M : SL(2,ℂ)) :
|
||||
altRightRightToMatrix (TensorProduct.map (altRightHanded.ρ M) (rightHanded.ρ M) v) =
|
||||
((M.1⁻¹).conjTranspose) * altRightRightToMatrix v * (M.1.map star)ᵀ := by
|
||||
nth_rewrite 1 [altRightRightToMatrix]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
|
||||
trans (LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)) ((LinearMap.toMatrix
|
||||
(altRightBasis.tensorProduct rightBasis) (altRightBasis.tensorProduct rightBasis)
|
||||
(TensorProduct.map (altRightHanded.ρ M) (rightHanded.ρ M)))
|
||||
*ᵥ ((Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2))
|
||||
((altRightBasis.tensorProduct rightBasis).repr (v))))
|
||||
· apply congrArg
|
||||
have h1 := (LinearMap.toMatrix_mulVec_repr (altRightBasis.tensorProduct rightBasis)
|
||||
(altRightBasis.tensorProduct rightBasis)
|
||||
(TensorProduct.map (altRightHanded.ρ M) (rightHanded.ρ M)) v)
|
||||
erw [h1]
|
||||
rfl
|
||||
rw [TensorProduct.toMatrix_map]
|
||||
funext i j
|
||||
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
|
||||
((LinearMap.toMatrix altRightBasis altRightBasis) (altRightHanded.ρ M))
|
||||
((LinearMap.toMatrix rightBasis rightBasis) (rightHanded.ρ M)) (i, j) k)
|
||||
* altRightRightToMatrix v k.1 k.2) = _
|
||||
erw [Finset.sum_product]
|
||||
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
|
||||
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2,
|
||||
(↑M)⁻¹ᴴ i x1 * altRightRightToMatrix v x1 x) * (M.1.map star) j x
|
||||
= ∑ x : Fin 2, ∑ x1 : Fin 2, ((↑M)⁻¹ᴴ i x1 * altRightRightToMatrix v x1 x) *
|
||||
(M.1.map star) j x := by
|
||||
congr
|
||||
funext x
|
||||
rw [Finset.sum_mul]
|
||||
erw [h1]
|
||||
rw [Finset.sum_comm]
|
||||
congr
|
||||
funext x
|
||||
congr
|
||||
funext x1
|
||||
simp only [altRightBasis_ρ_apply, rightBasis_ρ_apply, transpose_apply,
|
||||
Action.instMonoidalCategory_tensorObj_V]
|
||||
ring
|
||||
|
||||
lemma altLeftAltRightToMatrix_ρ (v : (altLeftHanded ⊗ altRightHanded).V) (M : SL(2,ℂ)) :
|
||||
altLeftAltRightToMatrix (TensorProduct.map (altLeftHanded.ρ M) (altRightHanded.ρ M) v) =
|
||||
(M.1⁻¹)ᵀ * altLeftAltRightToMatrix v * ((M.1⁻¹).conjTranspose)ᵀ := by
|
||||
nth_rewrite 1 [altLeftAltRightToMatrix]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
|
||||
trans (LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)) ((LinearMap.toMatrix
|
||||
(altLeftBasis.tensorProduct altRightBasis) (altLeftBasis.tensorProduct altRightBasis)
|
||||
(TensorProduct.map (altLeftHanded.ρ M) (altRightHanded.ρ M)))
|
||||
*ᵥ ((Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2))
|
||||
((altLeftBasis.tensorProduct altRightBasis).repr (v))))
|
||||
· apply congrArg
|
||||
have h1 := (LinearMap.toMatrix_mulVec_repr (altLeftBasis.tensorProduct altRightBasis)
|
||||
(altLeftBasis.tensorProduct altRightBasis)
|
||||
(TensorProduct.map (altLeftHanded.ρ M) (altRightHanded.ρ M)) v)
|
||||
erw [h1]
|
||||
rfl
|
||||
rw [TensorProduct.toMatrix_map]
|
||||
funext i j
|
||||
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
|
||||
((LinearMap.toMatrix altLeftBasis altLeftBasis) (altLeftHanded.ρ M))
|
||||
((LinearMap.toMatrix altRightBasis altRightBasis) (altRightHanded.ρ M)) (i, j) k)
|
||||
* altLeftAltRightToMatrix v k.1 k.2) = _
|
||||
erw [Finset.sum_product]
|
||||
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
|
||||
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (M.1)⁻¹ x1 i * altLeftAltRightToMatrix v x1 x) *
|
||||
(M.1)⁻¹ᴴ j x = ∑ x : Fin 2, ∑ x1 : Fin 2,
|
||||
((M.1)⁻¹ x1 i * altLeftAltRightToMatrix v x1 x) * (M.1)⁻¹ᴴ j x:= by
|
||||
congr
|
||||
funext x
|
||||
rw [Finset.sum_mul]
|
||||
erw [h1]
|
||||
rw [Finset.sum_comm]
|
||||
congr
|
||||
funext x
|
||||
congr
|
||||
funext x1
|
||||
simp only [altLeftBasis_ρ_apply, altRightBasis_ρ_apply, transpose_apply,
|
||||
Action.instMonoidalCategory_tensorObj_V]
|
||||
ring
|
||||
|
||||
lemma leftRightToMatrix_ρ (v : (leftHanded ⊗ rightHanded).V) (M : SL(2,ℂ)) :
|
||||
leftRightToMatrix (TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M) v) =
|
||||
M.1 * leftRightToMatrix v * (M.1)ᴴ := by
|
||||
nth_rewrite 1 [leftRightToMatrix]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
|
||||
trans (LinearEquiv.curry ℂ ℂ (Fin 2) (Fin 2)) ((LinearMap.toMatrix
|
||||
(leftBasis.tensorProduct rightBasis) (leftBasis.tensorProduct rightBasis)
|
||||
(TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M)))
|
||||
*ᵥ ((Finsupp.linearEquivFunOnFinite ℂ ℂ (Fin 2 × Fin 2))
|
||||
((leftBasis.tensorProduct rightBasis).repr (v))))
|
||||
· apply congrArg
|
||||
have h1 := (LinearMap.toMatrix_mulVec_repr (leftBasis.tensorProduct rightBasis)
|
||||
(leftBasis.tensorProduct rightBasis)
|
||||
(TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M)) v)
|
||||
erw [h1]
|
||||
rfl
|
||||
rw [TensorProduct.toMatrix_map]
|
||||
funext i j
|
||||
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
|
||||
((LinearMap.toMatrix leftBasis leftBasis) (leftHanded.ρ M))
|
||||
((LinearMap.toMatrix rightBasis rightBasis) (rightHanded.ρ M)) (i, j) k)
|
||||
* leftRightToMatrix v k.1 k.2) = _
|
||||
erw [Finset.sum_product]
|
||||
simp_rw [kroneckerMap_apply, Matrix.mul_apply]
|
||||
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, M.1 i x1 * leftRightToMatrix v x1 x) * (M.1)ᴴ x j
|
||||
= ∑ x : Fin 2, ∑ x1 : Fin 2, (M.1 i x1 * leftRightToMatrix v x1 x) * (M.1)ᴴ x j := by
|
||||
congr
|
||||
funext x
|
||||
rw [Finset.sum_mul]
|
||||
erw [h1]
|
||||
rw [Finset.sum_comm]
|
||||
congr
|
||||
funext x
|
||||
congr
|
||||
funext x1
|
||||
simp only [leftBasis_ρ_apply, rightBasis_ρ_apply, transpose_apply,
|
||||
Action.instMonoidalCategory_tensorObj_V]
|
||||
rw [Matrix.conjTranspose]
|
||||
simp only [RCLike.star_def, map_apply, transpose_apply]
|
||||
ring
|
||||
|
||||
/-!
|
||||
|
||||
## The symm version of the group actions.
|
||||
|
||||
-/
|
||||
|
||||
lemma leftLeftToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ℂ) (M : SL(2,ℂ)) :
|
||||
TensorProduct.map (leftHanded.ρ M) (leftHanded.ρ M) (leftLeftToMatrix.symm v) =
|
||||
leftLeftToMatrix.symm (M.1 * v * (M.1)ᵀ) := by
|
||||
have h1 := leftLeftToMatrix_ρ (leftLeftToMatrix.symm v) M
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
|
||||
rw [← h1]
|
||||
simp
|
||||
|
||||
lemma altLeftaltLeftToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ℂ) (M : SL(2,ℂ)) :
|
||||
TensorProduct.map (altLeftHanded.ρ M) (altLeftHanded.ρ M) (altLeftaltLeftToMatrix.symm v) =
|
||||
altLeftaltLeftToMatrix.symm ((M.1⁻¹)ᵀ * v * (M.1⁻¹)) := by
|
||||
have h1 := altLeftaltLeftToMatrix_ρ (altLeftaltLeftToMatrix.symm v) M
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
|
||||
rw [← h1]
|
||||
simp
|
||||
|
||||
lemma leftAltLeftToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ℂ) (M : SL(2,ℂ)) :
|
||||
TensorProduct.map (leftHanded.ρ M) (altLeftHanded.ρ M) (leftAltLeftToMatrix.symm v) =
|
||||
leftAltLeftToMatrix.symm (M.1 * v * (M.1⁻¹)) := by
|
||||
have h1 := leftAltLeftToMatrix_ρ (leftAltLeftToMatrix.symm v) M
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
|
||||
rw [← h1]
|
||||
simp
|
||||
|
||||
lemma altLeftLeftToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ℂ) (M : SL(2,ℂ)) :
|
||||
TensorProduct.map (altLeftHanded.ρ M) (leftHanded.ρ M) (altLeftLeftToMatrix.symm v) =
|
||||
altLeftLeftToMatrix.symm ((M.1⁻¹)ᵀ * v * (M.1)ᵀ) := by
|
||||
have h1 := altLeftLeftToMatrix_ρ (altLeftLeftToMatrix.symm v) M
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
|
||||
rw [← h1]
|
||||
simp
|
||||
|
||||
lemma rightRightToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ℂ) (M : SL(2,ℂ)) :
|
||||
TensorProduct.map (rightHanded.ρ M) (rightHanded.ρ M) (rightRightToMatrix.symm v) =
|
||||
rightRightToMatrix.symm ((M.1.map star) * v * ((M.1.map star))ᵀ) := by
|
||||
have h1 := rightRightToMatrix_ρ (rightRightToMatrix.symm v) M
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
|
||||
rw [← h1]
|
||||
simp
|
||||
|
||||
lemma altRightAltRightToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ℂ) (M : SL(2,ℂ)) :
|
||||
TensorProduct.map (altRightHanded.ρ M) (altRightHanded.ρ M) (altRightAltRightToMatrix.symm v) =
|
||||
altRightAltRightToMatrix.symm (((M.1⁻¹).conjTranspose) * v * ((M.1⁻¹).conjTranspose)ᵀ) := by
|
||||
have h1 := altRightAltRightToMatrix_ρ (altRightAltRightToMatrix.symm v) M
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
|
||||
rw [← h1]
|
||||
simp
|
||||
|
||||
lemma rightAltRightToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ℂ) (M : SL(2,ℂ)) :
|
||||
TensorProduct.map (rightHanded.ρ M) (altRightHanded.ρ M) (rightAltRightToMatrix.symm v) =
|
||||
rightAltRightToMatrix.symm ((M.1.map star) * v * (((M.1⁻¹).conjTranspose)ᵀ)) := by
|
||||
have h1 := rightAltRightToMatrix_ρ (rightAltRightToMatrix.symm v) M
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
|
||||
rw [← h1]
|
||||
simp
|
||||
|
||||
lemma altRightRightToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ℂ) (M : SL(2,ℂ)) :
|
||||
TensorProduct.map (altRightHanded.ρ M) (rightHanded.ρ M) (altRightRightToMatrix.symm v) =
|
||||
altRightRightToMatrix.symm (((M.1⁻¹).conjTranspose) * v * (M.1.map star)ᵀ) := by
|
||||
have h1 := altRightRightToMatrix_ρ (altRightRightToMatrix.symm v) M
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
|
||||
rw [← h1]
|
||||
simp
|
||||
|
||||
lemma altLeftAltRightToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ℂ) (M : SL(2,ℂ)) :
|
||||
TensorProduct.map (altLeftHanded.ρ M) (altRightHanded.ρ M) (altLeftAltRightToMatrix.symm v) =
|
||||
altLeftAltRightToMatrix.symm ((M.1⁻¹)ᵀ * v * ((M.1⁻¹).conjTranspose)ᵀ) := by
|
||||
have h1 := altLeftAltRightToMatrix_ρ (altLeftAltRightToMatrix.symm v) M
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
|
||||
rw [← h1]
|
||||
simp
|
||||
|
||||
lemma leftRightToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ℂ) (M : SL(2,ℂ)) :
|
||||
TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M) (leftRightToMatrix.symm v) =
|
||||
leftRightToMatrix.symm (M.1 * v * (M.1)ᴴ) := by
|
||||
have h1 := leftRightToMatrix_ρ (leftRightToMatrix.symm v) M
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
|
||||
rw [← h1]
|
||||
simp
|
||||
|
||||
open SpaceTime
|
||||
|
||||
lemma altLeftAltRightToMatrix_ρ_symm_selfAdjoint (v : Matrix (Fin 2) (Fin 2) ℂ)
|
||||
(hv : IsSelfAdjoint v) (M : SL(2,ℂ)) :
|
||||
TensorProduct.map (altLeftHanded.ρ M) (altRightHanded.ρ M) (altLeftAltRightToMatrix.symm v) =
|
||||
altLeftAltRightToMatrix.symm
|
||||
(SL2C.toLinearMapSelfAdjointMatrix (M.transpose⁻¹) ⟨v, hv⟩) := by
|
||||
rw [altLeftAltRightToMatrix_ρ_symm]
|
||||
apply congrArg
|
||||
simp only [ MonoidHom.coe_mk, OneHom.coe_mk,
|
||||
SL2C.toLinearMapSelfAdjointMatrix_apply_coe, SpecialLinearGroup.coe_inv,
|
||||
SpecialLinearGroup.coe_transpose]
|
||||
congr
|
||||
· rw [SL2C.inverse_coe]
|
||||
simp only [SpecialLinearGroup.coe_inv]
|
||||
rw [@adjugate_transpose]
|
||||
· rw [SL2C.inverse_coe]
|
||||
simp only [SpecialLinearGroup.coe_inv]
|
||||
rw [← @adjugate_transpose]
|
||||
rfl
|
||||
|
||||
lemma leftRightToMatrix_ρ_symm_selfAdjoint (v : Matrix (Fin 2) (Fin 2) ℂ)
|
||||
(hv : IsSelfAdjoint v) (M : SL(2,ℂ)) :
|
||||
TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M) (leftRightToMatrix.symm v) =
|
||||
leftRightToMatrix.symm
|
||||
(SL2C.toLinearMapSelfAdjointMatrix M ⟨v, hv⟩) := by
|
||||
rw [leftRightToMatrix_ρ_symm]
|
||||
rfl
|
||||
|
||||
end
|
||||
end Fermion
|
363
HepLean/Lorentz/Weyl/Unit.lean
Normal file
363
HepLean/Lorentz/Weyl/Unit.lean
Normal file
|
@ -0,0 +1,363 @@
|
|||
/-
|
||||
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.Basic
|
||||
import HepLean.Lorentz.Weyl.Contraction
|
||||
import Mathlib.LinearAlgebra.TensorProduct.Matrix
|
||||
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
|
||||
|
||||
/-- The left-alt-left unit `δᵃₐ` as an element of `(leftHanded ⊗ altLeftHanded).V`. -/
|
||||
def leftAltLeftUnitVal : (leftHanded ⊗ altLeftHanded).V :=
|
||||
leftAltLeftToMatrix.symm 1
|
||||
|
||||
/-- 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]
|
||||
|
||||
/-- 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 := {
|
||||
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
|
||||
ext x : 2
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
|
||||
Function.comp_apply]
|
||||
let x' : ℂ := x
|
||||
change x' • leftAltLeftUnitVal =
|
||||
(TensorProduct.map (leftHanded.ρ M) (altLeftHanded.ρ M)) (x' • leftAltLeftUnitVal)
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
|
||||
apply congrArg
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, leftAltLeftUnitVal]
|
||||
erw [leftAltLeftToMatrix_ρ_symm]
|
||||
apply congrArg
|
||||
simp
|
||||
|
||||
lemma leftAltLeftUnit_apply_one : leftAltLeftUnit.hom (1 : ℂ) = leftAltLeftUnitVal := by
|
||||
change leftAltLeftUnit.hom.toFun (1 : ℂ) = leftAltLeftUnitVal
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
leftAltLeftUnit, AddHom.toFun_eq_coe, AddHom.coe_mk, one_smul]
|
||||
|
||||
/-- The alt-left-left unit `δₐᵃ` as an element of `(altLeftHanded ⊗ leftHanded).V`. -/
|
||||
def altLeftLeftUnitVal : (altLeftHanded ⊗ leftHanded).V :=
|
||||
altLeftLeftToMatrix.symm 1
|
||||
|
||||
/-- 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]
|
||||
|
||||
/-- 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 := {
|
||||
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
|
||||
ext x : 2
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
|
||||
Function.comp_apply]
|
||||
let x' : ℂ := x
|
||||
change x' • altLeftLeftUnitVal =
|
||||
(TensorProduct.map (altLeftHanded.ρ M) (leftHanded.ρ M)) (x' • altLeftLeftUnitVal)
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
|
||||
apply congrArg
|
||||
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]
|
||||
|
||||
lemma altLeftLeftUnit_apply_one : altLeftLeftUnit.hom (1 : ℂ) = altLeftLeftUnitVal := by
|
||||
change altLeftLeftUnit.hom.toFun (1 : ℂ) = altLeftLeftUnitVal
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
altLeftLeftUnit, AddHom.toFun_eq_coe, AddHom.coe_mk, one_smul]
|
||||
|
||||
/-- The right-alt-right unit `δ^{dot a}_{dot a}` as an element of
|
||||
`(rightHanded ⊗ altRightHanded).V`. -/
|
||||
def rightAltRightUnitVal : (rightHanded ⊗ altRightHanded).V :=
|
||||
rightAltRightToMatrix.symm 1
|
||||
|
||||
/-- 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]
|
||||
|
||||
/-- 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 := {
|
||||
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
|
||||
ext x : 2
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
|
||||
Function.comp_apply]
|
||||
let x' : ℂ := x
|
||||
change x' • rightAltRightUnitVal =
|
||||
(TensorProduct.map (rightHanded.ρ M) (altRightHanded.ρ M)) (x' • rightAltRightUnitVal)
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
|
||||
apply congrArg
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, rightAltRightUnitVal]
|
||||
erw [rightAltRightToMatrix_ρ_symm]
|
||||
apply congrArg
|
||||
simp only [RCLike.star_def, mul_one]
|
||||
symm
|
||||
refine transpose_eq_one.mp ?h.h.h.a
|
||||
simp only [transpose_mul, transpose_transpose]
|
||||
change (M.1)⁻¹ᴴ * (M.1)ᴴ = 1
|
||||
rw [@conjTranspose_nonsing_inv]
|
||||
simp
|
||||
|
||||
lemma rightAltRightUnit_apply_one : rightAltRightUnit.hom (1 : ℂ) = rightAltRightUnitVal := by
|
||||
change rightAltRightUnit.hom.toFun (1 : ℂ) = rightAltRightUnitVal
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
rightAltRightUnit, AddHom.toFun_eq_coe, AddHom.coe_mk, one_smul]
|
||||
|
||||
/-- The alt-right-right unit `δ_{dot a}^{dot a}` as an element of
|
||||
`(rightHanded ⊗ altRightHanded).V`. -/
|
||||
def altRightRightUnitVal : (altRightHanded ⊗ rightHanded).V :=
|
||||
altRightRightToMatrix.symm 1
|
||||
|
||||
/-- 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]
|
||||
|
||||
/-- 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 := {
|
||||
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
|
||||
ext x : 2
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
|
||||
Function.comp_apply]
|
||||
let x' : ℂ := x
|
||||
change x' • altRightRightUnitVal =
|
||||
(TensorProduct.map (altRightHanded.ρ M) (rightHanded.ρ M)) (x' • altRightRightUnitVal)
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
|
||||
apply congrArg
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, altRightRightUnitVal]
|
||||
erw [altRightRightToMatrix_ρ_symm]
|
||||
apply congrArg
|
||||
simp only [mul_one, RCLike.star_def]
|
||||
symm
|
||||
change (M.1)⁻¹ᴴ * (M.1)ᴴ = 1
|
||||
rw [@conjTranspose_nonsing_inv]
|
||||
simp
|
||||
|
||||
lemma altRightRightUnit_apply_one : altRightRightUnit.hom (1 : ℂ) = altRightRightUnitVal := by
|
||||
change altRightRightUnit.hom.toFun (1 : ℂ) = altRightRightUnitVal
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
altRightRightUnit, AddHom.toFun_eq_coe, AddHom.coe_mk, one_smul]
|
||||
|
||||
/-!
|
||||
|
||||
## Contraction of the units
|
||||
|
||||
-/
|
||||
|
||||
/-- Contraction on the right with `altLeftLeftUnit` does nothing. -/
|
||||
lemma contr_altLeftLeftUnit (x : leftHanded) :
|
||||
(λ_ 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, 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. -/
|
||||
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 :
|
||||
(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 Fermion
|
Loading…
Add table
Add a link
Reference in a new issue