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-11-18 15:42:37 +00:00
|
|
|
|
import HepLean.Tensors.TensorSpecies.Contractions.ContrMap
|
2024-10-07 12:20:53 +00:00
|
|
|
|
/-!
|
|
|
|
|
|
2024-11-13 14:24:59 +00:00
|
|
|
|
# Tensor trees
|
2024-10-31 07:42:50 +00:00
|
|
|
|
|
|
|
|
|
- Tensor trees provide an abstract way to represent tensor expressions.
|
|
|
|
|
- Their nodes are either tensors or operations between tensors.
|
|
|
|
|
- Every tensor tree has associated with an underlying tensor.
|
|
|
|
|
- This is not a one-to-one correspondence. Lots tensor trees represent the same tensor.
|
|
|
|
|
In the same way that lots of tensor expressions represent the same tensor.
|
|
|
|
|
- Define a sub-tensor tree as a node of a tensor tree and all child nodes thereof. One
|
|
|
|
|
can replace sub-tensor tree with another tensor tree which has the same underlying tensor
|
|
|
|
|
without changing the underlying tensor of the parent tensor tree. These appear as the e.g.
|
|
|
|
|
the lemmas `contr_tensor_eq` in what follows.
|
2024-10-07 12:20:53 +00:00
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
open IndexNotation
|
|
|
|
|
open CategoryTheory
|
2024-10-09 16:57:41 +00:00
|
|
|
|
open MonoidalCategory
|
2024-10-07 12:20:53 +00:00
|
|
|
|
|
2024-10-15 06:08:56 +00:00
|
|
|
|
noncomputable section
|
|
|
|
|
|
2024-10-08 07:52:55 +00:00
|
|
|
|
/-- A syntax tree for tensor expressions. -/
|
2024-10-31 14:03:31 +00:00
|
|
|
|
inductive TensorTree (S : TensorSpecies) : {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-11-06 11:41:59 +00:00
|
|
|
|
/-- A node corresponding to the scalar multiple of a tensor by a element of the field. -/
|
2024-11-04 15:51:03 +00:00
|
|
|
|
| smul {n : ℕ} {c : Fin n → S.C} : S.k → TensorTree S c → TensorTree S c
|
|
|
|
|
/-- A node corresponding to negation of a tensor. -/
|
|
|
|
|
| neg {n : ℕ} {c : Fin n → S.C} : TensorTree S c → TensorTree S c
|
2024-10-16 16:38:36 +00:00
|
|
|
|
/-- 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-11-06 11:41:59 +00:00
|
|
|
|
/-- A node corresponding to the action of a group element on a tensor. -/
|
2024-11-04 15:51:03 +00:00
|
|
|
|
| action {n : ℕ} {c : Fin n → S.C} : S.G → 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
|
2024-10-31 14:03:31 +00:00
|
|
|
|
/-- A node corresponding to the product of two tensors. -/
|
2024-10-07 12:20:53 +00:00
|
|
|
|
| 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-31 14:03:31 +00:00
|
|
|
|
/-- A node corresponding to the contraction of indices of a tensor. -/
|
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 →
|
2024-11-05 10:46:18 +00:00
|
|
|
|
TensorTree S (c ∘ i.succAbove ∘ j.succAbove)
|
2024-10-31 14:03:31 +00:00
|
|
|
|
/-- A node corresponding to the evaluation of an index of a tensor. -/
|
|
|
|
|
| eval {n : ℕ} {c : Fin n.succ → S.C} : (i : Fin n.succ) → (x : ℕ) → TensorTree S c →
|
2024-11-05 10:46:18 +00:00
|
|
|
|
TensorTree S (c ∘ i.succAbove)
|
2024-10-07 12:20:53 +00:00
|
|
|
|
|
|
|
|
|
namespace TensorTree
|
|
|
|
|
|
2024-10-21 12:24:17 +00:00
|
|
|
|
variable {S : TensorSpecies} {n : ℕ} {c : Fin n → S.C} (T : TensorTree S c)
|
2024-10-07 12:20:53 +00:00
|
|
|
|
|
|
|
|
|
open MonoidalCategory
|
|
|
|
|
open TensorProduct
|
|
|
|
|
|
2024-10-22 07:11:44 +00:00
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
## Composite nodes
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
/-- A node consisting of a single vector. -/
|
2024-11-05 14:37:10 +00:00
|
|
|
|
def vecNode {c : S.C} (v : S.FD.obj (Discrete.mk c)) : TensorTree S ![c] :=
|
2024-10-22 07:11:44 +00:00
|
|
|
|
perm (OverColor.mkIso (by
|
|
|
|
|
ext x; fin_cases x; rfl)).hom
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(tensorNode ((OverColor.forgetLiftApp S.FD c).symm.hom.hom v))
|
2024-10-22 07:11:44 +00:00
|
|
|
|
|
2024-10-22 06:42:06 +00:00
|
|
|
|
/-- The node `vecNode` of a tensor tree, with all arguments explicit. -/
|
|
|
|
|
abbrev vecNodeE (S : TensorSpecies) (c1 : S.C)
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(v : (S.FD.obj (Discrete.mk c1)).V) :
|
2024-10-22 06:42:06 +00:00
|
|
|
|
TensorTree S ![c1] := vecNode v
|
|
|
|
|
|
2024-10-22 07:11:44 +00:00
|
|
|
|
/-- A node consisting of a two tensor. -/
|
2024-11-05 14:37:10 +00:00
|
|
|
|
def twoNode {c1 c2 : S.C} (t : (S.FD.obj (Discrete.mk c1) ⊗
|
|
|
|
|
S.FD.obj (Discrete.mk c2)).V) :
|
2024-10-22 07:11:44 +00:00
|
|
|
|
TensorTree S ![c1, c2] :=
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(tensorNode ((OverColor.Discrete.pairIsoSep S.FD).hom.hom t))
|
2024-10-22 07:11:44 +00:00
|
|
|
|
|
2024-10-16 16:38:36 +00:00
|
|
|
|
/-- The node `twoNode` of a tensor tree, with all arguments explicit. -/
|
2024-10-21 12:24:17 +00:00
|
|
|
|
abbrev twoNodeE (S : TensorSpecies) (c1 c2 : S.C)
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(v : (S.FD.obj (Discrete.mk c1) ⊗ S.FD.obj (Discrete.mk c2)).V) :
|
2024-10-16 16:38:36 +00:00
|
|
|
|
TensorTree S ![c1, c2] := twoNode v
|
|
|
|
|
|
2024-10-22 14:19:43 +00:00
|
|
|
|
/-- A node consisting of a three tensor. -/
|
2024-11-05 14:37:10 +00:00
|
|
|
|
def threeNode {c1 c2 c3 : S.C} (t : (S.FD.obj (Discrete.mk c1) ⊗
|
|
|
|
|
S.FD.obj (Discrete.mk c2) ⊗ S.FD.obj (Discrete.mk c3)).V) :
|
2024-10-22 14:19:43 +00:00
|
|
|
|
TensorTree S ![c1, c2, c3] :=
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(tensorNode ((OverColor.Discrete.tripleIsoSep S.FD).hom.hom t))
|
2024-10-22 14:19:43 +00:00
|
|
|
|
|
|
|
|
|
/-- The node `threeNode` of a tensor tree, with all arguments explicit. -/
|
|
|
|
|
abbrev threeNodeE (S : TensorSpecies) (c1 c2 c3 : S.C)
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(v : (S.FD.obj (Discrete.mk c1) ⊗ S.FD.obj (Discrete.mk c2) ⊗
|
|
|
|
|
S.FD.obj (Discrete.mk c3)).V) :
|
2024-10-22 14:19:43 +00:00
|
|
|
|
TensorTree S ![c1, c2, c3] := threeNode v
|
|
|
|
|
|
2024-10-22 07:11:44 +00:00
|
|
|
|
/-- A general constant node. -/
|
|
|
|
|
def constNode {n : ℕ} {c : Fin n → S.C} (T : 𝟙_ (Rep S.k S.G) ⟶ S.F.obj (OverColor.mk c)) :
|
|
|
|
|
TensorTree S c := tensorNode (T.hom (1 : S.k))
|
|
|
|
|
|
|
|
|
|
/-- A constant vector. -/
|
2024-11-05 14:37:10 +00:00
|
|
|
|
def constVecNode {c : S.C} (v : 𝟙_ (Rep S.k S.G) ⟶ S.FD.obj (Discrete.mk c)) :
|
2024-10-22 07:11:44 +00:00
|
|
|
|
TensorTree S ![c] := vecNode (v.hom (1 : S.k))
|
|
|
|
|
|
2024-10-22 14:19:43 +00:00
|
|
|
|
/-- A constant two tensor (e.g. metric and unit). -/
|
2024-10-22 07:11:44 +00:00
|
|
|
|
def constTwoNode {c1 c2 : S.C}
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(v : 𝟙_ (Rep S.k S.G) ⟶ S.FD.obj (Discrete.mk c1) ⊗ S.FD.obj (Discrete.mk c2)) :
|
2024-10-22 07:11:44 +00:00
|
|
|
|
TensorTree S ![c1, c2] := twoNode (v.hom (1 : S.k))
|
|
|
|
|
|
2024-10-22 14:19:43 +00:00
|
|
|
|
/-- The node `constTwoNode` of a tensor tree, with all arguments explicit. -/
|
2024-10-21 12:24:17 +00:00
|
|
|
|
abbrev constTwoNodeE (S : TensorSpecies) (c1 c2 : S.C)
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(v : 𝟙_ (Rep S.k S.G) ⟶ S.FD.obj (Discrete.mk c1) ⊗ S.FD.obj (Discrete.mk c2)) :
|
2024-10-16 16:38:36 +00:00
|
|
|
|
TensorTree S ![c1, c2] := constTwoNode v
|
|
|
|
|
|
2024-10-22 14:19:43 +00:00
|
|
|
|
/-- A constant three tensor (e.g. Pauli matrices). -/
|
|
|
|
|
def constThreeNode {c1 c2 c3 : S.C}
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(v : 𝟙_ (Rep S.k S.G) ⟶ S.FD.obj (Discrete.mk c1) ⊗ S.FD.obj (Discrete.mk c2) ⊗
|
|
|
|
|
S.FD.obj (Discrete.mk c3)) : TensorTree S ![c1, c2, c3] :=
|
2024-10-22 14:19:43 +00:00
|
|
|
|
threeNode (v.hom (1 : S.k))
|
|
|
|
|
|
|
|
|
|
/-- The node `constThreeNode` of a tensor tree, with all arguments explicit. -/
|
|
|
|
|
abbrev constThreeNodeE (S : TensorSpecies) (c1 c2 c3 : S.C)
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(v : 𝟙_ (Rep S.k S.G) ⟶ S.FD.obj (Discrete.mk c1) ⊗ S.FD.obj (Discrete.mk c2) ⊗
|
|
|
|
|
S.FD.obj (Discrete.mk c3)) : TensorTree S ![c1, c2, c3] :=
|
2024-10-22 14:19:43 +00:00
|
|
|
|
constThreeNode v
|
|
|
|
|
|
2024-10-22 07:11:44 +00:00
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
## Other operations.
|
2024-10-16 16:38:36 +00:00
|
|
|
|
|
2024-10-22 07:11:44 +00:00
|
|
|
|
-/
|
2024-10-08 07:52:55 +00:00
|
|
|
|
/-- The number of nodes in a tensor tree. -/
|
2024-11-04 15:51:03 +00:00
|
|
|
|
def size {n : ℕ} {c : Fin n → S.C} : TensorTree S c → ℕ := fun
|
2024-10-07 12:20:53 +00:00
|
|
|
|
| tensorNode _ => 1
|
|
|
|
|
| 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-29 12:32:33 +00:00
|
|
|
|
| action _ t => t.size + 1
|
2024-10-07 12:20:53 +00:00
|
|
|
|
|
|
|
|
|
noncomputable section
|
|
|
|
|
|
2024-10-31 14:13:35 +00:00
|
|
|
|
/-- The underlying tensor a tensor tree corresponds to. -/
|
2024-11-04 15:51:03 +00:00
|
|
|
|
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-11-04 15:51:03 +00:00
|
|
|
|
| smul a t => a • t.tensor
|
|
|
|
|
| neg t => - t.tensor
|
2024-10-07 12:20:53 +00:00
|
|
|
|
| add t1 t2 => t1.tensor + t2.tensor
|
2024-11-04 15:51:03 +00:00
|
|
|
|
| action g t => (S.F.obj (OverColor.mk _)).ρ g t.tensor
|
2024-10-07 12:20:53 +00:00
|
|
|
|
| perm σ t => (S.F.map σ).hom t.tensor
|
|
|
|
|
| prod t1 t2 => (S.F.map (OverColor.equivToIso finSumFinEquiv).hom).hom
|
2024-12-10 13:44:39 +00:00
|
|
|
|
((Functor.LaxMonoidal.μ 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-11-02 08:50:17 +00:00
|
|
|
|
| eval i e t => (S.evalMap i (Fin.ofNat' _ e)) t.tensor
|
2024-10-07 12:20:53 +00:00
|
|
|
|
|
2024-10-22 10:47:37 +00:00
|
|
|
|
/-- Takes a tensor tree based on `Fin 0`, into the field `S.k`. -/
|
2024-10-22 10:33:28 +00:00
|
|
|
|
def field {c : Fin 0 → S.C} (t : TensorTree S c) : S.k := S.castFin0ToField t.tensor
|
|
|
|
|
|
2024-10-17 11:43:33 +00:00
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
## Tensor on different nodes.
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
@[simp]
|
2024-10-19 15:26:57 +00:00
|
|
|
|
lemma tensorNode_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}
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(v : 𝟙_ (Rep S.k S.G) ⟶ S.FD.obj (Discrete.mk c1) ⊗ S.FD.obj (Discrete.mk c2)) :
|
2024-10-19 09:47:23 +00:00
|
|
|
|
(constTwoNode v).tensor =
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(OverColor.Discrete.pairIsoSep S.FD).hom.hom (v.hom (1 : S.k)) :=
|
2024-10-17 11:43:33 +00:00
|
|
|
|
rfl
|
|
|
|
|
|
2024-10-23 06:50:55 +00:00
|
|
|
|
@[simp]
|
|
|
|
|
lemma constThreeNode_tensor {c1 c2 c3 : S.C}
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(v : 𝟙_ (Rep S.k S.G) ⟶ S.FD.obj (Discrete.mk c1) ⊗ S.FD.obj (Discrete.mk c2) ⊗
|
|
|
|
|
S.FD.obj (Discrete.mk c3)) :
|
2024-10-23 06:50:55 +00:00
|
|
|
|
(constThreeNode v).tensor =
|
2024-11-05 14:37:10 +00:00
|
|
|
|
(OverColor.Discrete.tripleIsoSep S.FD).hom.hom (v.hom (1 : S.k)) :=
|
2024-10-23 06:50:55 +00:00
|
|
|
|
rfl
|
|
|
|
|
|
2024-10-19 09:47:23 +00:00
|
|
|
|
lemma prod_tensor {c1 : Fin n → S.C} {c2 : Fin m → S.C} (t1 : TensorTree S c1)
|
|
|
|
|
(t2 : TensorTree S c2) :
|
2024-10-17 11:43:33 +00:00
|
|
|
|
(prod t1 t2).tensor = (S.F.map (OverColor.equivToIso finSumFinEquiv).hom).hom
|
2024-12-10 13:44:39 +00:00
|
|
|
|
((Functor.LaxMonoidal.μ S.F _ _).hom (t1.tensor ⊗ₜ t2.tensor)) := rfl
|
2024-10-17 11:43:33 +00:00
|
|
|
|
|
|
|
|
|
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
|
|
|
|
|
|
2024-10-19 09:47:23 +00:00
|
|
|
|
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
|
2024-10-17 11:43:33 +00:00
|
|
|
|
|
|
|
|
|
lemma neg_tensor (t : TensorTree S c) : (neg t).tensor = - t.tensor := rfl
|
|
|
|
|
|
2024-10-22 10:47:37 +00:00
|
|
|
|
lemma eval_tensor {n : ℕ} {c : Fin n.succ → S.C} (i : Fin n.succ) (e : ℕ) (t : TensorTree S c) :
|
2024-11-02 08:50:17 +00:00
|
|
|
|
(eval i e t).tensor = (S.evalMap i (Fin.ofNat' (S.repDim (c i)) e)) t.tensor := rfl
|
2024-10-21 16:07:39 +00:00
|
|
|
|
|
2024-10-23 15:19:41 +00:00
|
|
|
|
lemma smul_tensor {c : Fin n → S.C} (a : S.k) (T : TensorTree S c) :
|
2024-10-24 06:10:08 +00:00
|
|
|
|
(smul a T).tensor = a • T.tensor:= rfl
|
2024-10-29 12:32:33 +00:00
|
|
|
|
|
|
|
|
|
lemma action_tensor {c : Fin n → S.C} (g : S.G) (T : TensorTree S c) :
|
|
|
|
|
(action g T).tensor = (S.F.obj (OverColor.mk c)).ρ g T.tensor := rfl
|
|
|
|
|
|
2024-10-17 11:43:33 +00:00
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
## 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}
|
2024-10-19 09:19:29 +00:00
|
|
|
|
{T1 T1' : TensorTree S c} { T2 : TensorTree S c1}
|
2024-10-17 11:43:33 +00:00
|
|
|
|
(h : T1.tensor = T1'.tensor) :
|
|
|
|
|
(prod T1 T2).tensor = (prod T1' T2).tensor := by
|
2024-10-19 09:19:29 +00:00
|
|
|
|
simp only [prod_tensor, Functor.id_obj, OverColor.mk_hom, Action.instMonoidalCategory_tensorObj_V,
|
|
|
|
|
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
|
|
|
|
Action.FunctorCategoryEquivalence.functor_obj_obj]
|
2024-10-17 11:43:33 +00:00
|
|
|
|
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
|
2024-10-19 09:19:29 +00:00
|
|
|
|
simp only [prod_tensor, Functor.id_obj, OverColor.mk_hom, Action.instMonoidalCategory_tensorObj_V,
|
|
|
|
|
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
|
|
|
|
Action.FunctorCategoryEquivalence.functor_obj_obj]
|
2024-10-17 11:43:33 +00:00
|
|
|
|
rw [h]
|
|
|
|
|
|
2024-10-21 05:38:22 +00:00
|
|
|
|
lemma perm_tensor_eq {n m : ℕ} {c : Fin n → S.C} {c1 : Fin m → S.C}
|
|
|
|
|
{σ : (OverColor.mk c) ⟶ (OverColor.mk c1)} {T1 T2 : TensorTree S c}
|
|
|
|
|
(h : T1.tensor = T2.tensor) :
|
|
|
|
|
(perm σ T1).tensor = (perm σ T2).tensor := by
|
|
|
|
|
simp only [perm_tensor]
|
|
|
|
|
rw [h]
|
|
|
|
|
|
2024-10-22 11:49:58 +00:00
|
|
|
|
lemma add_tensor_eq_fst {T1 T1' T2 : TensorTree S c} (h : T1.tensor = T1'.tensor) :
|
|
|
|
|
(add T1 T2).tensor = (add T1' T2).tensor := by
|
|
|
|
|
simp only [add_tensor]
|
|
|
|
|
rw [h]
|
|
|
|
|
|
|
|
|
|
lemma add_tensor_eq_snd {T1 T2 T2' : TensorTree S c} (h : T2.tensor = T2'.tensor) :
|
|
|
|
|
(add T1 T2).tensor = (add T1 T2').tensor := by
|
|
|
|
|
simp only [add_tensor]
|
|
|
|
|
rw [h]
|
|
|
|
|
|
|
|
|
|
lemma add_tensor_eq {T1 T1' T2 T2' : TensorTree S c} (h1 : T1.tensor = T1'.tensor)
|
|
|
|
|
(h2 : T2.tensor = T2'.tensor) :
|
|
|
|
|
(add T1 T2).tensor = (add T1' T2').tensor := by
|
|
|
|
|
simp only [add_tensor]
|
|
|
|
|
rw [h1, h2]
|
|
|
|
|
|
|
|
|
|
lemma neg_tensor_eq {T1 T2 : TensorTree S c} (h : T1.tensor = T2.tensor) :
|
|
|
|
|
(neg T1).tensor = (neg T2).tensor := by
|
|
|
|
|
simp only [neg_tensor]
|
|
|
|
|
rw [h]
|
|
|
|
|
|
2024-10-23 15:19:41 +00:00
|
|
|
|
lemma smul_tensor_eq {T1 T2 : TensorTree S c} {a : S.k} (h : T1.tensor = T2.tensor) :
|
|
|
|
|
(smul a T1).tensor = (smul a T2).tensor := by
|
|
|
|
|
simp only [smul_tensor]
|
|
|
|
|
rw [h]
|
|
|
|
|
|
2024-10-29 12:32:33 +00:00
|
|
|
|
lemma action_tensor_eq {T1 T2 : TensorTree S c} {g : S.G} (h : T1.tensor = T2.tensor) :
|
|
|
|
|
(action g T1).tensor = (action g T2).tensor := by
|
|
|
|
|
simp only [action_tensor]
|
|
|
|
|
rw [h]
|
|
|
|
|
|
2024-10-25 13:50:19 +00:00
|
|
|
|
lemma smul_mul_eq {T1 : TensorTree S c} {a b : S.k} (h : a = b) :
|
|
|
|
|
(smul a T1).tensor = (smul b T1).tensor := by
|
|
|
|
|
rw [h]
|
|
|
|
|
|
2024-10-23 19:43:58 +00:00
|
|
|
|
lemma eq_tensorNode_of_eq_tensor {T1 : TensorTree S c} {t : S.F.obj (OverColor.mk c)}
|
|
|
|
|
(h : T1.tensor = t) : T1.tensor = (tensorNode t).tensor := by
|
|
|
|
|
simpa using h
|
|
|
|
|
|
2024-10-25 13:50:19 +00:00
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
## The zero tensor tree
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
/-- The zero tensor tree. -/
|
|
|
|
|
def zeroTree {n : ℕ} {c : Fin n → S.C} : TensorTree S c := tensorNode 0
|
|
|
|
|
|
|
|
|
|
@[simp]
|
|
|
|
|
lemma zeroTree_tensor {n : ℕ} {c : Fin n → S.C} : (zeroTree (c := c)).tensor = 0 := by
|
|
|
|
|
rfl
|
|
|
|
|
|
2024-10-25 13:54:58 +00:00
|
|
|
|
lemma zero_smul {T1 : TensorTree S c} :
|
2024-10-25 13:50:19 +00:00
|
|
|
|
(smul 0 T1).tensor = zeroTree.tensor := by
|
|
|
|
|
simp only [smul_tensor, _root_.zero_smul, zeroTree_tensor]
|
|
|
|
|
|
2024-10-25 13:54:58 +00:00
|
|
|
|
lemma smul_zero {a : S.k} : (smul a (zeroTree (c := c))).tensor = zeroTree.tensor := by
|
2024-10-25 13:50:19 +00:00
|
|
|
|
simp only [smul_tensor, zeroTree_tensor, _root_.smul_zero]
|
|
|
|
|
|
|
|
|
|
lemma zero_add {T1 : TensorTree S c} : (add zeroTree T1).tensor = T1.tensor := by
|
|
|
|
|
simp only [add_tensor, zeroTree_tensor, _root_.zero_add]
|
|
|
|
|
|
|
|
|
|
lemma add_zero {T1 : TensorTree S c} : (add T1 zeroTree).tensor = T1.tensor := by
|
|
|
|
|
simp only [add_tensor, zeroTree_tensor, _root_.add_zero]
|
|
|
|
|
|
|
|
|
|
lemma perm_zero {n m : ℕ} {c : Fin n → S.C} {c1 : Fin m → S.C} (σ : (OverColor.mk c) ⟶
|
|
|
|
|
(OverColor.mk c1)) : (perm σ zeroTree).tensor = zeroTree.tensor := by
|
|
|
|
|
simp only [perm_tensor, zeroTree_tensor, map_zero]
|
|
|
|
|
|
2024-10-25 13:54:58 +00:00
|
|
|
|
lemma neg_zero : (neg (zeroTree (c := c))).tensor = zeroTree.tensor := by
|
2024-10-25 13:50:19 +00:00
|
|
|
|
simp only [neg_tensor, zeroTree_tensor, _root_.neg_zero]
|
|
|
|
|
|
2024-10-25 13:54:58 +00:00
|
|
|
|
lemma contr_zero {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)} : (contr i j h zeroTree).tensor = zeroTree.tensor := by
|
2024-10-25 13:50:19 +00:00
|
|
|
|
simp only [contr_tensor, zeroTree_tensor, map_zero]
|
|
|
|
|
|
|
|
|
|
lemma zero_prod {n m : ℕ} {c : Fin n → S.C} {c1 : Fin m → S.C} (t : TensorTree S c1) :
|
|
|
|
|
(prod (zeroTree (c := c)) t).tensor = zeroTree.tensor := by
|
|
|
|
|
simp only [prod_tensor, Functor.id_obj, OverColor.mk_hom, Action.instMonoidalCategory_tensorObj_V,
|
|
|
|
|
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
|
|
|
|
Action.FunctorCategoryEquivalence.functor_obj_obj, zeroTree_tensor, zero_tmul, map_zero]
|
|
|
|
|
|
|
|
|
|
lemma prod_zero {n m : ℕ} {c : Fin n → S.C} {c1 : Fin m → S.C} (t : TensorTree S c) :
|
|
|
|
|
(prod t (zeroTree (c := c1))).tensor = zeroTree.tensor := by
|
|
|
|
|
simp only [prod_tensor, Functor.id_obj, OverColor.mk_hom, Action.instMonoidalCategory_tensorObj_V,
|
|
|
|
|
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
|
|
|
|
Action.FunctorCategoryEquivalence.functor_obj_obj, zeroTree_tensor, tmul_zero, map_zero]
|
|
|
|
|
|
2024-10-21 12:34:42 +00:00
|
|
|
|
/-- A structure containing a pair of indices (i, j) to be contracted in a tensor.
|
|
|
|
|
This is used in some proofs of node identities for tensor trees. -/
|
|
|
|
|
structure ContrPair {n : ℕ} (c : Fin n.succ.succ → S.C) where
|
|
|
|
|
/-- The first index in the pair, appearing on the left in the contraction
|
|
|
|
|
node `contr i j h _`. -/
|
|
|
|
|
i : Fin n.succ.succ
|
|
|
|
|
/-- The second index in the pair, appearing on the right in the contraction
|
|
|
|
|
node `contr i j h _`. -/
|
|
|
|
|
j : Fin n.succ
|
|
|
|
|
/-- A proof that the two indices can be contracted. -/
|
|
|
|
|
h : c (i.succAbove j) = S.τ (c i)
|
|
|
|
|
|
2024-10-21 13:40:23 +00:00
|
|
|
|
namespace ContrPair
|
2024-10-27 17:07:45 +00:00
|
|
|
|
variable {n : ℕ} {c : Fin n.succ.succ → S.C} (q q' : ContrPair c)
|
2024-10-21 13:40:23 +00:00
|
|
|
|
|
|
|
|
|
lemma ext (hi : q.i = q'.i) (hj : q.j = q'.j) : q = q' := by
|
|
|
|
|
cases q
|
|
|
|
|
cases q'
|
|
|
|
|
subst hi
|
|
|
|
|
subst hj
|
|
|
|
|
rfl
|
|
|
|
|
|
|
|
|
|
/-- The contraction map for a pair of indices. -/
|
|
|
|
|
def contrMap := S.contrMap c q.i q.j q.h
|
|
|
|
|
|
|
|
|
|
end ContrPair
|
2024-10-07 12:20:53 +00:00
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
end TensorTree
|
2024-10-15 06:08:56 +00:00
|
|
|
|
|
|
|
|
|
end
|