chore: Add doc strings

This commit is contained in:
jstoobysmith 2024-10-19 10:07:03 +00:00
parent 1f3ba14462
commit 48bec8c891
6 changed files with 27 additions and 0 deletions

View file

@ -263,12 +263,17 @@ 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.
e.g. [0, 1, 2, 3] becomes [(0, 1), (2, 3)]. -/
def toPairs (l : List ) : List ( × ) :=
match l with
| x1 :: x2 :: xs => (x1, x2) :: toPairs xs
| [] => []
| [x] => [(x, 0)]
/-- Adjusts a list `List ( × )` by subtracting from each natrual number the number
of elements before it in the list which are less then itself. This is used
to form a list of pairs which can be used for contracting indices. -/
def contrListAdjust (l : List ( × )) : List ( × ) :=
let l' := l.bind (fun p => [p.1, p.2])
let l'' := List.mapAccumr
@ -367,6 +372,7 @@ partial def syntaxFull (stx : Syntax) : TermElabM Term := do
end ProdNode
/-- 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
@ -402,6 +408,8 @@ partial def getIndicesRight (stx : Syntax) : TermElabM (List (TSyntax `indexExpr
| _ =>
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
let l1' ← l1.mapM (fun x => indexToIdent x)
let l2' ← l2.mapM (fun x => indexToIdent x)