refactor: Lint

This commit is contained in:
jstoobysmith 2024-10-21 06:53:58 +00:00
parent c820a90939
commit 2975e08f85
6 changed files with 10 additions and 15 deletions

View file

@ -183,7 +183,6 @@ def termNodeSyntax (T : Term) : TermElabM Term := do
let strType := toString type
let n := (String.splitOn strType "CategoryTheory.MonoidalCategoryStruct.tensorObj").length
let const := (String.splitOn strType "Quiver.Hom").length
println! "n: {n}, const: {const}"
match n, const with
| 1, 1 =>
match type with
@ -266,7 +265,7 @@ def withoutContr (stx : Syntax) : TermElabM (List (TSyntax `indexExpr)) := do
let indFilt : List (TSyntax `indexExpr) := ind.filter (fun x => ¬ indexExprIsNum x)
return ind.filter (fun x => indFilt.count x ≤ 1)
/-- Takes a list and puts conseutive elements into pairs.
/-- Takes a list and puts conseutive elements into pairs.
e.g. [0, 1, 2, 3] becomes [(0, 1), (2, 3)]. -/
def toPairs (l : List ) : List ( × ) :=
match l with
@ -364,8 +363,6 @@ partial def syntaxFull (stx : Syntax) : TermElabM Term := do
| `(tensorExpr| $_:term | $[$args]*) => TensorNode.syntaxFull stx
| `(tensorExpr| $a:tensorExpr ⊗ $b:tensorExpr) => do
let prodSyntax := prodSyntax (← syntaxFull a) (← syntaxFull b)
println! (← getContrPos stx)
println! TensorNode.contrListAdjust (← getContrPos stx)
let contrSyntax := contrSyntax (← getContrPos stx) prodSyntax
return contrSyntax
| `(tensorExpr| ($a:tensorExpr)) => do
@ -383,7 +380,6 @@ def negSyntax (T1 : Term) : Term :=
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