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:
parent
6e31281a5b
commit
656a3e422f
49 changed files with 484 additions and 472 deletions
|
@ -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.toFin2ℂEquiv.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.toFin2ℂEquiv.symm (!![0, 1; -1, 0] *ᵥ M.1 *ᵥ ψ.val) =
|
||||
AltLeftHandedModule.toFin2ℂEquiv.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.toFin2ℂEquiv.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.toFin2ℂEquiv.symm (!![0, -1; 1, 0] *ᵥ (M.1⁻¹)ᵀ *ᵥ ψ.val) =
|
||||
LeftHandedModule.toFin2ℂEquiv.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]
|
||||
|
|
|
@ -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
|
||||
|
|
|
@ -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]
|
||||
|
||||
/-!
|
||||
|
||||
|
|
|
@ -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]
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue