refactor: Lint
This commit is contained in:
parent
833a570ce8
commit
7d983f5b4b
5 changed files with 23 additions and 20 deletions
|
@ -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
|
||||
|
|
|
@ -145,10 +145,11 @@ lemma contr_contrCoUnit (x : complexCo) :
|
|||
Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
|
||||
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 complexCo) (z : CoeSort.coe complexContr) :
|
||||
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']
|
||||
|
@ -178,11 +179,14 @@ lemma contr_coContrUnit (x : complexContr) :
|
|||
Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
|
||||
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
|
||||
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
|
||||
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 :
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue