Tail recursive implementations for List
definitions. #
Many of the proofs require theorems about Array
,
so these are in a separate file to minimize imports.
Tail recursive version of erase
.
Equations
- l.setTR n a = List.setTR.go l a l n #[]
Instances For
Auxiliary for setTR
: setTR.go l a xs n acc = acc.toList ++ set xs a
,
unless n ≥ l.length
in which case it returns l
Equations
- List.setTR.go l a [] x✝ x = l
- List.setTR.go l a (head :: xs) 0 x = x.toListAppend (a :: xs)
- List.setTR.go l a (x_3 :: xs) n.succ x = List.setTR.go l a xs n (x.push x_3)
Instances For
Tail recursive version of erase
.
Equations
- l.eraseTR a = List.eraseTR.go l a l #[]
Instances For
Auxiliary for eraseTR
: eraseTR.go l a xs acc = acc.toList ++ erase xs a
,
unless a
is not present in which case it returns l
Equations
- List.eraseTR.go l a [] x = l
- List.eraseTR.go l a (x_2 :: xs) x = bif x_2 == a then x.toListAppend xs else List.eraseTR.go l a xs (x.push x_2)
Instances For
Tail recursive version of eraseIdx
.
Equations
- l.eraseIdxTR n = List.eraseIdxTR.go l l n #[]
Instances For
Auxiliary for eraseIdxTR
: eraseIdxTR.go l n xs acc = acc.toList ++ eraseIdx xs a
,
unless a
is not present in which case it returns l
Equations
- List.eraseIdxTR.go l [] x✝ x = l
- List.eraseIdxTR.go l (head :: xs) 0 x = x.toListAppend xs
- List.eraseIdxTR.go l (x_3 :: xs) n.succ x = List.eraseIdxTR.go l xs n (x.push x_3)
Instances For
Tail recursive version of bind
.
Equations
- as.bindTR f = List.bindTR.go f as #[]
Instances For
Auxiliary for bind
: bind.go f as = acc.toList ++ bind f as
Equations
- List.bindTR.go f [] x = x.toList
- List.bindTR.go f (x_2 :: xs) x = List.bindTR.go f xs (x ++ f x_2)
Instances For
Tail recursive version of filterMap
.
Equations
- List.filterMapTR f l = List.filterMapTR.go f l #[]
Instances For
Auxiliary for filterMap
: filterMap.go f l = acc.toList ++ filterMap f l
Equations
- List.filterMapTR.go f [] x = x.toList
- List.filterMapTR.go f (x_2 :: xs) x = match f x_2 with | none => List.filterMapTR.go f xs x | some b => List.filterMapTR.go f xs (x.push b)
Instances For
Tail recursive version of replace
.
Equations
- l.replaceTR b c = List.replaceTR.go l b c l #[]
Instances For
Auxiliary for replace
: replace.go l b c xs acc = acc.toList ++ replace xs b c
,
unless b
is not found in xs
in which case it returns l
.
Equations
- List.replaceTR.go l b c [] x = l
- List.replaceTR.go l b c (x_2 :: xs) x = bif x_2 == b then x.toListAppend (c :: xs) else List.replaceTR.go l b c xs (x.push x_2)
Instances For
Tail recursive version of take
.
Equations
- List.takeTR n l = List.takeTR.go l l n #[]
Instances For
Auxiliary for take
: take.go l xs n acc = acc.toList ++ take n xs
,
unless n ≥ xs.length
in which case it returns l
.
Equations
- List.takeTR.go l [] x✝ x = l
- List.takeTR.go l (head :: xs) 0 x = x.toList
- List.takeTR.go l (x_3 :: xs) n.succ x = List.takeTR.go l xs n (x.push x_3)
Instances For
Tail recursive version of takeWhile
.
Equations
- List.takeWhileTR p l = List.takeWhileTR.go p l l #[]
Instances For
Auxiliary for takeWhile
: takeWhile.go p l xs acc = acc.toList ++ takeWhile p xs
,
unless no element satisfying p
is found in xs
in which case it returns l
.
Equations
- List.takeWhileTR.go p l [] x = l
- List.takeWhileTR.go p l (x_2 :: xs) x = bif p x_2 then List.takeWhileTR.go p l xs (x.push x_2) else x.toList
Instances For
Tail recursive version of foldr
.
Equations
- List.foldrTR f init l = Array.foldr f init (List.toArray l) (List.toArray l).size
Instances For
Tail recursive version of zipWith
.
Equations
- List.zipWithTR f as bs = List.zipWithTR.go f as bs #[]
Instances For
Auxiliary for zipWith
: zipWith.go f as bs acc = acc.toList ++ zipWith f as bs
Equations
- List.zipWithTR.go f (a :: as) (b :: bs) x = List.zipWithTR.go f as bs (x.push (f a b))
- List.zipWithTR.go f x✝¹ x✝ x = x.toList
Instances For
Tail recursive version of dropLast
.
Equations
- l.dropLastTR = (List.toArray l).pop.toList
Instances For
Tail recursive version of intercalate
.
Equations
- sep.intercalateTR x = match x with | [] => [] | [x] => x | x :: xs => List.intercalateTR.go (List.toArray sep) x xs #[]
Instances For
Auxiliary for intercalateTR
:
intercalateTR.go sep x xs acc = acc.toList ++ intercalate sep.toList (x::xs)
Equations
- List.intercalateTR.go sep x✝ [] x = x.toListAppend x✝
- List.intercalateTR.go sep x✝ (y :: xs) x = List.intercalateTR.go sep y xs (x ++ x✝ ++ sep)