IdrisDoc: Data.Vect

Data.Vect

zipWith3 : (a -> b -> c -> d) -> (xs : Vect n a) -> (ys : Vect n b) -> (zs : Vect n c) -> Vect n d

Combine three equal-length vectors into a vector with some function

zipWith : (f : a -> b -> c) -> (xs : Vect n a) -> (ys : Vect n b) -> Vect n c

Combine two equal-length vectors pairwise with some function.

f

the function to combine elements with

xs

the first vector of elements

ys

the second vector of elements

zip3 : (xs : Vect n a) -> (ys : Vect n b) -> (zs : Vect n c) -> Vect n (a, b, c)

Combine three equal-length vectors elementwise into a vector of tuples

zip : (xs : Vect n a) -> (ys : Vect n b) -> Vect n (a, b)

Combine two equal-length vectors pairwise

vectToMaybe : Vect len elem -> Maybe elem
vectNilRightNeutral : (xs : Vect n a) -> xs ++ [] = xs
vectInjective2 : (x :: xs = y :: ys) -> xs = ys
vectInjective1 : (x :: xs = y :: ys) -> x = y
vectConsCong : (x : elem) -> (xs : Vect len elem) -> (ys : Vect m elem) -> (xs = ys) -> x :: xs = x :: ys
vectAppendAssociative : (xs : Vect xLen elem) -> (ys : Vect yLen elem) -> (zs : Vect zLen elem) -> xs ++ ys ++ zs = (xs ++ ys) ++ zs
updateAt : (i : Fin len) -> (f : elem -> elem) -> (xs : Vect len elem) -> Vect len elem

Replace the element at a particular index with the result of applying a function to it

i

the index to replace at

f

the update function

xs

the vector to replace in

unzip3 : (xs : Vect n (a, b, c)) -> (Vect n a, Vect n b, Vect n c)

Convert a vector of three-tuples to a triplet of vectors

unzip : (xs : Vect n (a, b)) -> (Vect n a, Vect n b)

Convert a vector of pairs to a pair of vectors

transpose : Vect m (Vect n elem) -> Vect n (Vect m elem)

Transpose a Vect of Vects, turning rows into columns and vice versa.

As the types ensure rectangularity, this is an involution, unlike Prelude.List.transpose.

takeWhile : (p : elem -> Bool) -> Vect len elem -> (q : Nat ** Vect q elem)

Take the longest prefix of a Vect such that all elements satisfy some
Boolean predicate.

p

the predicate

take : (n : Nat) -> Vect (n + m) elem -> Vect n elem

Get the first n elements of a Vect

n

the number of elements to take

tail : Vect (S len) elem -> Vect len elem

All but the first element of the vector

splitAt : (n : Nat) -> (xs : Vect (n + m) elem) -> (Vect n elem, Vect m elem)

A tuple where the first element is a Vect of the n first elements and
the second element is a Vect of the remaining elements of the original Vect
It is equivalent to (take n xs, drop n xs)

n

the index to split at

xs

the Vect to split in two

scanl1 : (elem -> elem -> elem) -> Vect len elem -> Vect len elem

The scanl1 function is a variant of scanl that doesn't require an explicit
starting value.
It assumes the first element of the vector to be the starting value and then
starts the fold with the element following it.

scanl : (res -> elem -> res) -> res -> Vect len elem -> Vect (S len) res

The scanl function is similar to foldl, but returns all the intermediate
accumulator states in the form of a vector.

reverse : Vect len elem -> Vect len elem

Reverse the order of the elements of a vector

replicate : (len : Nat) -> (x : elem) -> Vect len elem

Repeate some value some number of times.

len

the number of times to repeat it

x

the value to repeat

replaceElem : (xs : Vect k t) -> Elem x xs -> (y : t) -> (ys : Vect k t ** Elem y ys)
replaceByElem : (xs : Vect k t) -> Elem x xs -> t -> Vect k t
replaceAt : Fin len -> elem -> Vect len elem -> Vect len elem

Replace an element at a particlar index with another

range : Vect len (Fin len)
partition : (elem -> Bool) -> Vect len elem -> ((p : Nat ** Vect p elem), (q : Nat ** Vect q elem))
overLength : (len : Nat) -> (xs : Vect m a) -> Maybe (p : Nat ** Vect (plus p len) a)

If the given Vect is at least the required length, return a Vect with
at least that length in its type, otherwise return Nothing

len

the required length

xs

the vector with the desired length

nubBy : (elem -> elem -> Bool) -> Vect len elem -> (p : Nat ** Vect p elem)

Make the elements of some vector unique by some test

nub : Eq elem => Vect len elem -> (p : Nat ** Vect p elem)

Make the elements of some vector unique by the default Boolean equality

noEmptyElem : Elem x [] -> Void

Nothing can be in an empty Vect

neitherHereNorThere : Not (x = y) -> Not (Elem x xs) -> Not (Elem x (y :: xs))

An item not in the head and not in the tail is not in the Vect at all

mergeBy : (elem -> elem -> Ordering) -> (xs : Vect n elem) -> (ys : Vect m elem) -> Vect (n + m) elem

Merge two ordered vectors

merge : Ord elem => Vect n elem -> Vect m elem -> Vect (n + m) elem
maybeToVect : Maybe elem -> (p : Nat ** Vect p elem)
mapMaybe : (f : a -> Maybe b) -> (xs : Vect len a) -> (m : Nat ** Vect m b)

Map a partial function across a vector, returning those elements for which
the function had a value.

The first projection of the resulting pair (ie the length) will always be
at most the length of the input vector. This is not, however, guaranteed
by the type.

f

the partial function (expressed by returning Maybe)

xs

the vector to check for results

mapElem : Elem x xs -> Elem (f x) (map f xs)
lookupBy : (p : key -> key -> Bool) -> (e : key) -> (xs : Vect n (key, val)) -> Maybe val

Find the association of some key with a user-provided comparison

p

the comparison operator for keys (True if they match)

e

the key to look for

lookup : Eq key => key -> Vect n (key, val) -> Maybe val

Find the assocation of some key using the default Boolean equality test

length : (xs : Vect len elem) -> Nat

Calculate the length of a Vect.

Note: this is only useful if you don't already statically know the length
and you want to avoid matching the implicit argument for erasure reasons.

len

the length (provably equal to the return value)

xs

the vector

last : Vect (S len) elem -> elem

The last element of the vector

isSuffixOfBy : (elem -> elem -> Bool) -> Vect m elem -> Vect len elem -> Bool
isSuffixOf : Eq elem => Vect m elem -> Vect len elem -> Bool
isPrefixOfBy : (elem -> elem -> Bool) -> Vect m elem -> Vect len elem -> Bool
isPrefixOf : Eq elem => Vect m elem -> Vect len elem -> Bool
isElem : DecEq a => (x : a) -> (xs : Vect n a) -> Dec (Elem x xs)

A decision procedure for Elem

intersperse : (sep : elem) -> (xs : Vect len elem) -> Vect (len + pred len) elem

Alternate an element between the other elements of a vector

sep

the element to intersperse

xs

the vector to separate with sep

insertAt : Fin (S len) -> elem -> Vect len elem -> Vect (S len) elem

Insert an element at a particular index

init : Vect (S len) elem -> Vect len elem

All but the last element of the vector

index : Fin len -> Vect len elem -> elem

Extract a particular element from a vector

head : Vect (S len) elem -> elem

Only the first element of the vector

hasAnyBy : (p : elem -> elem -> Bool) -> (elems : Vect m elem) -> (xs : Vect len elem) -> Bool

Check if any element of xs is found in elems by a user-provided comparison

p

the comparison operator

elems

the vector to search

xs

what to search for

hasAny : Eq elem => Vect m elem -> Vect len elem -> Bool

Check if any element of xs is found in elems using the default Boolean equality test

fromList' : Vect len elem -> (l : List elem) -> Vect (length l + len) elem
fromList : (l : List elem) -> Vect (length l) elem

Convert a list to a vector.

The length of the list should be statically known.

foldrImpl : (t -> acc -> acc) -> acc -> (acc -> acc) -> Vect n t -> acc
foldr1 : (t -> t -> t) -> Vect (S n) t -> t

Foldr without seeding the accumulator

foldl1 : (t -> t -> t) -> Vect (S n) t -> t

Foldl without seeding the accumulator

findIndices : (elem -> Bool) -> Vect m elem -> List (Fin m)

Find the indices of all elements that satisfy some test

findIndex : (elem -> Bool) -> Vect len elem -> Maybe (Fin len)

Find the index of the first element of the vector that satisfies some test

find : (p : elem -> Bool) -> (xs : Vect len elem) -> Maybe elem

Find the first element of the vector that satisfies some test

p

the test to satisfy

filter : (elem -> Bool) -> Vect len elem -> (p : Nat ** Vect p elem)

Find all elements of a vector that satisfy some test

exactLength : (len : Nat) -> (xs : Vect m a) -> Maybe (Vect len a)

If the given Vect is the required length, return a Vect with that
length in its type, otherwise return Nothing

len

the required length

xs

the vector with the desired length

elemIndicesBy : (elem -> elem -> Bool) -> elem -> Vect m elem -> List (Fin m)
elemIndices : Eq elem => elem -> Vect m elem -> List (Fin m)
elemIndexBy : (elem -> elem -> Bool) -> elem -> Vect m elem -> Maybe (Fin m)
elemIndex : Eq elem => elem -> Vect m elem -> Maybe (Fin m)
elemBy : (p : elem -> elem -> Bool) -> (e : elem) -> (xs : Vect len elem) -> Bool

Search for an item using a user-provided test

p

the equality test

e

the item to search for

xs

the vector to search in

elem : Eq elem => (x : elem) -> (xs : Vect len elem) -> Bool

Use the default Boolean equality on elements to search for an item

x

what to search for

xs

where to search

dropWhile : (p : elem -> Bool) -> Vect len elem -> (q : Nat ** Vect q elem)

Remove the longest prefix of a Vect such that all removed elements satisfy some
Boolean predicate.

p

the predicate

dropElem : (xs : Vect (S k) t) -> (p : Elem x xs) -> Vect k t

Remove the element at the given position.

xs

The vector to be removed from

p

A proof that the element to be removed is in the vector

drop : (n : Nat) -> Vect (n + m) elem -> Vect m elem

Remove the first n elements of a Vect

n

the number of elements to remove

diag : Vect len (Vect len elem) -> Vect len elem
deleteBy : (elem -> elem -> Bool) -> elem -> Vect len elem -> (p : Nat ** Vect p elem)
deleteAt : Fin (S len) -> Vect (S len) elem -> Vect len elem

Construct a new vector consisting of all but the indicated element

delete : Eq elem => elem -> Vect len elem -> (p : Nat ** Vect p elem)
concat : (xss : Vect m (Vect n elem)) -> Vect (m * n) elem

Flatten a vector of equal-length vectors

catMaybes : Vect n (Maybe elem) -> (p : Nat ** Vect p elem)
data Vect : (len : Nat) -> (elem : Type) -> Type

Vectors: Generic lists with explicit length in the type

len

the length of the list

elem

the type of elements

Nil : Vect 0 elem

Empty vector

(::) : (x : elem) -> (xs : Vect len elem) -> Vect (S len) elem

A non-empty vector of length S len, consisting of a head element and
the rest of the list, of length len.

Fixity
Left associative, precedence 7
data Elem : a -> Vect k a -> Type

A proof that some element is found in a vector

Here : Elem x (x :: xs)
There : (later : Elem x xs) -> Elem x (y :: xs)
(++) : (xs : Vect m elem) -> (ys : Vect n elem) -> Vect (m + n) elem

Append two vectors

Fixity
Left associative, precedence 7