refactor: Creation of composite nodes

This commit is contained in:
jstoobysmith 2024-10-22 07:11:44 +00:00
parent 8fa8a51367
commit ccb9623bd6
6 changed files with 43 additions and 70 deletions

View file

@ -32,6 +32,9 @@ noncomputable section
namespace Fermion
example : 0 < complexLorentzTensor.repDim (![Color.down] 0):= by decide
lemma coMetric_expand : {Lorentz.coMetric | μ ν}ᵀ.tensor =
(PiTensorProduct.tprod (fun i => Fin.cases (Lorentz.complexCoBasis (Sum.inl 0))
(fun i => Fin.cases (Lorentz.complexCoBasis (Sum.inl 0)) (fun i => i.elim0) i) i) :
@ -146,9 +149,12 @@ lemma symm_contr_antiSymm {S : (Lorentz.complexCo ⊗ Lorentz.complexCo).V}
rw [antiSymm_contr_symm hA hs]
rfl
variable (p : Lorentz.complexCo) (q : Lorentz.complexContr)
lemma contr_rank_1_expand (p : Lorentz.complexCo) (q : Lorentz.complexContr) :
{(p | μ ⊗ q | μ) = p | 0}ᵀ := by
{p | μ ⊗ q | μ = p | 0}ᵀ := by
sorry
end Fermion
end

View file

@ -765,7 +765,7 @@ between the objects obtained by applying the lift of `F` and that obtained by ap
`F`. -/
def forgetLiftApp (c : C) : (lift.obj F).obj (OverColor.mk (fun (_ : Fin 1) => c))
≅ F.obj (Discrete.mk c) :=
Action.mkIso (forgetLiftAppV F c).toModuleIso
Action.mkIso (forgetLiftAppV F c).toModuleIso
(fun g => by
refine LinearMap.ext (fun x => ?_)
simp only [forgetLiftAppV, Fin.isValue, LinearEquiv.toModuleIso_hom]

View file

@ -494,30 +494,6 @@ end TensorSpecies
inductive TensorTree (S : TensorSpecies) : ∀ {n : }, (Fin n → S.C) → Type where
/-- A general tensor node. -/
| tensorNode {n : } {c : Fin n → S.C} (T : S.F.obj (OverColor.mk c)) : TensorTree S c
/-- 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}
(v : (S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2)).V) :
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. -/
| add {n : } {c : Fin n → S.C} : TensorTree S c → TensorTree S c → TensorTree S c
/-- A node corresponding to the permutation of indices of a tensor. -/
@ -542,37 +518,60 @@ variable {S : TensorSpecies} {n : } {c : Fin n → S.C} (T : TensorTree S c)
open MonoidalCategory
open TensorProduct
/-!
## Composite nodes
-/
/-- A node consisting of a single vector. -/
def vecNode {c : S.C} (v : S.FDiscrete.obj (Discrete.mk c)) : TensorTree S ![c] :=
perm (OverColor.mkIso (by
ext x; fin_cases x; rfl)).hom
(tensorNode ((OverColor.forgetLiftApp S.FDiscrete c).symm.hom.hom v))
/-- The node `vecNode` of a tensor tree, with all arguments explicit. -/
abbrev vecNodeE (S : TensorSpecies) (c1 : S.C)
(v : (S.FDiscrete.obj (Discrete.mk c1)).V) :
TensorTree S ![c1] := vecNode v
/-- A node consisting of a two tensor. -/
def twoNode {c1 c2 : S.C} (t : (S.FDiscrete.obj (Discrete.mk c1) ⊗
S.FDiscrete.obj (Discrete.mk c2)).V) :
TensorTree S ![c1, c2] :=
(tensorNode ((OverColor.Discrete.pairIsoSep S.FDiscrete).hom.hom t))
/-- The node `twoNode` of a tensor tree, with all arguments explicit. -/
abbrev twoNodeE (S : TensorSpecies) (c1 c2 : S.C)
(v : (S.FDiscrete.obj (Discrete.mk c1) ⊗ S.FDiscrete.obj (Discrete.mk c2)).V) :
TensorTree S ![c1, c2] := twoNode v
/-- 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. -/
def constVecNode {c : S.C} (v : 𝟙_ (Rep S.k S.G) ⟶ S.FDiscrete.obj (Discrete.mk c)) :
TensorTree S ![c] := vecNode (v.hom (1 : S.k))
/-- A constant two tensor (e.g. metric and unit). -/
def 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] := twoNode (v.hom (1 : S.k))
/-- The node `constTwoNodeE` of a tensor tree, with all arguments explicit. -/
abbrev constTwoNodeE (S : TensorSpecies) (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] := constTwoNode v
/-- The node `constThreeNodeE` of a tensor tree, with all arguments explicit. -/
abbrev constThreeNodeE (S : TensorSpecies) (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] :=
constThreeNode v
/-!
## Other operations.
-/
/-- The number of nodes in a tensor tree. -/
def size : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → := fun
| tensorNode _ => 1
| vecNode _ => 1
| twoNode _ => 1
| threeNode _ => 1
| constNode _ => 1
| constVecNode _ => 1
| constTwoNode _ => 1
| constThreeNode _ => 1
| add t1 t2 => t1.size + t2.size + 1
| perm _ t => t.size + 1
| neg t => t.size + 1
@ -587,8 +586,6 @@ noncomputable section
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
| tensorNode t => t
| twoNode t => (OverColor.Discrete.pairIsoSep S.FDiscrete).hom.hom t
| constTwoNode t => (OverColor.Discrete.pairIsoSep S.FDiscrete).hom.hom (t.hom (1 : S.k))
| add t1 t2 => t1.tensor + t2.tensor
| perm σ t => (S.F.map σ).hom t.tensor
| neg t => - t.tensor
@ -597,7 +594,6 @@ 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
| _ => 0
/-!

View file

@ -20,20 +20,6 @@ namespace TensorTree
def dotString (m : ) (nt : ) : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → String := fun
| tensorNode _ =>
" node" ++ toString m ++ " [label=\"T" ++ toString nt ++ "\"];\n"
| vecNode T =>
" node" ++ toString m ++ " [label=\"vec " ++ toString nt ++ "\"];\n"
| twoNode T =>
" node" ++ toString m ++ " [label=\"vec2\", shape=box];\n"
| threeNode T =>
" node" ++ toString m ++ " [label=\"vec3\", shape=box];\n"
| constNode T =>
" node" ++ toString m ++ " [label=\"const " ++ toString nt ++ "\"];\n"
| constVecNode T =>
" node" ++ toString m ++ " [label=\"constVec " ++ toString nt ++ "\"];\n"
| constTwoNode T =>
" node" ++ toString m ++ " [label=\"constVec2\", shape=box];\n"
| constThreeNode T =>
" node" ++ toString m ++ " [label=\"constVec3\", shape=box];\n"
| add t1 t2 =>
let addNode := " node" ++ toString m ++ " [label=\"+\", shape=box];\n"
let edge1 := " node" ++ toString m ++ " -> node" ++ toString (m + 1) ++ ";\n"

View file

@ -223,7 +223,6 @@ def termNodeSyntax (T : Term) : TermElabM Term := do
mkIdent ``Fermion.Color.down, mkIdent ``Fermion.Color.down, T]
| _ =>
return Syntax.mkApp (mkIdent ``TensorTree.twoNode) #[T]
| 3, 1 => return Syntax.mkApp (mkIdent ``TensorTree.threeNode) #[T]
| 1, 2 => return Syntax.mkApp (mkIdent ``TensorTree.constVecNode) #[T]
| 2, 2 =>
match ← isDefEq type (← stringToType
@ -233,16 +232,6 @@ def termNodeSyntax (T : Term) : TermElabM Term := do
mkIdent ``Fermion.complexLorentzTensor, mkIdent ``Fermion.Color.down,
mkIdent ``Fermion.Color.down, T]
| _ => return Syntax.mkApp (mkIdent ``TensorTree.constTwoNode) #[T]
| 3, 2 =>
/- Specific types. -/
match ← isDefEq type (← stringToType
"𝟙_ (Rep SL(2, )) ⟶ Lorentz.complexContr ⊗ Fermion.leftHanded ⊗ Fermion.rightHanded") with
| true =>
return Syntax.mkApp (mkIdent ``TensorTree.constThreeNodeE) #[
mkIdent ``Fermion.complexLorentzTensor, mkIdent ``Fermion.Color.up,
mkIdent ``Fermion.Color.upL, mkIdent ``Fermion.Color.upR, T]
| _ =>
return Syntax.mkApp (mkIdent ``TensorTree.constThreeNode) #[T]
| _, _ => throwError "Could not create terminal node syntax (termNodeSyntax). "
/-- Adjusts a list `List ` by subtracting from each natrual number the number

View file

@ -38,10 +38,6 @@ informal_lemma constTwoNode_eq_twoNode where
math :≈ "A constTwoNode has equal tensor to the twoNode with the map evaluated at 1."
deps :≈ [``constTwoNode, ``twoNode]
informal_lemma constThreeNode_eq_threeNode where
math :≈ "A constThreeNode has equal tensor to the threeNode with the map evaluated at 1."
deps :≈ [``constThreeNode, ``threeNode]
/-!
## Negation