feat: addition elab and node identities
This commit is contained in:
parent
ecb2c7778c
commit
6fbace33da
5 changed files with 165 additions and 70 deletions
|
@ -95,7 +95,8 @@ lemma on_cubeTriLin (S : (PlusU1 n).Charges) :
|
|||
|
||||
lemma on_cubeTriLin_AFL (S : (PlusU1 n).LinSols) :
|
||||
cubeTriLin (Y n).val (Y n).val S.val = 0 := by
|
||||
rw [on_cubeTriLin, YYsol S]; simp
|
||||
rw [on_cubeTriLin, YYsol S]
|
||||
rfl
|
||||
|
||||
lemma on_cubeTriLin' (S : (PlusU1 n).Charges) :
|
||||
cubeTriLin (Y n).val S S = 6 * accQuad S := by
|
||||
|
@ -107,7 +108,8 @@ lemma on_cubeTriLin' (S : (PlusU1 n).Charges) :
|
|||
|
||||
lemma on_cubeTriLin'_ALQ (S : (PlusU1 n).QuadSols) :
|
||||
cubeTriLin (Y n).val S.val S.val = 0 := by
|
||||
rw [on_cubeTriLin', quadSol S]; simp
|
||||
rw [on_cubeTriLin', quadSol S]
|
||||
rfl
|
||||
|
||||
lemma add_AFL_cube (S : (PlusU1 n).LinSols) (a b : ℚ) :
|
||||
accCube (a • S.val + b • (Y n).val) =
|
||||
|
|
|
@ -115,7 +115,7 @@ lemma coMetric_0_0_field : {Lorentz.coMetric | 0 0}ᵀ.field = 1 := by
|
|||
set_option maxRecDepth 20000 in
|
||||
lemma contr_rank_2_symm {T1 : (Lorentz.complexContr ⊗ Lorentz.complexContr).V}
|
||||
{T2 : (Lorentz.complexCo ⊗ Lorentz.complexCo).V} :
|
||||
{(T1 | μ ν ⊗ T2 | μ ν) = (T2 | μ ν ⊗ T1 | μ ν)}ᵀ := by
|
||||
{T1 | μ ν ⊗ T2 | μ ν = T2 | μ ν ⊗ T1 | μ ν}ᵀ := by
|
||||
rw [perm_tensor_eq (contr_tensor_eq (contr_tensor_eq (prod_comm _ _ _ _)))]
|
||||
rw [perm_tensor_eq (contr_tensor_eq (perm_contr _ _))]
|
||||
rw [perm_tensor_eq (perm_contr _ _)]
|
||||
|
@ -134,7 +134,7 @@ lemma contr_rank_2_symm {T1 : (Lorentz.complexContr ⊗ Lorentz.complexContr).V}
|
|||
|
||||
lemma contr_rank_2_symm' {T1 : (Lorentz.complexCo ⊗ Lorentz.complexCo).V}
|
||||
{T2 : (Lorentz.complexContr ⊗ Lorentz.complexContr).V} :
|
||||
{(T1 | μ ν ⊗ T2 | μ ν) = (T2 | μ ν ⊗ T1 | μ ν)}ᵀ := by
|
||||
{T1 | μ ν ⊗ T2 | μ ν = T2 | μ ν ⊗ T1 | μ ν}ᵀ := by
|
||||
rw [perm_tensor_eq contr_rank_2_symm]
|
||||
rw [perm_perm]
|
||||
rw [perm_eq_id]
|
||||
|
@ -181,6 +181,13 @@ lemma symm_contr_antiSymm {S : (Lorentz.complexCo ⊗ Lorentz.complexCo).V}
|
|||
rw [contr_rank_2_symm', perm_tensor, antiSymm_contr_symm hA hs]
|
||||
rfl
|
||||
|
||||
lemma antiSymm_add_self {A : (Lorentz.complexContr ⊗ Lorentz.complexContr).V}
|
||||
(hA : {A | μ ν = - (A | ν μ)}ᵀ) :
|
||||
{A | μ ν + A | ν μ}ᵀ.tensor = 0 := by
|
||||
rw [← TensorTree.add_neg (twoNodeE complexLorentzTensor Color.up Color.up A)]
|
||||
apply TensorTree.add_tensor_eq_snd
|
||||
rw [neg_tensor_eq hA, neg_tensor_eq (neg_perm _ _), neg_neg]
|
||||
|
||||
end Fermion
|
||||
|
||||
end
|
||||
|
|
|
@ -684,6 +684,27 @@ lemma perm_tensor_eq {n m : ℕ} {c : Fin n → S.C} {c1 : Fin m → S.C}
|
|||
simp only [perm_tensor]
|
||||
rw [h]
|
||||
|
||||
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]
|
||||
|
||||
/-- 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
|
||||
|
|
|
@ -370,52 +370,11 @@ partial def syntaxFull (stx : Syntax) : TermElabM Term := do
|
|||
|
||||
end ProdNode
|
||||
|
||||
namespace negNode
|
||||
|
||||
/-- The syntax associated with a product of tensors. -/
|
||||
def negSyntax (T1 : Term) : Term :=
|
||||
Syntax.mkApp (mkIdent ``TensorTree.neg) #[T1]
|
||||
|
||||
end negNode
|
||||
|
||||
/-- Returns the full list of indices after contraction. TODO: Include evaluation. -/
|
||||
partial def getIndicesFull (stx : Syntax) : TermElabM (List (TSyntax `indexExpr)) := do
|
||||
match stx with
|
||||
| `(tensorExpr| $_:term | $[$args]*) => do
|
||||
return (← TensorNode.withoutContr stx)
|
||||
| `(tensorExpr| $_:tensorExpr ⊗ $_:tensorExpr) => do
|
||||
return (← ProdNode.withoutContr stx)
|
||||
| `(tensorExpr| ($a:tensorExpr)) => do
|
||||
return (← getIndicesFull a)
|
||||
| `(tensorExpr| -$a:tensorExpr) => do
|
||||
return (← getIndicesFull a)
|
||||
| _ =>
|
||||
throwError "Unsupported tensor expression syntax in getIndicesProd: {stx}"
|
||||
|
||||
namespace Equality
|
||||
|
||||
/-!
|
||||
|
||||
## For equality.
|
||||
## Permutation constructions
|
||||
|
||||
-/
|
||||
|
||||
/-- Gets the indices associated with the LHS of an equality. -/
|
||||
partial def getIndicesLeft (stx : Syntax) : TermElabM (List (TSyntax `indexExpr)) := do
|
||||
match stx with
|
||||
| `(tensorExpr| $a:tensorExpr = $_:tensorExpr) => do
|
||||
return (← getIndicesFull a)
|
||||
| _ =>
|
||||
throwError "Unsupported tensor expression syntax in getIndicesProd: {stx}"
|
||||
|
||||
/-- Gets the indices associated with the RHS of an equality. -/
|
||||
partial def getIndicesRight (stx : Syntax) : TermElabM (List (TSyntax `indexExpr)) := do
|
||||
match stx with
|
||||
| `(tensorExpr| $_:tensorExpr = $a:tensorExpr) => do
|
||||
return (← getIndicesFull a)
|
||||
| _ =>
|
||||
throwError "Unsupported tensor expression syntax in getIndicesProd: {stx}"
|
||||
|
||||
/-- Given two lists of indices returns the `List (ℕ)` representing the how one list
|
||||
permutes into the other. -/
|
||||
def getPermutation (l1 l2 : List (TSyntax `indexExpr)) : TermElabM (List (ℕ)) := do
|
||||
|
@ -445,6 +404,82 @@ def getPermutationSyntax (l1 l2 : List (TSyntax `indexExpr)) : TermElabM Term :=
|
|||
let stx := Syntax.mkApp (mkIdent ``finMapToEquiv) #[P1, P2]
|
||||
return stx
|
||||
|
||||
|
||||
namespace negNode
|
||||
|
||||
/-- The syntax associated with a product of tensors. -/
|
||||
def negSyntax (T1 : Term) : Term :=
|
||||
Syntax.mkApp (mkIdent ``TensorTree.neg) #[T1]
|
||||
|
||||
end negNode
|
||||
|
||||
/-- Returns the full list of indices after contraction. TODO: Include evaluation. -/
|
||||
partial def getIndicesFull (stx : Syntax) : TermElabM (List (TSyntax `indexExpr)) := do
|
||||
match stx with
|
||||
| `(tensorExpr| $_:term | $[$args]*) => do
|
||||
return (← TensorNode.withoutContr stx)
|
||||
| `(tensorExpr| $_:tensorExpr ⊗ $_:tensorExpr) => do
|
||||
return (← ProdNode.withoutContr stx)
|
||||
| `(tensorExpr| ($a:tensorExpr)) => do
|
||||
return (← getIndicesFull a)
|
||||
| `(tensorExpr| -$a:tensorExpr) => do
|
||||
return (← getIndicesFull a)
|
||||
| `(tensorExpr| $a:tensorExpr + $_:tensorExpr) => do
|
||||
return (← getIndicesFull a)
|
||||
| _ =>
|
||||
throwError "Unsupported tensor expression syntax in getIndicesProd: {stx}"
|
||||
|
||||
namespace Add
|
||||
|
||||
/-- Gets the indices associated with the LHS of an addition. -/
|
||||
partial def getIndicesLeft (stx : Syntax) : TermElabM (List (TSyntax `indexExpr)) := do
|
||||
match stx with
|
||||
| `(tensorExpr| $a:tensorExpr + $_:tensorExpr) => do
|
||||
return (← getIndicesFull a)
|
||||
| _ =>
|
||||
throwError "Unsupported tensor expression syntax in Add.getIndicesLeft: {stx}"
|
||||
|
||||
|
||||
/-- Gets the indices associated with the RHS of an addition. -/
|
||||
partial def getIndicesRight (stx : Syntax) : TermElabM (List (TSyntax `indexExpr)) := do
|
||||
match stx with
|
||||
| `(tensorExpr| $_:tensorExpr + $a:tensorExpr) => do
|
||||
return (← getIndicesFull a)
|
||||
| _ =>
|
||||
throwError "Unsupported tensor expression syntax in Add.getIndicesRight: {stx}"
|
||||
|
||||
/-- The syntax for a equality of tensor trees. -/
|
||||
def addSyntax (permSyntax : Term) (T1 T2 : Term) : TermElabM Term := do
|
||||
let P := Syntax.mkApp (mkIdent ``OverColor.equivToHomEq) #[permSyntax]
|
||||
let RHS := Syntax.mkApp (mkIdent ``TensorTree.perm) #[P, T2]
|
||||
return Syntax.mkApp (mkIdent ``add) #[T1, RHS]
|
||||
|
||||
end Add
|
||||
|
||||
namespace Equality
|
||||
|
||||
/-!
|
||||
|
||||
## For equality.
|
||||
|
||||
-/
|
||||
|
||||
/-- Gets the indices associated with the LHS of an equality. -/
|
||||
partial def getIndicesLeft (stx : Syntax) : TermElabM (List (TSyntax `indexExpr)) := do
|
||||
match stx with
|
||||
| `(tensorExpr| $a:tensorExpr = $_:tensorExpr) => do
|
||||
return (← getIndicesFull a)
|
||||
| _ =>
|
||||
throwError "Unsupported tensor expression syntax in getIndicesProd: {stx}"
|
||||
|
||||
/-- Gets the indices associated with the RHS of an equality. -/
|
||||
partial def getIndicesRight (stx : Syntax) : TermElabM (List (TSyntax `indexExpr)) := do
|
||||
match stx with
|
||||
| `(tensorExpr| $_:tensorExpr = $a:tensorExpr) => do
|
||||
return (← getIndicesFull a)
|
||||
| _ =>
|
||||
throwError "Unsupported tensor expression syntax in getIndicesProd: {stx}"
|
||||
|
||||
/-- The syntax for a equality of tensor trees. -/
|
||||
def equalSyntax (permSyntax : Term) (T1 T2 : Term) : TermElabM Term := do
|
||||
let X1 := Syntax.mkApp (mkIdent ``TensorTree.tensor) #[T1]
|
||||
|
@ -453,6 +488,8 @@ def equalSyntax (permSyntax : Term) (T1 T2 : Term) : TermElabM Term := do
|
|||
let X2 := Syntax.mkApp (mkIdent ``TensorTree.tensor) #[X2']
|
||||
return Syntax.mkApp (mkIdent ``Eq) #[X1, X2]
|
||||
|
||||
end Equality
|
||||
|
||||
/-- Creates the syntax associated with a tensor node. -/
|
||||
partial def syntaxFull (stx : Syntax) : TermElabM Term := do
|
||||
match stx with
|
||||
|
@ -464,11 +501,17 @@ partial def syntaxFull (stx : Syntax) : TermElabM Term := do
|
|||
return (← syntaxFull a)
|
||||
| `(tensorExpr| -$a:tensorExpr) => do
|
||||
return negNode.negSyntax (← syntaxFull a)
|
||||
| `(tensorExpr| $a:tensorExpr = $b:tensorExpr) => do
|
||||
let indicesLeft ← getIndicesLeft stx
|
||||
let indicesRight ← getIndicesRight stx
|
||||
| `(tensorExpr| $a + $b) => do
|
||||
let indicesLeft ← Add.getIndicesLeft stx
|
||||
let indicesRight ← Add.getIndicesRight stx
|
||||
let permSyntax ← getPermutationSyntax indicesLeft indicesRight
|
||||
let equalSyntax ← equalSyntax permSyntax (← syntaxFull a) (← syntaxFull b)
|
||||
let addSyntax ← Add.addSyntax permSyntax (← syntaxFull a) (← syntaxFull b)
|
||||
return addSyntax
|
||||
| `(tensorExpr| $a:tensorExpr = $b:tensorExpr) => do
|
||||
let indicesLeft ← Equality.getIndicesLeft stx
|
||||
let indicesRight ← Equality.getIndicesRight stx
|
||||
let permSyntax ← getPermutationSyntax indicesLeft indicesRight
|
||||
let equalSyntax ← Equality.equalSyntax permSyntax (← syntaxFull a) (← syntaxFull b)
|
||||
return equalSyntax
|
||||
| _ =>
|
||||
throwError "Unsupported tensor expression syntax in elaborateTensorNode: {stx}"
|
||||
|
@ -486,24 +529,4 @@ elab_rules (kind:=tensorExprSyntax) : term
|
|||
let tensorTree ← elaborateTensorNode e
|
||||
return tensorTree
|
||||
|
||||
variable {S : TensorSpecies} {c4 : Fin 4 → S.C} (T4 : S.F.obj (OverColor.mk c4))
|
||||
{c5 : Fin 5 → S.C} (T5 : S.F.obj (OverColor.mk c5)) (a : S.k)
|
||||
|
||||
variable (𝓣 : TensorTree S c4)
|
||||
|
||||
/-!
|
||||
|
||||
# Checks
|
||||
|
||||
-/
|
||||
|
||||
/-
|
||||
#tensor_dot {T4 | i j τ(l) d ⊗ T5 | i j k m m}ᵀ.dot
|
||||
|
||||
#check {T4 | i j l d ⊗ T5 | i j k a b}ᵀ
|
||||
|
||||
#check {(T4 | i j l a ⊗ T5 | i j k c d) ⊗ T5 | i1 i2 i3 e f}ᵀ
|
||||
-/
|
||||
end Equality
|
||||
|
||||
end TensorTree
|
||||
|
|
|
@ -75,6 +75,16 @@ lemma neg_perm {n m : ℕ} {c : Fin n → S.C} {c1 : Fin m → S.C}
|
|||
(perm σ (neg t)).tensor = (neg (perm σ t)).tensor := by
|
||||
simp only [perm_tensor, neg_tensor, map_neg]
|
||||
|
||||
@[simp]
|
||||
lemma neg_add (t : TensorTree S c) : (add (neg t) t).tensor = 0 := by
|
||||
rw [add_tensor, neg_tensor]
|
||||
simp only [neg_add_cancel]
|
||||
|
||||
@[simp]
|
||||
lemma add_neg (t : TensorTree S c) : (add t (neg t)).tensor = 0 := by
|
||||
rw [add_tensor, neg_tensor]
|
||||
simp only [add_neg_cancel]
|
||||
|
||||
/-!
|
||||
|
||||
## Basic perm identities
|
||||
|
@ -97,4 +107,36 @@ lemma perm_eq_id {n : ℕ} {c : Fin n → S.C} (σ : (OverColor.mk c) ⟶ (OverC
|
|||
(h : σ = 𝟙 _) (t : TensorTree S c) : (perm σ t).tensor = t.tensor := by
|
||||
simp [perm_tensor, h]
|
||||
|
||||
/-!
|
||||
|
||||
## Additive identities
|
||||
|
||||
These identities are related to the fact that all the maps are linear.
|
||||
|
||||
-/
|
||||
|
||||
/-- The addition node is commutative. -/
|
||||
lemma add_comm (t1 t2 : TensorTree S c) : (add t1 t2).tensor = (add t2 t1).tensor := by
|
||||
simp only [add_tensor]
|
||||
exact AddCommMagma.add_comm t1.tensor t2.tensor
|
||||
|
||||
/-- The addition node is associative. -/
|
||||
lemma add_assoc (t1 t2 t3 : TensorTree S c) :
|
||||
(add (add t1 t2) t3).tensor = (add t1 (add t2 t3)).tensor := by
|
||||
simp only [add_tensor]
|
||||
exact _root_.add_assoc t1.tensor t2.tensor t3.tensor
|
||||
|
||||
/-- When the same permutation acts on both arguments of an addition, the permutation
|
||||
can be moved out of the addition. -/
|
||||
lemma add_perm {n : ℕ} {c : Fin n → S.C} {c1 : Fin n → S.C}
|
||||
(σ : (OverColor.mk c) ⟶ (OverColor.mk c1)) (t t1 : TensorTree S c) :
|
||||
(add (perm σ t) (perm σ t1)).tensor = (perm σ (add t t1)).tensor := by
|
||||
simp only [add_tensor, perm_tensor, map_add]
|
||||
|
||||
/-- When the same evaluation acts on both arguments of an addition, the evaluation
|
||||
can be moved out of the addition. -/
|
||||
lemma add_eval {n : ℕ} {c : Fin n.succ → S.C} (i : Fin n.succ) (e : ℕ) (t t1 : TensorTree S c) :
|
||||
(add (eval i e t) (eval i e t1)).tensor = (eval i e (add t t1)).tensor := by
|
||||
simp only [add_tensor, eval_tensor, Nat.succ_eq_add_one, map_add]
|
||||
|
||||
end TensorTree
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue