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

@ -42,7 +42,7 @@ lemma contrMetricVal_expand_tmul : contrMetricVal =
/-- The metric `ηᵃᵃ` as a morphism `𝟙_ (Rep SL(2,)) ⟶ complexContr ⊗ complexContr`,
making its invariance under the action of `SL(2,)`. -/
def contrMetric : 𝟙_ (Rep SL(2,)) ⟶ complexContr ⊗ complexContr where
hom := {
hom := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • contrMetricVal,
@ -52,13 +52,13 @@ def contrMetric : 𝟙_ (Rep SL(2,)) ⟶ complexContr ⊗ complexContr wh
simp only [smul_smul]
rfl}
comm M := by
ext x : 2
refine ModuleCat.hom_ext ?_
refine LinearMap.ext fun x : => ?_
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.hom_comp,
Function.comp_apply]
let x' : := x
change x' • contrMetricVal =
(TensorProduct.map (complexContr.ρ M) (complexContr.ρ M)) (x' • contrMetricVal)
change x • contrMetricVal =
(TensorProduct.map (complexContr.ρ M) (complexContr.ρ M)) (x • contrMetricVal)
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp only [Action.instMonoidalCategory_tensorObj_V, contrMetricVal]
@ -67,9 +67,8 @@ def contrMetric : 𝟙_ (Rep SL(2,)) ⟶ complexContr ⊗ complexContr wh
simp only [LorentzGroup.toComplex_mul_minkowskiMatrix_mul_transpose]
lemma contrMetric_apply_one : contrMetric.hom (1 : ) = contrMetricVal := by
change contrMetric.hom.toFun (1 : ) = contrMetricVal
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
contrMetric, AddHom.toFun_eq_coe, AddHom.coe_mk, one_smul]
change contrMetric.hom.hom.toFun (1 : ) = contrMetricVal
simp only [contrMetric, one_smul]
/-- The metric `ηᵢᵢ` as an element of `(complexCo ⊗ complexCo).V`. -/
def coMetricVal : (complexCo ⊗ complexCo).V :=
@ -95,7 +94,7 @@ lemma coMetricVal_expand_tmul : coMetricVal =
/-- The metric `ηᵢᵢ` as a morphism `𝟙_ (Rep SL(2,)) ⟶ complexCo ⊗ complexCo`,
making its invariance under the action of `SL(2,)`. -/
def coMetric : 𝟙_ (Rep SL(2,)) ⟶ complexCo ⊗ complexCo where
hom := {
hom := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • coMetricVal,
@ -105,13 +104,13 @@ def coMetric : 𝟙_ (Rep SL(2,)) ⟶ complexCo ⊗ complexCo where
simp only [smul_smul]
rfl}
comm M := by
ext x : 2
refine ModuleCat.hom_ext ?_
refine LinearMap.ext fun x : => ?_
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.hom_comp,
Function.comp_apply]
let x' : := x
change x' • coMetricVal =
(TensorProduct.map (complexCo.ρ M) (complexCo.ρ M)) (x' • coMetricVal)
change x • coMetricVal =
(TensorProduct.map (complexCo.ρ M) (complexCo.ρ M)) (x • coMetricVal)
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp only [Action.instMonoidalCategory_tensorObj_V, coMetricVal]
@ -122,9 +121,8 @@ def coMetric : 𝟙_ (Rep SL(2,)) ⟶ complexCo ⊗ complexCo where
LorentzGroup.toComplex_transpose_mul_minkowskiMatrix_mul_self]
lemma coMetric_apply_one : coMetric.hom (1 : ) = coMetricVal := by
change coMetric.hom.toFun (1 : ) = coMetricVal
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
coMetric, AddHom.toFun_eq_coe, AddHom.coe_mk, one_smul]
change coMetric.hom.hom.toFun (1 : ) = coMetricVal
simp only [coMetric, one_smul]
/-!