2024-10-07 12:20:53 +00:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
2024-10-10 08:57:22 +00:00
import HepLean.Tensors.OverColor.Iso
2024-10-15 06:08:56 +00:00
import HepLean.Tensors.OverColor.Discrete
import HepLean.Tensors.OverColor.Lift
2024-10-09 07:42:56 +00:00
import Mathlib.CategoryTheory.Monoidal.NaturalTransformation
2024-10-17 11:43:33 +00:00
import LLMLean
2024-10-07 12:20:53 +00:00
/-!
## Tensor trees
-/
open IndexNotation
open CategoryTheory
2024-10-09 16:57:41 +00:00
open MonoidalCategory
2024-10-07 12:20:53 +00:00
2024-10-08 07:52:55 +00:00
/-- The sturcture of a type of tensors e.g. Lorentz tensors, Einstien tensors,
complex Lorentz tensors.
Note: This structure is not fully defined yet. -/
2024-10-07 12:20:53 +00:00
structure TensorStruct where
2024-10-08 07:52:55 +00:00
/-- The colors of indices e.g. up or down. -/
2024-10-07 12:20:53 +00:00
C : Type
2024-10-08 07:52:55 +00:00
/-- The symmetry group acting on these tensor e.g. the Lorentz group or SL(2,ℂ ). -/
2024-10-07 12:20:53 +00:00
G : Type
2024-10-08 07:52:55 +00:00
/-- An instance of `G` as a group. -/
2024-10-07 12:20:53 +00:00
G_group : Group G
2024-10-08 07:52:55 +00:00
/-- The field over which we want to consider the tensors to live in, usually `ℝ ` or `ℂ `. -/
2024-10-07 12:20:53 +00:00
k : Type
2024-10-08 07:52:55 +00:00
/-- An instance of `k` as a commutative ring. -/
2024-10-07 12:20:53 +00:00
k_commRing : CommRing k
2024-10-08 07:52:55 +00:00
/-- A `MonoidalFunctor` from `OverColor C` giving the rep corresponding to a map of colors
`X → C`. -/
2024-10-15 06:08:56 +00:00
FDiscrete : Discrete C ⥤ Rep k G
2024-10-08 07:52:55 +00:00
/-- A map from `C` to `C`. An involution. -/
2024-10-07 12:20:53 +00:00
τ : C → C
2024-10-15 06:08:56 +00:00
τ_involution : Function.Involutive τ
/-- The natural transformation describing contraction. -/
2024-10-16 16:38:36 +00:00
contr : OverColor.Discrete.pairτ FDiscrete τ ⟶ 𝟙 _ (Discrete C ⥤ Rep k G)
2024-10-15 06:08:56 +00:00
/-- The natural transformation describing the metric. -/
2024-10-16 16:38:36 +00:00
metric : 𝟙 _ (Discrete C ⥤ Rep k G) ⟶ OverColor.Discrete.pair FDiscrete
2024-10-15 06:08:56 +00:00
/-- The natural transformation describing the unit. -/
unit : 𝟙 _ (Discrete C ⥤ Rep k G) ⟶ OverColor.Discrete.τPair FDiscrete τ
2024-10-08 07:52:55 +00:00
/-- A specification of the dimension of each color in C. This will be used for explicit
evaluation of tensors. -/
evalNo : C → ℕ
2024-10-07 12:20:53 +00:00
2024-10-15 06:08:56 +00:00
noncomputable section
2024-10-07 12:20:53 +00:00
namespace TensorStruct
variable (S : TensorStruct)
instance : CommRing S.k := S.k_commRing
instance : Group S.G := S.G_group
2024-10-15 06:08:56 +00:00
/-- The lift of the functor `S.F` to a monoidal functor. -/
def F : MonoidalFunctor (OverColor S.C) (Rep S.k S.G) := (OverColor.lift).obj S.FDiscrete
2024-10-16 16:38:36 +00:00
/-
2024-10-15 06:08:56 +00:00
def metric (c : S.C) : S.F.obj (OverColor.mk ![c, c]) :=
(OverColor.Discrete.pairIso S.FDiscrete c).hom.hom <|
(S.metricNat.app (Discrete.mk c)).hom (1 : S.k)
2024-10-16 16:38:36 +00:00
-/
2024-10-15 06:08:56 +00:00
2024-10-16 16:38:36 +00:00
/-- The isomorphism of objects in `Rep S.k S.G` given an `i` in `Fin n.succ.succ` and
2024-10-16 16:42:20 +00:00
a `j` in `Fin n.succ` allowing us to undertake contraction. -/
2024-10-15 06:08:56 +00:00
def contrIso {n : ℕ } (c : Fin n.succ.succ → S.C)
(i : Fin n.succ.succ) (j : Fin n.succ) (h : c (i.succAbove j) = S.τ (c i)) :
2024-10-16 16:38:36 +00:00
S.F.obj (OverColor.mk c) ≅ ((OverColor.Discrete.pairτ S.FDiscrete S.τ).obj
(Discrete.mk (c i))) ⊗
2024-10-15 06:08:56 +00:00
(OverColor.lift.obj S.FDiscrete).obj (OverColor.mk (c ∘ i.succAbove ∘ j.succAbove)) :=
2024-10-17 17:11:06 +00:00
(S.F.mapIso (OverColor.equivToIso (HepLean.Fin.finExtractTwo i j))).trans <|
(S.F.mapIso (OverColor.mkSum (c ∘ (HepLean.Fin.finExtractTwo i j).symm))).trans <|
2024-10-15 06:08:56 +00:00
(S.F.μIso _ _).symm.trans <| by
refine tensorIso ?_ (S.F.mapIso (OverColor.mkIso (by ext x; simp)))
2024-10-17 17:11:06 +00:00
apply (S.F.mapIso (OverColor.mkSum (((c ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inl)))).trans
2024-10-15 06:08:56 +00:00
apply (S.F.μIso _ _).symm.trans
apply tensorIso ?_ ?_
· symm
apply (OverColor.forgetLiftApp S.FDiscrete (c i)).symm.trans
apply S.F.mapIso
apply OverColor.mkIso
funext x
fin_cases x
rfl
· symm
apply (OverColor.forgetLiftApp S.FDiscrete (S.τ (c i))).symm.trans
apply S.F.mapIso
apply OverColor.mkIso
funext x
fin_cases x
simp [h]
2024-10-17 17:11:06 +00:00
open OverColor
lemma perm_contr_cond {n : ℕ } {c : Fin n.succ.succ.succ → S.C} {c1 : Fin n.succ.succ.succ → S.C}
{i : Fin n.succ.succ.succ} {j : Fin n.succ.succ}
(h : c1 (i.succAbove j) = S.τ (c1 i)) (σ : (OverColor.mk c) ⟶ (OverColor.mk c1)) :
c (Fin.succAbove ((Hom.toEquiv σ ).symm i) ((Hom.toEquiv (extractOne i σ )).symm j)) =
S.τ (c ((Hom.toEquiv σ ).symm i)) := by
have h1 := Hom.toEquiv_comp_apply σ
simp at h1
rw [h1, h1]
simp
rw [← h]
congr
simp [HepLean.Fin.finExtractOnePerm, HepLean.Fin.finExtractOnPermHom]
erw [Equiv.apply_symm_apply]
rw [HepLean.Fin.succsAbove_predAboveI]
erw [Equiv.apply_symm_apply]
simp
erw [Equiv.apply_eq_iff_eq]
exact (Fin.succAbove_ne i j).symm
open OverColor in
lemma contrIso_comm_map {n : ℕ } {c c1 : Fin n.succ.succ.succ → S.C}
{i : Fin n.succ.succ.succ} {j : Fin n.succ.succ}
{h : c1 (i.succAbove j) = S.τ (c1 i)}
(σ : (OverColor.mk c) ⟶ (OverColor.mk c1)) :
(S.F.map σ ) ≫ (S.contrIso c1 i j h).hom =
(S.contrIso c ((OverColor.Hom.toEquiv σ ).symm i)
(((Hom.toEquiv (extractOne i σ ))).symm j) (S.perm_contr_cond h σ )).hom ≫
(((Discrete.pairτ S.FDiscrete S.τ).map (Discrete.eqToHom (Hom.toEquiv_comp_inv_apply σ i)
: (Discrete.mk (c ((Hom.toEquiv σ ).symm i))) ⟶ (Discrete.mk (c1 i)) )) ⊗ (S.F.map (extractTwo i j σ ))) := by
ext Z
simp
rw [contrIso]
simp
have h1 : ((S.F.map (mkSum (c1 ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm)).hom).hom ((S.F.map (equivToIso (HepLean.Fin.finExtractTwo i j)).hom).hom ((S.F.map σ ).hom Z)))
= ((S.F.map σ ) ≫ (S.F.map (equivToIso (HepLean.Fin.finExtractTwo i j)).hom) ≫ (S.F.map (mkSum (c1 ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm)).hom)).hom Z := by
rfl
have h1' : ((S.F.map (mkSum (c1 ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm)).hom).hom ((S.F.map (equivToIso (HepLean.Fin.finExtractTwo i j)).hom).hom ((S.F.map σ ).hom Z)))
= ((S.F.map (σ ≫ (equivToIso (HepLean.Fin.finExtractTwo i j)).hom ≫ (mkSum (c1 ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm)).hom))).hom Z := by
rw [h1]
simp
rw [h1']
rw [extractTwo_finExtractTwo]
simp
sorry
2024-10-16 16:38:36 +00:00
/--
`contrMap` is a function that takes a natural number `n`, a function `c` from
`Fin n.succ.succ` to `S.C`, an index `i` of type `Fin n.succ.succ`, an index `j` of type
`Fin n.succ`, and a proof `h` that `c (i.succAbove j) = S.τ (c i)`. It returns a morphism
corresponding to the contraction of the `i`th index with the `i.succAbove j` index.
--/
def contrMap {n : ℕ } (c : Fin n.succ.succ → S.C)
2024-10-15 06:08:56 +00:00
(i : Fin n.succ.succ) (j : Fin n.succ) (h : c (i.succAbove j) = S.τ (c i)) :
S.F.obj (OverColor.mk c) ⟶
(OverColor.lift.obj S.FDiscrete).obj (OverColor.mk (c ∘ i.succAbove ∘ j.succAbove)) :=
(S.contrIso c i j h).hom ≫
(tensorHom (S.contr.app (Discrete.mk (c i))) (𝟙 _)) ≫
(MonoidalCategory.leftUnitor _).hom
2024-10-07 12:20:53 +00:00
end TensorStruct
2024-10-08 07:52:55 +00:00
/-- A syntax tree for tensor expressions. -/
2024-10-07 12:20:53 +00:00
inductive TensorTree (S : TensorStruct) : ∀ {n : ℕ }, (Fin n → S.C) → Type where
2024-10-16 16:38:36 +00:00
/-- A general tensor node. -/
2024-10-08 07:52:55 +00:00
| tensorNode {n : ℕ } {c : Fin n → S.C} (T : S.F.obj (OverColor.mk c)) : TensorTree S c
2024-10-16 16:38:36 +00:00
/-- A node consisting of a single vector. -/
| vecNode {c : S.C} (v : S.FDiscrete.obj (Discrete.mk c)) : TensorTree S ![c]
/-- A node consisting of a two tensor. -/
| twoNode {c1 c2 : S.C}
2024-10-17 11:43:33 +00:00
(v : (S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2)).V) :
2024-10-16 16:38:36 +00:00
TensorTree S ![c1, c2]
/-- A node consisting of a three tensor. -/
| threeNode {c1 c2 c3 : S.C}
(v : S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2) ⊗
S.FDiscrete.obj (Discrete.mk c3)) : TensorTree S ![c1, c2, c3]
/-- A general constant node. -/
| constNode {n : ℕ } {c : Fin n → S.C} (T : 𝟙 _ (Rep S.k S.G) ⟶ S.F.obj (OverColor.mk c)) :
TensorTree S c
/-- A constant vector. -/
| constVecNode {c : S.C} (v : 𝟙 _ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c)) :
TensorTree S ![c]
/-- A constant two tensor (e.g. metric and unit). -/
| constTwoNode {c1 c2 : S.C}
(v : 𝟙 _ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2)) :
TensorTree S ![c1, c2]
/-- A constant three tensor (e.g. Pauli-matrices). -/
| constThreeNode {c1 c2 c3 : S.C}
(v : 𝟙 _ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2) ⊗
S.FDiscrete.obj (Discrete.mk c3)) : TensorTree S ![c1, c2, c3]
/-- A node corresponding to the addition of two tensors. -/
2024-10-07 12:20:53 +00:00
| add {n : ℕ } {c : Fin n → S.C} : TensorTree S c → TensorTree S c → TensorTree S c
2024-10-16 16:38:36 +00:00
/-- A node corresponding to the permutation of indices of a tensor. -/
2024-10-07 12:20:53 +00:00
| perm {n m : ℕ } {c : Fin n → S.C} {c1 : Fin m → S.C}
(σ : (OverColor.mk c) ⟶ (OverColor.mk c1)) (t : TensorTree S c) : TensorTree S c1
| prod {n m : ℕ } {c : Fin n → S.C} {c1 : Fin m → S.C}
(t : TensorTree S c) (t1 : TensorTree S c1) : TensorTree S (Sum.elim c c1 ∘ finSumFinEquiv.symm)
2024-10-08 15:45:51 +00:00
| smul {n : ℕ } {c : Fin n → S.C} : S.k → TensorTree S c → TensorTree S c
2024-10-17 11:43:33 +00:00
/-- The negative of a node. -/
| neg {n : ℕ } {c : Fin n → S.C} : TensorTree S c → TensorTree S c
2024-10-16 16:38:36 +00:00
| contr {n : ℕ } {c : Fin n.succ.succ → S.C} : (i : Fin n.succ.succ) →
(j : Fin n.succ) → (h : c (i.succAbove j) = S.τ (c i)) → TensorTree S c →
TensorTree S (c ∘ Fin.succAbove i ∘ Fin.succAbove j)
2024-10-08 07:26:23 +00:00
| eval {n : ℕ } {c : Fin n.succ → S.C} :
(i : Fin n.succ) → (x : Fin (S.evalNo (c i))) → TensorTree S c →
2024-10-08 07:52:55 +00:00
TensorTree S (c ∘ Fin.succAbove i)
2024-10-07 12:20:53 +00:00
namespace TensorTree
variable {S : TensorStruct} {n : ℕ } {c : Fin n → S.C} (T : TensorTree S c)
open MonoidalCategory
open TensorProduct
2024-10-16 16:38:36 +00:00
/-- The node `twoNode` of a tensor tree, with all arguments explicit. -/
abbrev twoNodeE (S : TensorStruct) (c1 c2 : S.C)
2024-10-17 11:43:33 +00:00
(v : (S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2)).V) :
2024-10-16 16:38:36 +00:00
TensorTree S ![c1, c2] := twoNode v
/-- The node `constTwoNodeE` of a tensor tree, with all arguments explicit. -/
abbrev constTwoNodeE (S : TensorStruct) (c1 c2 : S.C)
2024-10-16 16:42:20 +00:00
(v : 𝟙 _ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2)) :
2024-10-16 16:38:36 +00:00
TensorTree S ![c1, c2] := constTwoNode v
/-- The node `constThreeNodeE` of a tensor tree, with all arguments explicit. -/
2024-10-16 16:42:20 +00:00
abbrev constThreeNodeE (S : TensorStruct) (c1 c2 c3 : S.C)
(v : 𝟙 _ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2) ⊗
2024-10-16 16:38:36 +00:00
S.FDiscrete.obj (Discrete.mk c3)) : TensorTree S ![c1, c2, c3] :=
constThreeNode v
2024-10-08 07:52:55 +00:00
/-- The number of nodes in a tensor tree. -/
def size : ∀ {n : ℕ } {c : Fin n → S.C}, TensorTree S c → ℕ := fun
2024-10-07 12:20:53 +00:00
| tensorNode _ => 1
2024-10-16 16:38:36 +00:00
| vecNode _ => 1
| twoNode _ => 1
| threeNode _ => 1
| constNode _ => 1
| constVecNode _ => 1
| constTwoNode _ => 1
| constThreeNode _ => 1
2024-10-07 12:20:53 +00:00
| add t1 t2 => t1.size + t2.size + 1
| perm _ t => t.size + 1
2024-10-17 11:43:33 +00:00
| neg t => t.size + 1
2024-10-08 15:45:51 +00:00
| smul _ t => t.size + 1
2024-10-07 12:20:53 +00:00
| prod t1 t2 => t1.size + t2.size + 1
2024-10-16 16:38:36 +00:00
| contr _ _ _ t => t.size + 1
2024-10-08 07:26:23 +00:00
| eval _ _ t => t.size + 1
2024-10-07 12:20:53 +00:00
noncomputable section
2024-10-08 07:52:55 +00:00
/-- The underlying tensor a tensor tree corresponds to.
Note: This function is not fully defined yet. -/
def tensor : ∀ {n : ℕ } {c : Fin n → S.C}, TensorTree S c → S.F.obj (OverColor.mk c) := fun
2024-10-07 12:20:53 +00:00
| tensorNode t => t
2024-10-17 11:43:33 +00:00
| constTwoNode t => (OverColor.Discrete.pairIsoSep S.FDiscrete).hom.hom (t.hom (1 : S.k))
2024-10-07 12:20:53 +00:00
| add t1 t2 => t1.tensor + t2.tensor
| perm σ t => (S.F.map σ ).hom t.tensor
2024-10-17 11:43:33 +00:00
| neg t => - t.tensor
2024-10-08 15:45:51 +00:00
| smul a t => a • t.tensor
2024-10-07 12:20:53 +00:00
| prod t1 t2 => (S.F.map (OverColor.equivToIso finSumFinEquiv).hom).hom
((S.F.μ _ _).hom (t1.tensor ⊗ₜ t2.tensor))
2024-10-16 16:42:20 +00:00
| contr i j h t => (S.contrMap _ i j h).hom t.tensor
2024-10-07 12:20:53 +00:00
| _ => 0
2024-10-17 11:43:33 +00:00
/-!
## Tensor on different nodes.
-/
@[simp]
lemma tensoreNode_tensor {c : Fin n → S.C} (T : S.F.obj (OverColor.mk c)) :
2024-10-07 12:20:53 +00:00
(tensorNode T).tensor = T := rfl
2024-10-17 11:43:33 +00:00
@[simp]
lemma constTwoNode_tensor {c1 c2 : S.C}
(v : 𝟙 _ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2)) :
(constTwoNode v).tensor = (OverColor.Discrete.pairIsoSep S.FDiscrete).hom.hom (v.hom (1 : S.k)) :=
rfl
lemma prod_tensor {c1 : Fin n → S.C} {c2 : Fin m → S.C} (t1 : TensorTree S c1) (t2 : TensorTree S c2) :
(prod t1 t2).tensor = (S.F.map (OverColor.equivToIso finSumFinEquiv).hom).hom
((S.F.μ _ _).hom (t1.tensor ⊗ₜ t2.tensor)) := rfl
lemma add_tensor (t1 t2 : TensorTree S c) : (add t1 t2).tensor = t1.tensor + t2.tensor := rfl
lemma perm_tensor (σ : (OverColor.mk c) ⟶ (OverColor.mk c1)) (t : TensorTree S c) :
(perm σ t).tensor = (S.F.map σ ).hom t.tensor := rfl
lemma contr_tensor {n : ℕ } {c : Fin n.succ.succ → S.C} {i : Fin n.succ.succ} {j : Fin n.succ} {h : c (i.succAbove j) = S.τ (c i)}
(t : TensorTree S c) : (contr i j h t).tensor = (S.contrMap c i j h).hom t.tensor := rfl
lemma neg_tensor (t : TensorTree S c) : (neg t).tensor = - t.tensor := rfl
/-!
## Equality of tensors and rewrites.
-/
lemma contr_tensor_eq {n : ℕ } {c : Fin n.succ.succ → S.C} {T1 T2 : TensorTree S c}
(h : T1.tensor = T2.tensor) {i : Fin n.succ.succ} {j : Fin n.succ}
{h' : c (i.succAbove j) = S.τ (c i)} :
(contr i j h' T1).tensor = (contr i j h' T2).tensor := by
simp only [Nat.succ_eq_add_one, contr_tensor]
rw [h]
lemma prod_tensor_eq_fst {n m : ℕ } {c : Fin n → S.C} {c1 : Fin m → S.C}
{T1 T1' : TensorTree S c} { T2 : TensorTree S c1}
(h : T1.tensor = T1'.tensor) :
(prod T1 T2).tensor = (prod T1' T2).tensor := by
simp [prod_tensor]
rw [h]
lemma prod_tensor_eq_snd {n m : ℕ } {c : Fin n → S.C} {c1 : Fin m → S.C}
{T1 : TensorTree S c} {T2 T2' : TensorTree S c1}
(h : T2.tensor = T2'.tensor) :
(prod T1 T2).tensor = (prod T1 T2').tensor := by
simp [prod_tensor]
rw [h]
/-!
## Negation lemmas
We define the simp lemmas here so that negation is always moved to the top of the tree.
-/
@[simp]
lemma neg_neg (t : TensorTree S c) : (neg (neg t)).tensor = t.tensor := by
simp only [neg_tensor, _root_.neg_neg]
@[simp]
lemma neg_fst_prod {c1 : Fin n → S.C} {c2 : Fin m → S.C} (T1 : TensorTree S c1)
(T2 : TensorTree S c2) :
(prod (neg T1) T2).tensor = (neg (prod T1 T2)).tensor := by
simp only [prod_tensor, Functor.id_obj, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, neg_tensor, neg_tmul, map_neg]
@[simp]
lemma neg_snd_prod {c1 : Fin n → S.C} {c2 : Fin m → S.C} (T1 : TensorTree S c1)
(T2 : TensorTree S c2) :
(prod T1 (neg T2)).tensor = (neg (prod T1 T2)).tensor := by
simp only [prod_tensor, Functor.id_obj, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, neg_tensor, tmul_neg, map_neg]
@[simp]
lemma neg_contr {n : ℕ } {c : Fin n.succ.succ → S.C} {i : Fin n.succ.succ} {j : Fin n.succ} {h : c (i.succAbove j) = S.τ (c i)}
(t : TensorTree S c) : (contr i j h (neg t)).tensor = (neg (contr i j h t)).tensor := by
simp only [Nat.succ_eq_add_one, contr_tensor, neg_tensor, map_neg]
2024-10-17 17:11:06 +00:00
lemma neg_perm {n m : ℕ } {c : Fin n → S.C} {c1 : Fin m → S.C}
(σ : (OverColor.mk c) ⟶ (OverColor.mk c1)) (t : TensorTree S c) :
(perm σ (neg t)).tensor = (neg (perm σ t)).tensor := by
simp only [perm_tensor, neg_tensor, map_neg]
/-!
## Permutation lemmas
-/
lemma perm_contr {n : ℕ } {c : Fin n.succ.succ.succ → S.C} {c1 : Fin n.succ.succ.succ → S.C}
{i : Fin n.succ.succ.succ} {j : Fin n.succ.succ}
{h : c1 (i.succAbove j) = S.τ (c1 i)} (t : TensorTree S c)
(σ : (OverColor.mk c) ⟶ (OverColor.mk c1)) :
((contr i j h (perm σ t))).tensor
= (perm (extractTwo i j σ ) (contr ((OverColor.Hom.toEquiv σ ).symm i)
(((Hom.toEquiv (extractOne i σ ))).symm j) (perm_contr_cond h σ ) t)).tensor := by
rw [contr_tensor, perm_tensor]
rw [TensorStruct.contrMap]
change (
(S.contr.app { as := c1 i } ⊗
𝟙 ((OverColor.lift.obj S.FDiscrete).obj (OverColor.mk (c1 ∘ i.succAbove ∘ j.succAbove)))) ≫
(λ_ ((OverColor.lift.obj S.FDiscrete).obj (OverColor.mk (c1 ∘ i.succAbove ∘ j.succAbove)))).hom).hom
((S.contrIso c1 i j h).hom.hom ((S.F.map σ ).hom t.tensor)) = _
2024-10-07 12:20:53 +00:00
end
end TensorTree
2024-10-15 06:08:56 +00:00
end