refactor: Rename TensorSpeciesStruct to TensorSpecies
This commit is contained in:
parent
ef0d857cb7
commit
271745c11a
9 changed files with 25 additions and 26 deletions
|
@ -18,9 +18,8 @@ open CategoryTheory
|
|||
open MonoidalCategory
|
||||
|
||||
/-- The sturcture of a type of tensors e.g. Lorentz tensors, Einstien tensors,
|
||||
complex Lorentz tensors.
|
||||
Note: This structure is not fully defined yet. -/
|
||||
structure TensorSpeciesStruct where
|
||||
complex Lorentz tensors. -/
|
||||
structure TensorSpecies where
|
||||
/-- The colors of indices e.g. up or down. -/
|
||||
C : Type
|
||||
/-- The symmetry group acting on these tensor e.g. the Lorentz group or SL(2,ℂ). -/
|
||||
|
@ -55,10 +54,10 @@ structure TensorSpeciesStruct where
|
|||
|
||||
noncomputable section
|
||||
|
||||
namespace TensorSpeciesStruct
|
||||
namespace TensorSpecies
|
||||
open OverColor
|
||||
|
||||
variable (S : TensorSpeciesStruct)
|
||||
variable (S : TensorSpecies)
|
||||
|
||||
instance : CommRing S.k := S.k_commRing
|
||||
|
||||
|
@ -355,10 +354,10 @@ lemma contrMap_tprod {n : ℕ} (c : Fin n.succ.succ → S.C)
|
|||
simp
|
||||
exact h1' h1
|
||||
|
||||
end TensorSpeciesStruct
|
||||
end TensorSpecies
|
||||
|
||||
/-- A syntax tree for tensor expressions. -/
|
||||
inductive TensorTree (S : TensorSpeciesStruct) : ∀ {n : ℕ}, (Fin n → S.C) → Type where
|
||||
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. -/
|
||||
|
@ -404,23 +403,23 @@ inductive TensorTree (S : TensorSpeciesStruct) : ∀ {n : ℕ}, (Fin n → S.C)
|
|||
|
||||
namespace TensorTree
|
||||
|
||||
variable {S : TensorSpeciesStruct} {n : ℕ} {c : Fin n → S.C} (T : TensorTree S c)
|
||||
variable {S : TensorSpecies} {n : ℕ} {c : Fin n → S.C} (T : TensorTree S c)
|
||||
|
||||
open MonoidalCategory
|
||||
open TensorProduct
|
||||
|
||||
/-- The node `twoNode` of a tensor tree, with all arguments explicit. -/
|
||||
abbrev twoNodeE (S : TensorSpeciesStruct) (c1 c2 : S.C)
|
||||
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
|
||||
|
||||
/-- The node `constTwoNodeE` of a tensor tree, with all arguments explicit. -/
|
||||
abbrev constTwoNodeE (S : TensorSpeciesStruct) (c1 c2 : S.C)
|
||||
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 : TensorSpeciesStruct) (c1 c2 c3 : S.C)
|
||||
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
|
||||
|
|
|
@ -487,7 +487,7 @@ elab_rules (kind:=tensorExprSyntax) : term
|
|||
let tensorTree ← elaborateTensorNode e
|
||||
return tensorTree
|
||||
|
||||
variable {S : TensorSpeciesStruct} {c4 : Fin 4 → S.C} (T4 : S.F.obj (OverColor.mk c4))
|
||||
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)
|
||||
|
|
|
@ -22,7 +22,7 @@ open TensorProduct
|
|||
|
||||
namespace TensorTree
|
||||
|
||||
variable {S : TensorSpeciesStruct}
|
||||
variable {S : TensorSpecies}
|
||||
|
||||
/-!
|
||||
|
||||
|
|
|
@ -21,7 +21,7 @@ open HepLean.Fin
|
|||
|
||||
namespace TensorTree
|
||||
|
||||
variable {S : TensorSpeciesStruct}
|
||||
variable {S : TensorSpecies}
|
||||
|
||||
/-- A structure containing two pairs of indices (i, j) and (k, l) to be sequentially
|
||||
contracted in a tensor. -/
|
||||
|
@ -168,7 +168,7 @@ lemma contrSwapHom_contrMapSnd_tprod (x : (i : (𝟭 Type).obj (OverColor.mk c).
|
|||
(x (q.swap.i.succAbove (q.swap.j.succAbove (q.swap.k.succAbove q.swap.l))))))) •
|
||||
((lift.obj S.FDiscrete).map q.contrSwapHom).hom ((PiTensorProduct.tprod S.k) fun k =>
|
||||
x (q.swap.i.succAbove (q.swap.j.succAbove (q.swap.k.succAbove (q.swap.l.succAbove k)))))) := by
|
||||
rw [contrMapSnd,TensorSpeciesStruct.contrMap_tprod]
|
||||
rw [contrMapSnd,TensorSpecies.contrMap_tprod]
|
||||
change ((lift.obj S.FDiscrete).map q.contrSwapHom).hom
|
||||
(_ • ((PiTensorProduct.tprod S.k) fun k =>
|
||||
x (q.swap.i.succAbove (q.swap.j.succAbove
|
||||
|
@ -214,7 +214,7 @@ lemma contrMapFst_contrMapSnd_swap :
|
|||
(S.F.map q.contrSwapHom).hom
|
||||
(q.swap.contrMapSnd.hom ((S.contrMap c q.swap.i q.swap.j _).hom
|
||||
((PiTensorProduct.tprod S.k) x)))
|
||||
rw [TensorSpeciesStruct.contrMap_tprod, TensorSpeciesStruct.contrMap_tprod]
|
||||
rw [TensorSpecies.contrMap_tprod, TensorSpecies.contrMap_tprod]
|
||||
simp only [Nat.succ_eq_add_one, Monoidal.tensorUnit_obj, Action.instMonoidalCategory_tensorUnit_V,
|
||||
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
||||
Action.FunctorCategoryEquivalence.functor_obj_obj, Functor.comp_obj, Discrete.functor_obj_eq_as,
|
||||
|
@ -226,7 +226,7 @@ lemma contrMapFst_contrMapSnd_swap :
|
|||
((lift.obj S.FDiscrete).map q.contrSwapHom).hom
|
||||
(q.swap.contrMapSnd.hom ((PiTensorProduct.tprod S.k)
|
||||
fun k => x (q.swap.i.succAbove (q.swap.j.succAbove k))))
|
||||
rw [contrMapSnd, TensorSpeciesStruct.contrMap_tprod, q.contrSwapHom_contrMapSnd_tprod]
|
||||
rw [contrMapSnd, TensorSpecies.contrMap_tprod, q.contrSwapHom_contrMapSnd_tprod]
|
||||
rw [smul_smul, smul_smul]
|
||||
congr 1
|
||||
/- The contractions. -/
|
||||
|
|
|
@ -8,7 +8,7 @@ import HepLean.Tensors.Tree.Basic
|
|||
|
||||
# The commutativity of Permutations and contractions.
|
||||
|
||||
There is very likely a better way to do this using `TensorSpeciesStruct.contrMap_tprod`.
|
||||
There is very likely a better way to do this using `TensorSpecies.contrMap_tprod`.
|
||||
|
||||
-/
|
||||
|
||||
|
@ -18,10 +18,10 @@ open MonoidalCategory
|
|||
open OverColor
|
||||
open HepLean.Fin
|
||||
|
||||
namespace TensorSpeciesStruct
|
||||
namespace TensorSpecies
|
||||
noncomputable section
|
||||
|
||||
variable (S : TensorSpeciesStruct)
|
||||
variable (S : TensorSpecies)
|
||||
|
||||
lemma contrFin1Fin1_naturality {n : ℕ} {c c1 : Fin n.succ.succ → S.C}
|
||||
{i : Fin n.succ.succ} {j : Fin n.succ} (h : c1 (i.succAbove j) = S.τ (c1 i))
|
||||
|
@ -236,11 +236,11 @@ lemma contrMap_naturality {n : ℕ} {c c1 : Fin n.succ.succ → S.C}
|
|||
rfl
|
||||
|
||||
end
|
||||
end TensorSpeciesStruct
|
||||
end TensorSpecies
|
||||
|
||||
namespace TensorTree
|
||||
|
||||
variable {S : TensorSpeciesStruct}
|
||||
variable {S : TensorSpecies}
|
||||
|
||||
/-- Permuting indices, and then contracting is equivalent to contracting and then permuting,
|
||||
once care is taking about ensuring one is contracting the same idices. -/
|
||||
|
|
|
@ -19,7 +19,7 @@ open HepLean.Fin
|
|||
|
||||
namespace TensorTree
|
||||
|
||||
variable {S : TensorSpeciesStruct} {n n' n2 : ℕ}
|
||||
variable {S : TensorSpecies} {n n' n2 : ℕ}
|
||||
{c : Fin n → S.C} {c' : Fin n' → S.C} (c2 : Fin n2 → S.C)
|
||||
(σ : OverColor.mk c ⟶ OverColor.mk c')
|
||||
|
||||
|
|
|
@ -19,7 +19,7 @@ open HepLean.Fin
|
|||
|
||||
namespace TensorTree
|
||||
|
||||
variable {S : TensorSpeciesStruct} {n n2 : ℕ}
|
||||
variable {S : TensorSpecies} {n n2 : ℕ}
|
||||
(c : Fin n → S.C) (c2 : Fin n2 → S.C)
|
||||
|
||||
/-- The permutation that arises when moving a commuting terms in a `prod` node.
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue