refactor: Lint

This commit is contained in:
jstoobysmith 2024-10-08 07:52:55 +00:00
parent 93431bda47
commit 48a69b56a8
4 changed files with 58 additions and 18 deletions

View file

@ -126,7 +126,7 @@ end colorFun
/-- The functor between `OverColor Color` and `Rep SL(2, )` taking a map of colors
to the corresponding tensor product representation. -/
@[simps!]
@[simps! obj_V_carrier]
def colorFun : OverColor Color ⥤ Rep SL(2, ) where
obj := colorFun.obj'
map := colorFun.map'
@ -165,7 +165,7 @@ open CategoryTheory
open MonoidalCategory
@[simp]
lemma obj_ρ_empty (g : SL(2, )) : (colorFun.obj (𝟙_ (OverColor Color))).ρ g = LinearMap.id := by
lemma obj_ρ_empty (g : SL(2, )) : (colorFun.obj (𝟙_ (OverColor Color))).ρ g = LinearMap.id := by
erw [colorFun.obj'_ρ]
ext x
refine PiTensorProduct.induction_on' x (fun r x => ?_) <| fun x y hx hy => by

View file

@ -15,7 +15,6 @@ import HepLean.SpaceTime.LorentzVector.Complex
## Over category.
-/
namespace IndexNotation
@ -106,7 +105,7 @@ instance (C : Type) : MonoidalCategoryStruct (OverColor C) where
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr x => rfl)).symm,
hom_inv_id := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
hom_inv_id := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
@ -226,10 +225,13 @@ instance (C : Type) : SymmetricCategory (OverColor C) where
end monoidal
/-- Make an object of `OverColor C` from a map `X → C`. -/
def mk (f : X → C) : OverColor C := Over.mk f
open MonoidalCategory
/-- The isomorphism between `c : X → C` and `c ∘ e.symm` as objects in `OverColor C` for an
equivalence `e`. -/
def equivToIso {c : X → C} (e : X ≃ Y) : mk c ≅ mk (c ∘ e.symm) := {
hom := Over.isoMk e.toIso ((Iso.eq_inv_comp e.toIso).mp rfl),
inv := (Over.isoMk e.toIso ((Iso.eq_inv_comp e.toIso).mp rfl)).symm,

View file

@ -13,15 +13,28 @@ import HepLean.Tensors.ColorCat.Basic
open IndexNotation
open CategoryTheory
/-- 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 TensorStruct 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,). -/
G : Type
/-- An instance of `G` as a group. -/
G_group : Group G
/-- The field over which we want to consider the tensors to live in, usually `` or ``. -/
k : Type
/-- An instance of `k` as a commutative ring. -/
k_commRing : CommRing k
/-- A `MonoidalFunctor` from `OverColor C` giving the rep corresponding to a map of colors
`X → C`. -/
F : MonoidalFunctor (OverColor C) (Rep k G)
/-- A map from `C` to `C`. An involution. -/
τ : C → C
evalNo : C → N
/-- A specification of the dimension of each color in C. This will be used for explicit
evaluation of tensors. -/
evalNo : C →
namespace TensorStruct
@ -33,8 +46,9 @@ instance : Group S.G := S.G_group
end TensorStruct
/-- A syntax tree for tensor expressions. -/
inductive TensorTree (S : TensorStruct) : ∀ {n : }, (Fin n → S.C) → Type where
| tensorNode {n : } {c : Fin n → S.C} (T: S.F.obj (OverColor.mk c)): TensorTree S c
| tensorNode {n : } {c : Fin n → S.C} (T : S.F.obj (OverColor.mk c)) : TensorTree S c
| add {n : } {c : Fin n → S.C} : TensorTree S c → TensorTree S c → TensorTree S c
| 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
@ -46,11 +60,11 @@ inductive TensorTree (S : TensorStruct) : ∀ {n : }, (Fin n → S.C) → Typ
TensorTree S (Sum.elim (c ∘ Fin.succAbove i) (c1 ∘ Fin.succAbove j) ∘ finSumFinEquiv.symm)
| contr {n : } {c : Fin n.succ.succ → S.C} : (i : Fin n.succ.succ) →
(j : Fin n.succ) → TensorTree S c → TensorTree S (c ∘ Fin.succAbove i ∘ Fin.succAbove j)
| jiggle {n : } {c : Fin n → S.C} : (i : Fin n) → TensorTree S c
TensorTree S (Function.update c i (S.τ (c i)))
| jiggle {n : } {c : Fin n → S.C} : (i : Fin n) → TensorTree S c →
TensorTree S (Function.update c i (S.τ (c i)))
| eval {n : } {c : Fin n.succ → S.C} :
(i : Fin n.succ) → (x : Fin (S.evalNo (c i))) → TensorTree S c →
TensorTree S (c ∘ Fin.succAbove i)
TensorTree S (c ∘ Fin.succAbove i)
namespace TensorTree
@ -59,7 +73,8 @@ variable {S : TensorStruct} {n : } {c : Fin n → S.C} (T : TensorTree S c)
open MonoidalCategory
open TensorProduct
def size : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → := fun
/-- The number of nodes in a tensor tree. -/
def size : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → := fun
| tensorNode _ => 1
| add t1 t2 => t1.size + t2.size + 1
| perm _ t => t.size + 1
@ -70,10 +85,11 @@ def size : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → := fun
| jiggle _ t => t.size + 1
| eval _ _ t => t.size + 1
noncomputable section
def tensor : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → S.F.obj (OverColor.mk c) := fun
/-- The underlying tensor a tensor tree corresponds to.
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
| add t1 t2 => t1.tensor + t2.tensor
| perm σ t => (S.F.map σ).hom t.tensor

View file

@ -27,35 +27,44 @@ open Lean Meta Elab Tactic
-/
/-- A syntax category for indices of tensor expressions. -/
declare_syntax_cat indexExpr
/-- A basic index is a ident. -/
syntax ident : indexExpr
/-- An index can be a num, which will be used to evaluate the tensor. -/
syntax num : indexExpr
/-- Notation to discribe the jiggle of a tensor index. -/
syntax "τ(" ident ")" : indexExpr
/-- Bool which is ture if an index is a num. -/
def indexExprIsNum (stx : Syntax) : Bool :=
match stx with
| `(indexExpr|$_:num) => true
| _ => false
/-- If an index is a num - the undelrying natural number. -/
def indexToNum (stx : Syntax) : TermElabM Nat :=
match stx with
| `(indexExpr|$a:num) =>
match a.raw.isNatLit? with
| some n => return n
| none => throwError "Expected a natural number literal."
| none => throwError "Expected a natural number literal."
| _ =>
throwError "Unsupported tensor expression syntax (indexToNum): {stx}"
throwError "Unsupported tensor expression syntax in indexToNum: {stx}"
/-- When an index is not a num, the corresponding ident. -/
def indexToIdent (stx : Syntax) : TermElabM Ident :=
match stx with
| `(indexExpr|$a:ident) => return a
| `(indexExpr| τ($a:ident)) => return a
| _ =>
throwError "Unsupported tensor expression syntax (indexToIdent): {stx}"
throwError "Unsupported tensor expression syntax in indexToIdent: {stx}"
/-- Takes a pair ``a b : × TSyntax `indexExpr``. If `a.1 < b.1` and `a.2 = b.2` then
outputs `some (a.1, b.1)`, otherwise `none`. -/
def indexPosEq (a b : × TSyntax `indexExpr) : TermElabM (Option ( × )) := do
let a' ← indexToIdent a.2
let b' ← indexToIdent b.2
@ -64,6 +73,7 @@ def indexPosEq (a b : × TSyntax `indexExpr) : TermElabM (Option ( ×
else
return none
/-- Bool which is true if an index is of the form τ(i) that is, to be dualed. -/
def indexToDual (stx : Syntax) : Bool :=
match stx with
| `(indexExpr| τ($_)) => true
@ -73,12 +83,17 @@ def indexToDual (stx : Syntax) : Bool :=
## Tensor expressions
-/
/-- A syntax category for tensor expressions. -/
declare_syntax_cat tensorExpr
/-- The syntax for a tensor node. -/
syntax term "|" (ppSpace indexExpr)* : tensorExpr
/-- The syntax for tensor prod two tensor nodes. -/
syntax tensorExpr "⊗" tensorExpr : tensorExpr
/-- Allowing brackets to be used in a tensor expression. -/
syntax "(" tensorExpr ")" : tensorExpr
/-!
@ -102,7 +117,7 @@ partial def getIndicesNode (stx : Syntax) : TermElabM (List (TSyntax `indexExpr)
| `(indexExpr|$t:indexExpr) => pure t
return indices
| _ =>
throwError "Unsupported tensor expression syntax (getIndicesNode): {stx}"
throwError "Unsupported tensor expression syntax in getIndicesNode: {stx}"
/-- The positions in getIndicesNode which get evaluated, and the value they take. -/
partial def getEvalPos (stx : Syntax) : TermElabM (List ( × )) := do
@ -112,6 +127,7 @@ partial def getEvalPos (stx : Syntax) : TermElabM (List ( × )) := do
let evals2 ← (evals.mapM (fun x => indexToNum x.2))
return List.zip (evals.map (fun x => x.1)) evals2
/-- For each element of `l : List ( × )` applies `TensorTree.eval` to the given term. -/
def evalSyntax (l : List ( × )) (T : Term) : Term :=
l.foldl (fun T' (x1, x2) => Syntax.mkApp (mkIdent ``TensorTree.eval)
#[Syntax.mkNumLit (toString x1), Syntax.mkNumLit (toString x2), T']) T
@ -124,29 +140,34 @@ partial def getDualPos (stx : Syntax) : TermElabM (List ) := do
let duals := indEnum.filter (fun x => indexToDual x.2)
return duals.map (fun x => x.1)
/-- For each element of `l : List ` applies `TensorTree.jiggle` to the given term. -/
def dualSyntax (l : List ) (T : Term) : Term :=
l.foldl (fun T' x => Syntax.mkApp (mkIdent ``TensorTree.jiggle)
#[Syntax.mkNumLit (toString x), T']) T
/-- The pairs of positions in getIndicesNode which get contracted. -/
partial def getContrPos (stx : Syntax) : TermElabM (List ( × )) := do
let ind ← getIndicesNode stx
let indFilt : List (TSyntax `indexExpr) := ind.filter (fun x => ¬ indexExprIsNum x)
let indEnum := indFilt.enum
let bind := List.bind indEnum (fun a => indEnum.map (fun b => (a, b)))
let filt ← bind.filterMapM (fun x => indexPosEq x.1 x.2)
if ¬ ((filt.map Prod.fst).Nodup ∧ (filt.map Prod.snd).Nodup) then
if ¬ ((filt.map Prod.fst).Nodup ∧ (filt.map Prod.snd).Nodup) then
throwError "To many contractions"
return filt
/-- The list of indices after contraction. -/
def withoutContr (stx : Syntax) : TermElabM (List (TSyntax `indexExpr)) := do
let ind ← getIndicesNode stx
let indFilt : List (TSyntax `indexExpr) := ind.filter (fun x => ¬ indexExprIsNum x)
return ind.filter (fun x => indFilt.count x ≤ 1)
/-- For each element of `l : List ( × )` applies `TensorTree.contr` to the given term. -/
def contrSyntax (l : List ( × )) (T : Term) : Term :=
l.foldl (fun T' (x0, x1) => Syntax.mkApp (mkIdent ``TensorTree.contr)
#[Syntax.mkNumLit (toString x1), Syntax.mkNumLit (toString x0), T']) T
/-- An elaborator for tensor nodes. This is to be generalized. -/
def elaborateTensorNode (stx : Syntax) : TermElabM Expr := do
match stx with
| `(tensorExpr| $T:term | $[$args]*) => do
@ -157,8 +178,9 @@ def elaborateTensorNode (stx : Syntax) : TermElabM Expr := do
let tensorExpr ← elabTerm contrSyntax none
return tensorExpr
| _ =>
throwError "Unsupported tensor expression syntax (elaborateTensorNode): {stx}"
throwError "Unsupported tensor expression syntax in elaborateTensorNode: {stx}"
/-- Syntax turning a tensor expression into a term. -/
syntax (name := tensorExprSyntax) "{" tensorExpr "}ᵀ" : term
elab_rules (kind:=tensorExprSyntax) : term