Module Util.List
include CList.S
val length : 'a list -> int
val compare_lengths : 'a list -> 'b list -> int
val compare_length_with : 'a list -> int -> int
val cons : 'a -> 'a list -> 'a list
val hd : 'a list -> 'a
val tl : 'a list -> 'a list
val nth : 'a list -> int -> 'a
val nth_opt : 'a list -> int -> 'a option
val rev : 'a list -> 'a list
val init : int -> (int -> 'a) -> 'a list
val append : 'a list -> 'a list -> 'a list
val rev_append : 'a list -> 'a list -> 'a list
val concat : 'a list list -> 'a list
val flatten : 'a list list -> 'a list
val equal : ('a -> 'a -> bool) -> 'a list -> 'a list -> bool
val compare : ('a -> 'a -> int) -> 'a list -> 'a list -> int
val iter : ('a -> unit) -> 'a list -> unit
val iteri : (int -> 'a -> unit) -> 'a list -> unit
val map : ('a -> 'b) -> 'a list -> 'b list
val mapi : (int -> 'a -> 'b) -> 'a list -> 'b list
val rev_map : ('a -> 'b) -> 'a list -> 'b list
val filter_map : ('a -> 'b option) -> 'a list -> 'b list
val concat_map : ('a -> 'b list) -> 'a list -> 'b list
val fold_left_map : ('a -> 'b -> 'a * 'c) -> 'a -> 'b list -> 'a * 'c list
val fold_left : ('a -> 'b -> 'a) -> 'a -> 'b list -> 'a
val fold_right : ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b
val iter2 : ('a -> 'b -> unit) -> 'a list -> 'b list -> unit
val map2 : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list
val rev_map2 : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list
val fold_left2 : ('a -> 'b -> 'c -> 'a) -> 'a -> 'b list -> 'c list -> 'a
val fold_right2 : ('a -> 'b -> 'c -> 'c) -> 'a list -> 'b list -> 'c -> 'c
val for_all : ('a -> bool) -> 'a list -> bool
val exists : ('a -> bool) -> 'a list -> bool
val for_all2 : ('a -> 'b -> bool) -> 'a list -> 'b list -> bool
val exists2 : ('a -> 'b -> bool) -> 'a list -> 'b list -> bool
val mem : 'a -> 'a list -> bool
val memq : 'a -> 'a list -> bool
val find : ('a -> bool) -> 'a list -> 'a
val find_opt : ('a -> bool) -> 'a list -> 'a option
val find_map : ('a -> 'b option) -> 'a list -> 'b option
val filter : ('a -> bool) -> 'a list -> 'a list
val find_all : ('a -> bool) -> 'a list -> 'a list
val filteri : (int -> 'a -> bool) -> 'a list -> 'a list
val partition : ('a -> bool) -> 'a list -> 'a list * 'a list
val partition_map : ('a -> ('b, 'c) Stdlib.Either.t) -> 'a list -> 'b list * 'c list
val assoc : 'a -> ('a * 'b) list -> 'b
val assoc_opt : 'a -> ('a * 'b) list -> 'b option
val assq : 'a -> ('a * 'b) list -> 'b
val assq_opt : 'a -> ('a * 'b) list -> 'b option
val mem_assoc : 'a -> ('a * 'b) list -> bool
val mem_assq : 'a -> ('a * 'b) list -> bool
val remove_assoc : 'a -> ('a * 'b) list -> ('a * 'b) list
val remove_assq : 'a -> ('a * 'b) list -> ('a * 'b) list
val split : ('a * 'b) list -> 'a list * 'b list
val combine : 'a list -> 'b list -> ('a * 'b) list
val sort : ('a -> 'a -> int) -> 'a list -> 'a list
val stable_sort : ('a -> 'a -> int) -> 'a list -> 'a list
val fast_sort : ('a -> 'a -> int) -> 'a list -> 'a list
val sort_uniq : ('a -> 'a -> int) -> 'a list -> 'a list
val merge : ('a -> 'a -> int) -> 'a list -> 'a list -> 'a list
val to_seq : 'a list -> 'a Stdlib.Seq.t
val of_seq : 'a Stdlib.Seq.t -> 'a list
Equality, testing
val mem_f : 'a CList.eq -> 'a -> 'a list -> bool
Same as
List.mem
, for some specific equality
val for_all2eq : ('a -> 'b -> bool) -> 'a list -> 'b list -> bool
Same as
List.for_all2
but returningfalse
when of different length
Creating lists
val interval : int -> int -> int list
interval i j
creates the list[i; i + 1; ...; j]
, or[]
whenj <= i
.
val make : int -> 'a -> 'a list
make n x
returns a list made ofn
timesx
. RaiseInvalid_argument _
ifn
is negative.
Lists as arrays
Filtering
val filter : ('a -> bool) -> 'a list -> 'a list
Like OCaml
List.filter
but tail-recursive and physically returns the original list if the predicate holds for all elements.
val filter2 : ('a -> 'b -> bool) -> 'a list -> 'b list -> 'a list * 'b list
Like
List.filter
but with 2 arguments, raiseInvalid_argument _
if the lists are not of same length.
val filteri : (int -> 'a -> bool) -> 'a list -> 'a list
Like
List.filter
but with an index starting from0
val filter_with : bool list -> 'a list -> 'a list
filter_with bl l
selects elements ofl
whose corresponding element inbl
istrue
. RaiseInvalid_argument _
if sizes differ.
Applying functorially
val map_left : ('a -> 'b) -> 'a list -> 'b list
As
map
but ensures the left-to-right order of evaluation.
val map_i : (int -> 'a -> 'b) -> int -> 'a list -> 'b list
Like OCaml
List.mapi
but tail-recursive. Alternatively, likemap
but with an index
val map2_i : (int -> 'a -> 'b -> 'c) -> int -> 'a list -> 'b list -> 'c list
Like
map2
but with an index
val map3 : ('a -> 'b -> 'c -> 'd) -> 'a list -> 'b list -> 'c list -> 'd list
Like
map
but for 3 lists.
val map4 : ('a -> 'b -> 'c -> 'd -> 'e) -> 'a list -> 'b list -> 'c list -> 'd list -> 'e list
Like
map
but for 4 lists.
val map_of_array : ('a -> 'b) -> 'a array -> 'b list
map_of_array f a
behaves asList.map f (Array.to_list a)
val map_append : ('a -> 'b list) -> 'a list -> 'b list
map_append f [x1; ...; xn]
returnsf x1 @ ... @ f xn
.
val map_append2 : ('a -> 'b -> 'c list) -> 'a list -> 'b list -> 'c list
Like
map_append
but for two lists; raisesInvalid_argument _
if the two lists do not have the same length.
Finding position
val index : 'a CList.eq -> 'a -> 'a list -> int
index
returns the 1st index of an element in a list (counting from 1).
val safe_index : 'a CList.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 CList.eq -> 'a -> 'a list -> int
index0
behaves asindex
except that it starts counting at 0.
Folding
val fold_left_until : ('c -> 'a -> 'c CSig.until) -> 'c -> 'a list -> 'c
acts like
fold_left f acc s
whilef
returnsCont acc'
; it stops returningc
as soon asf
returnsStop c
.
val fold_right_i : (int -> 'a -> 'b -> 'b) -> int -> 'a list -> 'b -> 'b
Like
List.fold_right
but with an index
val fold_left_i : (int -> 'a -> 'b -> 'a) -> int -> 'a -> 'b list -> 'a
Like
List.fold_left
but with an index
val fold_right_and_left : ('b -> 'a -> 'a list -> 'b) -> 'a list -> 'b -> 'b
fold_right_and_left f [a1;...;an] hd
isf (f (... (f (f hd an [an-1;...;a1]) an-1 [an-2;...;a1]) ...) a2 [a1]) a1 []
val fold_left3 : ('a -> 'b -> 'c -> 'd -> 'a) -> 'a -> 'b list -> 'c list -> 'd list -> 'a
Like
List.fold_left
but for 3 lists; raiseInvalid_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)
val fold_left_map : ('a -> 'b -> 'a * 'c) -> 'a -> 'b list -> 'a * 'c list
fold_left_map f e_0 [a1;...;an]
ise_n,[k_1...k_n]
where(e_i,k_i)
isf e_{i-1} ai
for each i<=n
val fold_right_map : ('b -> 'a -> 'c * 'a) -> 'b list -> 'a -> 'c list * 'a
Same, folding on the right
val fold_left2_map : ('a -> 'b -> 'c -> 'a * 'd) -> 'a -> 'b list -> 'c list -> 'a * 'd list
Same with two lists, folding on the left
val fold_right2_map : ('b -> 'c -> 'a -> 'd * 'a) -> 'b list -> 'c list -> 'a -> 'd list * 'a
Same with two lists, folding on the right
Splitting
val except : 'a CList.eq -> 'a -> 'a list -> 'a list
except eq a l
Remove all occurrences ofa
inl
val remove : 'a CList.eq -> 'a -> 'a list -> 'a list
Alias of
except
val remove_first : ('a -> bool) -> 'a list -> 'a list
Remove the first element satisfying a predicate, or raise
Not_found
val extract_first : ('a -> bool) -> 'a list -> 'a list * 'a
Remove and return the first element satisfying a predicate, or raise
Not_found
val find_map : ('a -> 'b option) -> 'a list -> 'b
Returns the first element that is mapped to
Some _
. RaiseNot_found
if there is none.
val goto : int -> 'a list -> 'a list * 'a list
goto i l
splitsl
into two lists(l1,l2)
such that(List.rev l1)++l2=l
andl1
has lengthi
. It raisesIndexOutOfRange
wheni
is negative or greater than the length ofl
.
val split_when : ('a -> bool) -> 'a list -> 'a list * 'a list
split_when p l
splitsl
into two lists(l1,a::l2)
such thatl1++(a::l2)=l
,p a=true
andp b = false
for every elementb
ofl1
. if there is no sucha
, then it returns(l,[])
instead.
val sep_last : 'a list -> 'a * 'a list
sep_last l
returns(a,l')
such thatl
isl'@[a]
. It raisesFailure _
if the list is empty.
val drop_last : 'a list -> 'a list
Remove the last element of the list. It raises
Failure _
if the list is empty. This is the second part ofsep_last
.
val last : 'a list -> 'a
Return the last element of the list. It raises
Failure _
if the list is empty. This is the first part ofsep_last
.
val lastn : int -> 'a list -> 'a list
lastn n l
returns then
last elements ofl
. It raisesFailure _
ifn
is less than 0 or larger than the length ofl
val chop : int -> 'a list -> 'a list * 'a list
chop i l
splitsl
into two lists(l1,l2)
such thatl1++l2=l
andl1
has lengthi
. It raisesFailure _
wheni
is negative or greater than the length ofl
.
val firstn : int -> 'a list -> 'a list
firstn n l
Returns then
first elements ofl
. It raisesFailure _
ifn
negative or too large. This is the first part ofchop
.
val skipn : int -> 'a list -> 'a list
skipn n l
drops then
first elements ofl
. It raisesFailure _
ifn
is less than 0 or larger than the length ofl
. This is the second part ofchop
.
val skipn_at_least : int -> 'a list -> 'a list
Same as
skipn
but returnsif
n
is larger than the length of the list.
val drop_prefix : 'a CList.eq -> 'a list -> 'a list -> 'a list
drop_prefix eq l1 l
returnsl2
ifl=l1++l2
else returnl
.
val insert : 'a CList.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 thatl1
isl1'\@l
andl2
isl2'\@l
andl
is maximal amongst all such decompositions
Association lists
val map_assoc : ('a -> 'b) -> ('c * 'a) list -> ('c * 'b) list
Applies a function on the codomain of an association list
val assoc_f : 'a CList.eq -> 'a -> ('a * 'b) list -> 'b
Like
List.assoc
but using the equality given as argument
val remove_assoc_f : 'a CList.eq -> 'a -> ('a * 'b) list -> ('a * 'b) list
Remove first matching element; unchanged if no such element
val mem_assoc_f : 'a CList.eq -> 'a -> ('a * 'b) list -> bool
Like
List.mem_assoc
but using the equality given as argument
val factorize_left : 'a CList.eq -> ('a * 'b) list -> ('a * 'b list) list
Create a list of associations from a list of pairs
Operations on lists of tuples
Operations on lists seen as sets, preserving uniqueness of elements
val add_set : 'a CList.eq -> 'a -> 'a list -> 'a list
add_set x l
addsx
inl
if it is not already there, or returnsl
otherwise.
val eq_set : 'a CList.eq -> 'a list CList.eq
Test equality up to permutation. It respects multiple occurrences and thus works also on multisets.
val subset : 'a list CList.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 CList.cmp -> 'a list -> 'a list -> 'a list
Merge two sorted lists and preserves the uniqueness property.
val intersect : 'a CList.eq -> 'a list -> 'a list -> 'a list
Return the intersection of two lists, assuming and preserving uniqueness of elements
val union : 'a CList.eq -> 'a list -> 'a list -> 'a list
Return the union of two lists, assuming and preserving uniqueness of elements
val subtract : 'a CList.eq -> 'a list -> 'a list -> 'a list
Remove from the first list all elements from the second list.
Uniqueness and duplication
val distinct_f : 'a CList.cmp -> 'a list -> bool
Like
distinct
but using the equality given as argument
val duplicates : 'a CList.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.
val uniquize_key : ('a -> 'b) -> 'a list -> 'a list
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.
val uniquize : 'a list -> 'a list
Return the list of elements without duplicates. This is the list unchanged if there was none.
val sort_uniquize : 'a CList.cmp -> 'a list -> 'a list
Return a sorted version of a list without duplicates according to some comparison function.
val min : 'a CList.cmp -> 'a list -> 'a
Return minimum element according to some comparison function.
- raises Not_found
on an empty list.
Cartesian product
val cartesian : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list
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.
val cartesians : ('a -> 'b -> 'b) -> 'b -> 'a list list -> 'b list
cartesians op init l
is an n-ary cartesian product: it builds the list of allop a1 .. (op an init) ..
fora1
, ...,an
in the product of the elements of the lists
val combinations : 'a list list -> 'a list list
combinations l
returns the list ofn_1
* ... *n_p
tuples[a11;...;ap1];...;[a1n_1;...;apn_pd]
wheneverl
is a list[a11;..;a1n_1];...;[ap1;apn_p]
; otherwise said, it iscartesians (::) [] l
val cartesians_filter : ('a -> 'b -> 'b option) -> 'b -> 'a list list -> 'b list
Like
cartesians op init l
but keep only the tuples for whichop
returnsSome _
on all the elements of the tuple.
module Smart : sig ... end
module type MonoS = sig ... end