refactor: Text based Lint
This commit is contained in:
parent
319089ad54
commit
7010a1dae2
12 changed files with 54 additions and 52 deletions
|
@ -116,15 +116,16 @@ lemma perm_eq_of_eq_perm {n m : ℕ} {c : Fin n → S.C} {c1 : Fin m → S.C}
|
|||
simp only [Iso.map_hom_inv_id, Action.id_hom, ModuleCat.id_apply]
|
||||
|
||||
lemma perm_eq_iff_eq_perm {n m : ℕ} {c : Fin n → S.C} {c1 : Fin m → S.C}
|
||||
(σ : (OverColor.mk c) ⟶ (OverColor.mk c1))
|
||||
(σ : (OverColor.mk c) ⟶ (OverColor.mk c1))
|
||||
{t : TensorTree S c} {t2 : TensorTree S c1} :
|
||||
(perm σ t).tensor = t2.tensor ↔ t.tensor =
|
||||
(perm (equivToHomEq (Hom.toEquiv σ).symm (fun x => Hom.toEquiv_comp_apply σ x)) t2).tensor := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· simp [perm_tensor, ← h]
|
||||
change _ = (S.F.map _ ≫ S.F.map _).hom _
|
||||
rw [← S.F.map_comp]
|
||||
have h1 : (σ ≫ equivToHomEq (Hom.toEquiv σ).symm (fun x => Hom.toEquiv_comp_apply σ x)) = 𝟙 _ := by
|
||||
have h1 : (σ ≫ equivToHomEq (Hom.toEquiv σ).symm
|
||||
(fun x => Hom.toEquiv_comp_apply σ x)) = 𝟙 _ := by
|
||||
apply Hom.ext
|
||||
ext x
|
||||
change (Hom.toEquiv σ).symm ((Hom.toEquiv σ) x) = x
|
||||
|
@ -134,7 +135,8 @@ lemma perm_eq_iff_eq_perm {n m : ℕ} {c : Fin n → S.C} {c1 : Fin m → S.C}
|
|||
· rw [perm_tensor, h]
|
||||
change (S.F.map _ ≫ S.F.map _).hom _ = _
|
||||
rw [← S.F.map_comp]
|
||||
have h1 : (equivToHomEq (Hom.toEquiv σ).symm (fun x => Hom.toEquiv_comp_apply σ x) ≫ σ) = 𝟙 _ := by
|
||||
have h1 : (equivToHomEq (Hom.toEquiv σ).symm
|
||||
(fun x => Hom.toEquiv_comp_apply σ x) ≫ σ) = 𝟙 _ := by
|
||||
apply Hom.ext
|
||||
ext x
|
||||
change (Hom.toEquiv σ) ((Hom.toEquiv σ).symm x) = x
|
||||
|
|
|
@ -30,21 +30,20 @@ variable {n m : ℕ}
|
|||
|
||||
lemma perm_congr {c1 : Fin n → S.C} {c2 : Fin m → S.C} {T T' : TensorTree S c1}
|
||||
{σ σ' : OverColor.mk c1 ⟶ OverColor.mk c2}
|
||||
(h : σ = σ') (hT : T.tensor = T'.tensor):
|
||||
(h : σ = σ') (hT : T.tensor = T'.tensor) :
|
||||
(perm σ T).tensor = (perm σ' T').tensor := by
|
||||
rw [h]
|
||||
simp only [perm_tensor, hT]
|
||||
|
||||
lemma perm_update {c1 : Fin n → S.C} {c2 : Fin m → S.C} {T : TensorTree S c1}
|
||||
lemma perm_update {c1 : Fin n → S.C} {c2 : Fin m → S.C} {T : TensorTree S c1}
|
||||
{σ : OverColor.mk c1 ⟶ OverColor.mk c2} (σ' : OverColor.mk c1 ⟶ OverColor.mk c2)
|
||||
(h : σ = σ') :
|
||||
(perm σ T).tensor = (perm σ' T).tensor := by rw [h]
|
||||
|
||||
lemma contr_congr {n : ℕ} {c : Fin n.succ.succ → S.C} {i : Fin n.succ.succ}
|
||||
(i' : Fin n.succ.succ) {j : Fin n.succ} (j' : Fin n.succ)
|
||||
{h : c (i.succAbove j) = S.τ (c i)} {t : TensorTree S c}
|
||||
(hi : i = i' := by decide) (hj : j = j' := by decide)
|
||||
:(contr i j h t).tensor =
|
||||
(i' : Fin n.succ.succ) {j : Fin n.succ} (j' : Fin n.succ){h : c (i.succAbove j) = S.τ (c i)}
|
||||
{t : TensorTree S c} (hi : i = i' := by decide) (hj : j = j' := by decide) :
|
||||
(contr i j h t).tensor =
|
||||
(perm (mkIso (by rw [hi, hj])).hom (contr i' j' (by rw [← hi, ← hj, h]) t)).tensor := by
|
||||
subst hi
|
||||
subst hj
|
||||
|
|
|
@ -285,15 +285,13 @@ def contrContrPerm {n : ℕ} {c : Fin n.succ.succ.succ.succ → S.C} {i : Fin n.
|
|||
{j : Fin n.succ.succ.succ} {k : Fin n.succ.succ} {l : Fin n.succ}
|
||||
(hij : c (i.succAbove j) = S.τ (c i)) (hkl : (c ∘ i.succAbove ∘ j.succAbove) (k.succAbove l) =
|
||||
S.τ ((c ∘ i.succAbove ∘ j.succAbove) k)) :
|
||||
OverColor.mk
|
||||
((c ∘
|
||||
(ContrQuartet.mk i j k l hij hkl).swapI.succAbove ∘
|
||||
(ContrQuartet.mk i j k l hij hkl).swapJ.succAbove) ∘
|
||||
OverColor.mk ((c ∘ (ContrQuartet.mk i j k l hij hkl).swapI.succAbove ∘
|
||||
(ContrQuartet.mk i j k l hij hkl).swapJ.succAbove) ∘
|
||||
(ContrQuartet.mk i j k l hij hkl).swapK.succAbove ∘
|
||||
(ContrQuartet.mk i j k l hij hkl).swapL.succAbove) ⟶
|
||||
OverColor.mk
|
||||
((c ∘ i.succAbove ∘ j.succAbove) ∘ k.succAbove ∘ l.succAbove)
|
||||
:= (ContrQuartet.mk i j k l hij hkl).contrSwapHom
|
||||
(ContrQuartet.mk i j k l hij hkl).swapL.succAbove) ⟶
|
||||
OverColor.mk
|
||||
((c ∘ i.succAbove ∘ j.succAbove) ∘ k.succAbove ∘ l.succAbove) :=
|
||||
(ContrQuartet.mk i j k l hij hkl).contrSwapHom
|
||||
|
||||
/-- Contraction nodes commute on adjusting indices. -/
|
||||
theorem contr_contr {n : ℕ} {c : Fin n.succ.succ.succ.succ → S.C} {i : Fin n.succ.succ.succ.succ}
|
||||
|
|
|
@ -264,8 +264,8 @@ lemma perm_contr_congr_mkIso_cond {n : ℕ} {c : Fin n.succ.succ → S.C} {c1 :
|
|||
{i' : Fin n.succ.succ} {j' : Fin n.succ}
|
||||
(hi : i' = ((Hom.toEquiv σ).symm i))
|
||||
(hj : j' = (((Hom.toEquiv (extractOne i σ))).symm j)) :
|
||||
c ∘ i'.succAbove ∘ j'.succAbove =
|
||||
c ∘ Fin.succAbove ((Hom.toEquiv σ).symm i) ∘ Fin.succAbove ((Hom.toEquiv (extractOne i σ)).symm j) := by
|
||||
c ∘ i'.succAbove ∘ j'.succAbove = c ∘ Fin.succAbove ((Hom.toEquiv σ).symm i) ∘
|
||||
Fin.succAbove ((Hom.toEquiv (extractOne i σ)).symm j) := by
|
||||
rw [hi, hj]
|
||||
|
||||
lemma perm_contr_congr_contr_cond {n : ℕ} {c : Fin n.succ.succ → S.C} {c1 : Fin n.succ.succ → S.C}
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue