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

@ -188,7 +188,7 @@ lemma altRightBasis_ρ_apply (M : SL(2,)) (i j : Fin 2) :
`altLeftHanded` defined by multiplying an element of
`leftHanded` by the matrix `εᵃ⁰ᵃ¹ = !![0, 1; -1, 0]]`. -/
def leftHandedToAlt : leftHanded ⟶ altLeftHanded where
hom := {
hom := ModuleCat.ofHom {
toFun := fun ψ => AltLeftHandedModule.toFin2Equiv.symm (!![0, 1; -1, 0] *ᵥ ψ.toFin2),
map_add' := by
intro ψ ψ'
@ -199,6 +199,7 @@ def leftHandedToAlt : leftHanded ⟶ altLeftHanded where
rfl}
comm := by
intro M
refine ModuleCat.hom_ext ?_
refine LinearMap.ext (fun ψ => ?_)
change AltLeftHandedModule.toFin2Equiv.symm (!![0, 1; -1, 0] *ᵥ M.1 *ᵥ ψ.val) =
AltLeftHandedModule.toFin2Equiv.symm ((M.1⁻¹)ᵀ *ᵥ !![0, 1; -1, 0] *ᵥ ψ.val)
@ -217,7 +218,7 @@ lemma leftHandedToAlt_hom_apply (ψ : leftHanded) :
`leftHanded` defined by multiplying an element of
altLeftHandedWeyl by the matrix `εₐ₁ₐ₂ = !![0, -1; 1, 0]`. -/
def leftHandedAltTo : altLeftHanded ⟶ leftHanded where
hom := {
hom := ModuleCat.ofHom {
toFun := fun ψ =>
LeftHandedModule.toFin2Equiv.symm (!![0, -1; 1, 0] *ᵥ ψ.toFin2),
map_add' := by
@ -231,6 +232,7 @@ def leftHandedAltTo : altLeftHanded ⟶ leftHanded where
rfl}
comm := by
intro M
refine ModuleCat.hom_ext ?_
refine LinearMap.ext (fun ψ => ?_)
change LeftHandedModule.toFin2Equiv.symm (!![0, -1; 1, 0] *ᵥ (M.1⁻¹)ᵀ *ᵥ ψ.val) =
LeftHandedModule.toFin2Equiv.symm (M.1 *ᵥ !![0, -1; 1, 0] *ᵥ ψ.val)
@ -253,8 +255,8 @@ def leftHandedAltEquiv : leftHanded ≅ altLeftHanded where
inv := leftHandedAltTo
hom_inv_id := by
ext ψ
simp only [Action.comp_hom, ModuleCat.coe_comp, Function.comp_apply, Action.id_hom,
ModuleCat.id_apply]
simp only [Action.comp_hom, ModuleCat.hom_comp, LinearMap.coe_comp, Function.comp_apply,
Action.id_hom, ModuleCat.hom_id, LinearMap.id_coe, id_eq]
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]
@ -262,8 +264,8 @@ def leftHandedAltEquiv : leftHanded ≅ altLeftHanded where
rfl
inv_hom_id := by
ext ψ
simp only [Action.comp_hom, ModuleCat.coe_comp, Function.comp_apply, Action.id_hom,
ModuleCat.id_apply]
simp only [Action.comp_hom, ModuleCat.hom_comp, LinearMap.coe_comp, Function.comp_apply,
Action.id_hom, ModuleCat.hom_id, LinearMap.id_coe, id_eq]
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]

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

View file

@ -81,7 +81,7 @@ lemma leftMetricVal_expand_tmul : leftMetricVal =
/-- 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 := {
hom := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • leftMetricVal,
@ -91,13 +91,13 @@ def leftMetric : 𝟙_ (Rep SL(2,)) ⟶ leftHanded ⊗ leftHanded 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' • leftMetricVal =
(TensorProduct.map (leftHanded.ρ M) (leftHanded.ρ M)) (x' • leftMetricVal)
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]
@ -108,9 +108,8 @@ def leftMetric : 𝟙_ (Rep SL(2,)) ⟶ leftHanded ⊗ leftHanded where
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]
change leftMetric.hom.hom.toFun (1 : ) = leftMetricVal
simp only [leftMetric, one_smul]
/-- The metric `εₐₐ` as an element of `(altLeftHanded ⊗ altLeftHanded).V`. -/
def altLeftMetricVal : (altLeftHanded ⊗ altLeftHanded).V :=
@ -129,7 +128,7 @@ lemma altLeftMetricVal_expand_tmul : altLeftMetricVal =
/-- 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 := {
hom := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • altLeftMetricVal,
@ -139,13 +138,13 @@ def altLeftMetric : 𝟙_ (Rep SL(2,)) ⟶ altLeftHanded ⊗ altLeftHande
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' • altLeftMetricVal =
(TensorProduct.map (altLeftHanded.ρ M) (altLeftHanded.ρ M)) (x' • altLeftMetricVal)
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]
@ -156,9 +155,8 @@ def altLeftMetric : 𝟙_ (Rep SL(2,)) ⟶ altLeftHanded ⊗ altLeftHande
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]
change altLeftMetric.hom.hom.toFun (1 : ) = altLeftMetricVal
simp only [altLeftMetric, one_smul]
/-- The metric `ε^{dot a}^{dot a}` as an element of `(rightHanded ⊗ rightHanded).V`. -/
def rightMetricVal : (rightHanded ⊗ rightHanded).V :=
@ -176,7 +174,7 @@ lemma rightMetricVal_expand_tmul : rightMetricVal =
/-- 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 := {
hom := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • rightMetricVal,
@ -186,13 +184,13 @@ def rightMetric : 𝟙_ (Rep SL(2,)) ⟶ rightHanded ⊗ rightHanded wher
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' • rightMetricVal =
(TensorProduct.map (rightHanded.ρ M) (rightHanded.ρ M)) (x' • rightMetricVal)
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]
@ -211,9 +209,8 @@ def rightMetric : 𝟙_ (Rep SL(2,)) ⟶ rightHanded ⊗ rightHanded wher
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]
change rightMetric.hom.hom.toFun (1 : ) = rightMetricVal
simp only [rightMetric, one_smul]
/-- The metric `ε_{dot a}_{dot a}` as an element of `(altRightHanded ⊗ altRightHanded).V`. -/
def altRightMetricVal : (altRightHanded ⊗ altRightHanded).V :=
@ -233,7 +230,7 @@ lemma altRightMetricVal_expand_tmul : altRightMetricVal =
`𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ altRightHanded`,
making manifest its invariance under the action of `SL(2,)`. -/
def altRightMetric : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ altRightHanded where
hom := {
hom := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • altRightMetricVal,
@ -243,13 +240,13 @@ def altRightMetric : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ altRightHa
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' • altRightMetricVal =
(TensorProduct.map (altRightHanded.ρ M) (altRightHanded.ρ M)) (x' • altRightMetricVal)
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]
@ -270,9 +267,8 @@ def altRightMetric : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ altRightHa
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]
change altRightMetric.hom.hom.toFun (1 : ) = altRightMetricVal
simp only [altRightMetric, one_smul]
/-!

View file

@ -36,7 +36,7 @@ lemma leftAltLeftUnitVal_expand_tmul : leftAltLeftUnitVal =
/-- 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 := {
hom := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • leftAltLeftUnitVal,
@ -46,13 +46,13 @@ def leftAltLeftUnit : 𝟙_ (Rep SL(2,)) ⟶ leftHanded ⊗ altLeftHanded
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' • leftAltLeftUnitVal =
(TensorProduct.map (leftHanded.ρ M) (altLeftHanded.ρ M)) (x' • leftAltLeftUnitVal)
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]
@ -61,9 +61,8 @@ def leftAltLeftUnit : 𝟙_ (Rep SL(2,)) ⟶ leftHanded ⊗ altLeftHanded
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]
change leftAltLeftUnit.hom.hom.toFun (1 : ) = leftAltLeftUnitVal
simp only [leftAltLeftUnit, one_smul]
/-- The alt-left-left unit `δₐᵃ` as an element of `(altLeftHanded ⊗ leftHanded).V`. -/
def altLeftLeftUnitVal : (altLeftHanded ⊗ leftHanded).V :=
@ -80,7 +79,7 @@ lemma altLeftLeftUnitVal_expand_tmul : altLeftLeftUnitVal =
/-- 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 := {
hom := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • altLeftLeftUnitVal,
@ -90,13 +89,13 @@ def altLeftLeftUnit : 𝟙_ (Rep SL(2,)) ⟶ altLeftHanded ⊗ leftHanded
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' • altLeftLeftUnitVal =
(TensorProduct.map (altLeftHanded.ρ M) (leftHanded.ρ M)) (x' • altLeftLeftUnitVal)
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]
@ -107,9 +106,8 @@ def altLeftLeftUnit : 𝟙_ (Rep SL(2,)) ⟶ altLeftHanded ⊗ leftHanded
/-- Applying the morphism `altLeftLeftUnit` to `1` returns `altLeftLeftUnitVal`. -/
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]
change altLeftLeftUnit.hom.hom.toFun (1 : ) = altLeftLeftUnitVal
simp only [altLeftLeftUnit, one_smul]
/-- The right-alt-right unit `δ^{dot a}_{dot a}` as an element of
`(rightHanded ⊗ altRightHanded).V`. -/
@ -128,7 +126,7 @@ lemma rightAltRightUnitVal_expand_tmul : rightAltRightUnitVal =
`𝟙_ (Rep SL(2,)) ⟶ rightHanded ⊗ altRightHanded`, manifesting
the invariance under the `SL(2,)` action. -/
def rightAltRightUnit : 𝟙_ (Rep SL(2,)) ⟶ rightHanded ⊗ altRightHanded where
hom := {
hom := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • rightAltRightUnitVal,
@ -138,13 +136,13 @@ def rightAltRightUnit : 𝟙_ (Rep SL(2,)) ⟶ rightHanded ⊗ altRightHa
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' • rightAltRightUnitVal =
(TensorProduct.map (rightHanded.ρ M) (altRightHanded.ρ M)) (x' • rightAltRightUnitVal)
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]
@ -159,9 +157,8 @@ def rightAltRightUnit : 𝟙_ (Rep SL(2,)) ⟶ rightHanded ⊗ altRightHa
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]
change rightAltRightUnit.hom.hom.toFun (1 : ) = rightAltRightUnitVal
simp only [rightAltRightUnit, one_smul]
/-- The alt-right-right unit `δ_{dot a}^{dot a}` as an element of
`(rightHanded ⊗ altRightHanded).V`. -/
@ -180,7 +177,7 @@ lemma altRightRightUnitVal_expand_tmul : altRightRightUnitVal =
`𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ rightHanded`, manifesting
the invariance under the `SL(2,)` action. -/
def altRightRightUnit : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ rightHanded where
hom := {
hom := ModuleCat.ofHom {
toFun := fun a =>
let a' : := a
a' • altRightRightUnitVal,
@ -190,13 +187,13 @@ def altRightRightUnit : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ rightHa
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' • altRightRightUnitVal =
(TensorProduct.map (altRightHanded.ρ M) (rightHanded.ρ M)) (x' • altRightRightUnitVal)
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]
@ -209,9 +206,8 @@ def altRightRightUnit : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ rightHa
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]
change altRightRightUnit.hom.hom.toFun (1 : ) = altRightRightUnitVal
simp only [altRightRightUnit, one_smul]
/-!
@ -240,7 +236,7 @@ lemma contr_altLeftLeftUnit (x : leftHanded) :
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.Iso.trans_hom, LinearEquiv.toModuleIso_hom_hom, ModuleCat.ofSelfIso_hom,
CategoryTheory.Category.comp_id, Action.instMonoidalCategory_tensorUnit_V, Fin.val_zero,
↓reduceIte, Fin.val_one, one_ne_zero, zero_tmul, map_zero, smul_zero, add_zero, zero_ne_one,
zero_add]