Util.List
include module type of Stdlib.List
val mem_f : 'a eq -> 'a -> 'a list -> bool
Same as List.mem
, for some specific equality
Same as List.for_all2
but returning false
when of different length
prefix_of eq l1 l2
returns true
if l1
is a prefix of l2
, false
otherwise. It uses eq
to compare elements
A more efficient variant of for_all2eq (fun _ _ -> true)
interval i j
creates the list [i; i + 1; ...; j]
, or []
when j <= i
.
make n x
returns a list made of n
times x
. Raise Invalid_argument _
if n
is negative.
init n f
constructs the list f 0; ... ; f (n - 1)
. Raise Invalid_argument _
if n
is negative
assign l i x
sets the i
-th element of l
to x
, starting from 0
. Raise Failure _
if i
is out of range.
Like OCaml List.filter
but tail-recursive and physically returns the original list if the predicate holds for all elements.
Like List.filter
but with 2 arguments, raise Invalid_argument _
if the lists are not of same length.
Like List.filter
but with an index starting from 0
filter_with bl l
selects elements of l
whose corresponding element in bl
is true
. Raise Invalid_argument _
if sizes differ.
Like map_filter
but with an index starting from 0
Like List.partition
but with an index starting from 0
As map
but ensures the left-to-right order of evaluation.
Like OCaml List.mapi
but tail-recursive. Alternatively, like map
but with an index
Like map2
but with an index
Like map
but for 3 lists.
Like map
but for 4 lists.
map_of_array f a
behaves as List.map f (Array.to_list a)
map_append f [x1; ...; xn]
returns f x1 @ ... @ f xn
.
Like map_append
but for two lists; raises Invalid_argument _
if the two lists do not have the same length.
extend l a [a1..an]
assumes that the number of true
in l
is n
; it extends a1..an
by inserting a
at the position of false
in l
val index : 'a eq -> 'a -> 'a list -> int
index
returns the 1st index of an element in a list (counting from 1).
val safe_index : 'a eq -> 'a -> 'a list -> int option
safe_index
returns the 1st index of an element in a list (counting from 1) and None otherwise.
val index0 : 'a eq -> 'a -> 'a list -> int
index0
behaves as index
except that it starts counting at 0.
val fold_left_until : ('c -> 'a -> 'c CSig.until) -> 'c -> 'a list -> 'c
acts like fold_left f acc s
while f
returns Cont acc'
; it stops returning c
as soon as f
returns Stop c
.
Like List.fold_right
but with an index
Like List.fold_left
but with an index
fold_right_and_left f [a1;...;an] hd
is f (f (... (f (f hd an [an-1;...;a1]) an-1 [an-2;...;a1]) ...) a2 [a1]) a1 []
Like List.fold_left
but for 3 lists; raise Invalid_argument _
if not all lists of the same size
val fold_left2_set : exn -> ('a -> 'b -> 'c -> 'b list -> 'c list -> 'a) -> 'a -> 'b list -> 'c list -> 'a
Fold sets, i.e. lists up to order; the folding function tells when elements match by returning a value and raising the given exception otherwise; sets should have the same size; raise the given exception if no pairing of the two sets is found;; complexity in O(n^2)
fold_left_map f e_0 [a1;...;an]
is e_n,[k_1...k_n]
where (e_i,k_i)
is f e_{i-1} ai
for each i<=n
Same, folding on the right
Same with two lists, folding on the left
Same with two lists, folding on the right
val fold_left3_map : ('a -> 'b -> 'c -> 'd -> 'a * 'e) -> 'a -> 'b list -> 'c list -> 'd list -> 'a * 'e list
Same with three lists, folding on the left
val fold_left4_map : ('a -> 'b -> 'c -> 'd -> 'e -> 'a * 'r) -> 'a -> 'b list -> 'c list -> 'd list -> 'e list -> 'a * 'r list
Same with four lists, folding on the left
val remove : 'a eq -> 'a -> 'a list -> 'a list
remove eq a l
Remove all occurrences of a
in l
Remove the first element satisfying a predicate, or raise Not_found
Remove and return the first element satisfying a predicate, or raise Not_found
Returns the first element that is mapped to Some _
. Raise Not_found
if there is none.
goto i l
splits l
into two lists (l1,l2)
such that (List.rev l1)++l2=l
and l1
has length i
. It raises IndexOutOfRange
when i
is negative or greater than the length of l
.
split_when p l
splits l
into two lists (l1,a::l2)
such that l1++(a::l2)=l
, p a=true
and p b = false
for every element b
of l1
. if there is no such a
, then it returns (l,[])
instead.
sep_first l
returns (a,l')
such that l
is a::l'
. It raises Failure _
if the list is empty.
sep_last l
returns (a,l')
such that l
is l'@[a]
. It raises Failure _
if the list is empty.
Remove the last element of the list. It raises Failure _
if the list is empty. This is the second part of sep_last
.
Return the last element of the list. It raises Failure _
if the list is empty. This is the first part of sep_last
.
lastn n l
returns the n
last elements of l
. It raises Failure _
if n
is less than 0 or larger than the length of l
chop i l
splits l
into two lists (l1,l2)
such that l1++l2=l
and l1
has length i
. It raises Failure _
when i
is negative or greater than the length of l
.
firstn n l
Returns the n
first elements of l
. It raises Failure _
if n
negative or too large. This is the first part of chop
.
skipn n l
drops the n
first elements of l
. It raises Failure _
if n
is less than 0 or larger than the length of l
. This is the second part of chop
.
Same as skipn
but returns if n
is larger than the length of the list.
val drop_prefix : 'a eq -> 'a list -> 'a list -> 'a list
drop_prefix eq l1 l
returns l2
if l=l1++l2
else return l
.
val insert : 'a eq -> 'a -> 'a list -> 'a list
Insert at the (first) position so that if the list is ordered wrt to the total order given as argument, the order is preserved
share_tails l1 l2
returns (l1',l2',l)
such that l1
is l1'\@l
and l2
is l2'\@l
and l
is maximal amongst all such decompositions
Applies a function on the codomain of an association list
val assoc_f : 'a eq -> 'a -> ('a * 'b) list -> 'b
Like List.assoc
but using the equality given as argument
val assoc_f_opt : 'a eq -> 'a -> ('a * 'b) list -> 'b option
Like List.assoc_opt
but using the equality given as argument
val remove_assoc_f : 'a eq -> 'a -> ('a * 'b) list -> ('a * 'b) list
Remove first matching element; unchanged if no such element
val mem_assoc_f : 'a eq -> 'a -> ('a * 'b) list -> bool
Like List.mem_assoc
but using the equality given as argument
val factorize_left : 'a eq -> ('a * 'b) list -> ('a * 'b list) list
Create a list of associations from a list of pairs
Like split
but for quads
val add_set : 'a eq -> 'a -> 'a list -> 'a list
add_set x l
adds x
in l
if it is not already there, or returns l
otherwise.
Test equality up to permutation. It respects multiple occurrences and thus works also on multisets.
val subset : 'a list eq
Tell if a list is a subset of another up to permutation. It expects each element to occur only once.
val merge_set : 'a cmp -> 'a list -> 'a list -> 'a list
Merge two sorted lists and preserves the uniqueness property.
val intersect : 'a eq -> 'a list -> 'a list -> 'a list
Return the intersection of two lists, assuming and preserving uniqueness of elements
val union : 'a eq -> 'a list -> 'a list -> 'a list
Return the union of two lists, assuming and preserving uniqueness of elements
val subtract : 'a eq -> 'a list -> 'a list -> 'a list
Remove from the first list all elements from the second list.
val distinct_f : 'a cmp -> 'a list -> bool
Like distinct
but using the equality given as argument
val duplicates : 'a eq -> 'a list -> 'a list
Return the list of unique elements which appear at least twice. Elements are kept in the order of their first appearance.
Return the list of elements without duplicates using the function to associate a comparison key to each element. This is the list unchanged if there was none.
Return the list of elements without duplicates. This is the list unchanged if there was none.
val sort_uniquize : 'a cmp -> 'a list -> 'a list
Return a sorted version of a list without duplicates according to some comparison function.
val min : 'a cmp -> 'a list -> 'a
Return minimum element according to some comparison function.
A generic binary cartesian product: for any operator (**), cartesian (**) [x1;x2] [y1;y2] = [x1**y1; x1**y2; x2**y1; x2**y1]
, and so on if there are more elements in the lists.
cartesians op init l
is an n-ary cartesian product: it builds the list of all op a1 .. (op an init) ..
for a1
, ..., an
in the product of the elements of the lists
combinations l
returns the list of n_1
* ... * n_p
tuples [a11;...;ap1];...;[a1n_1;...;apn_pd]
whenever l
is a list [a11;..;a1n_1];...;[ap1;apn_p]
; otherwise said, it is cartesians (::) [] l
Like cartesians op init l
but keep only the tuples for which op
returns Some _
on all the elements of the tuple.
module Smart : sig ... end
When returning a list of same type as the input, maximally shares the suffix of the output which is physically equal to the corresponding suffix of the input
module type MonoS = sig ... end