refactor: Lint

This commit is contained in:
jstoobysmith 2024-10-24 16:42:25 +00:00
parent 833a570ce8
commit 7d983f5b4b
5 changed files with 23 additions and 20 deletions

View file

@ -130,7 +130,6 @@ lemma coMetric_apply_one : coMetric.hom (1 : ) = coMetricVal := by
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
coMetric, AddHom.toFun_eq_coe, AddHom.coe_mk, one_smul]
/-!
## Contraction of metrics

View file

@ -148,7 +148,8 @@ lemma contr_contrCoUnit (x : complexCo) :
have h1' (x y : CoeSort.coe complexCo) (z : CoeSort.coe complexContr) :
(α_ complexCo.V complexContr.V complexCo.V).inv (x ⊗ₜ[] z ⊗ₜ[] y) = (x ⊗ₜ[] z) ⊗ₜ[] y := rfl
repeat rw [h1']
have h1'' ( y : CoeSort.coe complexCo) (z : CoeSort.coe complexCo ⊗[] CoeSort.coe complexContr) :
have h1'' (y : CoeSort.coe complexCo)
(z : CoeSort.coe complexCo ⊗[] CoeSort.coe complexContr) :
(coContrContraction.hom ▷ complexCo.V) (z ⊗ₜ[] y) = (coContrContraction.hom z) ⊗ₜ[] y := rfl
repeat rw (config := { transparency := .instances }) [h1'']
repeat rw [coContrContraction_basis']
@ -179,10 +180,13 @@ lemma contr_coContrUnit (x : complexContr) :
Finset.sum_singleton, Fin.sum_univ_three, tmul_add, add_tmul, smul_tmul, tmul_smul, map_add,
_root_.map_smul]
have h1' (x y : CoeSort.coe complexContr) (z : CoeSort.coe complexCo) :
(α_ complexContr.V complexCo.V complexContr.V).inv (x ⊗ₜ[] z ⊗ₜ[] y) = (x ⊗ₜ[] z) ⊗ₜ[] y := rfl
(α_ complexContr.V complexCo.V complexContr.V).inv
(x ⊗ₜ[] z ⊗ₜ[] y) = (x ⊗ₜ[] z) ⊗ₜ[] y := rfl
repeat rw [h1']
have h1'' ( y : CoeSort.coe complexContr) (z : CoeSort.coe complexContr ⊗[] CoeSort.coe complexCo) :
(contrCoContraction.hom ▷ complexContr.V) (z ⊗ₜ[] y) = (contrCoContraction.hom z) ⊗ₜ[] y := rfl
have h1'' (y : CoeSort.coe complexContr)
(z : CoeSort.coe complexContr ⊗[] CoeSort.coe complexCo) :
(contrCoContraction.hom ▷ complexContr.V) (z ⊗ₜ[] y) =
(contrCoContraction.hom z) ⊗ₜ[] y := rfl
repeat rw (config := { transparency := .instances }) [h1'']
repeat rw [contrCoContraction_basis']
simp only [Fin.isValue, Action.instMonoidalCategory_tensorUnit_V, ↓reduceIte, reduceCtorEq,
@ -199,7 +203,6 @@ lemma contr_coContrUnit (x : complexContr) :
-/
open CategoryTheory
lemma contrCoUnit_symm :

View file

@ -352,8 +352,8 @@ lemma rightAltContraction_apply_metric : (β_ rightHanded altRightHanded).hom.ho
((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
((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,
@ -374,7 +374,8 @@ lemma altRightContraction_apply_metric : (β_ altRightHanded rightHanded).hom.ho
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)
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

View file

@ -333,8 +333,8 @@ lemma contr_rightAltRightUnit (x : altRightHanded) :
open CategoryTheory
lemma altLeftLeftUnit_symm :
(altLeftLeftUnit.hom (1 : )) = (altLeftHanded ◁ 𝟙 _).hom ((β_ leftHanded altLeftHanded ).hom.hom
(leftAltLeftUnit.hom (1 : ))) := by
(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
@ -347,15 +347,15 @@ lemma leftAltLeftUnit_symm :
rfl
lemma altRightRightUnit_symm :
(altRightRightUnit.hom (1 : )) = (altRightHanded ◁ 𝟙 _).hom ((β_ rightHanded altRightHanded ).hom.hom
(rightAltRightUnit.hom (1 : ))) := by
(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
(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