feat: Add action node

This commit is contained in:
jstoobysmith 2024-10-29 12:32:33 +00:00
parent f7499f8d86
commit 202df13a17
5 changed files with 175 additions and 0 deletions

View file

@ -185,4 +185,63 @@ lemma pauliContrDown_eq_metric_mul_pauliContr :
apply perm_congr _ rfl
decide
/-!
## Group actions
-/
/-- The tensor `pauliContr` is invariant under the action of `SL(2,)`. -/
lemma action_pauliContr (g : SL(2,)) : {g •ₐ pauliContr | μ α β}ᵀ.tensor =
{pauliContr | μ α β}ᵀ.tensor := by
rw [tensorNode_pauliContr, constThreeNodeE]
rw [← action_constThreeNode _ g]
rfl
/-- The tensor `pauliCo` is invariant under the action of `SL(2,)`. -/
lemma action_pauliCo (g : SL(2,)) : {g •ₐ pauliCo | μ α β}ᵀ.tensor =
{pauliCo | μ α β}ᵀ.tensor := by
conv =>
lhs
rw [action_tensor_eq <| tensorNode_pauliCo]
rw [(contr_action _ _).symm]
rw [contr_tensor_eq <| (prod_action _ _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| action_constTwoNode _ _]
rw [contr_tensor_eq <| prod_tensor_eq_snd <| action_constThreeNode _ _]
rfl
/-- The tensor `pauliCoDown` is invariant under the action of `SL(2,)`. -/
lemma action_pauliCoDown (g : SL(2,)) : {g •ₐ pauliCoDown | μ α β}ᵀ.tensor =
{pauliCoDown | μ α β}ᵀ.tensor := by
conv =>
lhs
rw [action_tensor_eq <| tensorNode_pauliCoDown]
rw [(contr_action _ _).symm]
rw [contr_tensor_eq <| (prod_action _ _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| (contr_action _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| contr_tensor_eq <| (prod_action _ _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| contr_tensor_eq <| prod_tensor_eq_fst <|
action_pauliCo _]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| contr_tensor_eq <| prod_tensor_eq_snd <|
action_constTwoNode _ _]
rw [contr_tensor_eq <| prod_tensor_eq_snd <| action_constTwoNode _ _]
rfl
/-- The tensor `pauliContrDown` is invariant under the action of `SL(2,)`. -/
lemma action_pauliContrDown (g : SL(2,)) : {g •ₐ pauliContrDown | μ α β}ᵀ.tensor =
{pauliContrDown | μ α β}ᵀ.tensor := by
conv =>
lhs
rw [action_tensor_eq <| tensorNode_pauliContrDown]
rw [(contr_action _ _).symm]
rw [contr_tensor_eq <| (prod_action _ _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| (contr_action _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| contr_tensor_eq <| (prod_action _ _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| contr_tensor_eq <| prod_tensor_eq_fst <|
action_pauliContr _]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| contr_tensor_eq <| prod_tensor_eq_snd <|
action_constTwoNode _ _]
rw [contr_tensor_eq <| prod_tensor_eq_snd <| action_constTwoNode _ _]
rfl
end complexLorentzTensor

View file

@ -543,9 +543,13 @@ inductive TensorTree (S : TensorSpecies) : ∀ {n : }, (Fin n → S.C) → Ty
| smul {n : } {c : Fin n → S.C} : S.k → TensorTree S c → TensorTree S c
/-- The negative of a node. -/
| neg {n : } {c : Fin n → S.C} : TensorTree S c → TensorTree S c
/-- The contraction of indices. -/
| 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)
/-- The group action on a tensor. -/
| action {n : } {c : Fin n → S.C} : S.G → TensorTree S c → TensorTree S c
/-- The evaluation of an index-/
| eval {n : } {c : Fin n.succ → S.C} :
(i : Fin n.succ) → (x : ) → TensorTree S c →
TensorTree S (c ∘ Fin.succAbove i)
@ -642,6 +646,7 @@ def size : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → := fun
| prod t1 t2 => t1.size + t2.size + 1
| contr _ _ _ t => t.size + 1
| eval _ _ t => t.size + 1
| action _ t => t.size + 1
noncomputable section
@ -657,6 +662,7 @@ def tensor : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → S.F.obj (Over
((S.F.μ _ _).hom (t1.tensor ⊗ₜ t2.tensor))
| contr i j h t => (S.contrMap _ i j h).hom t.tensor
| eval i e t => (S.evalMap i (Fin.ofNat' e Fin.size_pos')) t.tensor
| action g t => (S.F.obj (OverColor.mk _)).ρ g t.tensor
/-- Takes a tensor tree based on `Fin 0`, into the field `S.k`. -/
def field {c : Fin 0 → S.C} (t : TensorTree S c) : S.k := S.castFin0ToField t.tensor
@ -707,6 +713,10 @@ lemma eval_tensor {n : } {c : Fin n.succ → S.C} (i : Fin n.succ) (e : )
lemma smul_tensor {c : Fin n → S.C} (a : S.k) (T : TensorTree S c) :
(smul a T).tensor = a • T.tensor:= rfl
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
/-!
## Equality of tensors and rewrites.
@ -770,6 +780,11 @@ lemma smul_tensor_eq {T1 T2 : TensorTree S c} {a : S.k} (h : T1.tensor = T2.tens
simp only [smul_tensor]
rw [h]
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]
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]

View file

@ -47,6 +47,10 @@ def dotString (m : ) (nt : ) : ∀ {n : } {c : Fin n → S.C}, TensorTr
let evalNode := " node" ++ toString m ++ " [label=\"eval\", shape=box];\n"
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
evalNode ++ dotString (m + 1) nt t1 ++ edge1
| action g t =>
let actionNode := " node" ++ toString m ++ " [label=\"action\", shape=box];\n"
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"
actionNode ++ dotString (m + 1) nt t ++ edge1
| contr i j _ t1 =>
let contrNode := " node" ++ toString m ++ " [label=\"contr " ++ toString i ++ " "
++ toString j ++ "\", shape=box];\n"

View file

@ -111,6 +111,9 @@ syntax "(" tensorExpr ")" : tensorExpr
/-- Scalar multiplication for tensors. -/
syntax term "•ₜ" tensorExpr : tensorExpr
/-- group action for tensors. -/
syntax term "•ₐ" tensorExpr : tensorExpr
/-- Negation of a tensor tree. -/
syntax "-" tensorExpr : tensorExpr
@ -440,6 +443,14 @@ def smulSyntax (c T : Term) : Term :=
end SMul
namespace Action
/-- The syntax associated with the group action of tensors. -/
def actionSyntax (c T : Term) : Term :=
Syntax.mkApp (mkIdent ``TensorTree.action) #[c, T]
end Action
namespace Add
/-- Gets the indices associated with the LHS of an addition. -/
@ -522,6 +533,8 @@ partial def syntaxFull (stx : Syntax) : TermElabM Term := do
return negNode.negSyntax (← syntaxFull a)
| `(tensorExpr| $c:term •ₜ $a:tensorExpr) => do
return SMul.smulSyntax c (← syntaxFull a)
| `(tensorExpr| $c:term •ₐ $a:tensorExpr) => do
return Action.actionSyntax c (← syntaxFull a)
| `(tensorExpr| $a + $b) => do
let indicesLeft ← Add.getIndicesLeft stx
let indicesRight ← Add.getIndicesRight stx

View file

@ -244,4 +244,88 @@ lemma prod_add_both {n m : } {c : Fin n → S.C} {c1 : Fin m → S.C}
rw [add_tensor_eq_fst (prod_add _ _ _)]
rw [add_tensor_eq_snd (prod_add _ _ _)]
/-!
## Nodes and the group action.
-/
lemma smul_action {n : } {c : Fin n → S.C} (g : S.G) (a : S.k) (t : TensorTree S c) :
(smul a (action g t)).tensor = (action g (smul a t)).tensor := by
simp only [smul_tensor, action_tensor, map_smul]
lemma contr_action {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)} (g : S.G) (t : TensorTree S c) :
(contr i j h (action g t)).tensor = (action g (contr i j h t)).tensor := by
simp only [Nat.succ_eq_add_one, contr_tensor, action_tensor]
change ((S.F.obj (OverColor.mk c)).ρ g ≫ (S.contrMap c i j h).hom) t.tensor = _
erw [(S.contrMap c i j h).comm g]
rfl
lemma prod_action {n n1 : } {c : Fin n → S.C} {c1 : Fin n1 → S.C} (g : S.G)
(t : TensorTree S c) (t1 : TensorTree S c1) :
(prod (action g t) (action g t1)).tensor = (action g (prod t t1)).tensor := by
simp only [prod_tensor, action_tensor, map_tmul]
change _ = ((S.F.map (equivToIso finSumFinEquiv).hom).hom ≫
(S.F.obj (OverColor.mk (Sum.elim c c1 ∘ ⇑finSumFinEquiv.symm))).ρ g)
(((S.F.μ (OverColor.mk c) (OverColor.mk c1)).hom (t.tensor ⊗ₜ[S.k] t1.tensor)))
erw [← (S.F.map (equivToIso finSumFinEquiv).hom).comm g]
simp
change _ = (S.F.map (equivToIso finSumFinEquiv).hom).hom
(((S.F.μ (OverColor.mk c) (OverColor.mk c1)).hom ≫ (S.F.obj (OverColor.mk (Sum.elim c c1))).ρ g)
(t.tensor ⊗ₜ[S.k] t1.tensor))
erw [← (S.F.μ (OverColor.mk c) (OverColor.mk c1)).comm g]
rfl
lemma add_action {n : } {c : Fin n → S.C} (g : S.G) (t t1 : TensorTree S c) :
(add (action g t) (action g t1)).tensor = (action g (add t t1)).tensor := by
simp only [add_tensor, action_tensor, map_add]
lemma perm_action {n m : } {c : Fin n → S.C} {c1 : Fin m → S.C}
(σ : (OverColor.mk c) ⟶ (OverColor.mk c1)) (g : S.G) (t : TensorTree S c) :
(perm σ (action g t)).tensor = (action g (perm σ t)).tensor := by
simp only [perm_tensor, action_tensor]
change (((S.F.obj (OverColor.mk c)).ρ g) ≫ (S.F.map σ).hom) t.tensor = _
erw [(S.F.map σ).comm g]
rfl
lemma neg_action {n : } {c : Fin n → S.C} (g : S.G) (t : TensorTree S c) :
(neg (action g t)).tensor = (action g (neg t)).tensor := by
simp only [neg_tensor, action_tensor, map_neg]
lemma action_action {n : } {c : Fin n → S.C} (g h : S.G) (t : TensorTree S c) :
(action g (action h t)).tensor = (action (g * h) t).tensor := by
simp only [action_tensor, map_mul, LinearMap.mul_apply]
lemma action_id {n : } {c : Fin n → S.C} (t : TensorTree S c) :
(action 1 t).tensor = t.tensor := by
simp only [action_tensor, map_one, LinearMap.one_apply]
lemma action_constTwoNode {c1 c2 : S.C}
(v : 𝟙_ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2))
(g : S.G) : (action g (constTwoNode v)).tensor = (constTwoNode v).tensor := by
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, action_tensor, constTwoNode_tensor,
Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V]
change ((Discrete.pairIsoSep S.FDiscrete).hom.hom ≫ (S.F.obj (OverColor.mk ![c1, c2])).ρ g)
((v.hom _)) = _
erw [← (Discrete.pairIsoSep S.FDiscrete).hom.comm g]
change ((v.hom ≫ (S.FDiscrete.obj { as := c1 } ⊗ S.FDiscrete.obj { as := c2 }).ρ g) ≫
(Discrete.pairIsoSep S.FDiscrete).hom.hom) _ =_
erw [← v.comm g]
simp
lemma action_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))
(g : S.G) : (action g (constThreeNode v)).tensor = (constThreeNode v).tensor := by
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, action_tensor, constThreeNode_tensor,
Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V]
change ((Discrete.tripleIsoSep S.FDiscrete).hom.hom ≫ (S.F.obj (OverColor.mk ![c1, c2, c3])).ρ g)
((v.hom _)) = _
erw [← (Discrete.tripleIsoSep S.FDiscrete).hom.comm g]
change ((v.hom ≫ (S.FDiscrete.obj { as := c1 } ⊗ S.FDiscrete.obj { as := c2 } ⊗
S.FDiscrete.obj { as := c3 }).ρ g) ≫ (Discrete.tripleIsoSep S.FDiscrete).hom.hom) _ =_
erw [← v.comm g]
simp
end TensorTree