chore: bump toolchain to v4.15.0

#281 adapt code to v4.15.0 and fix long heartbeats, e.g., toDualRep_apply_eq_contrOneTwoLeft.

---------

Co-authored-by: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com>
This commit is contained in:
KUO-TSAN HSU (Gordon) 2025-01-20 15:42:53 +08:00 committed by GitHub
parent 6e31281a5b
commit 656a3e422f
No known key found for this signature in database
GPG key ID: B5690EEEBB952194
49 changed files with 484 additions and 472 deletions

View file

@ -127,8 +127,8 @@ def altRightBi : altRightHanded →ₗ[] rightHanded →ₗ[] where
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
hom := ModuleCat.ofHom <| TensorProduct.lift leftAltBi
comm M := ModuleCat.hom_ext <| TensorProduct.ext' fun ψ φ => by
change (M.1 *ᵥ ψ.toFin2) ⬝ᵥ (M.1⁻¹ᵀ *ᵥ φ.toFin2) = ψ.toFin2 ⬝ᵥ φ.toFin2
rw [dotProduct_mulVec, vecMul_transpose, mulVec_mulVec]
simp
@ -153,8 +153,8 @@ lemma leftAltContraction_basis (i j : Fin 2) :
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
hom := ModuleCat.ofHom <| TensorProduct.lift altLeftBi
comm M := ModuleCat.hom_ext <| TensorProduct.ext' fun φ ψ => by
change (M.1⁻¹ᵀ *ᵥ φ.toFin2) ⬝ᵥ (M.1 *ᵥ ψ.toFin2) = φ.toFin2 ⬝ᵥ ψ.toFin2
rw [dotProduct_mulVec, mulVec_transpose, vecMul_vecMul]
simp
@ -181,8 +181,8 @@ The linear map from rightHandedWeyl ⊗ altRightHandedWeyl to given by
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
hom := ModuleCat.ofHom <| TensorProduct.lift rightAltBi
comm M := ModuleCat.hom_ext <| 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
@ -221,8 +221,8 @@ lemma rightAltContraction_basis (i j : Fin 2) :
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
hom := ModuleCat.ofHom <| TensorProduct.lift altRightBi
comm M := ModuleCat.hom_ext <| 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