aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--checker/check.ml8
-rw-r--r--checker/check.mllib1
-rw-r--r--checker/closure.ml4
-rw-r--r--checker/declarations.ml2
-rw-r--r--checker/indtypes.ml8
-rw-r--r--checker/inductive.ml14
-rw-r--r--checker/term.ml2
-rw-r--r--dev/printers.mllib1
-rw-r--r--grammar/grammar.mllib1
-rw-r--r--grammar/tacextend.ml46
-rw-r--r--grammar/vernacextend.ml42
-rw-r--r--ide/coqide.ml2
-rw-r--r--ide/wg_Command.ml2
-rw-r--r--ide/wg_Notebook.ml12
-rw-r--r--ide/wg_ProofView.ml2
-rw-r--r--interp/constrexpr_ops.ml2
-rw-r--r--interp/constrextern.ml18
-rw-r--r--interp/constrintern.ml40
-rw-r--r--interp/coqlib.ml4
-rw-r--r--interp/dumpglob.ml2
-rw-r--r--interp/implicit_quantifiers.ml2
-rw-r--r--interp/modintern.ml2
-rw-r--r--interp/notation.ml10
-rw-r--r--interp/notation_ops.ml36
-rw-r--r--interp/reserve.ml2
-rw-r--r--interp/topconstr.ml2
-rw-r--r--kernel/closure.ml4
-rw-r--r--kernel/declarations.ml6
-rw-r--r--kernel/indtypes.ml8
-rw-r--r--kernel/inductive.ml14
-rw-r--r--kernel/modops.ml4
-rw-r--r--kernel/names.ml2
-rw-r--r--kernel/pre_env.ml4
-rw-r--r--kernel/sign.ml4
-rw-r--r--kernel/subtyping.ml2
-rw-r--r--kernel/univ.ml20
-rw-r--r--lib/cList.ml718
-rw-r--r--lib/cList.mli188
-rw-r--r--lib/clib.mllib1
-rw-r--r--lib/gmapl.ml4
-rw-r--r--lib/util.ml535
-rw-r--r--lib/util.mli128
-rw-r--r--library/declare.ml2
-rw-r--r--library/heads.ml4
-rw-r--r--library/impargs.ml8
-rw-r--r--library/lib.ml2
-rw-r--r--library/libnames.ml10
-rw-r--r--library/library.ml2
-rw-r--r--parsing/egramcoq.ml6
-rw-r--r--parsing/g_tactic.ml410
-rw-r--r--parsing/g_xml.ml412
-rw-r--r--parsing/lexer.ml42
-rw-r--r--parsing/pcoq.ml42
-rw-r--r--plugins/cc/cctac.ml2
-rw-r--r--plugins/decl_mode/decl_interp.ml8
-rw-r--r--plugins/decl_mode/decl_proof_instr.ml18
-rw-r--r--plugins/extraction/common.ml4
-rw-r--r--plugins/extraction/extract_env.ml6
-rw-r--r--plugins/extraction/extraction.ml26
-rw-r--r--plugins/extraction/mlutil.ml18
-rw-r--r--plugins/extraction/modutil.ml2
-rw-r--r--plugins/extraction/ocaml.ml6
-rw-r--r--plugins/extraction/table.ml4
-rw-r--r--plugins/firstorder/formula.ml6
-rw-r--r--plugins/firstorder/instances.ml2
-rw-r--r--plugins/firstorder/rules.ml2
-rw-r--r--plugins/firstorder/sequent.ml4
-rw-r--r--plugins/firstorder/unify.ml2
-rw-r--r--plugins/funind/functional_principles_proofs.ml10
-rw-r--r--plugins/funind/functional_principles_types.ml6
-rw-r--r--plugins/funind/glob_term_to_relation.ml18
-rw-r--r--plugins/funind/indfun.ml16
-rw-r--r--plugins/funind/invfun.ml10
-rw-r--r--plugins/funind/merge.ml4
-rw-r--r--plugins/funind/recdef.ml8
-rw-r--r--plugins/nsatz/ideal.ml2
-rw-r--r--plugins/omega/coq_omega.ml2
-rw-r--r--plugins/omega/g_omega.ml42
-rw-r--r--plugins/omega/omega.ml18
-rw-r--r--plugins/romega/g_romega.ml42
-rw-r--r--plugins/romega/refl_omega.ml20
-rw-r--r--plugins/setoid_ring/newring.ml48
-rw-r--r--pretyping/cases.ml42
-rw-r--r--pretyping/cbv.ml4
-rw-r--r--pretyping/classops.ml2
-rw-r--r--pretyping/coercion.ml2
-rw-r--r--pretyping/detyping.ml18
-rw-r--r--pretyping/evarconv.ml8
-rw-r--r--pretyping/evarutil.ml26
-rw-r--r--pretyping/evd.ml8
-rw-r--r--pretyping/glob_ops.ml8
-rw-r--r--pretyping/indrec.ml6
-rw-r--r--pretyping/inductiveops.ml12
-rw-r--r--pretyping/matching.ml4
-rw-r--r--pretyping/patternops.ml14
-rw-r--r--pretyping/pretyping.ml4
-rw-r--r--pretyping/recordops.ml4
-rw-r--r--pretyping/redops.ml4
-rw-r--r--pretyping/reductionops.ml22
-rw-r--r--pretyping/tacred.ml28
-rw-r--r--pretyping/termops.ml14
-rw-r--r--pretyping/typeclasses.ml16
-rw-r--r--pretyping/typing.ml2
-rw-r--r--printing/ppconstr.ml8
-rw-r--r--printing/pptactic.ml2
-rw-r--r--printing/prettyp.ml4
-rw-r--r--printing/printer.ml4
-rw-r--r--proofs/clenv.ml4
-rw-r--r--proofs/logic.ml6
-rw-r--r--proofs/pfedit.ml2
-rw-r--r--proofs/proofview.ml2
-rw-r--r--proofs/redexpr.ml4
-rw-r--r--proofs/refiner.ml4
-rw-r--r--tactics/auto.ml36
-rw-r--r--tactics/autorewrite.ml2
-rw-r--r--tactics/class_tactics.ml410
-rw-r--r--tactics/eauto.ml48
-rw-r--r--tactics/elim.ml4
-rw-r--r--tactics/eqdecide.ml42
-rw-r--r--tactics/eqschemes.ml18
-rw-r--r--tactics/equality.ml12
-rw-r--r--tactics/hipattern.ml42
-rw-r--r--tactics/inv.ml4
-rw-r--r--tactics/leminv.ml2
-rw-r--r--tactics/refine.ml4
-rw-r--r--tactics/rewrite.ml42
-rw-r--r--tactics/tacinterp.ml24
-rw-r--r--tactics/tacticals.ml10
-rw-r--r--tactics/tactics.ml28
-rw-r--r--tactics/tauto.ml42
-rw-r--r--tools/coq_makefile.ml2
-rw-r--r--toplevel/command.ml52
-rw-r--r--toplevel/discharge.ml2
-rw-r--r--toplevel/himsg.ml4
-rw-r--r--toplevel/indschemes.ml10
-rw-r--r--toplevel/lemmas.ml16
-rw-r--r--toplevel/metasyntax.ml16
-rw-r--r--toplevel/obligations.ml20
-rw-r--r--toplevel/record.ml14
-rw-r--r--toplevel/vernacentries.ml28
140 files changed, 1501 insertions, 1246 deletions
diff --git a/checker/check.ml b/checker/check.ml
index 9e6b63537..5b58dc9ba 100644
--- a/checker/check.ml
+++ b/checker/check.ml
@@ -149,20 +149,20 @@ let canonical_path_name p =
let find_logical_path phys_dir =
let phys_dir = canonical_path_name phys_dir in
- match list_filter2 (fun p d -> p = phys_dir) !load_paths with
+ match List.filter2 (fun p d -> p = phys_dir) !load_paths with
| _,[dir] -> dir
| _,[] -> default_root_prefix
| _,l -> anomaly ("Two logical paths are associated to "^phys_dir)
let remove_load_path dir =
- load_paths := list_filter2 (fun p d -> p <> dir) !load_paths
+ load_paths := List.filter2 (fun p d -> p <> dir) !load_paths
let add_load_path (phys_path,coq_path) =
if !Flags.debug then
ppnl (str "path: " ++ pr_dirpath coq_path ++ str " ->" ++ spc() ++
str phys_path);
let phys_path = canonical_path_name phys_path in
- match list_filter2 (fun p d -> p = phys_path) !load_paths with
+ match List.filter2 (fun p d -> p = phys_path) !load_paths with
| _,[dir] ->
if coq_path <> dir
(* If this is not the default -I . to coqtop *)
@@ -185,7 +185,7 @@ let add_load_path (phys_path,coq_path) =
| _ -> anomaly ("Two logical paths are associated to "^phys_path)
let load_paths_of_dir_path dir =
- fst (list_filter2 (fun p d -> d = dir) !load_paths)
+ fst (List.filter2 (fun p d -> d = dir) !load_paths)
(************************************************************************)
(*s Locate absolute or partially qualified library names in the path *)
diff --git a/checker/check.mllib b/checker/check.mllib
index 15d7df1d3..b7e196d4b 100644
--- a/checker/check.mllib
+++ b/checker/check.mllib
@@ -7,6 +7,7 @@ Loc
Segmenttree
Unicodetable
Errors
+CList
Util
Option
Hashcons
diff --git a/checker/closure.ml b/checker/closure.ml
index 7f23ed201..c3351cea1 100644
--- a/checker/closure.ml
+++ b/checker/closure.ml
@@ -370,7 +370,7 @@ let rec compact_constr (lg, subs as s) c k =
match c with
Rel i ->
if i < k then c,s else
- (try Rel (k + lg - list_index (i-k+1) subs), (lg,subs)
+ (try Rel (k + lg - List.index (i-k+1) subs), (lg,subs)
with Not_found -> Rel (k+lg), (lg+1, (i-k+1)::subs))
| (Sort _|Var _|Meta _|Ind _|Const _|Construct _) -> c,s
| Evar(ev,v) ->
@@ -648,7 +648,7 @@ let rec get_args n tys f e stk =
let eargs = Array.sub l n (na-n) in
(Inl (subs_cons(args,e)), Zapp eargs :: s)
else (* more lambdas *)
- let etys = list_skipn na tys in
+ let etys = List.skipn na tys in
get_args (n-na) etys f (subs_cons(l,e)) s
| _ -> (Inr {norm=Cstr;term=FLambda(n,tys,f,e)}, stk)
diff --git a/checker/declarations.ml b/checker/declarations.ml
index ba56c243f..ad14a3bbe 100644
--- a/checker/declarations.ml
+++ b/checker/declarations.ml
@@ -546,7 +546,7 @@ let subst_rel_declaration sub (id,copt,t as x) =
let t' = subst_mps sub t in
if copt == copt' & t == t' then x else (id,copt',t')
-let subst_rel_context sub = list_smartmap (subst_rel_declaration sub)
+let subst_rel_context sub = List.smartmap (subst_rel_declaration sub)
type recarg =
| Norec
diff --git a/checker/indtypes.ml b/checker/indtypes.ml
index b11b79d1a..6d936796a 100644
--- a/checker/indtypes.ml
+++ b/checker/indtypes.ml
@@ -331,7 +331,7 @@ let check_correct_par (env,n,ntypes,_) hyps l largs =
recursive parameters *)
let check_rec_par (env,n,_,_) hyps nrecp largs =
- let (lpar,_) = list_chop nrecp largs in
+ let (lpar,_) = List.chop nrecp largs in
let rec find index =
function
| ([],_) -> ()
@@ -355,7 +355,7 @@ let abstract_mind_lc env ntyps npars lc =
lc
else
let make_abs =
- list_tabulate
+ List.tabulate
(function i -> lambda_implicit_lift npars (Rel (i+1))) ntyps
in
Array.map (substl make_abs) lc
@@ -432,7 +432,7 @@ let check_positivity_one (env, _,ntypes,_ as ienv) hyps nrecp (_,i as ind) indlc
let auxnpar = mib.mind_nparams_rec in
let nonrecpar = mib.mind_nparams - auxnpar in
let (lpar,auxlargs) =
- try list_chop auxnpar largs
+ try List.chop auxnpar largs
with Failure _ -> raise (IllFormedInd (LocalNonPos n)) in
(* If the inductive appears in the args (non params) then the
definition is not positive. *)
@@ -514,7 +514,7 @@ let check_positivity env_ar mind params nrecp inds =
let lparams = rel_context_length params in
let check_one i mip =
let ra_env =
- list_tabulate (fun _ -> (Norec,mk_norec)) lparams @ lra_ind in
+ List.tabulate (fun _ -> (Norec,mk_norec)) lparams @ lra_ind in
let ienv = (env_ar, 1+lparams, ntypes, ra_env) in
check_positivity_one ienv params nrecp (mind,i) mip.mind_nf_lc
in
diff --git a/checker/inductive.ml b/checker/inductive.ml
index 6e87fb365..f35d68ad1 100644
--- a/checker/inductive.ml
+++ b/checker/inductive.ml
@@ -58,7 +58,7 @@ let inductive_params (mib,_) = mib.mind_nparams
let ind_subst mind mib =
let ntypes = mib.mind_ntypes in
let make_Ik k = Ind (mind,ntypes-k-1) in
- list_tabulate make_Ik ntypes
+ List.tabulate make_Ik ntypes
(* Instantiate inductives in constructor type *)
let constructor_instantiate mind mib c =
@@ -255,7 +255,7 @@ let extended_rel_list n hyps =
reln [] 1 hyps
let build_dependent_inductive ind (_,mip) params =
- let realargs,_ = list_chop mip.mind_nrealargs_ctxt mip.mind_arity_ctxt in
+ let realargs,_ = List.chop mip.mind_nrealargs_ctxt mip.mind_arity_ctxt in
applist
(Ind ind,
List.map (lift mip.mind_nrealargs_ctxt) params
@@ -309,7 +309,7 @@ let build_branches_type ind (_,mip as specif) params dep p =
let (args,ccl) = decompose_prod_assum typi in
let nargs = rel_context_length args in
let (_,allargs) = decompose_app ccl in
- let (lparams,vargs) = list_chop (inductive_params specif) allargs in
+ let (lparams,vargs) = List.chop (inductive_params specif) allargs in
let cargs =
if dep then
let cstr = ith_constructor_of_inductive ind (i+1) in
@@ -331,7 +331,7 @@ let build_case_type dep p c realargs =
let type_case_branches env (ind,largs) (p,pj) c =
let specif = lookup_mind_specif env ind in
let nparams = inductive_params specif in
- let (params,realargs) = list_chop nparams largs in
+ let (params,realargs) = List.chop nparams largs in
let dep = is_correct_arity env c (p,pj) ind specif params in
let lc = build_branches_type ind specif params dep p in
let ty = build_case_type dep p c realargs in
@@ -444,7 +444,7 @@ let push_var renv (x,ty,spec) =
genv = spec:: renv.genv }
let assign_var_spec renv (i,spec) =
- { renv with genv = list_assign renv.genv (i-1) spec }
+ { renv with genv = List.assign renv.genv (i-1) spec }
let push_var_renv renv (x,ty) =
push_var renv (x,ty,Lazy.lazy_from_val Not_subterm)
@@ -527,7 +527,7 @@ let branches_specif renv c_spec ci =
vra
| Dead_code -> Array.create nca Dead_code
| _ -> Array.create nca Not_subterm) in
- list_tabulate (fun j -> lazy (Lazy.force lvra).(j)) nca)
+ List.tabulate (fun j -> lazy (Lazy.force lvra).(j)) nca)
car
(* [subterm_specif renv t] computes the recursive structure of [t] and
@@ -871,7 +871,7 @@ let check_one_cofix env nbfix def deftype =
let lra = vlra.(i-1) in
let mI = inductive_of_constructor cstr_kn in
let (mib,mip) = lookup_mind_specif env mI in
- let realargs = list_skipn mib.mind_nparams args in
+ let realargs = List.skipn mib.mind_nparams args in
let rec process_args_of_constr = function
| (t::lr), (rar::lrar) ->
if rar = mk_norec then
diff --git a/checker/term.ml b/checker/term.ml
index 955a61c14..a7fc0a8b8 100644
--- a/checker/term.ml
+++ b/checker/term.ml
@@ -346,7 +346,7 @@ let map_context f l =
if body_o' == body_o && typ' == typ then decl else
(n, body_o', typ')
in
- list_smartmap map_decl l
+ List.smartmap map_decl l
let map_rel_context = map_context
diff --git a/dev/printers.mllib b/dev/printers.mllib
index 45da00a72..0cf74e627 100644
--- a/dev/printers.mllib
+++ b/dev/printers.mllib
@@ -8,6 +8,7 @@ Loc
Segmenttree
Unicodetable
Errors
+CList
Util
Bigint
Hashcons
diff --git a/grammar/grammar.mllib b/grammar/grammar.mllib
index 006766e17..6849669db 100644
--- a/grammar/grammar.mllib
+++ b/grammar/grammar.mllib
@@ -8,6 +8,7 @@ Loc
Segmenttree
Unicodetable
Errors
+CList
Util
Bigint
Hashcons
diff --git a/grammar/tacextend.ml4 b/grammar/tacextend.ml4
index 87425ef54..f38479ac9 100644
--- a/grammar/tacextend.ml4
+++ b/grammar/tacextend.ml4
@@ -53,7 +53,7 @@ let rec extract_signature = function
let check_unicity s l =
let l' = List.map (fun (l,_) -> extract_signature l) l in
- if not (Util.list_distinct l') then
+ if not (Util.List.distinct l') then
Pp.msg_warning
(strbrk ("Two distinct rules of tactic entry "^s^" have the same "^
"non-terminals in the same order: put them in distinct tactic entries"))
@@ -152,10 +152,10 @@ let rec possibly_empty_subentries loc = function
else possibly_empty_subentries loc l
let possibly_atomic loc prods =
- let l = list_map_filter (function
+ let l = List.map_filter (function
| GramTerminal s :: l, _ -> Some (s,l)
| _ -> None) prods in
- possibly_empty_subentries loc (list_factorize_left l)
+ possibly_empty_subentries loc (List.factorize_left l)
let declare_tactic loc s cl =
let se = mlexpr_of_string s in
diff --git a/grammar/vernacextend.ml4 b/grammar/vernacextend.ml4
index 029755f08..3074337f6 100644
--- a/grammar/vernacextend.ml4
+++ b/grammar/vernacextend.ml4
@@ -30,7 +30,7 @@ let rec make_let e = function
let check_unicity s l =
let l' = List.map (fun (_,l,_) -> extract_signature l) l in
- if not (Util.list_distinct l') then
+ if not (Util.List.distinct l') then
Pp.msg_warning
(strbrk ("Two distinct rules of entry "^s^" have the same "^
"non-terminals in the same order: put them in distinct vernac entries"))
diff --git a/ide/coqide.ml b/ide/coqide.ml
index b2fd0c1d7..1655ffb8a 100644
--- a/ide/coqide.ml
+++ b/ide/coqide.ml
@@ -1292,7 +1292,7 @@ let load_file handler f =
try
Minilib.log "Loading file starts";
let is_f = CUnix.same_file f in
- if not (Util.list_fold_left_i
+ if not (Util.List.fold_left_i
(fun i found x -> if found then found else
let {analyzed_view=av} = x in
(match av#filename with
diff --git a/ide/wg_Command.ml b/ide/wg_Command.ml
index d09adb474..5a72669b8 100644
--- a/ide/wg_Command.ml
+++ b/ide/wg_Command.ml
@@ -55,7 +55,7 @@ class command_window coqtop =
let remove_cb () =
let index = notebook#current_page in
let () = notebook#remove_page index in
- views := Util.list_filter_i (fun i x -> i <> index) !views
+ views := Util.List.filter_i (fun i x -> i <> index) !views
in
let _ =
toolbar#insert_button
diff --git a/ide/wg_Notebook.ml b/ide/wg_Notebook.ml
index df3d8df05..c75b371fe 100644
--- a/ide/wg_Notebook.ml
+++ b/ide/wg_Notebook.ml
@@ -16,14 +16,14 @@ object(self)
(* XXX - Temporary hack to compile with archaic lablgtk *)
ignore (super#append_page ?tab_label ?menu_label page);
let real_pos = super#page_num page in
- let lower,higher = Util.list_chop real_pos term_list in
+ let lower,higher = Util.List.chop real_pos term_list in
term_list <- lower@[term]@higher;
real_pos
(* XXX - Temporary hack to compile with archaic lablgtk
method insert_term ?(build=default_build) ?pos (term:'a) =
let tab_label,menu_label,page = build term in
let real_pos = super#insert_page ?tab_label ?menu_label ?pos page in
- let lower,higher = Util.list_chop real_pos term_list in
+ let lower,higher = Util.List.chop real_pos term_list in
term_list <- lower@[term]@higher;
real_pos
*)
@@ -32,26 +32,26 @@ object(self)
(* XXX - Temporary hack to compile with archaic lablgtk *)
ignore (super#prepend_page ?tab_label ?menu_label page);
let real_pos = super#page_num page in
- let lower,higher = Util.list_chop real_pos term_list in
+ let lower,higher = Util.List.chop real_pos term_list in
term_list <- lower@[term]@higher;
real_pos
method set_term (term:'a) =
let tab_label,menu_label,page = make_page term in
let real_pos = super#current_page in
- term_list <- Util.list_map_i (fun i x -> if i = real_pos then term else x) 0 term_list;
+ term_list <- Util.List.map_i (fun i x -> if i = real_pos then term else x) 0 term_list;
super#set_page ?tab_label ?menu_label page
method get_nth_term i =
List.nth term_list i
method term_num p =
- Util.list_index0 p term_list
+ Util.List.index0 p term_list
method pages = term_list
method remove_page index =
- term_list <- Util.list_filter_i (fun i x -> if i = index then kill_page x; i <> index) term_list;
+ term_list <- Util.List.filter_i (fun i x -> if i = index then kill_page x; i <> index) term_list;
super#remove_page index
method current_term =
diff --git a/ide/wg_ProofView.ml b/ide/wg_ProofView.ml
index af9dcba3a..5da0eeceb 100644
--- a/ide/wg_ProofView.ml
+++ b/ide/wg_ProofView.ml
@@ -100,7 +100,7 @@ let mode_tactic sel_cb (proof:GText.view) goals hints = match goals with
proof#buffer#insert (goal_str i goals_cnt);
proof#buffer#insert (g ^ "\n")
in
- let () = Util.list_fold_left_i fold_goal 2 () rem_goals in
+ let () = Util.List.fold_left_i fold_goal 2 () rem_goals in
ignore(proof#buffer#place_cursor
~where:(proof#buffer#end_iter#backward_to_tag_toggle
diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml
index 950edc5a3..306fbfad9 100644
--- a/interp/constrexpr_ops.ml
+++ b/interp/constrexpr_ops.ml
@@ -75,7 +75,7 @@ let local_binder_loc = function
let local_binders_loc bll =
if bll = [] then Loc.ghost else
- Loc.merge (local_binder_loc (List.hd bll)) (local_binder_loc (list_last bll))
+ Loc.merge (local_binder_loc (List.hd bll)) (local_binder_loc (List.last bll))
(** Pseudo-constructors *)
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index 272845d07..0ce8a8aac 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -257,7 +257,7 @@ let is_same_type c d =
(* mapping patterns to cases_pattern_expr *)
let add_patt_for_params ind l =
- Util.list_addn (fst (Inductiveops.inductive_nargs ind)) (CPatAtom (Loc.ghost,None)) l
+ Util.List.addn (fst (Inductiveops.inductive_nargs ind)) (CPatAtom (Loc.ghost,None)) l
let drop_implicits_in_patt cst nb_expl args =
let impl_st = (implicits_of_global cst) in
@@ -276,7 +276,7 @@ let drop_implicits_in_patt cst nb_expl args =
in
if nb_expl = 0 then aux impl_data
else
- let imps = list_skipn_at_least nb_expl (select_stronger_impargs impl_st) in
+ let imps = List.skipn_at_least nb_expl (select_stronger_impargs impl_st) in
impls_fit [] (imps,args)
let has_curly_brackets ntn =
@@ -355,7 +355,7 @@ let pattern_printable_in_both_syntax (ind,_ as c) =
let nb_params = Inductiveops.inductive_nparams ind in
List.exists (fun (_,impls) ->
(List.length impls >= nb_params) &&
- let params,args = Util.list_chop nb_params impls in
+ let params,args = Util.List.chop nb_params impls in
(List.for_all is_status_implicit params)&&(List.for_all (fun x -> not (is_status_implicit x)) args)
) impl_st
@@ -570,8 +570,8 @@ let explicitize loc inctx impl (cf,f) args =
let f' = match f with CRef f -> f | _ -> assert false in
CAppExpl (loc,(ip,f'),args)
else
- let (args1,args2) = list_chop i args in
- let (impl1,impl2) = if impl=[] then [],[] else list_chop i impl in
+ let (args1,args2) = List.chop i args in
+ let (impl1,impl2) = if impl=[] then [],[] else List.chop i impl in
let args1 = exprec 1 (args1,impl1) in
let args2 = exprec (i+1) (args2,impl2) in
CApp (loc,(Some (List.length args1),f),args1@args2)
@@ -617,7 +617,7 @@ let rec remove_coercions inctx = function
(try match Classops.hide_coercion r with
| Some n when n < nargs && (inctx or n+1 < nargs) ->
(* We skip a coercion *)
- let l = list_skipn n args in
+ let l = List.skipn n args in
let (a,l) = match l with a::l -> (a,l) | [] -> assert false in
(* Recursively remove the head coercions *)
let a' = remove_coercions true a in
@@ -893,17 +893,17 @@ and extern_symbol (tmp_scope,scopes as allscopes) vars t = function
let (t,args,argsscopes,argsimpls) = match t,n with
| GApp (_,f,args), Some n
when List.length args >= n ->
- let args1, args2 = list_chop n args in
+ let args1, args2 = List.chop n args in
let subscopes, impls =
match f with
| GRef (_,ref) ->
let subscopes =
- try list_skipn n (find_arguments_scope ref) with _ -> [] in
+ try List.skipn n (find_arguments_scope ref) with _ -> [] in
let impls =
let impls =
select_impargs_size
(List.length args) (implicits_of_global ref) in
- try list_skipn n impls with _ -> [] in
+ try List.skipn n impls with _ -> [] in
subscopes,impls
| _ ->
[], [] in
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index 86998c210..7ad1327fd 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -193,7 +193,7 @@ let empty_internalization_env = Idmap.empty
let compute_explicitable_implicit imps = function
| Inductive params ->
(* In inductive types, the parameters are fixed implicit arguments *)
- let sub_impl,_ = list_chop (List.length params) imps in
+ let sub_impl,_ = List.chop (List.length params) imps in
let sub_impl' = List.filter is_status_implicit sub_impl in
List.map name_of_implicit sub_impl'
| Recursive | Method | Variable ->
@@ -206,7 +206,7 @@ let compute_internalization_data env ty typ impl =
(ty, expls_impl, impl, compute_arguments_scope typ)
let compute_internalization_env env ty =
- list_fold_left3
+ List.fold_left3
(fun map id typ impl -> Idmap.add id (compute_internalization_data env ty typ impl) map)
empty_internalization_env
@@ -674,7 +674,7 @@ let find_appl_head_data = function
when l <> [] & Flags.version_strictly_greater Flags.V8_2 ->
let n = List.length l in
x,List.map (drop_first_implicits n) (implicits_of_global ref),
- list_skipn_at_least n (find_arguments_scope ref),[]
+ List.skipn_at_least n (find_arguments_scope ref),[]
| x -> x,[],[],[]
let error_not_enough_arguments loc =
@@ -708,7 +708,7 @@ let intern_qualid loc qid intern env lvar args =
let (ids,c) = Syntax_def.search_syntactic_definition sp in
let nids = List.length ids in
if List.length args < nids then error_not_enough_arguments loc;
- let args1,args2 = list_chop nids args in
+ let args1,args2 = List.chop nids args in
check_no_explicitation args1;
let subst = make_subst ids (List.map fst args1) in
subst_aconstr_in_glob_constr loc intern lvar (subst,[],[]) ([],env) c, args2
@@ -766,15 +766,15 @@ let find_remaining_scopes pl1 pl2 ref =
let len_pl2 = List.length pl2 in
let impl_list = if len_pl1 = 0
then select_impargs_size len_pl2 impls_st
- else list_skipn_at_least len_pl1 (select_stronger_impargs impls_st) in
+ else List.skipn_at_least len_pl1 (select_stronger_impargs impls_st) in
let allscs = find_arguments_scope ref in
- let scope_list = list_skipn_at_least len_pl1 allscs in
+ let scope_list = List.skipn_at_least len_pl1 allscs in
let rec aux = function
|[],l -> l
|_,[] -> []
|h::t,_::tt when is_status_implicit h -> aux (t,tt)
|_::t,h::tt -> h :: aux (t,tt)
- in ((try list_firstn len_pl1 allscs with Failure _ -> simple_adjust_scopes len_pl1 allscs),
+ in ((try List.firstn len_pl1 allscs with Failure _ -> simple_adjust_scopes len_pl1 allscs),
simple_adjust_scopes len_pl2 (aux (impl_list,scope_list)))
let product_of_cases_patterns ids idspl =
@@ -800,7 +800,7 @@ let rec has_duplicate = function
| x::l -> if List.mem x l then (Some x) else has_duplicate l
let loc_of_lhs lhs =
- Loc.merge (fst (List.hd lhs)) (fst (list_last lhs))
+ Loc.merge (fst (List.hd lhs)) (fst (List.last lhs))
let check_linearity lhs ids =
match has_duplicate ids with
@@ -815,7 +815,7 @@ let check_number_of_pattern loc n l =
if n<>p then raise (InternalizationError (loc,BadPatternsNumber (n,p)))
let check_or_pat_variables loc ids idsl =
- if List.exists (fun ids' -> not (list_eq_set ids ids')) idsl then
+ if List.exists (fun ids' -> not (List.eq_set ids ids')) idsl then
user_err_loc (loc, "", str
"The components of this disjunctive pattern must bind the same variables.")
@@ -832,7 +832,7 @@ let check_constructor_length env loc cstr len_pl pl0 =
let add_implicits_check_length fail nargs nargs_with_letin impls_st len_pl1 pl2 =
let impl_list = if len_pl1 = 0
then select_impargs_size (List.length pl2) impls_st
- else list_skipn_at_least len_pl1 (select_stronger_impargs impls_st) in
+ else List.skipn_at_least len_pl1 (select_stronger_impargs impls_st) in
let remaining_args = List.fold_left (fun i x -> if is_status_implicit x then i else succ i) in
let rec aux i = function
|[],l -> let args_len = List.length l + List.length impl_list + len_pl1 in
@@ -870,7 +870,7 @@ let chop_params_pattern loc ind args with_letin =
then fst (Inductiveops.inductive_nargs ind)
else Inductiveops.inductive_nparams ind in
assert (nparams <= List.length args);
- let params,args = list_chop nparams args in
+ let params,args = List.chop nparams args in
List.iter (function PatVar(_,Anonymous) -> ()
| PatVar (loc',_) | PatCstr(loc',_,_,_) -> error_parameter_not_implicit loc') params;
args
@@ -883,7 +883,7 @@ let find_constructor loc add_params ref =
|Some nb_args -> let nb = if nb_args = Inductiveops.constructor_nrealhyps c
then fst (Inductiveops.inductive_nargs ind)
else Inductiveops.inductive_nparams ind in
- Util.list_make nb ([],[([],PatVar(Loc.ghost,Anonymous))])
+ Util.List.make nb ([],[([],PatVar(Loc.ghost,Anonymous))])
|None -> []) cstr
let find_pattern_variable = function
@@ -1033,7 +1033,7 @@ let drop_notations_pattern looked_for =
looked_for g;
let nvars = List.length vars in
if List.length pats < nvars then error_not_enough_arguments loc;
- let pats1,pats2 = list_chop nvars pats in
+ let pats1,pats2 = List.chop nvars pats in
let subst = List.map2 (fun (id,scl) a -> (id,(a,scl))) vars pats1 in
let idspl1 = List.map (in_not loc env (subst,[]) []) args in
let (_,argscs) = find_remaining_scopes pats1 pats2 g in
@@ -1283,7 +1283,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
let lf = List.map (fun ((_, id),_,_,_,_) -> id) dl in
let dl = Array.of_list dl in
let n =
- try list_index0 iddef lf
+ try List.index0 iddef lf
with Not_found ->
raise (InternalizationError (locid,UnboundFixName (false,iddef)))
in
@@ -1308,7 +1308,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
in
((n, ro), List.rev rbl, intern_type env' ty, env')) dl in
let idl = array_map2 (fun (_,_,_,_,bd) (a,b,c,env') ->
- let env'' = list_fold_left_i (fun i en name ->
+ let env'' = List.fold_left_i (fun i en name ->
let (_,bli,tyi,_) = idl_temp.(i) in
let fix_args = (List.map (fun (_,(na, bk, _, _)) -> (build_impls bk na)) bli) in
push_name_env lvar (impls_type_list ~args:fix_args tyi)
@@ -1324,7 +1324,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
let lf = List.map (fun ((_, id),_,_,_) -> id) dl in
let dl = Array.of_list dl in
let n =
- try list_index0 iddef lf
+ try List.index0 iddef lf
with Not_found ->
raise (InternalizationError (locid,UnboundFixName (true,iddef)))
in
@@ -1335,7 +1335,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
(List.rev rbl,
intern_type env' ty,env')) dl in
let idl = array_map2 (fun (_,_,_,bd) (b,c,env') ->
- let env'' = list_fold_left_i (fun i en name ->
+ let env'' = List.fold_left_i (fun i en name ->
let (bli,tyi,_) = idl_tmp.(i) in
let cofix_args = List.map (fun (_, (na, bk, _, _)) -> (build_impls bk na)) bli in
push_name_env lvar (impls_type_list ~args:cofix_args tyi)
@@ -1407,7 +1407,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
match cargs with
| None -> user_err_loc (loc, "intern", str"No constructor inference.")
| Some (n, constrname, args) ->
- let pars = list_make n (CHole (loc, None)) in
+ let pars = List.make n (CHole (loc, None)) in
let app = CAppExpl (loc, (None, constrname), List.rev_append pars args) in
intern env app
end
@@ -1440,7 +1440,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
List.map (fun id -> GVar (Loc.ghost,id),(Name id,None)) thevars, (* "match v1,..,vn" *)
[Loc.ghost,[],thepats, (* "|p1,..,pn" *)
Option.cata (intern_type env') (GHole(Loc.ghost,Evar_kinds.CasesType)) rtnpo; (* "=> P" is there were a P "=> _" else *)
- Loc.ghost,[],list_make (List.length thepats) (PatVar(Loc.ghost,Anonymous)), (* "|_,..,_" *)
+ Loc.ghost,[],List.make (List.length thepats) (PatVar(Loc.ghost,Anonymous)), (* "|_,..,_" *)
GHole(Loc.ghost,Evar_kinds.ImpossibleCase) (* "=> _" *)]))
in
let eqns' = List.map (intern_eqn (List.length tms) env) eqns in
@@ -1553,7 +1553,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
((fresh,c)::match_acc) ((cases_pattern_loc c,Name fresh)::var_acc)
|_ -> assert false in
let _,args_rel =
- list_chop nparams (List.rev mip.Declarations.mind_arity_ctxt) in
+ List.chop nparams (List.rev mip.Declarations.mind_arity_ctxt) in
canonize_args args_rel l (Idset.elements forbidden_names_for_gen) [] [] in
match_to_do, Some (cases_pattern_expr_loc t,ind,List.rev_map snd nal)
| None ->
diff --git a/interp/coqlib.ml b/interp/coqlib.ml
index 94cb91749..531ca5bf4 100644
--- a/interp/coqlib.ml
+++ b/interp/coqlib.ml
@@ -45,7 +45,7 @@ let gen_constant_in_modules locstr dirs s =
let dirs = List.map make_dir dirs in
let qualid = qualid_of_string s in
let all = Nametab.locate_extended_all qualid in
- let all = list_uniquize (list_map_filter global_of_extended all) in
+ let all = List.uniquize (List.map_filter global_of_extended all) in
let these = List.filter (has_suffix_in_dirs dirs) all in
match these with
| [x] -> constr_of_global x
@@ -73,7 +73,7 @@ let check_required_library d =
if not (Library.library_is_loaded dir) then
if not current_dir then
(* Loading silently ...
- let m, prefix = list_sep_last d' in
+ let m, prefix = List.sep_last d' in
read_library
(Loc.ghost,make_qualid (make_dirpath (List.rev prefix)) m)
*)
diff --git a/interp/dumpglob.ml b/interp/dumpglob.ml
index c9e18c15e..101645dfc 100644
--- a/interp/dumpglob.ml
+++ b/interp/dumpglob.ml
@@ -174,7 +174,7 @@ let dump_constraint ((loc, n), _, _) sec ty =
let dump_modref loc mp ty =
if dump () then
let (dp, l) = Lib.split_modpath mp in
- let l = if l = [] then l else Util.list_drop_last l in
+ let l = if l = [] then l else Util.List.drop_last l in
let fp = Names.string_of_dirpath dp in
let mp = Names.string_of_dirpath (Names.make_dirpath l) in
let bl,el = interval loc in
diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml
index 3057cef34..c4cc38a56 100644
--- a/interp/implicit_quantifiers.ml
+++ b/interp/implicit_quantifiers.ml
@@ -316,6 +316,6 @@ let implicits_of_glob_constr ?(with_products=true) l =
| GLetIn (loc, na, t, b) -> aux i b
| GRec (_, fix_kind, nas, args, tys, bds) ->
let nb = match fix_kind with |GFix (_, n) -> n | GCoFix n -> n in
- list_fold_left_i (fun i l (na,bk,_,_) -> add_impl i na bk l) i (aux (List.length args.(nb) + i) bds.(nb)) args.(nb)
+ List.fold_left_i (fun i l (na,bk,_,_) -> add_impl i na bk l) i (aux (List.length args.(nb) + i) bds.(nb)) args.(nb)
| _ -> []
in aux 1 l
diff --git a/interp/modintern.ml b/interp/modintern.ml
index 1e7d84c96..ba5c68a5d 100644
--- a/interp/modintern.ml
+++ b/interp/modintern.ml
@@ -77,7 +77,7 @@ let lookup_qualid (modtype:bool) qid =
in
let rec find_module_prefix dir n =
if n<0 then raise Not_found;
- let dir',dir'' = list_chop n dir in
+ let dir',dir'' = List.chop n dir in
let id',dir''' =
match dir'' with
| hd::tl -> hd,tl
diff --git a/interp/notation.ml b/interp/notation.ml
index 6a574a56e..46a06b700 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -125,7 +125,7 @@ let open_scope i (_,(local,op,sc)) =
| _ -> sc
in
scope_stack :=
- if op then sc :: !scope_stack else list_except sc !scope_stack
+ if op then sc :: !scope_stack else List.except sc !scope_stack
let cache_scope o =
open_scope 1 o
@@ -285,7 +285,7 @@ let check_required_module loc sc (sp,d) =
with Not_found ->
user_err_loc (loc,"prim_token_interpreter",
str ("Cannot interpret in "^sc^" without requiring first module "
- ^(list_last d)^"."))
+ ^(List.last d)^"."))
(* Look if some notation or numeral printer in [scope] can be used in
the scope stack [scopes], and if yes, using delimiters or not *)
@@ -402,7 +402,7 @@ let interp_notation loc ntn local_scopes =
let isGApp = function GApp _ -> true | _ -> false
let uninterp_notations c =
- list_map_append (fun key -> Gmapl.find key !notations_key_table)
+ List.map_append (fun key -> Gmapl.find key !notations_key_table)
(glob_constr_keys c)
let uninterp_cases_pattern_notations c =
@@ -554,7 +554,7 @@ let subst_arguments_scope (subst,(req,r,scl,cls)) =
match subst_scope_class subst cl with
| Some cl' as ocl' when cl' != cl -> ocl'
| _ -> ocl in
- let cls' = list_smartmap subst_cl cls in
+ let cls' = List.smartmap subst_cl cls in
let scl' = merge_scope (List.map find_scope_class_opt cls') scl in
let scl'' = List.map (Option.map Declaremods.subst_scope) scl' in
(ArgsScopeNoDischarge,r',scl'',cls')
@@ -576,7 +576,7 @@ let rebuild_arguments_scope (req,r,l,_) =
(* Add to the manually given scopes the one found automatically
for the extra parameters of the section *)
let l',cls = compute_arguments_scope_full (Global.type_of_global r) in
- let l1,_ = list_chop (List.length l' - List.length l) l' in
+ let l1,_ = List.chop (List.length l' - List.length l) l' in
(req,r,l1@l,cls)
type arguments_scope_obj =
diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml
index bd3ba4274..13356445f 100644
--- a/interp/notation_ops.ml
+++ b/interp/notation_ops.ml
@@ -33,7 +33,7 @@ let rec cases_pattern_fold_map loc g e = function
let e',na' = g e na in e', PatVar (loc,na')
| PatCstr (_,cstr,patl,na) ->
let e',na' = g e na in
- let e',patl' = list_fold_map (cases_pattern_fold_map loc g) e patl in
+ let e',patl' = List.fold_map (cases_pattern_fold_map loc g) e patl in
e', PatCstr (loc,cstr,patl',na')
let rec subst_glob_vars l = function
@@ -86,18 +86,18 @@ let glob_constr_of_notation_constr_with_binders loc g f e = function
let fold (idl,e) na = let (e,na) = g e na in ((name_cons na idl,e),na) in
let eqnl' = List.map (fun (patl,rhs) ->
let ((idl,e),patl) =
- list_fold_map (cases_pattern_fold_map loc fold) ([],e) patl in
+ List.fold_map (cases_pattern_fold_map loc fold) ([],e) patl in
(loc,idl,patl,f e rhs)) eqnl in
GCases (loc,sty,Option.map (f e') rtntypopt,tml',eqnl')
| NLetTuple (nal,(na,po),b,c) ->
- let e',nal = list_fold_map g e nal in
+ let e',nal = List.fold_map g e nal in
let e'',na = g e na in
GLetTuple (loc,nal,(na,Option.map (f e'') po),f e b,f e' c)
| NIf (c,(na,po),b1,b2) ->
let e',na = g e na in
GIf (loc,f e c,(na,Option.map (f e') po),f e b1,f e b2)
| NRec (fk,idl,dll,tl,bl) ->
- let e,dll = array_fold_map (list_fold_map (fun e (na,oc,b) ->
+ let e,dll = array_fold_map (List.fold_map (fun e (na,oc,b) ->
let e,na = g e na in
(e,(na,Explicit,Option.map (f e) oc,f e b)))) e dll in
let e',idl = array_fold_map (to_id g) e idl in
@@ -144,7 +144,7 @@ let on_true_do b f c = if b then (f c; b) else b
let compare_glob_constr f add t1 t2 = match t1,t2 with
| GRef (_,r1), GRef (_,r2) -> eq_gr r1 r2
| GVar (_,v1), GVar (_,v2) -> on_true_do (v1 = v2) add (Name v1)
- | GApp (_,f1,l1), GApp (_,f2,l2) -> f f1 f2 & list_for_all2eq f l1 l2
+ | GApp (_,f1,l1), GApp (_,f2,l2) -> f f1 f2 & List.for_all2eq f l1 l2
| GLambda (_,na1,bk1,ty1,c1), GLambda (_,na2,bk2,ty2,c2) when na1 = na2 && bk1 = bk2 -> on_true_do (f ty1 ty2 & f c1 c2) add na1
| GProd (_,na1,bk1,ty1,c1), GProd (_,na2,bk2,ty2,c2) when na1 = na2 && bk1 = bk2 ->
on_true_do (f ty1 ty2 & f c1 c2) add na1
@@ -182,7 +182,7 @@ let compare_recursive_parts found f (iterator,subc) =
(* We found the pattern, but there are extra arguments *)
(* (this allows e.g. alternative (recursive) notation of application) *)
assert (!terminator = None); terminator := Some term;
- list_for_all2eq aux l1 l2
+ List.for_all2eq aux l1 l2
| GVar (_,x), GVar (_,y) when x<>y ->
(* We found the position where it differs *)
let lassoc = (!terminator <> None) in
@@ -335,7 +335,7 @@ let rec subst_pat subst pat =
| PatVar _ -> pat
| PatCstr (loc,((kn,i),j),cpl,n) ->
let kn' = subst_ind subst kn
- and cpl' = list_smartmap (subst_pat subst) cpl in
+ and cpl' = List.smartmap (subst_pat subst) cpl in
if kn' == kn && cpl' == cpl then pat else
PatCstr (loc,((kn',i),j),cpl',n)
@@ -350,7 +350,7 @@ let rec subst_notation_constr subst bound raw =
| NApp (r,rl) ->
let r' = subst_notation_constr subst bound r
- and rl' = list_smartmap (subst_notation_constr subst bound) rl in
+ and rl' = List.smartmap (subst_notation_constr subst bound) rl in
if r' == r && rl' == rl then raw else
NApp(r',rl')
@@ -386,7 +386,7 @@ let rec subst_notation_constr subst bound raw =
| NCases (sty,rtntypopt,rl,branches) ->
let rtntypopt' = Option.smartmap (subst_notation_constr subst bound) rtntypopt
- and rl' = list_smartmap
+ and rl' = List.smartmap
(fun (a,(n,signopt) as x) ->
let a' = subst_notation_constr subst bound a in
let signopt' = Option.map (fun ((indkn,i),nal as z) ->
@@ -394,9 +394,9 @@ let rec subst_notation_constr subst bound raw =
if indkn == indkn' then z else ((indkn',i),nal)) signopt in
if a' == a && signopt' == signopt then x else (a',(n,signopt')))
rl
- and branches' = list_smartmap
+ and branches' = List.smartmap
(fun (cpl,r as branch) ->
- let cpl' = list_smartmap (subst_pat subst) cpl
+ let cpl' = List.smartmap (subst_pat subst) cpl
and r' = subst_notation_constr subst bound r in
if cpl' == cpl && r' == r then branch else
(cpl',r'))
@@ -423,7 +423,7 @@ let rec subst_notation_constr subst bound raw =
| NRec (fk,idl,dll,tl,bl) ->
let dll' =
- array_smartmap (list_smartmap (fun (na,oc,b as x) ->
+ array_smartmap (List.smartmap (fun (na,oc,b as x) ->
let oc' = Option.smartmap (subst_notation_constr subst bound) oc in
let b' = subst_notation_constr subst bound b in
if oc' == oc && b' == b then x else (na,oc',b'))) dll in
@@ -619,9 +619,9 @@ let rec match_ inner u alp (tmetas,blmetas as metas) sigma a1 a2 =
let n1 = List.length l1 and n2 = List.length l2 in
let f1,l1,f2,l2 =
if n1 < n2 then
- let l21,l22 = list_chop (n2-n1) l2 in f1,l1, NApp (f2,l21), l22
+ let l21,l22 = List.chop (n2-n1) l2 in f1,l1, NApp (f2,l21), l22
else if n1 > n2 then
- let l11,l12 = list_chop (n1-n2) l1 in GApp (loc,f1,l11),l12, f2,l2
+ let l11,l12 = List.chop (n1-n2) l1 in GApp (loc,f1,l11),l12, f2,l2
else f1,l1, f2, l2 in
let may_use_eta = does_not_come_from_already_eta_expanded_var f1 in
List.fold_left2 (match_ may_use_eta u alp metas)
@@ -710,7 +710,7 @@ and match_equations u alp metas sigma (_,_,patl1,rhs1) (patl2,rhs2) =
match_in u alp metas sigma rhs1 rhs2
let match_notation_constr u c (metas,pat) =
- let vars = list_split_by (fun (_,(_,x)) -> x <> NtnTypeBinderList) metas in
+ let vars = List.split_by (fun (_,(_,x)) -> x <> NtnTypeBinderList) metas in
let vars = (List.map fst (fst vars), List.map fst (snd vars)) in
let terms,termlists,binders = match_ false u [] vars ([],[],[]) c pat in
(* Reorder canonically the substitution *)
@@ -734,7 +734,7 @@ let match_notation_constr u c (metas,pat) =
let add_patterns_for_params ind l =
let mib,_ = Global.lookup_inductive ind in
let nparams = mib.Declarations.mind_nparams in
- Util.list_addn nparams (PatVar (Loc.ghost,Anonymous)) l
+ Util.List.addn nparams (PatVar (Loc.ghost,Anonymous)) l
let bind_env_cases_pattern (sigma,sigmalist,x as fullsigma) var v =
try
@@ -758,7 +758,7 @@ let rec match_cases_pattern metas sigma a1 a2 =
then
raise No_match
else
- let l1',more_args = Util.list_chop le2 l1 in
+ let l1',more_args = Util.List.chop le2 l1 in
(List.fold_left2 (match_cases_pattern_no_more_args metas) sigma l1' l2),(le2,more_args)
| r1, NList (x,_,iter,termin,lassoc) ->
(match_alist (fun (metas,_) -> match_cases_pattern_no_more_args metas)
@@ -781,7 +781,7 @@ let match_ind_pattern metas sigma ind pats a2 =
then
raise No_match
else
- let l1',more_args = Util.list_chop le2 pats in
+ let l1',more_args = Util.List.chop le2 pats in
(List.fold_left2 (match_cases_pattern_no_more_args metas) sigma l1' l2),(le2,more_args)
|_ -> raise No_match
diff --git a/interp/reserve.ml b/interp/reserve.ml
index 259ffa171..7a2d4bfe7 100644
--- a/interp/reserve.ml
+++ b/interp/reserve.ml
@@ -77,7 +77,7 @@ let revert_reserved_type t =
try
let l = Gmapl.find (constr_key t) !reserve_revtable in
let t = Detyping.detype false [] [] t in
- list_try_find
+ List.try_find
(fun (pat,id) ->
try let _ = Notation_ops.match_notation_constr false t ([],pat) in Name id
with Notation_ops.No_match -> failwith "") l
diff --git a/interp/topconstr.ml b/interp/topconstr.ml
index e717e1747..abdde1f54 100644
--- a/interp/topconstr.ml
+++ b/interp/topconstr.ml
@@ -158,7 +158,7 @@ let split_at_annot bl na =
| Some (loc, id) ->
let rec aux acc = function
| LocalRawAssum (bls, k, t) as x :: rest ->
- let l, r = list_split_when (fun (loc, na) -> na = Name id) bls in
+ let l, r = List.split_when (fun (loc, na) -> na = Name id) bls in
if r = [] then aux (x :: acc) rest
else
(List.rev (if l = [] then acc else LocalRawAssum (l, k, t) :: acc),
diff --git a/kernel/closure.ml b/kernel/closure.ml
index 94980b596..210dbb02b 100644
--- a/kernel/closure.ml
+++ b/kernel/closure.ml
@@ -455,7 +455,7 @@ let rec compact_constr (lg, subs as s) c k =
match kind_of_term c with
Rel i ->
if i < k then c,s else
- (try mkRel (k + lg - list_index (i-k+1) subs), (lg,subs)
+ (try mkRel (k + lg - List.index (i-k+1) subs), (lg,subs)
with Not_found -> mkRel (k+lg), (lg+1, (i-k+1)::subs))
| (Sort _|Var _|Meta _|Ind _|Const _|Construct _) -> c,s
| Evar(ev,v) ->
@@ -734,7 +734,7 @@ let rec get_args n tys f e stk =
let eargs = Array.sub l n (na-n) in
(Inl (subs_cons(args,e)), Zapp eargs :: s)
else (* more lambdas *)
- let etys = list_skipn na tys in
+ let etys = List.skipn na tys in
get_args (n-na) etys f (subs_cons(l,e)) s
| _ -> (Inr {norm=Cstr;term=FLambda(n,tys,f,e)}, stk)
diff --git a/kernel/declarations.ml b/kernel/declarations.ml
index 5f677d056..a2ce4f270 100644
--- a/kernel/declarations.ml
+++ b/kernel/declarations.ml
@@ -110,7 +110,7 @@ let subst_rel_declaration sub (id,copt,t as x) =
let t' = subst_mps sub t in
if copt == copt' & t == t' then x else (id,copt',t')
-let subst_rel_context sub = list_smartmap (subst_rel_declaration sub)
+let subst_rel_context sub = List.smartmap (subst_rel_declaration sub)
(* TODO: these substitution functions could avoid duplicating things
when the substitution have preserved all the fields *)
@@ -141,11 +141,11 @@ let hcons_rel_decl ((n,oc,t) as d) =
and t' = hcons_types t
in if n' == n && oc' == oc && t' == t then d else (n',oc',t')
-let hcons_rel_context l = list_smartmap hcons_rel_decl l
+let hcons_rel_context l = List.smartmap hcons_rel_decl l
let hcons_polyarity ar =
{ poly_param_levels =
- list_smartmap (Option.smartmap hcons_univ) ar.poly_param_levels;
+ List.smartmap (Option.smartmap hcons_univ) ar.poly_param_levels;
poly_level = hcons_univ ar.poly_level }
let hcons_const_type = function
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index 6c7f4408f..b93b9f19b 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -351,7 +351,7 @@ let check_correct_par (env,n,ntypes,_) hyps l largs =
let compute_rec_par (env,n,_,_) hyps nmr largs =
if nmr = 0 then 0 else
(* start from 0, hyps will be in reverse order *)
- let (lpar,_) = list_chop nmr largs in
+ let (lpar,_) = List.chop nmr largs in
let rec find k index =
function
([],_) -> nmr
@@ -375,7 +375,7 @@ let abstract_mind_lc env ntyps npars lc =
lc
else
let make_abs =
- list_tabulate
+ List.tabulate
(function i -> lambda_implicit_lift npars (mkRel (i+1))) ntyps
in
Array.map (substl make_abs) lc
@@ -457,7 +457,7 @@ let check_positivity_one (env,_,ntypes,_ as ienv) hyps (_,i as ind) nargs lcname
let auxnpar = mib.mind_nparams_rec in
let nonrecpar = mib.mind_nparams - auxnpar in
let (lpar,auxlargs) =
- try list_chop auxnpar largs
+ try List.chop auxnpar largs
with Failure _ -> raise (IllFormedInd (LocalNonPos n)) in
(* If the inductive appears in the args (non params) then the
definition is not positive. *)
@@ -539,7 +539,7 @@ let check_positivity kn env_ar params inds =
let nmr = rel_context_nhyps params in
let check_one i (_,lcnames,lc,(sign,_)) =
let ra_env =
- list_tabulate (fun _ -> (Norec,mk_norec)) lparams @ lra_ind in
+ List.tabulate (fun _ -> (Norec,mk_norec)) lparams @ lra_ind in
let ienv = (env_ar, 1+lparams, ntypes, ra_env) in
let nargs = rel_context_nhyps sign - nmr in
check_positivity_one ienv params (kn,i) nargs lcnames lc
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index 58689a926..d86abb435 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -54,7 +54,7 @@ let inductive_params (mib,_) = mib.mind_nparams
let ind_subst mind mib =
let ntypes = mib.mind_ntypes in
let make_Ik k = mkInd (mind,ntypes-k-1) in
- list_tabulate make_Ik ntypes
+ List.tabulate make_Ik ntypes
(* Instantiate inductives in constructor type *)
let constructor_instantiate mind mib c =
@@ -272,7 +272,7 @@ let extended_rel_list n hyps =
reln [] 1 hyps
let build_dependent_inductive ind (_,mip) params =
- let realargs,_ = list_chop mip.mind_nrealargs_ctxt mip.mind_arity_ctxt in
+ let realargs,_ = List.chop mip.mind_nrealargs_ctxt mip.mind_arity_ctxt in
applist
(mkInd ind,
List.map (lift mip.mind_nrealargs_ctxt) params
@@ -327,7 +327,7 @@ let build_branches_type ind (_,mip as specif) params p =
let (args,ccl) = decompose_prod_assum typi in
let nargs = rel_context_length args in
let (_,allargs) = decompose_app ccl in
- let (lparams,vargs) = list_chop (inductive_params specif) allargs in
+ let (lparams,vargs) = List.chop (inductive_params specif) allargs in
let cargs =
let cstr = ith_constructor_of_inductive ind (i+1) in
let dep_cstr = applist (mkConstruct cstr,lparams@(local_rels args)) in
@@ -344,7 +344,7 @@ let build_case_type n p c realargs =
let type_case_branches env (ind,largs) pj c =
let specif = lookup_mind_specif env ind in
let nparams = inductive_params specif in
- let (params,realargs) = list_chop nparams largs in
+ let (params,realargs) = List.chop nparams largs in
let p = pj.uj_val in
let univ = is_correct_arity env c pj ind specif params in
let lc = build_branches_type ind specif params p in
@@ -451,7 +451,7 @@ let push_var renv (x,ty,spec) =
genv = spec:: renv.genv }
let assign_var_spec renv (i,spec) =
- { renv with genv = list_assign renv.genv (i-1) spec }
+ { renv with genv = List.assign renv.genv (i-1) spec }
let push_var_renv renv (x,ty) =
push_var renv (x,ty,lazy Not_subterm)
@@ -521,7 +521,7 @@ let branches_specif renv c_spec ci =
vra
| Dead_code -> Array.create nca Dead_code
| _ -> Array.create nca Not_subterm) in
- list_tabulate (fun j -> lazy (Lazy.force lvra).(j)) nca)
+ List.tabulate (fun j -> lazy (Lazy.force lvra).(j)) nca)
car
(* [subterm_specif renv t] computes the recursive structure of [t] and
@@ -870,7 +870,7 @@ let check_one_cofix env nbfix def deftype =
let lra = vlra.(i-1) in
let mI = inductive_of_constructor cstr_kn in
let (mib,mip) = lookup_mind_specif env mI in
- let realargs = list_skipn mib.mind_nparams args in
+ let realargs = List.skipn mib.mind_nparams args in
let rec process_args_of_constr = function
| (t::lr), (rar::lrar) ->
if rar = mk_norec then
diff --git a/kernel/modops.ml b/kernel/modops.ml
index 5eddb6c84..19075058a 100644
--- a/kernel/modops.ml
+++ b/kernel/modops.ml
@@ -562,7 +562,7 @@ and clean_expr l = function
if str_clean == str && sig_clean = sigt.typ_expr then
s else SEBfunctor (mbid,{sigt with typ_expr=sig_clean},str_clean)
| SEBstruct str as s->
- let str_clean = Util.list_smartmap (clean_struct l) str in
+ let str_clean = Util.List.smartmap (clean_struct l) str in
if str_clean == str then s else SEBstruct(str_clean)
| str -> str
@@ -572,7 +572,7 @@ let rec collect_mbid l = function
if str_clean == str then s else
SEBfunctor (mbid,sigt,str_clean)
| SEBstruct str as s->
- let str_clean = Util.list_smartmap (clean_struct l) str in
+ let str_clean = Util.List.smartmap (clean_struct l) str in
if str_clean == str then s else SEBstruct(str_clean)
| _ -> anomaly "Modops:the evaluation of the structure failed "
diff --git a/kernel/names.ml b/kernel/names.ml
index d7b9516ae..aa5313842 100644
--- a/kernel/names.ml
+++ b/kernel/names.ml
@@ -369,7 +369,7 @@ module Hdir = Hashcons.Make(
struct
type t = dir_path
type u = identifier -> identifier
- let hash_sub hident d = list_smartmap hident d
+ let hash_sub hident d = List.smartmap hident d
let rec equal d1 d2 =
(d1==d2) ||
match (d1,d2) with
diff --git a/kernel/pre_env.ml b/kernel/pre_env.ml
index 3d2f19aac..476a92bed 100644
--- a/kernel/pre_env.ml
+++ b/kernel/pre_env.ml
@@ -93,8 +93,8 @@ let lookup_rel_val n env =
let env_of_rel n env =
{ env with
- env_rel_context = Util.list_skipn n env.env_rel_context;
- env_rel_val = Util.list_skipn n env.env_rel_val;
+ env_rel_context = Util.List.skipn n env.env_rel_context;
+ env_rel_val = Util.List.skipn n env.env_rel_val;
env_nb_rel = env.env_nb_rel - n
}
diff --git a/kernel/sign.ml b/kernel/sign.ml
index a654bc04d..269f7a5d9 100644
--- a/kernel/sign.ml
+++ b/kernel/sign.ml
@@ -34,7 +34,7 @@ let rec lookup_named id = function
| [] -> raise Not_found
let named_context_length = List.length
-let named_context_equal = list_equal eq_named_declaration
+let named_context_equal = List.equal eq_named_declaration
let vars_of_named_context = List.map (fun (id,_,_) -> id)
@@ -61,7 +61,7 @@ let map_context f l =
if body_o' == body_o && typ' == typ then decl else
(n, body_o', typ')
in
- list_smartmap map_decl l
+ List.smartmap map_decl l
let map_rel_context = map_context
let map_named_context = map_context
diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml
index 5fec0c2c7..c590a5101 100644
--- a/kernel/subtyping.ml
+++ b/kernel/subtyping.ml
@@ -198,7 +198,7 @@ let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2
check (fun mib ->
let nparamdecls = List.length mib.mind_params_ctxt in
let names = names_prod_letin (mib.mind_packets.(0).mind_user_lc.(0)) in
- snd (list_chop nparamdecls names))
+ snd (List.chop nparamdecls names))
(fun x -> RecordProjectionsExpected x);
end;
(* we first check simple things *)
diff --git a/kernel/univ.ml b/kernel/univ.ml
index 1f864926a..9334a06d1 100644
--- a/kernel/univ.ml
+++ b/kernel/univ.ml
@@ -130,12 +130,12 @@ let sup u v =
if UniverseLevel.equal u v then Atom u else Max ([u;v],[])
| u, Max ([],[]) -> u
| Max ([],[]), v -> v
- | Atom u, Max (gel,gtl) -> Max (list_add_set u gel,gtl)
- | Max (gel,gtl), Atom v -> Max (list_add_set v gel,gtl)
+ | Atom u, Max (gel,gtl) -> Max (List.add_set u gel,gtl)
+ | Max (gel,gtl), Atom v -> Max (List.add_set v gel,gtl)
| Max (gel,gtl), Max (gel',gtl') ->
- let gel'' = list_union gel gel' in
- let gtl'' = list_union gtl gtl' in
- Max (list_subtract gel'' gtl'',gtl'')
+ let gel'' = List.union gel gel' in
+ let gtl'' = List.union gtl gtl' in
+ Max (List.subtract gel'' gtl'',gtl'')
(* Comparison on this type is pointer equality *)
type canonical_arc =
@@ -423,7 +423,7 @@ let merge g arcu arcv =
(* redirected to it *)
let redirect (g,w,w') arcv =
let g' = enter_equiv_arc arcv.univ arcu.univ g in
- (g',list_unionq arcv.lt w,arcv.le@w')
+ (g',List.unionq arcv.lt w,arcv.le@w')
in
let (g',w,w') = List.fold_left redirect (g,[],[]) v in
let g_arcu = (g',arcu) in
@@ -759,7 +759,7 @@ let make_max = function
let remove_large_constraint u = function
| Atom u' as x -> if u = u' then Max ([],[]) else x
- | Max (le,lt) -> make_max (list_remove u le,lt)
+ | Max (le,lt) -> make_max (List.remove u le,lt)
let is_direct_constraint u = function
| Atom u' -> u = u'
@@ -900,8 +900,8 @@ module Huniv =
match u, v with
| Atom u, Atom v -> u == v
| Max (gel,gtl), Max (gel',gtl') ->
- (list_for_all2eq (==) gel gel') &&
- (list_for_all2eq (==) gtl gtl')
+ (List.for_all2eq (==) gel gel') &&
+ (List.for_all2eq (==) gtl gtl')
| _ -> false
let hash = Hashtbl.hash
end)
@@ -928,7 +928,7 @@ module Hconstraints =
let hash_sub huc s =
Constraint.fold (fun x -> Constraint.add (huc x)) s Constraint.empty
let equal s s' =
- list_for_all2eq (==)
+ List.for_all2eq (==)
(Constraint.elements s)
(Constraint.elements s')
let hash = Hashtbl.hash
diff --git a/lib/cList.ml b/lib/cList.ml
new file mode 100644
index 000000000..e3f35dcac
--- /dev/null
+++ b/lib/cList.ml
@@ -0,0 +1,718 @@
+(***********************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
+(* \VV/ *************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(***********************************************************************)
+
+module type S =
+sig
+ val length : 'a list -> int
+ val hd : 'a list -> 'a
+ val tl : 'a list -> 'a list
+ val nth : 'a list -> int -> 'a
+ val rev : 'a list -> '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 iter : ('a -> unit) -> 'a list -> unit
+ val map : ('a -> 'b) -> 'a list -> 'b list
+ val rev_map : ('a -> 'b) -> 'a list -> 'b 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 filter : ('a -> bool) -> 'a list -> 'a list
+ val find_all : ('a -> bool) -> 'a list -> 'a list
+ val partition : ('a -> bool) -> 'a list -> 'a list * 'a list
+ val assoc : 'a -> ('a * 'b) list -> 'b
+ val assq : 'a -> ('a * 'b) list -> 'b
+ 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 merge : ('a -> 'a -> int) -> 'a list -> 'a list -> 'a list
+end
+
+module type ExtS =
+sig
+ include S
+ val compare : ('a -> 'a -> int) -> 'a list -> 'a list -> int
+ val equal : ('a -> 'a -> bool) -> 'a list -> 'a list -> bool
+ val add_set : 'a -> 'a list -> 'a list
+ val eq_set : 'a list -> 'a list -> bool
+ val intersect : 'a list -> 'a list -> 'a list
+ val union : 'a list -> 'a list -> 'a list
+ val unionq : 'a list -> 'a list -> 'a list
+ val subtract : 'a list -> 'a list -> 'a list
+ val subtractq : 'a list -> 'a list -> 'a list
+
+ (** [tabulate f n] builds [[f 0; ...; f (n-1)]] *)
+ val tabulate : (int -> 'a) -> int -> 'a list
+ val make : int -> 'a -> 'a list
+ val assign : 'a list -> int -> 'a -> 'a list
+ val distinct : 'a list -> bool
+ val duplicates : 'a list -> 'a list
+ val filter2 : ('a -> 'b -> bool) -> 'a list * 'b list -> 'a list * 'b list
+ val map_filter : ('a -> 'b option) -> 'a list -> 'b list
+ val map_filter_i : (int -> 'a -> 'b option) -> 'a list -> 'b list
+ val filter_with : bool list -> 'a list -> 'a list
+ val filter_along : ('a -> bool) -> 'a list -> 'b list -> 'b list
+
+ (** [smartmap f [a1...an] = List.map f [a1...an]] but if for all i
+ [ f ai == ai], then [smartmap f l==l] *)
+ val smartmap : ('a -> 'a) -> 'a list -> 'a list
+ val map_left : ('a -> 'b) -> 'a list -> 'b list
+ val map_i : (int -> 'a -> 'b) -> int -> 'a list -> 'b list
+ val map2_i :
+ (int -> 'a -> 'b -> 'c) -> int -> 'a list -> 'b list -> 'c list
+ val map3 :
+ ('a -> 'b -> 'c -> 'd) -> 'a list -> 'b list -> 'c list -> 'd list
+ val map4 :
+ ('a -> 'b -> 'c -> 'd -> 'e) -> 'a list -> 'b list -> 'c list -> 'd list -> 'e list
+ val map_to_array : ('a -> 'b) -> 'a list -> 'b array
+ val filter_i :
+ (int -> 'a -> bool) -> 'a list -> 'a list
+
+ (** [smartfilter f [a1...an] = List.filter f [a1...an]] but if for all i
+ [f ai = true], then [smartfilter f l==l] *)
+ val smartfilter : ('a -> bool) -> 'a list -> 'a list
+
+ (** [index] returns the 1st index of an element in a list (counting from 1) *)
+ val index : 'a -> 'a list -> int
+ val index_f : ('a -> 'a -> bool) -> 'a -> 'a list -> int
+
+ (** [unique_index x l] returns [Not_found] if [x] doesn't occur exactly once *)
+ val unique_index : 'a -> 'a list -> int
+
+ (** [index0] behaves as [index] except that it starts counting at 0 *)
+ val index0 : 'a -> 'a list -> int
+ val index0_f : ('a -> 'a -> bool) -> 'a -> 'a list -> int
+ val iter3 : ('a -> 'b -> 'c -> unit) -> 'a list -> 'b list -> 'c list -> unit
+ val iter_i : (int -> 'a -> unit) -> 'a list -> unit
+ val fold_right_i : (int -> 'a -> 'b -> 'b) -> int -> 'a list -> 'b -> 'b
+ val fold_left_i : (int -> 'a -> 'b -> 'a) -> int -> 'a -> 'b list -> 'a
+ val fold_right_and_left :
+ ('a -> 'b -> 'b list -> 'a) -> 'b list -> 'a -> 'a
+ val fold_left3 : ('a -> 'b -> 'c -> 'd -> 'a) -> 'a -> 'b list -> 'c list -> 'd list -> 'a
+ val for_all_i : (int -> 'a -> bool) -> int -> 'a list -> bool
+ val except : 'a -> 'a list -> 'a list
+ val remove : 'a -> 'a list -> 'a list
+ val remove_first : 'a -> 'a list -> 'a list
+ val remove_assoc_in_triple : 'a -> ('a * 'b * 'c) list -> ('a * 'b * 'c) list
+ val assoc_snd_in_triple : 'a -> ('a * 'b * 'c) list -> 'b
+ val for_all2eq : ('a -> 'b -> bool) -> 'a list -> 'b list -> bool
+ val sep_last : 'a list -> 'a * 'a list
+ val try_find_i : (int -> 'a -> 'b) -> int -> 'a list -> 'b
+ val try_find : ('a -> 'b) -> 'a list -> 'b
+ val uniquize : 'a list -> 'a list
+
+ (** merges two sorted lists and preserves the uniqueness property: *)
+ val merge_uniq : ('a -> 'a -> int) -> 'a list -> 'a list -> 'a list
+ val subset : 'a list -> 'a list -> bool
+ val chop : int -> 'a list -> 'a list * 'a list
+ (* former [split_at] was a duplicate of [chop] *)
+ val split_when : ('a -> bool) -> 'a list -> 'a list * 'a list
+ val split_by : ('a -> bool) -> 'a list -> 'a list * 'a list
+ val split3 : ('a * 'b * 'c) list -> 'a list * 'b list * 'c list
+ val partition_by : ('a -> 'a -> bool) -> 'a list -> 'a list list
+ val firstn : int -> 'a list -> 'a list
+ val last : 'a list -> 'a
+ val lastn : int -> 'a list -> 'a list
+ val skipn : int -> 'a list -> 'a list
+ val skipn_at_least : int -> 'a list -> 'a list
+ val addn : int -> 'a -> 'a list -> 'a list
+ val prefix_of : 'a list -> 'a list -> bool
+
+ (** [drop_prefix p l] returns [t] if [l=p++t] else return [l] *)
+ val drop_prefix : 'a list -> 'a list -> 'a list
+ val drop_last : 'a list -> 'a list
+
+ (** [map_append f [x1; ...; xn]] returns [(f x1)@(f x2)@...@(f xn)] *)
+ val map_append : ('a -> 'b list) -> 'a list -> 'b list
+ val join_map : ('a -> 'b list) -> 'a list -> 'b list
+
+ (** raises [Invalid_argument] if the two lists don't have the same length *)
+ val map_append2 : ('a -> 'b -> 'c list) -> 'a list -> 'b list -> 'c list
+ val share_tails : 'a list -> 'a list -> 'a list * 'a list * 'a list
+
+ (** [fold_map f e_0 [l_1...l_n] = e_n,[k_1...k_n]]
+ where [(e_i,k_i)=f e_{i-1} l_i] *)
+ val fold_map : ('a -> 'b -> 'a * 'c) -> 'a -> 'b list -> 'a * 'c list
+ val fold_map' : ('b -> 'a -> 'c * 'a) -> 'b list -> 'a -> 'c list * 'a
+ val map_assoc : ('a -> 'b) -> ('c * 'a) list -> ('c * 'b) list
+ val assoc_f : ('a -> 'a -> bool) -> 'a -> ('a * 'b) list -> 'b
+
+ (** A generic 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 cartesian : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list
+
+ (** [cartesians] is an n-ary cartesian product: it iterates
+ [cartesian] over a list of lists. *)
+ val cartesians : ('a -> 'b -> 'b) -> 'b -> 'a list list -> 'b list
+
+ (** combinations [[a;b];[c;d]] returns [[a;c];[a;d];[b;c];[b;d]] *)
+ val combinations : 'a list list -> 'a list list
+ val combine3 : 'a list -> 'b list -> 'c list -> ('a * 'b * 'c) list
+
+ (** Keep only those products that do not return None *)
+ val cartesian_filter :
+ ('a -> 'b -> 'c option) -> 'a list -> 'b list -> 'c list
+ val cartesians_filter :
+ ('a -> 'b -> 'b option) -> 'b -> 'a list list -> 'b list
+
+ val union_map : ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b
+ val factorize_left : ('a * 'b) list -> ('a * 'b list) list
+end
+
+include List
+
+(* Lists *)
+
+let rec compare cmp l1 l2 =
+ if l1 == l2 then 0 else
+ match l1,l2 with
+ [], [] -> 0
+ | _::_, [] -> 1
+ | [], _::_ -> -1
+ | x1::l1, x2::l2 ->
+ (match cmp x1 x2 with
+ | 0 -> compare cmp l1 l2
+ | c -> c)
+
+let rec equal cmp l1 l2 =
+ l1 == l2 ||
+ match l1, l2 with
+ | [], [] -> true
+ | x1 :: l1, x2 :: l2 ->
+ cmp x1 x2 && equal cmp l1 l2
+ | _ -> false
+
+let intersect l1 l2 =
+ List.filter (fun x -> List.mem x l2) l1
+
+let union l1 l2 =
+ let rec urec = function
+ | [] -> l2
+ | a::l -> if List.mem a l2 then urec l else a::urec l
+ in
+ urec l1
+
+let unionq l1 l2 =
+ let rec urec = function
+ | [] -> l2
+ | a::l -> if List.memq a l2 then urec l else a::urec l
+ in
+ urec l1
+
+let subtract l1 l2 =
+ if l2 = [] then l1 else List.filter (fun x -> not (List.mem x l2)) l1
+
+let subtractq l1 l2 =
+ if l2 = [] then l1 else List.filter (fun x -> not (List.memq x l2)) l1
+
+let tabulate f len =
+ let rec tabrec n =
+ if n = len then [] else (f n)::(tabrec (n+1))
+ in
+ tabrec 0
+
+let addn n v =
+ let rec aux n l =
+ if n = 0 then l
+ else aux (pred n) (v::l)
+ in
+ if n < 0 then invalid_arg "List.addn"
+ else aux n
+
+let make n v = addn n v []
+
+let assign l n e =
+ let rec assrec stk = function
+ | ((h::t), 0) -> List.rev_append stk (e::t)
+ | ((h::t), n) -> assrec (h::stk) (t, n-1)
+ | ([], _) -> failwith "List.assign"
+ in
+ assrec [] (l,n)
+
+let rec smartmap f l = match l with
+ [] -> l
+ | h::tl ->
+ let h' = f h and tl' = smartmap f tl in
+ if h'==h && tl'==tl then l
+ else h'::tl'
+
+let map_left f = (* ensures the order in case of side-effects *)
+ let rec map_rec = function
+ | [] -> []
+ | x::l -> let v = f x in v :: map_rec l
+ in
+ map_rec
+
+let map_i f =
+ let rec map_i_rec i = function
+ | [] -> []
+ | x::l -> let v = f i x in v :: map_i_rec (i+1) l
+ in
+ map_i_rec
+
+let map2_i f i l1 l2 =
+ let rec map_i i = function
+ | ([], []) -> []
+ | ((h1::t1), (h2::t2)) -> let v = f i h1 h2 in v :: map_i (succ i) (t1,t2)
+ | (_, _) -> invalid_arg "map2_i"
+ in
+ map_i i (l1,l2)
+
+let map3 f l1 l2 l3 =
+ let rec map = function
+ | ([], [], []) -> []
+ | ((h1::t1), (h2::t2), (h3::t3)) -> let v = f h1 h2 h3 in v::map (t1,t2,t3)
+ | (_, _, _) -> invalid_arg "map3"
+ in
+ map (l1,l2,l3)
+
+let map4 f l1 l2 l3 l4 =
+ let rec map = function
+ | ([], [], [], []) -> []
+ | ((h1::t1), (h2::t2), (h3::t3), (h4::t4)) -> let v = f h1 h2 h3 h4 in v::map (t1,t2,t3,t4)
+ | (_, _, _, _) -> invalid_arg "map4"
+ in
+ map (l1,l2,l3,l4)
+
+let map_to_array f l =
+ Array.of_list (List.map f l)
+
+let rec smartfilter f l = match l with
+ [] -> l
+ | h::tl ->
+ let tl' = smartfilter f tl in
+ if f h then
+ if tl' == tl then l
+ else h :: tl'
+ else tl'
+
+let index_f f x =
+ let rec index_x n = function
+ | y::l -> if f x y then n else index_x (succ n) l
+ | [] -> raise Not_found
+ in
+ index_x 1
+
+let index0_f f x l = index_f f x l - 1
+
+let index x =
+ let rec index_x n = function
+ | y::l -> if x = y then n else index_x (succ n) l
+ | [] -> raise Not_found
+ in
+ index_x 1
+
+let index0 x l = index x l - 1
+
+let unique_index x =
+ let rec index_x n = function
+ | y::l ->
+ if x = y then
+ if List.mem x l then raise Not_found
+ else n
+ else index_x (succ n) l
+ | [] -> raise Not_found
+ in index_x 1
+
+let fold_right_i f i l =
+ let rec it_f i l a = match l with
+ | [] -> a
+ | b::l -> f (i-1) b (it_f (i-1) l a)
+ in
+ it_f (List.length l + i) l
+
+let fold_left_i f =
+ let rec it_list_f i a = function
+ | [] -> a
+ | b::l -> it_list_f (i+1) (f i a b) l
+ in
+ it_list_f
+
+let rec fold_left3 f accu l1 l2 l3 =
+ match (l1, l2, l3) with
+ ([], [], []) -> accu
+ | (a1::l1, a2::l2, a3::l3) -> fold_left3 f (f accu a1 a2 a3) l1 l2 l3
+ | (_, _, _) -> invalid_arg "List.fold_left3"
+
+(* [fold_right_and_left f [a1;...;an] hd =
+ f (f (... (f (f hd
+ an
+ [an-1;...;a1])
+ an-1
+ [an-2;...;a1])
+ ...)
+ a2
+ [a1])
+ a1
+ []] *)
+
+let fold_right_and_left f l hd =
+ let rec aux tl = function
+ | [] -> hd
+ | a::l -> let hd = aux (a::tl) l in f hd a tl
+ in aux [] l
+
+let iter3 f l1 l2 l3 =
+ let rec iter = function
+ | ([], [], []) -> ()
+ | ((h1::t1), (h2::t2), (h3::t3)) -> f h1 h2 h3; iter (t1,t2,t3)
+ | (_, _, _) -> invalid_arg "map3"
+ in
+ iter (l1,l2,l3)
+
+let iter_i f l = fold_left_i (fun i _ x -> f i x) 0 () l
+
+let for_all_i p =
+ let rec for_all_p i = function
+ | [] -> true
+ | a::l -> p i a && for_all_p (i+1) l
+ in
+ for_all_p
+
+let except x l = List.filter (fun y -> not (x = y)) l
+
+let remove = except (* Alias *)
+
+let rec remove_first a = function
+ | b::l when a = b -> l
+ | b::l -> b::remove_first a l
+ | [] -> raise Not_found
+
+let rec remove_assoc_in_triple x = function
+ | [] -> []
+ | (y,_,_ as z)::l -> if x = y then l else z::remove_assoc_in_triple x l
+
+let rec assoc_snd_in_triple x = function
+ [] -> raise Not_found
+ | (a,b,_)::l -> if Pervasives.compare a x = 0 then b else assoc_snd_in_triple x l
+
+let add_set x l = if List.mem x l then l else x::l
+
+let eq_set l1 l2 =
+ let rec aux l1 = function
+ | [] -> l1 = []
+ | a::l2 -> aux (remove_first a l1) l2 in
+ try aux l1 l2 with Not_found -> false
+
+let for_all2eq f l1 l2 =
+ try List.for_all2 f l1 l2 with Invalid_argument _ -> false
+
+let filter_i p =
+ let rec filter_i_rec i = function
+ | [] -> []
+ | x::l -> let l' = filter_i_rec (succ i) l in if p i x then x::l' else l'
+ in
+ filter_i_rec 0
+
+let rec sep_last = function
+ | [] -> failwith "sep_last"
+ | hd::[] -> (hd,[])
+ | hd::tl -> let (l,tl) = sep_last tl in (l,hd::tl)
+
+let try_find_i f =
+ let rec try_find_f n = function
+ | [] -> failwith "try_find_i"
+ | h::t -> try f n h with Failure _ -> try_find_f (n+1) t
+ in
+ try_find_f
+
+let try_find f =
+ let rec try_find_f = function
+ | [] -> failwith "try_find"
+ | h::t -> try f h with Failure _ -> try_find_f t
+ in
+ try_find_f
+
+let uniquize l =
+ let visited = Hashtbl.create 23 in
+ let rec aux acc = function
+ | h::t -> if Hashtbl.mem visited h then aux acc t else
+ begin
+ Hashtbl.add visited h h;
+ aux (h::acc) t
+ end
+ | [] -> List.rev acc
+ in aux [] l
+
+let distinct l =
+ let visited = Hashtbl.create 23 in
+ let rec loop = function
+ | h::t ->
+ if Hashtbl.mem visited h then false
+ else
+ begin
+ Hashtbl.add visited h h;
+ loop t
+ end
+ | [] -> true
+ in loop l
+
+let rec merge_uniq cmp l1 l2 =
+ match l1, l2 with
+ | [], l2 -> l2
+ | l1, [] -> l1
+ | h1 :: t1, h2 :: t2 ->
+ let c = cmp h1 h2 in
+ if c = 0
+ then h1 :: merge_uniq cmp t1 t2
+ else if c <= 0
+ then h1 :: merge_uniq cmp t1 l2
+ else h2 :: merge_uniq cmp l1 t2
+
+let rec duplicates = function
+ | [] -> []
+ | x::l ->
+ let l' = duplicates l in
+ if List.mem x l then add_set x l' else l'
+
+let rec filter2 f = function
+ | [], [] as p -> p
+ | d::dp, l::lp ->
+ let (dp',lp' as p) = filter2 f (dp,lp) in
+ if f d l then d::dp', l::lp' else p
+ | _ -> invalid_arg "List.filter2"
+
+let rec map_filter f = function
+ | [] -> []
+ | x::l ->
+ let l' = map_filter f l in
+ match f x with None -> l' | Some y -> y::l'
+
+let map_filter_i f =
+ let rec aux i = function
+ | [] -> []
+ | x::l ->
+ let l' = aux (succ i) l in
+ match f i x with None -> l' | Some y -> y::l'
+ in aux 0
+
+let filter_along f filter l =
+ snd (filter2 (fun b c -> f b) (filter,l))
+
+let filter_with filter l =
+ filter_along (fun x -> x) filter l
+
+let subset l1 l2 =
+ let t2 = Hashtbl.create 151 in
+ List.iter (fun x -> Hashtbl.add t2 x ()) l2;
+ let rec look = function
+ | [] -> true
+ | x::ll -> try Hashtbl.find t2 x; look ll with Not_found -> false
+ in
+ look l1
+
+(* [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] *)
+
+let chop n l =
+ let rec chop_aux i acc = function
+ | tl when i=0 -> (List.rev acc, tl)
+ | h::t -> chop_aux (pred i) (h::acc) t
+ | [] -> failwith "List.chop"
+ in
+ chop_aux n [] 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 *)
+let split_when p =
+ let rec split_when_loop x y =
+ match y with
+ | [] -> (List.rev x,[])
+ | (a::l) -> if (p a) then (List.rev x,y) else split_when_loop (a::x) l
+ in
+ split_when_loop []
+
+(* [split_by p l] splits [l] into two lists [(l1,l2)] such that elements of
+ [l1] satisfy [p] and elements of [l2] do not; order is preserved *)
+let split_by p =
+ let rec split_by_loop = function
+ | [] -> ([],[])
+ | a::l ->
+ let (l1,l2) = split_by_loop l in if p a then (a::l1,l2) else (l1,a::l2)
+ in
+ split_by_loop
+
+let rec split3 = function
+ | [] -> ([], [], [])
+ | (x,y,z)::l ->
+ let (rx, ry, rz) = split3 l in (x::rx, y::ry, z::rz)
+
+let rec insert_in_class f a = function
+ | [] -> [[a]]
+ | (b::_ as l)::classes when f a b -> (a::l)::classes
+ | l::classes -> l :: insert_in_class f a classes
+
+let partition_by f l =
+ List.fold_right (insert_in_class f) l []
+
+let firstn n l =
+ let rec aux acc = function
+ | (0, l) -> List.rev acc
+ | (n, (h::t)) -> aux (h::acc) (pred n, t)
+ | _ -> failwith "firstn"
+ in
+ aux [] (n,l)
+
+let rec last = function
+ | [] -> failwith "List.last"
+ | [x] -> x
+ | _ :: l -> last l
+
+let lastn n l =
+ let len = List.length l in
+ let rec aux m l =
+ if m = n then l else aux (m - 1) (List.tl l)
+ in
+ if len < n then failwith "lastn" else aux len l
+
+let rec skipn n l = match n,l with
+ | 0, _ -> l
+ | _, [] -> failwith "List.skipn"
+ | n, _::l -> skipn (pred n) l
+
+let skipn_at_least n l =
+ try skipn n l with Failure _ -> []
+
+let prefix_of prefl l =
+ let rec prefrec = function
+ | (h1::t1, h2::t2) -> h1 = h2 && prefrec (t1,t2)
+ | ([], _) -> true
+ | (_, _) -> false
+ in
+ prefrec (prefl,l)
+
+let drop_prefix p l =
+(* if l=p++t then return t else l *)
+ let rec drop_prefix_rec = function
+ | ([], tl) -> Some tl
+ | (_, []) -> None
+ | (h1::tp, h2::tl) ->
+ if h1 = h2 then drop_prefix_rec (tp,tl) else None
+ in
+ match drop_prefix_rec (p,l) with
+ | Some r -> r
+ | None -> l
+
+let map_append f l = List.flatten (List.map f l)
+let join_map = map_append (* Alias *)
+
+let map_append2 f l1 l2 = List.flatten (List.map2 f l1 l2)
+
+let share_tails l1 l2 =
+ let rec shr_rev acc = function
+ | ((x1::l1), (x2::l2)) when x1 == x2 -> shr_rev (x1::acc) (l1,l2)
+ | (l1,l2) -> (List.rev l1, List.rev l2, acc)
+ in
+ shr_rev [] (List.rev l1, List.rev l2)
+
+let rec fold_map f e = function
+ | [] -> (e,[])
+ | h::t ->
+ let e',h' = f e h in
+ let e'',t' = fold_map f e' t in
+ e'',h'::t'
+
+(* (* tail-recursive version of the above function *)
+let fold_map f e l =
+ let g (e,b') h =
+ let (e',h') = f e h in
+ (e',h'::b')
+ in
+ let (e',lrev) = List.fold_left g (e,[]) l in
+ (e',List.rev lrev)
+*)
+
+(* The same, based on fold_right, with the effect accumulated on the right *)
+let fold_map' f l e =
+ List.fold_right (fun x (l,e) -> let (y,e) = f x e in (y::l,e)) l ([],e)
+
+let map_assoc f = List.map (fun (x,a) -> (x,f a))
+
+let rec assoc_f f a = function
+ | (x, e) :: xs -> if f a x then e else assoc_f f a xs
+ | [] -> raise Not_found
+
+(* Specification:
+ - =p= is set equality (double inclusion)
+ - f such that \forall l acc, (f l acc) =p= append (f l []) acc
+ - let g = fun x -> f x [] in
+ - union_map f l acc =p= append (flatten (map g l)) acc
+ *)
+let union_map f l acc =
+ List.fold_left
+ (fun x y -> f y x)
+ acc
+ l
+
+(* A generic 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. *)
+
+let cartesian op l1 l2 =
+ map_append (fun x -> List.map (op x) l2) l1
+
+(* [cartesians] is an n-ary cartesian product: it iterates
+ [cartesian] over a list of lists. *)
+
+let cartesians op init ll =
+ List.fold_right (cartesian op) ll [init]
+
+(* combinations [[a;b];[c;d]] gives [[a;c];[a;d];[b;c];[b;d]] *)
+
+let combinations l = cartesians (fun x l -> x::l) [] l
+
+let rec combine3 x y z =
+ match x, y, z with
+ | [], [], [] -> []
+ | (x :: xs), (y :: ys), (z :: zs) ->
+ (x, y, z) :: combine3 xs ys zs
+ | _, _, _ -> raise (Invalid_argument "List.combine3")
+
+(* Keep only those products that do not return None *)
+
+let cartesian_filter op l1 l2 =
+ map_append (fun x -> map_filter (op x) l2) l1
+
+(* Keep only those products that do not return None *)
+
+let cartesians_filter op init ll =
+ List.fold_right (cartesian_filter op) ll [init]
+
+(* Drop the last element of a list *)
+
+let rec drop_last = function [] -> assert false | hd :: [] -> [] | hd :: tl -> hd :: drop_last tl
+
+(* Factorize lists of pairs according to the left argument *)
+let rec factorize_left = function
+ | (a,b)::l ->
+ let al,l' = split_by (fun (a',b) -> a=a') l in
+ (a,(b::List.map snd al)) :: factorize_left l'
+ | [] ->
+ []
+
diff --git a/lib/cList.mli b/lib/cList.mli
new file mode 100644
index 000000000..c530a2081
--- /dev/null
+++ b/lib/cList.mli
@@ -0,0 +1,188 @@
+(***********************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
+(* \VV/ *************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(***********************************************************************)
+
+(** FIXME: From OCaml 3.12 onwards, that would only be a [module type of] *)
+
+module type S =
+sig
+ val length : 'a list -> int
+ val hd : 'a list -> 'a
+ val tl : 'a list -> 'a list
+ val nth : 'a list -> int -> 'a
+ val rev : 'a list -> '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 iter : ('a -> unit) -> 'a list -> unit
+ val map : ('a -> 'b) -> 'a list -> 'b list
+ val rev_map : ('a -> 'b) -> 'a list -> 'b 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 filter : ('a -> bool) -> 'a list -> 'a list
+ val find_all : ('a -> bool) -> 'a list -> 'a list
+ val partition : ('a -> bool) -> 'a list -> 'a list * 'a list
+ val assoc : 'a -> ('a * 'b) list -> 'b
+ val assq : 'a -> ('a * 'b) list -> 'b
+ 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 merge : ('a -> 'a -> int) -> 'a list -> 'a list -> 'a list
+end
+
+module type ExtS =
+sig
+ include S
+ val compare : ('a -> 'a -> int) -> 'a list -> 'a list -> int
+ val equal : ('a -> 'a -> bool) -> 'a list -> 'a list -> bool
+ val add_set : 'a -> 'a list -> 'a list
+ val eq_set : 'a list -> 'a list -> bool
+ val intersect : 'a list -> 'a list -> 'a list
+ val union : 'a list -> 'a list -> 'a list
+ val unionq : 'a list -> 'a list -> 'a list
+ val subtract : 'a list -> 'a list -> 'a list
+ val subtractq : 'a list -> 'a list -> 'a list
+
+ (** [tabulate f n] builds [[f 0; ...; f (n-1)]] *)
+ val tabulate : (int -> 'a) -> int -> 'a list
+ val make : int -> 'a -> 'a list
+ val assign : 'a list -> int -> 'a -> 'a list
+ val distinct : 'a list -> bool
+ val duplicates : 'a list -> 'a list
+ val filter2 : ('a -> 'b -> bool) -> 'a list * 'b list -> 'a list * 'b list
+ val map_filter : ('a -> 'b option) -> 'a list -> 'b list
+ val map_filter_i : (int -> 'a -> 'b option) -> 'a list -> 'b list
+ val filter_with : bool list -> 'a list -> 'a list
+ val filter_along : ('a -> bool) -> 'a list -> 'b list -> 'b list
+
+ (** [smartmap f [a1...an] = List.map f [a1...an]] but if for all i
+ [ f ai == ai], then [smartmap f l==l] *)
+ val smartmap : ('a -> 'a) -> 'a list -> 'a list
+ val map_left : ('a -> 'b) -> 'a list -> 'b list
+ val map_i : (int -> 'a -> 'b) -> int -> 'a list -> 'b list
+ val map2_i :
+ (int -> 'a -> 'b -> 'c) -> int -> 'a list -> 'b list -> 'c list
+ val map3 :
+ ('a -> 'b -> 'c -> 'd) -> 'a list -> 'b list -> 'c list -> 'd list
+ val map4 :
+ ('a -> 'b -> 'c -> 'd -> 'e) -> 'a list -> 'b list -> 'c list -> 'd list -> 'e list
+ val map_to_array : ('a -> 'b) -> 'a list -> 'b array
+ val filter_i :
+ (int -> 'a -> bool) -> 'a list -> 'a list
+
+ (** [smartfilter f [a1...an] = List.filter f [a1...an]] but if for all i
+ [f ai = true], then [smartfilter f l==l] *)
+ val smartfilter : ('a -> bool) -> 'a list -> 'a list
+
+ (** [index] returns the 1st index of an element in a list (counting from 1) *)
+ val index : 'a -> 'a list -> int
+ val index_f : ('a -> 'a -> bool) -> 'a -> 'a list -> int
+
+ (** [unique_index x l] returns [Not_found] if [x] doesn't occur exactly once *)
+ val unique_index : 'a -> 'a list -> int
+
+ (** [index0] behaves as [index] except that it starts counting at 0 *)
+ val index0 : 'a -> 'a list -> int
+ val index0_f : ('a -> 'a -> bool) -> 'a -> 'a list -> int
+ val iter3 : ('a -> 'b -> 'c -> unit) -> 'a list -> 'b list -> 'c list -> unit
+ val iter_i : (int -> 'a -> unit) -> 'a list -> unit
+ val fold_right_i : (int -> 'a -> 'b -> 'b) -> int -> 'a list -> 'b -> 'b
+ val fold_left_i : (int -> 'a -> 'b -> 'a) -> int -> 'a -> 'b list -> 'a
+ val fold_right_and_left :
+ ('a -> 'b -> 'b list -> 'a) -> 'b list -> 'a -> 'a
+ val fold_left3 : ('a -> 'b -> 'c -> 'd -> 'a) -> 'a -> 'b list -> 'c list -> 'd list -> 'a
+ val for_all_i : (int -> 'a -> bool) -> int -> 'a list -> bool
+ val except : 'a -> 'a list -> 'a list
+ val remove : 'a -> 'a list -> 'a list
+ val remove_first : 'a -> 'a list -> 'a list
+ val remove_assoc_in_triple : 'a -> ('a * 'b * 'c) list -> ('a * 'b * 'c) list
+ val assoc_snd_in_triple : 'a -> ('a * 'b * 'c) list -> 'b
+ val for_all2eq : ('a -> 'b -> bool) -> 'a list -> 'b list -> bool
+ val sep_last : 'a list -> 'a * 'a list
+ val try_find_i : (int -> 'a -> 'b) -> int -> 'a list -> 'b
+ val try_find : ('a -> 'b) -> 'a list -> 'b
+ val uniquize : 'a list -> 'a list
+
+ (** merges two sorted lists and preserves the uniqueness property: *)
+ val merge_uniq : ('a -> 'a -> int) -> 'a list -> 'a list -> 'a list
+ val subset : 'a list -> 'a list -> bool
+ val chop : int -> 'a list -> 'a list * 'a list
+ (* former [split_at] was a duplicate of [chop] *)
+ val split_when : ('a -> bool) -> 'a list -> 'a list * 'a list
+ val split_by : ('a -> bool) -> 'a list -> 'a list * 'a list
+ val split3 : ('a * 'b * 'c) list -> 'a list * 'b list * 'c list
+ val partition_by : ('a -> 'a -> bool) -> 'a list -> 'a list list
+ val firstn : int -> 'a list -> 'a list
+ val last : 'a list -> 'a
+ val lastn : int -> 'a list -> 'a list
+ val skipn : int -> 'a list -> 'a list
+ val skipn_at_least : int -> 'a list -> 'a list
+ val addn : int -> 'a -> 'a list -> 'a list
+ val prefix_of : 'a list -> 'a list -> bool
+
+ (** [drop_prefix p l] returns [t] if [l=p++t] else return [l] *)
+ val drop_prefix : 'a list -> 'a list -> 'a list
+ val drop_last : 'a list -> 'a list
+
+ (** [map_append f [x1; ...; xn]] returns [(f x1)@(f x2)@...@(f xn)] *)
+ val map_append : ('a -> 'b list) -> 'a list -> 'b list
+ val join_map : ('a -> 'b list) -> 'a list -> 'b list
+
+ (** raises [Invalid_argument] if the two lists don't have the same length *)
+ val map_append2 : ('a -> 'b -> 'c list) -> 'a list -> 'b list -> 'c list
+ val share_tails : 'a list -> 'a list -> 'a list * 'a list * 'a list
+
+ (** [fold_map f e_0 [l_1...l_n] = e_n,[k_1...k_n]]
+ where [(e_i,k_i)=f e_{i-1} l_i] *)
+ val fold_map : ('a -> 'b -> 'a * 'c) -> 'a -> 'b list -> 'a * 'c list
+ val fold_map' : ('b -> 'a -> 'c * 'a) -> 'b list -> 'a -> 'c list * 'a
+ val map_assoc : ('a -> 'b) -> ('c * 'a) list -> ('c * 'b) list
+ val assoc_f : ('a -> 'a -> bool) -> 'a -> ('a * 'b) list -> 'b
+
+ (** A generic 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 cartesian : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list
+
+ (** [cartesians] is an n-ary cartesian product: it iterates
+ [cartesian] over a list of lists. *)
+ val cartesians : ('a -> 'b -> 'b) -> 'b -> 'a list list -> 'b list
+
+ (** combinations [[a;b];[c;d]] returns [[a;c];[a;d];[b;c];[b;d]] *)
+ val combinations : 'a list list -> 'a list list
+ val combine3 : 'a list -> 'b list -> 'c list -> ('a * 'b * 'c) list
+
+ (** Keep only those products that do not return None *)
+ val cartesian_filter :
+ ('a -> 'b -> 'c option) -> 'a list -> 'b list -> 'c list
+ val cartesians_filter :
+ ('a -> 'b -> 'b option) -> 'b -> 'a list list -> 'b list
+
+ val union_map : ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b
+ val factorize_left : ('a * 'b) list -> ('a * 'b list) list
+end
+
+include ExtS
diff --git a/lib/clib.mllib b/lib/clib.mllib
index 5974fe517..5b0e14f1e 100644
--- a/lib/clib.mllib
+++ b/lib/clib.mllib
@@ -4,6 +4,7 @@ Coq_config
Segmenttree
Unicodetable
Deque
+CList
Util
Serialize
Xml_utils
diff --git a/lib/gmapl.ml b/lib/gmapl.ml
index 3c89ea68d..5e8f27dc7 100644
--- a/lib/gmapl.ml
+++ b/lib/gmapl.ml
@@ -20,7 +20,7 @@ let fold = Gmap.fold
let add x y m =
try
let l = Gmap.find x m in
- Gmap.add x (y::list_except y l) m
+ Gmap.add x (y::List.except y l) m
with Not_found ->
Gmap.add x [y] m
@@ -29,6 +29,6 @@ let find x m =
let remove x y m =
let l = Gmap.find x m in
- Gmap.add x (if List.mem y l then list_subtract l [y] else l) m
+ Gmap.add x (if List.mem y l then List.subtract l [y] else l) m
diff --git a/lib/util.ml b/lib/util.ml
index 1365f53ce..8916400ac 100644
--- a/lib/util.ml
+++ b/lib/util.ml
@@ -363,536 +363,7 @@ let ascii_of_ident s =
done with End_of_input -> () end;
!out
-(* Lists *)
-
-let rec list_compare cmp l1 l2 =
- if l1 == l2 then 0 else
- match l1,l2 with
- [], [] -> 0
- | _::_, [] -> 1
- | [], _::_ -> -1
- | x1::l1, x2::l2 ->
- (match cmp x1 x2 with
- | 0 -> list_compare cmp l1 l2
- | c -> c)
-
-let rec list_equal cmp l1 l2 =
- l1 == l2 ||
- match l1, l2 with
- | [], [] -> true
- | x1 :: l1, x2 :: l2 ->
- cmp x1 x2 && list_equal cmp l1 l2
- | _ -> false
-
-let list_intersect l1 l2 =
- List.filter (fun x -> List.mem x l2) l1
-
-let list_union l1 l2 =
- let rec urec = function
- | [] -> l2
- | a::l -> if List.mem a l2 then urec l else a::urec l
- in
- urec l1
-
-let list_unionq l1 l2 =
- let rec urec = function
- | [] -> l2
- | a::l -> if List.memq a l2 then urec l else a::urec l
- in
- urec l1
-
-let list_subtract l1 l2 =
- if l2 = [] then l1 else List.filter (fun x -> not (List.mem x l2)) l1
-
-let list_subtractq l1 l2 =
- if l2 = [] then l1 else List.filter (fun x -> not (List.memq x l2)) l1
-
-let list_tabulate f len =
- let rec tabrec n =
- if n = len then [] else (f n)::(tabrec (n+1))
- in
- tabrec 0
-
-let list_addn n v =
- let rec aux n l =
- if n = 0 then l
- else aux (pred n) (v::l)
- in
- if n < 0 then invalid_arg "list_addn"
- else aux n
-
-let list_make n v = list_addn n v []
-
-let list_assign l n e =
- let rec assrec stk = function
- | ((h::t), 0) -> List.rev_append stk (e::t)
- | ((h::t), n) -> assrec (h::stk) (t, n-1)
- | ([], _) -> failwith "list_assign"
- in
- assrec [] (l,n)
-
-let rec list_smartmap f l = match l with
- [] -> l
- | h::tl ->
- let h' = f h and tl' = list_smartmap f tl in
- if h'==h && tl'==tl then l
- else h'::tl'
-
-let list_map_left f = (* ensures the order in case of side-effects *)
- let rec map_rec = function
- | [] -> []
- | x::l -> let v = f x in v :: map_rec l
- in
- map_rec
-
-let list_map_i f =
- let rec map_i_rec i = function
- | [] -> []
- | x::l -> let v = f i x in v :: map_i_rec (i+1) l
- in
- map_i_rec
-
-let list_map2_i f i l1 l2 =
- let rec map_i i = function
- | ([], []) -> []
- | ((h1::t1), (h2::t2)) -> let v = f i h1 h2 in v :: map_i (succ i) (t1,t2)
- | (_, _) -> invalid_arg "map2_i"
- in
- map_i i (l1,l2)
-
-let list_map3 f l1 l2 l3 =
- let rec map = function
- | ([], [], []) -> []
- | ((h1::t1), (h2::t2), (h3::t3)) -> let v = f h1 h2 h3 in v::map (t1,t2,t3)
- | (_, _, _) -> invalid_arg "map3"
- in
- map (l1,l2,l3)
-
-let list_map4 f l1 l2 l3 l4 =
- let rec map = function
- | ([], [], [], []) -> []
- | ((h1::t1), (h2::t2), (h3::t3), (h4::t4)) -> let v = f h1 h2 h3 h4 in v::map (t1,t2,t3,t4)
- | (_, _, _, _) -> invalid_arg "map4"
- in
- map (l1,l2,l3,l4)
-
-let list_map_to_array f l =
- Array.of_list (List.map f l)
-
-let rec list_smartfilter f l = match l with
- [] -> l
- | h::tl ->
- let tl' = list_smartfilter f tl in
- if f h then
- if tl' == tl then l
- else h :: tl'
- else tl'
-
-let list_index_f f x =
- let rec index_x n = function
- | y::l -> if f x y then n else index_x (succ n) l
- | [] -> raise Not_found
- in
- index_x 1
-
-let list_index0_f f x l = list_index_f f x l - 1
-
-let list_index x =
- let rec index_x n = function
- | y::l -> if x = y then n else index_x (succ n) l
- | [] -> raise Not_found
- in
- index_x 1
-
-let list_index0 x l = list_index x l - 1
-
-let list_unique_index x =
- let rec index_x n = function
- | y::l ->
- if x = y then
- if List.mem x l then raise Not_found
- else n
- else index_x (succ n) l
- | [] -> raise Not_found
- in index_x 1
-
-let list_fold_right_i f i l =
- let rec it_list_f i l a = match l with
- | [] -> a
- | b::l -> f (i-1) b (it_list_f (i-1) l a)
- in
- it_list_f (List.length l + i) l
-
-let list_fold_left_i f =
- let rec it_list_f i a = function
- | [] -> a
- | b::l -> it_list_f (i+1) (f i a b) l
- in
- it_list_f
-
-let rec list_fold_left3 f accu l1 l2 l3 =
- match (l1, l2, l3) with
- ([], [], []) -> accu
- | (a1::l1, a2::l2, a3::l3) -> list_fold_left3 f (f accu a1 a2 a3) l1 l2 l3
- | (_, _, _) -> invalid_arg "list_fold_left3"
-
-(* [list_fold_right_and_left f [a1;...;an] hd =
- f (f (... (f (f hd
- an
- [an-1;...;a1])
- an-1
- [an-2;...;a1])
- ...)
- a2
- [a1])
- a1
- []] *)
-
-let list_fold_right_and_left f l hd =
- let rec aux tl = function
- | [] -> hd
- | a::l -> let hd = aux (a::tl) l in f hd a tl
- in aux [] l
-
-let list_iter3 f l1 l2 l3 =
- let rec iter = function
- | ([], [], []) -> ()
- | ((h1::t1), (h2::t2), (h3::t3)) -> f h1 h2 h3; iter (t1,t2,t3)
- | (_, _, _) -> invalid_arg "map3"
- in
- iter (l1,l2,l3)
-
-let list_iter_i f l = list_fold_left_i (fun i _ x -> f i x) 0 () l
-
-let list_for_all_i p =
- let rec for_all_p i = function
- | [] -> true
- | a::l -> p i a && for_all_p (i+1) l
- in
- for_all_p
-
-let list_except x l = List.filter (fun y -> not (x = y)) l
-
-let list_remove = list_except (* Alias *)
-
-let rec list_remove_first a = function
- | b::l when a = b -> l
- | b::l -> b::list_remove_first a l
- | [] -> raise Not_found
-
-let rec list_remove_assoc_in_triple x = function
- | [] -> []
- | (y,_,_ as z)::l -> if x = y then l else z::list_remove_assoc_in_triple x l
-
-let rec list_assoc_snd_in_triple x = function
- [] -> raise Not_found
- | (a,b,_)::l -> if compare a x = 0 then b else list_assoc_snd_in_triple x l
-
-let list_add_set x l = if List.mem x l then l else x::l
-
-let list_eq_set l1 l2 =
- let rec aux l1 = function
- | [] -> l1 = []
- | a::l2 -> aux (list_remove_first a l1) l2 in
- try aux l1 l2 with Not_found -> false
-
-let list_for_all2eq f l1 l2 =
- try List.for_all2 f l1 l2 with Invalid_argument _ -> false
-
-let list_filter_i p =
- let rec filter_i_rec i = function
- | [] -> []
- | x::l -> let l' = filter_i_rec (succ i) l in if p i x then x::l' else l'
- in
- filter_i_rec 0
-
-let rec list_sep_last = function
- | [] -> failwith "sep_last"
- | hd::[] -> (hd,[])
- | hd::tl -> let (l,tl) = list_sep_last tl in (l,hd::tl)
-
-let list_try_find_i f =
- let rec try_find_f n = function
- | [] -> failwith "try_find_i"
- | h::t -> try f n h with Failure _ -> try_find_f (n+1) t
- in
- try_find_f
-
-let list_try_find f =
- let rec try_find_f = function
- | [] -> failwith "try_find"
- | h::t -> try f h with Failure _ -> try_find_f t
- in
- try_find_f
-
-let list_uniquize l =
- let visited = Hashtbl.create 23 in
- let rec aux acc = function
- | h::t -> if Hashtbl.mem visited h then aux acc t else
- begin
- Hashtbl.add visited h h;
- aux (h::acc) t
- end
- | [] -> List.rev acc
- in aux [] l
-
-let list_distinct l =
- let visited = Hashtbl.create 23 in
- let rec loop = function
- | h::t ->
- if Hashtbl.mem visited h then false
- else
- begin
- Hashtbl.add visited h h;
- loop t
- end
- | [] -> true
- in loop l
-
-let rec list_merge_uniq cmp l1 l2 =
- match l1, l2 with
- | [], l2 -> l2
- | l1, [] -> l1
- | h1 :: t1, h2 :: t2 ->
- let c = cmp h1 h2 in
- if c = 0
- then h1 :: list_merge_uniq cmp t1 t2
- else if c <= 0
- then h1 :: list_merge_uniq cmp t1 l2
- else h2 :: list_merge_uniq cmp l1 t2
-
-let rec list_duplicates = function
- | [] -> []
- | x::l ->
- let l' = list_duplicates l in
- if List.mem x l then list_add_set x l' else l'
-
-let rec list_filter2 f = function
- | [], [] as p -> p
- | d::dp, l::lp ->
- let (dp',lp' as p) = list_filter2 f (dp,lp) in
- if f d l then d::dp', l::lp' else p
- | _ -> invalid_arg "list_filter2"
-
-let rec list_map_filter f = function
- | [] -> []
- | x::l ->
- let l' = list_map_filter f l in
- match f x with None -> l' | Some y -> y::l'
-
-let list_map_filter_i f =
- let rec aux i = function
- | [] -> []
- | x::l ->
- let l' = aux (succ i) l in
- match f i x with None -> l' | Some y -> y::l'
- in aux 0
-
-let list_filter_along f filter l =
- snd (list_filter2 (fun b c -> f b) (filter,l))
-
-let list_filter_with filter l =
- list_filter_along (fun x -> x) filter l
-
-let list_subset l1 l2 =
- let t2 = Hashtbl.create 151 in
- List.iter (fun x -> Hashtbl.add t2 x ()) l2;
- let rec look = function
- | [] -> true
- | x::ll -> try Hashtbl.find t2 x; look ll with Not_found -> false
- in
- look l1
-
-(* [list_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] *)
-
-let list_chop n l =
- let rec chop_aux i acc = function
- | tl when i=0 -> (List.rev acc, tl)
- | h::t -> chop_aux (pred i) (h::acc) t
- | [] -> failwith "list_chop"
- in
- chop_aux n [] l
-
-(* [list_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 *)
-let list_split_when p =
- let rec split_when_loop x y =
- match y with
- | [] -> (List.rev x,[])
- | (a::l) -> if (p a) then (List.rev x,y) else split_when_loop (a::x) l
- in
- split_when_loop []
-
-(* [list_split_by p l] splits [l] into two lists [(l1,l2)] such that elements of
- [l1] satisfy [p] and elements of [l2] do not; order is preserved *)
-let list_split_by p =
- let rec split_by_loop = function
- | [] -> ([],[])
- | a::l ->
- let (l1,l2) = split_by_loop l in if p a then (a::l1,l2) else (l1,a::l2)
- in
- split_by_loop
-
-let rec list_split3 = function
- | [] -> ([], [], [])
- | (x,y,z)::l ->
- let (rx, ry, rz) = list_split3 l in (x::rx, y::ry, z::rz)
-
-let rec list_insert_in_class f a = function
- | [] -> [[a]]
- | (b::_ as l)::classes when f a b -> (a::l)::classes
- | l::classes -> l :: list_insert_in_class f a classes
-
-let list_partition_by f l =
- List.fold_right (list_insert_in_class f) l []
-
-let list_firstn n l =
- let rec aux acc = function
- | (0, l) -> List.rev acc
- | (n, (h::t)) -> aux (h::acc) (pred n, t)
- | _ -> failwith "firstn"
- in
- aux [] (n,l)
-
-let rec list_last = function
- | [] -> failwith "list_last"
- | [x] -> x
- | _ :: l -> list_last l
-
-let list_lastn n l =
- let len = List.length l in
- let rec aux m l =
- if m = n then l else aux (m - 1) (List.tl l)
- in
- if len < n then failwith "lastn" else aux len l
-
-let rec list_skipn n l = match n,l with
- | 0, _ -> l
- | _, [] -> failwith "list_skipn"
- | n, _::l -> list_skipn (pred n) l
-
-let list_skipn_at_least n l =
- try list_skipn n l with Failure _ -> []
-
-let list_prefix_of prefl l =
- let rec prefrec = function
- | (h1::t1, h2::t2) -> h1 = h2 && prefrec (t1,t2)
- | ([], _) -> true
- | (_, _) -> false
- in
- prefrec (prefl,l)
-
-let list_drop_prefix p l =
-(* if l=p++t then return t else l *)
- let rec list_drop_prefix_rec = function
- | ([], tl) -> Some tl
- | (_, []) -> None
- | (h1::tp, h2::tl) ->
- if h1 = h2 then list_drop_prefix_rec (tp,tl) else None
- in
- match list_drop_prefix_rec (p,l) with
- | Some r -> r
- | None -> l
-
-let list_map_append f l = List.flatten (List.map f l)
-let list_join_map = list_map_append (* Alias *)
-
-let list_map_append2 f l1 l2 = List.flatten (List.map2 f l1 l2)
-
-let list_share_tails l1 l2 =
- let rec shr_rev acc = function
- | ((x1::l1), (x2::l2)) when x1 == x2 -> shr_rev (x1::acc) (l1,l2)
- | (l1,l2) -> (List.rev l1, List.rev l2, acc)
- in
- shr_rev [] (List.rev l1, List.rev l2)
-
-let rec list_fold_map f e = function
- | [] -> (e,[])
- | h::t ->
- let e',h' = f e h in
- let e'',t' = list_fold_map f e' t in
- e'',h'::t'
-
-(* (* tail-recursive version of the above function *)
-let list_fold_map f e l =
- let g (e,b') h =
- let (e',h') = f e h in
- (e',h'::b')
- in
- let (e',lrev) = List.fold_left g (e,[]) l in
- (e',List.rev lrev)
-*)
-
-(* The same, based on fold_right, with the effect accumulated on the right *)
-let list_fold_map' f l e =
- List.fold_right (fun x (l,e) -> let (y,e) = f x e in (y::l,e)) l ([],e)
-
-let list_map_assoc f = List.map (fun (x,a) -> (x,f a))
-
-let rec list_assoc_f f a = function
- | (x, e) :: xs -> if f a x then e else list_assoc_f f a xs
- | [] -> raise Not_found
-
-(* Specification:
- - =p= is set equality (double inclusion)
- - f such that \forall l acc, (f l acc) =p= append (f l []) acc
- - let g = fun x -> f x [] in
- - union_map f l acc =p= append (flatten (map g l)) acc
- *)
-let list_union_map f l acc =
- List.fold_left
- (fun x y -> f y x)
- acc
- l
-
-(* A generic cartesian product: for any operator (**),
- [list_cartesian (**) [x1;x2] [y1;y2] = [x1**y1; x1**y2; x2**y1; x2**y1]],
- and so on if there are more elements in the lists. *)
-
-let list_cartesian op l1 l2 =
- list_map_append (fun x -> List.map (op x) l2) l1
-
-(* [list_cartesians] is an n-ary cartesian product: it iterates
- [list_cartesian] over a list of lists. *)
-
-let list_cartesians op init ll =
- List.fold_right (list_cartesian op) ll [init]
-
-(* list_combinations [[a;b];[c;d]] gives [[a;c];[a;d];[b;c];[b;d]] *)
-
-let list_combinations l = list_cartesians (fun x l -> x::l) [] l
-
-let rec list_combine3 x y z =
- match x, y, z with
- | [], [], [] -> []
- | (x :: xs), (y :: ys), (z :: zs) ->
- (x, y, z) :: list_combine3 xs ys zs
- | _, _, _ -> raise (Invalid_argument "list_combine3")
-
-(* Keep only those products that do not return None *)
-
-let list_cartesian_filter op l1 l2 =
- list_map_append (fun x -> list_map_filter (op x) l2) l1
-
-(* Keep only those products that do not return None *)
-
-let list_cartesians_filter op init ll =
- List.fold_right (list_cartesian_filter op) ll [init]
-
-(* Drop the last element of a list *)
-
-let rec list_drop_last = function [] -> assert false | hd :: [] -> [] | hd :: tl -> hd :: list_drop_last tl
-
-(* Factorize lists of pairs according to the left argument *)
-let rec list_factorize_left = function
- | (a,b)::l ->
- let al,l' = list_split_by (fun (a',b) -> a=a') l in
- (a,(b::List.map snd al)) :: list_factorize_left l'
- | [] ->
- []
+module List : CList.ExtS = CList
(* Arrays *)
@@ -1212,10 +683,10 @@ let array_rev_to_list a =
tolist 0 []
let array_filter_along f filter v =
- Array.of_list (list_filter_along f filter (Array.to_list v))
+ Array.of_list (CList.filter_along f filter (Array.to_list v))
let array_filter_with filter v =
- Array.of_list (list_filter_with filter (Array.to_list v))
+ Array.of_list (CList.filter_with filter (Array.to_list v))
(* Matrices *)
diff --git a/lib/util.mli b/lib/util.mli
index 43d8ea1b0..a8ea6854c 100644
--- a/lib/util.mli
+++ b/lib/util.mli
@@ -71,133 +71,7 @@ val ascii_of_ident : string -> string
(** {6 Lists. } *)
-val list_compare : ('a -> 'a -> int) -> 'a list -> 'a list -> int
-val list_equal : ('a -> 'a -> bool) -> 'a list -> 'a list -> bool
-val list_add_set : 'a -> 'a list -> 'a list
-val list_eq_set : 'a list -> 'a list -> bool
-val list_intersect : 'a list -> 'a list -> 'a list
-val list_union : 'a list -> 'a list -> 'a list
-val list_unionq : 'a list -> 'a list -> 'a list
-val list_subtract : 'a list -> 'a list -> 'a list
-val list_subtractq : 'a list -> 'a list -> 'a list
-
-(** [list_tabulate f n] builds [[f 0; ...; f (n-1)]] *)
-val list_tabulate : (int -> 'a) -> int -> 'a list
-val list_make : int -> 'a -> 'a list
-val list_assign : 'a list -> int -> 'a -> 'a list
-val list_distinct : 'a list -> bool
-val list_duplicates : 'a list -> 'a list
-val list_filter2 : ('a -> 'b -> bool) -> 'a list * 'b list -> 'a list * 'b list
-val list_map_filter : ('a -> 'b option) -> 'a list -> 'b list
-val list_map_filter_i : (int -> 'a -> 'b option) -> 'a list -> 'b list
-val list_filter_with : bool list -> 'a list -> 'a list
-val list_filter_along : ('a -> bool) -> 'a list -> 'b list -> 'b list
-
-(** [list_smartmap f [a1...an] = List.map f [a1...an]] but if for all i
- [ f ai == ai], then [list_smartmap f l==l] *)
-val list_smartmap : ('a -> 'a) -> 'a list -> 'a list
-val list_map_left : ('a -> 'b) -> 'a list -> 'b list
-val list_map_i : (int -> 'a -> 'b) -> int -> 'a list -> 'b list
-val list_map2_i :
- (int -> 'a -> 'b -> 'c) -> int -> 'a list -> 'b list -> 'c list
-val list_map3 :
- ('a -> 'b -> 'c -> 'd) -> 'a list -> 'b list -> 'c list -> 'd list
-val list_map4 :
- ('a -> 'b -> 'c -> 'd -> 'e) -> 'a list -> 'b list -> 'c list -> 'd list -> 'e list
-val list_map_to_array : ('a -> 'b) -> 'a list -> 'b array
-val list_filter_i :
- (int -> 'a -> bool) -> 'a list -> 'a list
-
-(** [list_smartfilter f [a1...an] = List.filter f [a1...an]] but if for all i
- [f ai = true], then [list_smartfilter f l==l] *)
-val list_smartfilter : ('a -> bool) -> 'a list -> 'a list
-
-(** [list_index] returns the 1st index of an element in a list (counting from 1) *)
-val list_index : 'a -> 'a list -> int
-val list_index_f : ('a -> 'a -> bool) -> 'a -> 'a list -> int
-
-(** [list_unique_index x l] returns [Not_found] if [x] doesn't occur exactly once *)
-val list_unique_index : 'a -> 'a list -> int
-
-(** [list_index0] behaves as [list_index] except that it starts counting at 0 *)
-val list_index0 : 'a -> 'a list -> int
-val list_index0_f : ('a -> 'a -> bool) -> 'a -> 'a list -> int
-val list_iter3 : ('a -> 'b -> 'c -> unit) -> 'a list -> 'b list -> 'c list -> unit
-val list_iter_i : (int -> 'a -> unit) -> 'a list -> unit
-val list_fold_right_i : (int -> 'a -> 'b -> 'b) -> int -> 'a list -> 'b -> 'b
-val list_fold_left_i : (int -> 'a -> 'b -> 'a) -> int -> 'a -> 'b list -> 'a
-val list_fold_right_and_left :
- ('a -> 'b -> 'b list -> 'a) -> 'b list -> 'a -> 'a
-val list_fold_left3 : ('a -> 'b -> 'c -> 'd -> 'a) -> 'a -> 'b list -> 'c list -> 'd list -> 'a
-val list_for_all_i : (int -> 'a -> bool) -> int -> 'a list -> bool
-val list_except : 'a -> 'a list -> 'a list
-val list_remove : 'a -> 'a list -> 'a list
-val list_remove_first : 'a -> 'a list -> 'a list
-val list_remove_assoc_in_triple : 'a -> ('a * 'b * 'c) list -> ('a * 'b * 'c) list
-val list_assoc_snd_in_triple : 'a -> ('a * 'b * 'c) list -> 'b
-val list_for_all2eq : ('a -> 'b -> bool) -> 'a list -> 'b list -> bool
-val list_sep_last : 'a list -> 'a * 'a list
-val list_try_find_i : (int -> 'a -> 'b) -> int -> 'a list -> 'b
-val list_try_find : ('a -> 'b) -> 'a list -> 'b
-val list_uniquize : 'a list -> 'a list
-
-(** merges two sorted lists and preserves the uniqueness property: *)
-val list_merge_uniq : ('a -> 'a -> int) -> 'a list -> 'a list -> 'a list
-val list_subset : 'a list -> 'a list -> bool
-val list_chop : int -> 'a list -> 'a list * 'a list
-(* former [list_split_at] was a duplicate of [list_chop] *)
-val list_split_when : ('a -> bool) -> 'a list -> 'a list * 'a list
-val list_split_by : ('a -> bool) -> 'a list -> 'a list * 'a list
-val list_split3 : ('a * 'b * 'c) list -> 'a list * 'b list * 'c list
-val list_partition_by : ('a -> 'a -> bool) -> 'a list -> 'a list list
-val list_firstn : int -> 'a list -> 'a list
-val list_last : 'a list -> 'a
-val list_lastn : int -> 'a list -> 'a list
-val list_skipn : int -> 'a list -> 'a list
-val list_skipn_at_least : int -> 'a list -> 'a list
-val list_addn : int -> 'a -> 'a list -> 'a list
-val list_prefix_of : 'a list -> 'a list -> bool
-
-(** [list_drop_prefix p l] returns [t] if [l=p++t] else return [l] *)
-val list_drop_prefix : 'a list -> 'a list -> 'a list
-val list_drop_last : 'a list -> 'a list
-
-(** [map_append f [x1; ...; xn]] returns [(f x1)@(f x2)@...@(f xn)] *)
-val list_map_append : ('a -> 'b list) -> 'a list -> 'b list
-val list_join_map : ('a -> 'b list) -> 'a list -> 'b list
-
-(** raises [Invalid_argument] if the two lists don't have the same length *)
-val list_map_append2 : ('a -> 'b -> 'c list) -> 'a list -> 'b list -> 'c list
-val list_share_tails : 'a list -> 'a list -> 'a list * 'a list * 'a list
-
-(** [list_fold_map f e_0 [l_1...l_n] = e_n,[k_1...k_n]]
- where [(e_i,k_i)=f e_{i-1} l_i] *)
-val list_fold_map : ('a -> 'b -> 'a * 'c) -> 'a -> 'b list -> 'a * 'c list
-val list_fold_map' : ('b -> 'a -> 'c * 'a) -> 'b list -> 'a -> 'c list * 'a
-val list_map_assoc : ('a -> 'b) -> ('c * 'a) list -> ('c * 'b) list
-val list_assoc_f : ('a -> 'a -> bool) -> 'a -> ('a * 'b) list -> 'b
-
-(** A generic cartesian product: for any operator (**),
- [list_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 list_cartesian : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list
-
-(** [list_cartesians] is an n-ary cartesian product: it iterates
- [list_cartesian] over a list of lists. *)
-val list_cartesians : ('a -> 'b -> 'b) -> 'b -> 'a list list -> 'b list
-
-(** list_combinations [[a;b];[c;d]] returns [[a;c];[a;d];[b;c];[b;d]] *)
-val list_combinations : 'a list list -> 'a list list
-val list_combine3 : 'a list -> 'b list -> 'c list -> ('a * 'b * 'c) list
-
-(** Keep only those products that do not return None *)
-val list_cartesian_filter :
- ('a -> 'b -> 'c option) -> 'a list -> 'b list -> 'c list
-val list_cartesians_filter :
- ('a -> 'b -> 'b option) -> 'b -> 'a list list -> 'b list
-
-val list_union_map : ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b
-val list_factorize_left : ('a * 'b) list -> ('a * 'b list) list
+module List : CList.ExtS
(** {6 Arrays. } *)
diff --git a/library/declare.ml b/library/declare.ml
index 58ad1f00f..63a2c03a7 100644
--- a/library/declare.ml
+++ b/library/declare.ml
@@ -183,7 +183,7 @@ let declare_constant ?(internal = UserVerbose) id (cd,kind) =
(** Declaration of inductive blocks *)
let declare_inductive_argument_scopes kn mie =
- list_iter_i (fun i {mind_entry_consnames=lc} ->
+ List.iter_i (fun i {mind_entry_consnames=lc} ->
Notation.declare_ref_arguments_scope (IndRef (kn,i));
for j=1 to List.length lc do
Notation.declare_ref_arguments_scope (ConstructRef ((kn,i),j));
diff --git a/library/heads.ml b/library/heads.ml
index c86a91cc9..f3bcba770 100644
--- a/library/heads.ml
+++ b/library/heads.ml
@@ -113,8 +113,8 @@ let kind_of_head env t =
k+n-m,[],a
else
(* enough arguments to [cst] *)
- k,list_skipn n l,List.nth l (i-1) in
- let l' = list_tabulate (fun _ -> mkMeta 0) q @ rest in
+ k,List.skipn n l,List.nth l (i-1) in
+ let l' = List.tabulate (fun _ -> mkMeta 0) q @ rest in
aux k' l' a (with_subcase or with_case)
| ConstructorHead when with_case -> NotImmediatelyComputableHead
| x -> x
diff --git a/library/impargs.ml b/library/impargs.ml
index cd6365289..baf2e30ec 100644
--- a/library/impargs.ml
+++ b/library/impargs.ml
@@ -338,7 +338,7 @@ let set_manual_implicits env flags enriching autoimps l =
else l, None
with Not_found -> l, None
in
- if not (list_distinct l) then
+ if not (List.distinct l) then
error ("Some parameters are referred more than once.");
(* Compare with automatic implicits to recover printing data and names *)
let rec merge k l = function
@@ -435,7 +435,7 @@ let compute_global_implicits flags manual = function
(* Merge a manual explicitation with an implicit_status list *)
let merge_impls (cond,oldimpls) (_,newimpls) =
- let oldimpls,usersuffiximpls = list_chop (List.length newimpls) oldimpls in
+ let oldimpls,usersuffiximpls = List.chop (List.length newimpls) oldimpls in
cond, (List.map2 (fun orig ni ->
match orig with
| Some (_, Manual, _) -> orig
@@ -482,7 +482,7 @@ let subst_implicits_decl subst (r,imps as o) =
let r' = fst (subst_global subst r) in if r==r' then o else (r',imps)
let subst_implicits (subst,(req,l)) =
- (ImplLocal,list_smartmap (subst_implicits_decl subst) l)
+ (ImplLocal,List.smartmap (subst_implicits_decl subst) l)
let impls_of_context ctx =
List.rev_map
@@ -563,7 +563,7 @@ let rebuild_implicits (req,l) =
if flags.auto then
let newimpls = List.hd (compute_global_implicits flags [] ref) in
let p = List.length (snd newimpls) - userimplsize in
- let newimpls = on_snd (list_firstn p) newimpls in
+ let newimpls = on_snd (List.firstn p) newimpls in
[ref,List.map (fun o -> merge_impls o newimpls) oldimpls]
else
[ref,oldimpls]
diff --git a/library/lib.ml b/library/lib.ml
index 907d251e7..a8a9f0c26 100644
--- a/library/lib.ml
+++ b/library/lib.ml
@@ -48,7 +48,7 @@ let subst_objects subst seg =
if obj' == obj then node else
(id, obj')
in
- list_smartmap subst_one seg
+ List.smartmap subst_one seg
(*let load_and_subst_objects i prefix subst seg =
List.rev (List.fold_left (fun seg (id,obj as node) ->
diff --git a/library/libnames.ml b/library/libnames.ml
index 611cc9ad9..3555766f8 100644
--- a/library/libnames.ml
+++ b/library/libnames.ml
@@ -19,21 +19,21 @@ let pr_dirpath sl = (str (string_of_dirpath sl))
(* Pop the last n module idents *)
let pop_dirpath_n n dir =
- make_dirpath (list_skipn n (repr_dirpath dir))
+ make_dirpath (List.skipn n (repr_dirpath dir))
let pop_dirpath p = match repr_dirpath p with
| [] -> anomaly "dirpath_prefix: empty dirpath"
| _::l -> make_dirpath l
let is_dirpath_prefix_of d1 d2 =
- list_prefix_of (List.rev (repr_dirpath d1)) (List.rev (repr_dirpath d2))
+ List.prefix_of (List.rev (repr_dirpath d1)) (List.rev (repr_dirpath d2))
let chop_dirpath n d =
- let d1,d2 = list_chop n (List.rev (repr_dirpath d)) in
+ let d1,d2 = List.chop n (List.rev (repr_dirpath d)) in
make_dirpath (List.rev d1), make_dirpath (List.rev d2)
let drop_dirpath_prefix d1 d2 =
- let d = Util.list_drop_prefix (List.rev (repr_dirpath d1)) (List.rev (repr_dirpath d2)) in
+ let d = Util.List.drop_prefix (List.rev (repr_dirpath d1)) (List.rev (repr_dirpath d2)) in
make_dirpath (List.rev d)
let append_dirpath d1 d2 = make_dirpath (repr_dirpath d2 @ repr_dirpath d1)
@@ -116,7 +116,7 @@ let pr_path sp = str (string_of_path sp)
let restrict_path n sp =
let dir, s = repr_path sp in
- let dir' = list_firstn n (repr_dirpath dir) in
+ let dir' = List.firstn n (repr_dirpath dir) in
make_path (make_dirpath dir') s
(*s qualified names *)
diff --git a/library/library.ml b/library/library.ml
index a82b50481..bca44726d 100644
--- a/library/library.ml
+++ b/library/library.ml
@@ -598,7 +598,7 @@ let import_module export (loc,qid) =
let check_coq_overwriting p id =
let l = repr_dirpath p in
- if not !Flags.boot && l <> [] && string_of_id (list_last l) = "Coq" then
+ if not !Flags.boot && l <> [] && string_of_id (List.last l) = "Coq" then
errorlabstrm ""
(strbrk ("Cannot build module "^string_of_dirpath p^"."^string_of_id id^
": it starts with prefix \"Coq\" which is reserved for the Coq library."))
diff --git a/parsing/egramcoq.ml b/parsing/egramcoq.ml
index 488c425cc..b3fdd384a 100644
--- a/parsing/egramcoq.ml
+++ b/parsing/egramcoq.ml
@@ -97,7 +97,7 @@ let make_constr_action
failwith "Unexpected entry of type cases pattern")
| GramConstrListMark (n,b) :: tl ->
(* Rebuild expansions of ConstrList *)
- let heads,constrs = list_chop n constrs in
+ let heads,constrs = List.chop n constrs in
let constrlists =
if b then (heads@List.hd constrlists)::List.tl constrlists
else heads::constrlists
@@ -145,7 +145,7 @@ let make_cases_pattern_action
anomaly "Unexpected entry of type cases pattern or other")
| GramConstrListMark (n,b) :: tl ->
(* Rebuild expansions of ConstrList *)
- let heads,env = list_chop n env in
+ let heads,env = List.chop n env in
if b then
make (env,(heads@List.hd envlist)::List.tl envlist,hasbinders) tl
else
@@ -278,7 +278,7 @@ let freeze () = (!grammar_state, Lexer.freeze ())
(* We compare the current state of the grammar and the state to unfreeze,
by computing the longest common suffixes *)
let factorize_grams l1 l2 =
- if l1 == l2 then ([], [], l1) else list_share_tails l1 l2
+ if l1 == l2 then ([], [], l1) else List.share_tails l1 l2
let number_of_entries gcl =
List.fold_left (fun n (p,_) -> n + p) 0 gcl
diff --git a/parsing/g_tactic.ml4 b/parsing/g_tactic.ml4
index 77c86ad92..f62be2f5c 100644
--- a/parsing/g_tactic.ml4
+++ b/parsing/g_tactic.ml4
@@ -77,18 +77,18 @@ let check_for_coloneq =
Gram.Entry.of_parser "lpar_id_colon"
(fun strm ->
let rec skip_to_rpar p n =
- match get_tok (list_last (Stream.npeek n strm)) with
+ match get_tok (List.last (Stream.npeek n strm)) with
| KEYWORD "(" -> skip_to_rpar (p+1) (n+1)
| KEYWORD ")" -> if p=0 then n+1 else skip_to_rpar (p-1) (n+1)
| KEYWORD "." -> err ()
| _ -> skip_to_rpar p (n+1) in
let rec skip_names n =
- match get_tok (list_last (Stream.npeek n strm)) with
+ match get_tok (List.last (Stream.npeek n strm)) with
| IDENT _ | KEYWORD "_" -> skip_names (n+1)
| KEYWORD ":" -> skip_to_rpar 0 (n+1) (* skip a constr *)
| _ -> err () in
let rec skip_binders n =
- match get_tok (list_last (Stream.npeek n strm)) with
+ match get_tok (List.last (Stream.npeek n strm)) with
| KEYWORD "(" -> skip_binders (skip_names (n+1))
| IDENT _ | KEYWORD "_" -> skip_binders (n+1)
| KEYWORD ":=" -> ()
@@ -114,7 +114,7 @@ let mk_fix_tac (loc,id,bl,ann,ty) =
[([_],_,_)], None -> 1
| _, Some x ->
let ids = List.map snd (List.flatten (List.map pi1 bl)) in
- (try list_index (snd x) ids
+ (try List.index (snd x) ids
with Not_found -> error "No such fix variable.")
| _ -> error "Cannot guess decreasing argument of fix." in
(id,n,CProdN(loc,bl,ty))
@@ -164,7 +164,7 @@ let mkCLambdaN_simple bl c =
let loc = Loc.merge (fst (List.hd (pi1 (List.hd bl)))) (Constrexpr_ops.constr_loc c) in
mkCLambdaN_simple_loc loc bl c
-let loc_of_ne_list l = Loc.merge (fst (List.hd l)) (fst (list_last l))
+let loc_of_ne_list l = Loc.merge (fst (List.hd l)) (fst (List.last l))
let map_int_or_var f = function
| ArgArg x -> ArgArg (f x)
diff --git a/parsing/g_xml.ml4 b/parsing/g_xml.ml4
index 587e272dd..1527fa64b 100644
--- a/parsing/g_xml.ml4
+++ b/parsing/g_xml.ml4
@@ -148,17 +148,17 @@ let rec interp_xml_constr = function
| XmlTag (loc,"VAR",al,[]) ->
error "XML parser: unable to interp free variables"
| XmlTag (loc,"LAMBDA",al,(_::_ as xl)) ->
- let body,decls = list_sep_last xl in
+ let body,decls = List.sep_last xl in
let ctx = List.map interp_xml_decl decls in
List.fold_right (fun (na,t) b -> GLambda (loc, na, Explicit, t, b))
ctx (interp_xml_target body)
| XmlTag (loc,"PROD",al,(_::_ as xl)) ->
- let body,decls = list_sep_last xl in
+ let body,decls = List.sep_last xl in
let ctx = List.map interp_xml_decl decls in
List.fold_right (fun (na,t) b -> GProd (loc, na, Explicit, t, b))
ctx (interp_xml_target body)
| XmlTag (loc,"LETIN",al,(_::_ as xl)) ->
- let body,defs = list_sep_last xl in
+ let body,defs = List.sep_last xl in
let ctx = List.map interp_xml_def defs in
List.fold_right (fun (na,t) b -> GLetIn (loc, na, t, b))
ctx (interp_xml_target body)
@@ -176,7 +176,7 @@ let rec interp_xml_constr = function
let p = interp_xml_patternsType x in
let tm = interp_xml_inductiveTerm y in
let vars = compute_branches_lengths ind in
- let brs = list_map_i (fun i c -> (i,vars.(i),interp_xml_pattern c)) 0 yl
+ let brs = List.map_i (fun i c -> (i,vars.(i),interp_xml_pattern c)) 0 yl
in
let mat = simple_cases_matrix_of_branches ind brs in
let nparams,n = compute_inductive_nargs ind in
@@ -188,11 +188,11 @@ let rec interp_xml_constr = function
GRef (loc, ConstructRef (get_xml_constructor al))
| XmlTag (loc,"FIX",al,xl) ->
let li,lnct = List.split (List.map interp_xml_FixFunction xl) in
- let ln,lc,lt = list_split3 lnct in
+ let ln,lc,lt = List.split3 lnct in
let lctx = List.map (fun _ -> []) ln in
GRec (loc, GFix (Array.of_list li, get_xml_noFun al), Array.of_list ln, Array.of_list lctx, Array.of_list lc, Array.of_list lt)
| XmlTag (loc,"COFIX",al,xl) ->
- let ln,lc,lt = list_split3 (List.map interp_xml_CoFixFunction xl) in
+ let ln,lc,lt = List.split3 (List.map interp_xml_CoFixFunction xl) in
GRec (loc, GCoFix (get_xml_noFun al), Array.of_list ln, [||], Array.of_list lc, Array.of_list lt)
| XmlTag (loc,"CAST",al,[x1;x2]) ->
GCast (loc, interp_xml_term x1, CastConv (interp_xml_type x2))
diff --git a/parsing/lexer.ml4 b/parsing/lexer.ml4
index 14644339c..f277acfd1 100644
--- a/parsing/lexer.ml4
+++ b/parsing/lexer.ml4
@@ -402,7 +402,7 @@ and progress_utf8 last nj n c tt cs =
if n=1 then
update_longest_valid_token last (nj+n) tt cs
else
- match Util.list_skipn (nj+1) (Stream.npeek (nj+n) cs) with
+ match Util.List.skipn (nj+1) (Stream.npeek (nj+n) cs) with
| l when List.length l = n-1 ->
List.iter (check_utf8_trailing_byte cs) l;
let tt = List.fold_left (fun tt c -> CharMap.find c tt.branch) tt l in
diff --git a/parsing/pcoq.ml4 b/parsing/pcoq.ml4
index 4df8d9490..8fef987e6 100644
--- a/parsing/pcoq.ml4
+++ b/parsing/pcoq.ml4
@@ -538,7 +538,7 @@ let find_position_gen forpat ensure assoc lev =
Some (Level (constr_level n)), None, None, None
let remove_levels n =
- level_stack := list_skipn n !level_stack
+ level_stack := List.skipn n !level_stack
let rec list_mem_assoc_triple x = function
| [] -> false
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index 53c146bfc..423c95509 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -367,7 +367,7 @@ let discriminate_tac cstr p gls =
let build_term_to_complete uf meta pac =
let cinfo = get_constructor_info uf pac.cnode in
let real_args = List.map (fun i -> constr_of_term (term uf i)) pac.args in
- let dummy_args = List.rev (list_tabulate meta pac.arity) in
+ let dummy_args = List.rev (List.tabulate meta pac.arity) in
let all_args = List.rev_append real_args dummy_args in
applistc (mkConstruct cinfo.ci_constr) all_args
diff --git a/plugins/decl_mode/decl_interp.ml b/plugins/decl_mode/decl_interp.ml
index 5e128bc42..f2f978ccd 100644
--- a/plugins/decl_mode/decl_interp.ml
+++ b/plugins/decl_mode/decl_interp.ml
@@ -57,7 +57,7 @@ let intern_hyp iconstr globs = function
Hprop (intern_statement iconstr globs st)
let intern_hyps iconstr globs hyps =
- snd (list_fold_map (intern_hyp iconstr) globs hyps)
+ snd (List.fold_map (intern_hyp iconstr) globs hyps)
let intern_cut intern_it globs cut=
let nglobs,nstat=intern_it globs cut.cut_stat in
@@ -74,10 +74,10 @@ let intern_hyp_list args globs =
let intern_one globs (loc,(id,opttyp)) =
(add_var id globs),
(loc,(id,Option.map (intern_constr globs) opttyp)) in
- list_fold_map intern_one globs args
+ List.fold_map intern_one globs args
let intern_suffices_clause globs (hyps,c) =
- let nglobs,nhyps = list_fold_map (intern_hyp intern_constr) globs hyps in
+ let nglobs,nhyps = List.fold_map (intern_hyp intern_constr) globs hyps in
nglobs,(nhyps,intern_constr_or_thesis nglobs c)
let intern_fundecl args body globs=
@@ -340,7 +340,7 @@ let interp_cases info sigma env params (pat:cases_pattern_expr) hyps =
(fun (loc,(id,_)) ->
GVar (loc,id)) params in
let dum_args=
- list_tabulate
+ List.tabulate
(fun _ -> GHole (Loc.ghost,Evar_kinds.QuestionMark (Evar_kinds.Define false)))
oib.Declarations.mind_nrealargs in
glob_app(Loc.ghost,rind,rparams@rparams_rec@dum_args) in
diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml
index 31e15081b..1a0d05047 100644
--- a/plugins/decl_mode/decl_proof_instr.ml
+++ b/plugins/decl_mode/decl_proof_instr.ml
@@ -857,7 +857,7 @@ let build_per_info etype casee gls =
match etype with
ET_Induction -> mind.mind_nparams_rec,Some (snd ind)
| _ -> mind.mind_nparams,None in
- let params,real_args = list_chop nparams args in
+ let params,real_args = List.chop nparams args in
let abstract_obj c body =
let typ=pf_type_of gls c in
lambda_create env (typ,subst_term c body) in
@@ -953,7 +953,7 @@ let rec tree_of_pats ((id,_) as cpl) pats =
let nexti i ati =
if i = pred cnum then
let nargs =
- list_map_i (fun j a -> (a,ati.(j))) 0 args in
+ List.map_i (fun j a -> (a,ati.(j))) 0 args in
Some (Idset.singleton id,
tree_of_pats cpl (nargs::rest_args::stack))
else None
@@ -998,7 +998,7 @@ let rec add_branch ((id,_) as cpl) pats tree=
let nexti i ati =
if i = pred cnum then
let nargs =
- list_map_i (fun j a -> (a,ati.(j))) 0 args in
+ List.map_i (fun j a -> (a,ati.(j))) 0 args in
Some (Idset.add id ids,
add_branch cpl (nargs::rest_args::stack)
(skip_args t ids (Array.length ati)))
@@ -1013,7 +1013,7 @@ let rec add_branch ((id,_) as cpl) pats tree=
let mapi i ati bri =
if i = pred cnum then
let nargs =
- list_map_i (fun j a -> (a,ati.(j))) 0 args in
+ List.map_i (fun j a -> (a,ati.(j))) 0 args in
append_branch cpl 0
(nargs::rest_args::stack) bri
else bri in
@@ -1051,7 +1051,7 @@ let thesis_for obj typ per_info env=
errorlabstrm "thesis_for"
((Printer.pr_constr_env env obj) ++ spc () ++
str"cannot give an induction hypothesis (wrong inductive type).") in
- let params,args = list_chop per_info.per_nparams all_args in
+ let params,args = List.chop per_info.per_nparams all_args in
let _ = if not (List.for_all2 eq_constr params per_info.per_params) then
errorlabstrm "thesis_for"
((Printer.pr_constr_env env obj) ++ spc () ++
@@ -1182,7 +1182,7 @@ let hrec_for fix_id per_info gls obj_id =
let rc,hd1=decompose_prod typ in
let cind,all_args=decompose_app typ in
let ind = destInd cind in assert (ind=per_info.per_ind);
- let params,args= list_chop per_info.per_nparams all_args in
+ let params,args= List.chop per_info.per_nparams all_args in
assert begin
try List.for_all2 eq_constr params per_info.per_params with
Invalid_argument _ -> false end;
@@ -1203,8 +1203,8 @@ let rec execute_cases fix_name per_info tacnext args objs nhrec tree gls =
match List.assoc id args with
[None,br_args] ->
let all_metas =
- list_tabulate (fun n -> mkMeta (succ n)) (nparams + nhyps) in
- let param_metas,hyp_metas = list_chop nparams all_metas in
+ List.tabulate (fun n -> mkMeta (succ n)) (nparams + nhyps) in
+ let param_metas,hyp_metas = List.chop nparams all_metas in
tclTHEN
(tclDO nhrec introf)
(tacnext
@@ -1221,7 +1221,7 @@ let rec execute_cases fix_name per_info tacnext args objs nhrec tree gls =
let ctyp=pf_type_of gls casee in
let hd,all_args = decompose_app (special_whd gls ctyp) in
let _ = assert (destInd hd = ind) in (* just in case *)
- let params,real_args = list_chop nparams all_args in
+ let params,real_args = List.chop nparams all_args in
let abstract_obj c body =
let typ=pf_type_of gls c in
lambda_create env (typ,subst_term c body) in
diff --git a/plugins/extraction/common.ml b/plugins/extraction/common.ml
index b6b7e5833..8cceb2a11 100644
--- a/plugins/extraction/common.ml
+++ b/plugins/extraction/common.ml
@@ -349,7 +349,7 @@ let rec mp_renaming_fun mp = match mp with
| MPfile _ ->
assert (modular ()); (* see [at_toplevel] above *)
assert (get_phase () = Pre);
- let current_mpfile = (list_last (get_visible ())).mp in
+ let current_mpfile = (List.last (get_visible ())).mp in
if mp <> current_mpfile then mpfiles_add mp;
[string_of_modfile mp]
@@ -496,7 +496,7 @@ let fstlev_ks k = function
let pp_ocaml_local k prefix mp rls olab =
(* what is the largest prefix of [mp] that belongs to [visible]? *)
assert (k <> Mod || mp <> prefix); (* mp as whole module isn't in itself *)
- let rls' = list_skipn (mp_length prefix) rls in
+ let rls' = List.skipn (mp_length prefix) rls in
let k's = fstlev_ks k rls' in
(* Reference r / module path mp is of the form [<prefix>.s.<...>]. *)
if not (visible_clash prefix k's) then dottify rls'
diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml
index 588fe0cf2..b5cdec3ec 100644
--- a/plugins/extraction/extract_env.ml
+++ b/plugins/extraction/extract_env.ml
@@ -150,9 +150,9 @@ let factor_fix env l cb msb =
if n = 1 then [|l|], recd, msb
else begin
if List.length msb < n-1 then raise Impossible;
- let msb', msb'' = list_chop (n-1) msb in
+ let msb', msb'' = List.chop (n-1) msb in
let labels = Array.make n l in
- list_iter_i
+ List.iter_i
(fun j ->
function
| (l,SFBconst cb') ->
@@ -207,7 +207,7 @@ let env_for_mtb_with_def env mp seb idl =
in
let l = label_of_id (List.hd idl) in
let spot = function (l',SFBconst _) -> l = l' | _ -> false in
- let before = fst (list_split_when spot sig_b) in
+ let before = fst (List.split_when spot sig_b) in
Modops.add_signature mp before empty_delta_resolver env
(* From a [structure_body] (i.e. a list of [structure_field_body])
diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml
index 0556f6d77..e1bbcf9c7 100644
--- a/plugins/extraction/extraction.ml
+++ b/plugins/extraction/extraction.ml
@@ -193,7 +193,7 @@ let parse_ind_args si args relmax =
let oib_equal o1 o2 =
id_ord o1.mind_typename o2.mind_typename = 0 &&
- list_equal eq_rel_declaration o1.mind_arity_ctxt o2.mind_arity_ctxt &&
+ List.equal eq_rel_declaration o1.mind_arity_ctxt o2.mind_arity_ctxt &&
begin match o1.mind_arity, o2.mind_arity with
| Monomorphic {mind_user_arity=c1; mind_sort=s1},
Monomorphic {mind_user_arity=c2; mind_sort=s2} ->
@@ -206,10 +206,10 @@ let mib_equal m1 m2 =
m1.mind_record = m2.mind_record &&
m1.mind_finite = m2.mind_finite &&
m1.mind_ntypes = m2.mind_ntypes &&
- list_equal eq_named_declaration m1.mind_hyps m2.mind_hyps &&
+ List.equal eq_named_declaration m1.mind_hyps m2.mind_hyps &&
m1.mind_nparams = m2.mind_nparams &&
m1.mind_nparams_rec = m2.mind_nparams_rec &&
- list_equal eq_rel_declaration m1.mind_params_ctxt m2.mind_params_ctxt &&
+ List.equal eq_rel_declaration m1.mind_params_ctxt m2.mind_params_ctxt &&
m1.mind_constraints = m2.mind_constraints
(*S Extraction of a type. *)
@@ -439,7 +439,7 @@ and extract_ind env kn = (* kn is supposed to be in long form *)
| _ -> []
in
let field_names =
- list_skipn mib.mind_nparams (names_prod mip0.mind_user_lc.(0)) in
+ List.skipn mib.mind_nparams (names_prod mip0.mind_user_lc.(0)) in
assert (List.length field_names = List.length typ);
let projs = ref Cset.empty in
let mp,d,_ = repr_mind kn in
@@ -674,7 +674,7 @@ and extract_cst_app env mle mlt kn args =
let mla =
if not magic1 then
try
- let l,l' = list_chop (projection_arity (ConstRef kn)) mla in
+ let l,l' = List.chop (projection_arity (ConstRef kn)) mla in
if l' <> [] then (List.map (fun _ -> MLexn "Proj Args") l) @ l'
else mla
with _ -> mla
@@ -697,7 +697,7 @@ and extract_cst_app env mle mlt kn args =
(* Partially applied function with some logical arg missing.
We complete via eta and expunge logical args. *)
let ls' = ls-la in
- let s' = list_skipn la s in
+ let s' = List.skipn la s in
let mla = (List.map (ast_lift ls') mla) @ (eta_args_sign ls' s') in
let e = anonym_or_dummy_lams (mlapp head mla) s' in
put_magic_if magic2 (remove_n_lams (List.length optdummy) e)
@@ -729,7 +729,7 @@ and extract_cons_app env mle mlt (((kn,i) as ip,j) as cp) args =
let la = List.length args in
assert (la <= ls + params_nb);
let la' = max 0 (la - params_nb) in
- let args' = list_lastn la' args in
+ let args' = List.lastn la' args in
(* Now, we build the expected type of the constructor *)
let metas = List.map new_meta args' in
(* If stored and expected types differ, then magic! *)
@@ -758,7 +758,7 @@ and extract_cons_app env mle mlt (((kn,i) as ip,j) as cp) args =
then put_magic_if (magic2 && not magic1) (head mla)
else (* [ params_nb <= la <= ls + params_nb ] *)
let ls' = params_nb + ls - la in
- let s' = list_lastn ls' s in
+ let s' = List.lastn ls' s in
let mla = (List.map (ast_lift ls') mla) @ (eta_args_sign ls' s') in
put_magic_if magic2 (anonym_or_dummy_lams (head mla) s')
@@ -847,7 +847,7 @@ let decomp_lams_eta_n n m env c t =
let rels',c = decompose_lam c in
let d = n - m in
(* we'd better keep rels' as long as possible. *)
- let rels = (list_firstn d rels) @ rels' in
+ let rels = (List.firstn d rels) @ rels' in
let eta_args = List.rev_map mkRel (interval 1 d) in
rels, applist (lift d c,eta_args)
@@ -886,7 +886,7 @@ let extract_std_constant env kn body typ =
and m = nb_lam body in
if n <= m then decompose_lam_n n body
else
- let s,s' = list_chop m s in
+ let s,s' = List.chop m s in
if List.for_all ((=) Keep) s' &&
(lang () = Haskell || sign_kind s <> UnsafeLogicalSig)
then decompose_lam_n m body
@@ -895,7 +895,7 @@ let extract_std_constant env kn body typ =
(* Should we do one eta-expansion to avoid non-generalizable '_a ? *)
let rels, c =
let n = List.length rels in
- let s,s' = list_chop n s in
+ let s,s' = List.chop n s in
let k = sign_kind s in
let empty_s = (k = EmptySig || k = SafeLogicalSig) in
if lang () = Ocaml && empty_s && not (gentypvar_ok c)
@@ -904,8 +904,8 @@ let extract_std_constant env kn body typ =
else rels,c
in
let n = List.length rels in
- let s = list_firstn n s in
- let l,l' = list_chop n l in
+ let s = List.firstn n s in
+ let l,l' = List.chop n l in
let t' = type_recomp (l',t') in
(* The initial ML environment. *)
let mle = List.fold_left Mlenv.push_std_type Mlenv.empty l in
diff --git a/plugins/extraction/mlutil.ml b/plugins/extraction/mlutil.ml
index 3b89386d4..3716695b8 100644
--- a/plugins/extraction/mlutil.ml
+++ b/plugins/extraction/mlutil.ml
@@ -534,7 +534,7 @@ let is_regular_match br =
match pat with
| Pusual r -> r
| Pcons (r,l) ->
- if not (list_for_all_i (fun i -> (=) (Prel i)) 1 (List.rev l))
+ if not (List.for_all_i (fun i -> (=) (Prel i)) 1 (List.rev l))
then raise Impossible;
r
| _ -> raise Impossible
@@ -636,9 +636,9 @@ let eta_red e =
if m = n then
[], f, a
else if m < n then
- list_skipn m ids, f, a
+ List.skipn m ids, f, a
else (* m > n *)
- let a1,a2 = list_chop (m-n) a in
+ let a1,a2 = List.chop (m-n) a in
[], MLapp (f,a1), a2
in
let p = List.length args in
@@ -969,7 +969,7 @@ and simpl_case o typ br e =
else ([], Pwild, ast_pop f)
in
let brl = Array.to_list br in
- let brl_opt = list_filter_i (fun i _ -> not (Intset.mem i ints)) brl in
+ let brl_opt = List.filter_i (fun i _ -> not (Intset.mem i ints)) brl in
let brl_opt = brl_opt @ [last_br] in
MLcase (typ, e, Array.of_list brl_opt)
| None -> MLcase (typ, e, br)
@@ -1022,8 +1022,8 @@ let kill_dummy_lams c =
| Keep :: bl -> fst_kill (n+1) bl
in
let skip = max 0 ((fst_kill 0 bl) - 1) in
- let ids_skip, ids = list_chop skip ids in
- let _, bl = list_chop skip bl in
+ let ids_skip, ids = List.chop skip ids in
+ let _, bl = List.chop skip bl in
let c = named_lams ids_skip c in
let ids',c = kill_some_lams bl (ids,c) in
ids, named_lams ids' c
@@ -1051,7 +1051,7 @@ let case_expunge s e =
let m = List.length s in
let n = nb_lams e in
let p = if m <= n then collect_n_lams m e
- else eta_expansion_sign (list_skipn n s) (collect_lams e) in
+ else eta_expansion_sign (List.skipn n s) (collect_lams e) in
kill_some_lams (List.rev s) p
(*s [term_expunge] takes a function [fun idn ... id1 -> c]
@@ -1085,7 +1085,7 @@ let kill_dummy_args ids r t =
let a = List.map (killrec n) a in
let a = List.map (ast_lift k) a in
let a = select_via_bl bl (a @ (eta_args k)) in
- named_lams (list_firstn k ids) (MLapp (ast_lift k e, a))
+ named_lams (List.firstn k ids) (MLapp (ast_lift k e, a))
| e when found n e ->
let a = select_via_bl bl (eta_args m) in
named_lams ids (MLapp (ast_lift m e, a))
@@ -1164,7 +1164,7 @@ let general_optimize_fix f ids n args m c =
| MLrel j when v.(j-1)>=0 ->
if ast_occurs (j+1) c then raise Impossible else v.(j-1)<-(-i-1)
| _ -> raise Impossible
- in list_iter_i aux args;
+ in List.iter_i aux args;
let args_f = List.rev_map (fun i -> MLrel (i+m+1)) (Array.to_list v) in
let new_f = anonym_tmp_lams (MLapp (MLrel (n+m+1),args_f)) m in
let new_c = named_lams ids (normalize (MLapp ((ast_subst new_f c),args))) in
diff --git a/plugins/extraction/modutil.ml b/plugins/extraction/modutil.ml
index 1211bbf17..8659de03e 100644
--- a/plugins/extraction/modutil.ml
+++ b/plugins/extraction/modutil.ml
@@ -30,7 +30,7 @@ let se_iter do_decl do_spec =
| MTfunsig (_,mt,mt') -> mt_iter mt; mt_iter mt'
| MTwith (mt,ML_With_type(idl,l,t))->
let mp_mt = msid_of_mt mt in
- let l',idl' = list_sep_last idl in
+ let l',idl' = List.sep_last idl in
let mp_w =
List.fold_left (fun mp l -> MPdot(mp,label_of_id l)) mp_mt idl'
in
diff --git a/plugins/extraction/ocaml.ml b/plugins/extraction/ocaml.ml
index 34ae1f9ad..951049b7b 100644
--- a/plugins/extraction/ocaml.ml
+++ b/plugins/extraction/ocaml.ml
@@ -114,7 +114,7 @@ let pp_one_field r i = function
let pp_field r fields i = pp_one_field r i (List.nth fields i)
-let pp_fields r fields = list_map_i (pp_one_field r) 0 fields
+let pp_fields r fields = List.map_i (pp_one_field r) 0 fields
(*s Pretty-printing of types. [par] is a boolean indicating whether parentheses
are needed or not. *)
@@ -189,7 +189,7 @@ let rec pp_expr par env args =
hv 0 (apply2 (pp_letin pp_id pp_a1 pp_a2))
| MLglob r ->
(try
- let args = list_skipn (projection_arity r) args in
+ let args = List.skipn (projection_arity r) args in
let record = List.hd args in
pp_apply (record ++ str "." ++ pp_global Term r) par (List.tl args)
with _ -> apply (pp_global Term r))
@@ -642,7 +642,7 @@ and pp_module_type params = function
| MTwith(mt,ML_With_type(idl,vl,typ)) ->
let ids = pp_parameters (rename_tvars keywords vl) in
let mp_mt = msid_of_mt mt in
- let l,idl' = list_sep_last idl in
+ let l,idl' = List.sep_last idl in
let mp_w =
List.fold_left (fun mp l -> MPdot(mp,label_of_id l)) mp_mt idl'
in
diff --git a/plugins/extraction/table.ml b/plugins/extraction/table.ml
index 2dd07b2f2..6151abf6e 100644
--- a/plugins/extraction/table.ml
+++ b/plugins/extraction/table.ml
@@ -671,7 +671,7 @@ let add_implicits r l =
else err (int i ++ str " is not a valid argument number for " ++
safe_pr_global r)
| ArgId id ->
- (try list_index (Name id) names
+ (try List.index (Name id) names
with Not_found ->
err (str "No argument " ++ pr_id id ++ str " for " ++
safe_pr_global r))
@@ -877,7 +877,7 @@ let extract_inductive r s l optstr =
Lib.add_anonymous_leaf (in_customs (g,[],s));
Option.iter (fun s -> Lib.add_anonymous_leaf (in_custom_matchs (g,s)))
optstr;
- list_iter_i
+ List.iter_i
(fun j s ->
let g = ConstructRef (ip,succ j) in
Lib.add_anonymous_leaf (inline_extraction (true,[g]));
diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml
index 7abb4dc52..f4cc397bc 100644
--- a/plugins/firstorder/formula.ml
+++ b/plugins/firstorder/formula.ml
@@ -142,7 +142,7 @@ let build_atoms gl metagen side cciterm =
let g i _ (_,_,t) =
build_rec env polarity (lift i t) in
let f l =
- list_fold_left_i g (1-(List.length l)) () l in
+ List.fold_left_i g (1-(List.length l)) () l in
if polarity && (* we have a constant constructor *)
array_exists (function []->true|_->false) v
then trivial:=true;
@@ -152,7 +152,7 @@ let build_atoms gl metagen side cciterm =
let v =(ind_hyps 1 i l gl).(0) in
let g i _ (_,_,t) =
build_rec (var::env) polarity (lift i t) in
- list_fold_left_i g (2-(List.length l)) () v
+ List.fold_left_i g (2-(List.length l)) () v
| Forall(_,b)->
let var=mkMeta (metagen true) in
build_rec (var::env) polarity b
@@ -224,7 +224,7 @@ let build_formula side nam typ gl metagen=
| And(_,_,_) -> Rand
| Or(_,_,_) -> Ror
| Exists (i,l) ->
- let (_,_,d)=list_last (ind_hyps 0 i l gl).(0) in
+ let (_,_,d)=List.last (ind_hyps 0 i l gl).(0) in
Rexists(m,d,trivial)
| Forall (_,a) -> Rforall
| Arrow (a,b) -> Rarrow in
diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml
index e3367b4c2..414afad46 100644
--- a/plugins/firstorder/instances.ml
+++ b/plugins/firstorder/instances.ml
@@ -112,7 +112,7 @@ let mk_open_instance id gl m t=
match nam with
Name id -> id
| Anonymous -> dummy_bvid in
- let revt=substl (list_tabulate (fun i->mkRel (m-i)) m) t in
+ let revt=substl (List.tabulate (fun i->mkRel (m-i)) m) t in
let rec aux n avoid=
if n=0 then [] else
let nid=(fresh_id avoid var_id gl) in
diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml
index 37d9edef8..7acabaaa4 100644
--- a/plugins/firstorder/rules.ml
+++ b/plugins/firstorder/rules.ml
@@ -129,7 +129,7 @@ let ll_ind_tac ind largs backtrack id continue seq gl=
let head=mkApp ((lift p (constr_of_global id)),[|capply|]) in
it_mkLambda_or_LetIn head rc in
let lp=Array.length rcs in
- let newhyps=list_tabulate myterm lp in
+ let newhyps=List.tabulate myterm lp in
tclIFTHENELSE
(tclTHENLIST
[generalize newhyps;
diff --git a/plugins/firstorder/sequent.ml b/plugins/firstorder/sequent.ml
index 6d00e8d9f..c648939bb 100644
--- a/plugins/firstorder/sequent.ml
+++ b/plugins/firstorder/sequent.ml
@@ -191,9 +191,9 @@ let empty_seq depth=
depth=depth}
let expand_constructor_hints =
- list_map_append (function
+ List.map_append (function
| IndRef ind ->
- list_tabulate (fun i -> ConstructRef (ind,i+1))
+ List.tabulate (fun i -> ConstructRef (ind,i+1))
(Inductiveops.nconstructors ind)
| gr ->
[gr])
diff --git a/plugins/firstorder/unify.ml b/plugins/firstorder/unify.ml
index 33ea0ac8a..c3726f1a8 100644
--- a/plugins/firstorder/unify.ml
+++ b/plugins/firstorder/unify.ml
@@ -124,7 +124,7 @@ let unif_atoms i dom t1 t2=
| Not_found ->Some (Phantom dom)
let renum_metas_from k n t= (* requires n = max (free_rels t) *)
- let l=list_tabulate (fun i->mkMeta (k+i)) n in
+ let l=List.tabulate (fun i->mkMeta (k+i)) n in
substl l t
let more_general (m1,t1) (m2,t2)=
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index 8ea4be631..0796930fc 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -93,7 +93,7 @@ let observe_tac s = observe_tac_stream (str s)
let list_chop ?(msg="") n l =
try
- list_chop n l
+ List.chop n l
with Failure (msg') ->
failwith (msg ^ msg')
@@ -319,7 +319,7 @@ let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
)
in
let new_type_of_hyp,ctxt_size,witness_fun =
- list_fold_left_i
+ List.fold_left_i
(fun i (end_of_type,ctxt_size,witness_fun) ((x',b',t') as decl) ->
try
let witness = Intmap.find i sub in
@@ -1168,7 +1168,7 @@ let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams :
typess
in
let pte_to_fix,rev_info =
- list_fold_left_i
+ List.fold_left_i
(fun i (acc_map,acc_info) (pte,_,_) ->
let infos = info_array.(i) in
let type_args,_ = decompose_prod infos.types in
@@ -1557,7 +1557,7 @@ let prove_principle_for_gen
(* str "real_rec_arg_num := " ++ int real_rec_arg_num ++ fnl () ++ *)
(* str "npost_rec_arg := " ++ int npost_rec_arg ); *)
let (post_rec_arg,pre_rec_arg) =
- Util.list_chop npost_rec_arg princ_info.args
+ Util.List.chop npost_rec_arg princ_info.args
in
let rec_arg_id =
match List.rev post_rec_arg with
@@ -1624,7 +1624,7 @@ let prove_principle_for_gen
Elim.h_decompose_and (mkVar hid);
(fun g ->
let new_hyps = pf_ids_of_hyps g in
- tcc_list := List.rev (list_subtract new_hyps (hid::hyps));
+ tcc_list := List.rev (List.subtract new_hyps (hid::hyps));
if !tcc_list = []
then
begin
diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml
index e2dc149b0..f2eb4b23c 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -88,7 +88,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
Nameops.out_name x,None,compose_prod real_args (mkSort new_sort)
in
let new_predicates =
- list_map_i
+ List.map_i
change_predicate_sort
0
princ_type_info.predicates
@@ -252,7 +252,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
in
let pre_res =
replace_vars
- (list_map_i (fun i id -> (id, mkRel i)) 1 ptes_vars)
+ (List.map_i (fun i id -> (id, mkRel i)) 1 ptes_vars)
(lift (List.length ptes_vars) pre_res)
in
it_mkProd_or_LetIn
@@ -460,7 +460,7 @@ let get_funs_constant mp dp =
let first_params = List.hd l_params in
List.iter
(fun params ->
- if not (list_equal (fun (n1, c1) (n2, c2) -> n1 = n2 && eq_constr c1 c2) first_params params)
+ if not (List.equal (fun (n1, c1) (n2, c2) -> n1 = n2 && eq_constr c1 c2) first_params params)
then error "Not a mutal recursive block"
)
l_params
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index d11c01672..1c2193449 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -89,7 +89,7 @@ let combine_results
in (* and then we flatten the map *)
{
result = List.concat pre_result;
- to_avoid = list_union res1.to_avoid res2.to_avoid
+ to_avoid = List.union res1.to_avoid res2.to_avoid
}
@@ -269,7 +269,7 @@ let make_discr_match_el =
end
*)
let make_discr_match_brl i =
- list_map_i
+ List.map_i
(fun j (_,idl,patl,_) ->
if j=i
then (Loc.ghost,idl,patl, mkGRef (Lazy.force coq_True_ref))
@@ -659,7 +659,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
let case_pats = build_constructors_of_type ind [] in
assert (Array.length case_pats = 2);
let brl =
- list_map_i
+ List.map_i
(fun i x -> (Loc.ghost,[],[case_pats.(i)],x))
0
[lhs;rhs]
@@ -748,7 +748,7 @@ and build_entry_lc_from_case env funname make_discr
{
result = List.concat (List.map (fun r -> r.result) results);
to_avoid =
- List.fold_left (fun acc r -> list_union acc r.to_avoid) [] results
+ List.fold_left (fun acc r -> List.union acc r.to_avoid) [] results
}
and build_entry_lc_from_case_term env types funname make_discr patterns_to_prevent brl avoid
@@ -818,7 +818,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve
let those_pattern_preconds =
(List.flatten
(
- list_map3
+ List.map3
(fun pat e typ_as_constr ->
let this_pat_ids = ids_of_pat pat in
let typ = Detyping.detype false [] (Termops.names_of_rel_context new_env) typ_as_constr in
@@ -977,7 +977,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let mib,_ = Global.lookup_inductive ind in
let nparam = mib.Declarations.mind_nparams in
let params,arg' =
- ((Util.list_chop nparam args'))
+ ((Util.List.chop nparam args'))
in
let rt_typ =
GApp(Loc.ghost,
@@ -1000,7 +1000,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
match kind_of_term eq'_as_constr with
| App(_,[|_;_;ty;_|]) ->
let ty = Array.to_list (snd (destApp ty)) in
- let ty' = snd (Util.list_chop nparam ty) in
+ let ty' = snd (Util.List.chop nparam ty) in
List.fold_left2
(fun acc var_as_constr arg ->
if isRel var_as_constr
@@ -1238,7 +1238,7 @@ let compute_params_name relnames (args : (Names.name * Glob_term.glob_constr * b
let l = ref [] in
let _ =
try
- list_iter_i
+ List.iter_i
(fun i ((n,nt,is_defined) as param) ->
if array_for_all
(fun l ->
@@ -1362,7 +1362,7 @@ let do_build_inductive
in
let rel_arity i funargs = (* Reduilding arities (with parameters) *)
let rel_first_args :(Names.name * Glob_term.glob_constr * bool ) list =
- (snd (list_chop nrel_params funargs))
+ (snd (List.chop nrel_params funargs))
in
List.fold_right
(fun (n,t,is_defined) acc ->
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index 52562fb37..18b2bbe2f 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -21,7 +21,7 @@ let is_rec_info scheme_info =
Util.Intset.exists (fun i -> i >= min && i< max) free_rels_in_br
)
in
- Util.list_fold_left_i test_branche 1 false (List.rev scheme_info.Tactics.branches)
+ Util.List.fold_left_i test_branche 1 false (List.rev scheme_info.Tactics.branches)
let choose_dest_or_ind scheme_info =
if is_rec_info scheme_info
@@ -337,7 +337,7 @@ let generate_principle on_error
in
let funs_kn = Array.of_list (List.map fname_kn fix_rec_l) in
let _ =
- list_map_i
+ List.map_i
(fun i x ->
let princ = destConst (Indrec.lookup_eliminator (ind_kn,i) (InProp)) in
let princ_type = Typeops.type_of_constant (Global.env()) princ
@@ -392,7 +392,7 @@ let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas
if List.length names = 1 then 1
else error "Recursive argument must be specified"
| Some wf_arg ->
- list_index (Name wf_arg) names
+ List.index (Name wf_arg) names
in
let unbounded_eq =
let f_app_args =
@@ -516,7 +516,7 @@ let decompose_lambda_n_assum_constr_expr =
(n - nal_length)
(Constrexpr.CLambdaN(lambda_loc,bl,e'))
else
- let nal_keep,nal_expr = list_chop n nal in
+ let nal_keep,nal_expr = List.chop n nal in
(List.rev (Constrexpr.LocalRawAssum(nal_keep,bk,nal_type)::acc),
Constrexpr.CLambdaN(lambda_loc,(nal_expr,bk,nal_type)::bl,e')
)
@@ -548,7 +548,7 @@ let decompose_prod_n_assum_constr_expr =
(if bl = [] then e' else (Constrexpr.CLambdaN(lambda_loc,bl,e')))
else
(* let _ = Pp.msgnl (str "second case") in *)
- let nal_keep,nal_expr = list_chop n nal in
+ let nal_keep,nal_expr = List.chop n nal in
(List.rev (Constrexpr.LocalRawAssum(nal_keep,bk,nal_type)::acc),
Constrexpr.CLambdaN(lambda_loc,(nal_expr,bk,nal_type)::bl,e')
)
@@ -585,7 +585,7 @@ let rec rebuild_bl (aux,assoc) bl typ =
let lnal' = List.length nal' in
if lnal' >= lnal
then
- let old_nal',new_nal' = list_chop lnal nal' in
+ let old_nal',new_nal' = List.chop lnal nal' in
rebuild_bl ((LocalRawAssum(nal,bk,replace_vars_constr_expr assoc nal't)::aux),(make_assoc assoc old_nal' nal)) bl'
(if new_nal' = [] && rest = []
then typ'
@@ -593,7 +593,7 @@ let rec rebuild_bl (aux,assoc) bl typ =
then CProdN(Loc.ghost,rest,typ')
else CProdN(Loc.ghost,((new_nal',bk',nal't)::rest),typ'))
else
- let captured_nal,non_captured_nal = list_chop lnal' nal in
+ let captured_nal,non_captured_nal = List.chop lnal' nal in
rebuild_nal ((LocalRawAssum(captured_nal,bk,replace_vars_constr_expr assoc nal't)::aux), (make_assoc assoc nal' captured_nal))
bk bl' non_captured_nal (lnal - lnal') (CProdN(Loc.ghost,rest,typ'))
| _ -> assert false
@@ -789,7 +789,7 @@ let rec chop_n_arrow n t =
else
let new_t' =
Constrexpr.CProdN(Loc.ghost,
- ((snd (list_chop n nal)),k,t'')::nal_ta',t')
+ ((snd (List.chop n nal)),k,t'')::nal_ta',t')
in
raise (Stop new_t')
in
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index 2f6a6a7d7..1d1089a54 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -370,7 +370,7 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
in
(* in fact we must also add the parameters to the constructor args *)
let constructor_args g =
- let params_id = fst (list_chop princ_infos.nparams args_names) in
+ let params_id = fst (List.chop princ_infos.nparams args_names) in
(List.map mkVar params_id)@((constructor_args g))
in
(* We then get the constructor corresponding to this branch and
@@ -446,7 +446,7 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
)
lemmas_types_infos
in
- let param_names = fst (list_chop princ_infos.nparams args_names) in
+ let param_names = fst (List.chop princ_infos.nparams args_names) in
let params = List.map mkVar param_names in
let lemmas = Array.to_list (Array.map (fun c -> applist(c,params)) lemmas) in
(* The bindings of the principle
@@ -611,7 +611,7 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
in
(* in fact we must also add the parameters to the constructor args *)
let constructor_args =
- let params_id = fst (list_chop princ_infos.nparams args_names) in
+ let params_id = fst (List.chop princ_infos.nparams args_names) in
(List.map mkVar params_id)@(List.rev constructor_args)
in
(* We then get the constructor corresponding to this branch and
@@ -669,7 +669,7 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
g
in
(* end of branche proof *)
- let param_names = fst (list_chop princ_infos.nparams args_names) in
+ let param_names = fst (List.chop princ_infos.nparams args_names) in
let params = List.map mkVar param_names in
let lemmas = Array.to_list (Array.map (fun c -> applist(c,params)) lemmas) in
(* The bindings of the principle
@@ -996,7 +996,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
]
g
in
- let params_names = fst (list_chop princ_infos.nparams args_names) in
+ let params_names = fst (List.chop princ_infos.nparams args_names) in
let params = List.map mkVar params_names in
tclTHENSEQ
[ tclMAP h_intro (args_names@[res;hres]);
diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml
index 4fe22a354..21c0d86a4 100644
--- a/plugins/funind/merge.ml
+++ b/plugins/funind/merge.ml
@@ -165,11 +165,11 @@ let array_fold_lefti (f: int -> 'a -> 'b -> 'a) (acc:'a) (arr:'b array): 'a =
let res = f !i acc x in i := !i + 1; res)
acc arr
-(* Like list_chop but except that [i] is the size of the suffix of [l]. *)
+(* Like List.chop but except that [i] is the size of the suffix of [l]. *)
let list_chop_end i l =
let size_prefix = List.length l -i in
if size_prefix < 0 then failwith "list_chop_end"
- else list_chop size_prefix l
+ else List.chop size_prefix l
let list_fold_lefti (f: int -> 'a -> 'b -> 'a) (acc:'a) (arr:'b list): 'a =
let i = ref 0 in
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index 14d9b7fcb..8b0fc2739 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -705,7 +705,7 @@ let terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos g =
(try
(tclTHENS
destruct_tac
- (list_map_i (fun i e -> observe_tac (str "do treat case") (treat_case b to_thin_intro (next_step continuation_tac) ci.ci_cstr_ndecls.(i) e new_info)) 0 (Array.to_list l)
+ (List.map_i (fun i e -> observe_tac (str "do treat case") (treat_case b to_thin_intro (next_step continuation_tac) ci.ci_cstr_ndecls.(i) e new_info)) 0 (Array.to_list l)
))
with
| UserError("Refiner.thensn_tac3",_)
@@ -1014,7 +1014,7 @@ let compute_terminate_type nb_args func =
Array.of_list
(lift 5 a_arrow_b:: mkRel 3::
constr_of_global func::mkRel 1::
- List.rev (list_map_i (fun i _ -> mkRel (6+i)) 0 rev_args)
+ List.rev (List.map_i (fun i _ -> mkRel (6+i)) 0 rev_args)
)
)
in
@@ -1044,7 +1044,7 @@ let termination_proof_header is_mes input_type ids args_id relation
let nargs = List.length args_id in
let pre_rec_args =
List.rev_map
- mkVar (fst (list_chop (rec_arg_num - 1) args_id))
+ mkVar (fst (List.chop (rec_arg_num - 1) args_id))
in
let relation = substl pre_rec_args relation in
let input_type = substl pre_rec_args input_type in
@@ -1297,7 +1297,7 @@ let open_new_goal (build_proof:tactic -> tactic -> unit) using_lemmas ref_ goal_
(Elim.h_decompose_and (mkVar hid))
(fun g ->
let ids' = pf_ids_of_hyps g in
- lid := List.rev (list_subtract ids' ids);
+ lid := List.rev (List.subtract ids' ids);
if !lid = [] then lid := [hid];
tclIDTAC g
)
diff --git a/plugins/nsatz/ideal.ml b/plugins/nsatz/ideal.ml
index e0b27a2f5..94b79503a 100644
--- a/plugins/nsatz/ideal.ml
+++ b/plugins/nsatz/ideal.ml
@@ -214,7 +214,7 @@ let ppcm_mon m m' =
let repr p = p
let equal =
- Util.list_for_all2eq
+ Util.List.for_all2eq
(fun (c1,m1) (c2,m2) -> P.equal c1 c2 && m1=m2)
let hash p =
diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml
index ed06d6e11..ccf397eda 100644
--- a/plugins/omega/coq_omega.ml
+++ b/plugins/omega/coq_omega.ml
@@ -153,7 +153,7 @@ let tag_hypothesis,tag_of_hyp, hyp_of_tag =
let hide_constr,find_constr,clear_tables,dump_tables =
let l = ref ([]:(constr * (identifier * identifier * bool)) list) in
(fun h id eg b -> l := (h,(id,eg,b)):: !l),
- (fun h -> try list_assoc_f eq_constr h !l with Not_found -> failwith "find_contr"),
+ (fun h -> try List.assoc_f eq_constr h !l with Not_found -> failwith "find_contr"),
(fun () -> l := []),
(fun () -> !l)
diff --git a/plugins/omega/g_omega.ml4 b/plugins/omega/g_omega.ml4
index 62c0dc4a9..ee341cbc2 100644
--- a/plugins/omega/g_omega.ml4
+++ b/plugins/omega/g_omega.ml4
@@ -26,7 +26,7 @@ let omega_tactic l =
| "N" -> Tacinterp.interp <:tactic<zify_N>>
| "Z" -> Tacinterp.interp <:tactic<zify_op>>
| s -> Errors.error ("No Omega knowledge base for type "^s))
- (Util.list_uniquize (List.sort compare l))
+ (Util.List.uniquize (List.sort compare l))
in
tclTHEN
(tclREPEAT (tclPROGRESS (tclTHENLIST tacs)))
diff --git a/plugins/omega/omega.ml b/plugins/omega/omega.ml
index 98cad09e0..6abcc7b6f 100644
--- a/plugins/omega/omega.ml
+++ b/plugins/omega/omega.ml
@@ -352,11 +352,11 @@ let banerjee_step (new_eq_id,new_var_id,print_var) original l1 l2 =
let new_eq = List.hd (normalize new_eq) in
let eliminated_var, def = chop_var var new_eq.body in
let other_equations =
- Util.list_map_append
+ Util.List.map_append
(fun e ->
normalize (eliminate_with_in new_eq_id eliminated_var new_eq e)) l1 in
let inequations =
- Util.list_map_append
+ Util.List.map_append
(fun e ->
normalize (eliminate_with_in new_eq_id eliminated_var new_eq e)) l2 in
let original' = eliminate_with_in new_eq_id eliminated_var new_eq original in
@@ -368,9 +368,9 @@ let rec eliminate_one_equation ((new_eq_id,new_var_id,print_var) as new_ids) (e,
if !debug then display_system print_var (e::other);
try
let v,def = chop_factor_1 e.body in
- (Util.list_map_append
+ (Util.List.map_append
(fun e' -> normalize (eliminate_with_in new_eq_id v e e')) other,
- Util.list_map_append
+ Util.List.map_append
(fun e' -> normalize (eliminate_with_in new_eq_id v e e')) ineqs)
with FACTOR1 ->
eliminate_one_equation new_ids (banerjee_step new_ids e other ineqs)
@@ -523,7 +523,7 @@ let simplify ((new_eq_id,new_var_id,print_var) as new_ids) dark_shadow system =
failwith "disequation in simplify";
clear_history ();
List.iter (fun e -> add_event (HYP e)) system;
- let system = Util.list_map_append normalize system in
+ let system = Util.List.map_append normalize system in
let eqs,ineqs = List.partition (fun e -> e.kind=EQUA) system in
let simp_eq,simp_ineq = redundancy_elimination new_eq_id ineqs in
let system = (eqs @ simp_eq,simp_ineq) in
@@ -658,7 +658,7 @@ let simplify_strong ((new_eq_id,new_var_id,print_var) as new_ids) system =
| ([],ineqs,expl_map) -> ineqs,expl_map
in
try
- let system = Util.list_map_append normalize system in
+ let system = Util.List.map_append normalize system in
let eqs,ineqs = List.partition (fun e -> e.kind=EQUA) system in
let dise,ine = List.partition (fun e -> e.kind = DISE) ineqs in
let simp_eq,simp_ineq = redundancy_elimination new_eq_id ine in
@@ -700,13 +700,13 @@ let simplify_strong ((new_eq_id,new_var_id,print_var) as new_ids) system =
let s1,s2 =
List.partition (fun (_,_,decomp,_) -> sign decomp) systems in
let s1' =
- List.map (fun (dep,ro,dc,pa) -> (Util.list_except id dep,ro,dc,pa)) s1 in
+ List.map (fun (dep,ro,dc,pa) -> (Util.List.except id dep,ro,dc,pa)) s1 in
let s2' =
- List.map (fun (dep,ro,dc,pa) -> (Util.list_except id dep,ro,dc,pa)) s2 in
+ List.map (fun (dep,ro,dc,pa) -> (Util.List.except id dep,ro,dc,pa)) s2 in
let (r1,relie1) = solve s1'
and (r2,relie2) = solve s2' in
let (eq,id1,id2) = List.assoc id explode_map in
- [SPLIT_INEQ(eq,(id1,r1),(id2, r2))], eq.id :: Util.list_union relie1 relie2
+ [SPLIT_INEQ(eq,(id1,r1),(id2, r2))], eq.id :: Util.List.union relie1 relie2
with FULL_SOLUTION (x0,x1) -> (x0,x1)
in
let act,relie_on = solve all_solutions in
diff --git a/plugins/romega/g_romega.ml4 b/plugins/romega/g_romega.ml4
index a548d4d4a..5a843e2b7 100644
--- a/plugins/romega/g_romega.ml4
+++ b/plugins/romega/g_romega.ml4
@@ -19,7 +19,7 @@ let romega_tactic l =
| "N" -> Tacinterp.interp <:tactic<zify_N>>
| "Z" -> Tacinterp.interp <:tactic<zify_op>>
| s -> Errors.error ("No ROmega knowledge base for type "^s))
- (Util.list_uniquize (List.sort compare l))
+ (Util.List.uniquize (List.sort compare l))
in
tclTHEN
(tclREPEAT (tclPROGRESS (tclTHENLIST tacs)))
diff --git a/plugins/romega/refl_omega.ml b/plugins/romega/refl_omega.ml
index 50031052b..6c6e2c57f 100644
--- a/plugins/romega/refl_omega.ml
+++ b/plugins/romega/refl_omega.ml
@@ -220,7 +220,7 @@ let unintern_omega env id =
calcul des variables utiles. *)
let add_reified_atom t env =
- try list_index0_f Term.eq_constr t env.terms
+ try List.index0_f Term.eq_constr t env.terms
with Not_found ->
let i = List.length env.terms in
env.terms <- env.terms @ [t]; i
@@ -231,7 +231,7 @@ let get_reified_atom env =
(* \subsection{Gestion de l'environnement de proposition pour Omega} *)
(* ajout d'une proposition *)
let add_prop env t =
- try list_index0_f Term.eq_constr t env.props
+ try List.index0_f Term.eq_constr t env.props
with Not_found ->
let i = List.length env.props in env.props <- env.props @ [t]; i
@@ -412,7 +412,7 @@ pour faire des opérations de normalisation sur les équations. *)
(* Extraction des variables d'une équation. *)
(* Chaque fonction retourne une liste triée sans redondance *)
-let (@@) = list_merge_uniq compare
+let (@@) = List.merge_uniq compare
let rec vars_of_formula = function
| Oint _ -> []
@@ -812,7 +812,7 @@ let destructurate_hyps syst =
(i,t) :: l ->
let l_syst1 = destructurate_pos_hyp i [] [] t in
let l_syst2 = loop l in
- list_cartesian (@) l_syst1 l_syst2
+ List.cartesian (@) l_syst1 l_syst2
| [] -> [[]] in
loop syst
@@ -912,13 +912,13 @@ let add_stated_equations env tree =
(* PL: experimentally, the result order of the following function seems
_very_ crucial for efficiency. No idea why. Do not remove the List.rev
- or modify the current semantics of Util.list_union (some elements of first
+ or modify the current semantics of Util.List.union (some elements of first
arg, then second arg), unless you know what you're doing. *)
let rec get_eclatement env = function
i :: r ->
let l = try (get_equation env i).e_depends with Not_found -> [] in
- list_union (List.rev l) (get_eclatement env r)
+ List.union (List.rev l) (get_eclatement env r)
| [] -> []
let select_smaller l =
@@ -1031,7 +1031,7 @@ let mk_direction_list l =
(* \section{Rejouer l'historique} *)
let get_hyp env_hyp i =
- try list_index0 (CCEqua i) env_hyp
+ try List.index0 (CCEqua i) env_hyp
with Not_found -> failwith (Printf.sprintf "get_hyp %d" i)
let replay_history env env_hyp =
@@ -1198,8 +1198,8 @@ let resolution env full_reified_goal systems_list =
let useful_equa_id = equas_of_solution_tree solution_tree in
(* recupere explicitement ces equations *)
let equations = List.map (get_equation env) useful_equa_id in
- let l_hyps' = list_uniquize (List.map (fun e -> e.e_origin.o_hyp) equations) in
- let l_hyps = id_concl :: list_remove id_concl l_hyps' in
+ let l_hyps' = List.uniquize (List.map (fun e -> e.e_origin.o_hyp) equations) in
+ let l_hyps = id_concl :: List.remove id_concl l_hyps' in
let useful_hyps =
List.map (fun id -> List.assoc id full_reified_goal) l_hyps in
let useful_vars =
@@ -1253,7 +1253,7 @@ let resolution env full_reified_goal systems_list =
| ((O_left | O_mono) :: l) -> app coq_p_left [| loop l |]
| (O_right :: l) -> app coq_p_right [| loop l |] in
let correct_index =
- let i = list_index0 e.e_origin.o_hyp l_hyps in
+ let i = List.index0 e.e_origin.o_hyp l_hyps in
(* PL: it seems that additionnally introduced hyps are in the way during
normalization, hence this index shifting... *)
if i=0 then 0 else Pervasives.(+) i (List.length to_introduce)
diff --git a/plugins/setoid_ring/newring.ml4 b/plugins/setoid_ring/newring.ml4
index f540349ed..f838b56c6 100644
--- a/plugins/setoid_ring/newring.ml4
+++ b/plugins/setoid_ring/newring.ml4
@@ -87,7 +87,7 @@ let interp_map l c =
with Not_found -> None
let interp_map l t =
- try Some(list_assoc_f eq_constr t l) with Not_found -> None
+ try Some(List.assoc_f eq_constr t l) with Not_found -> None
let protect_maps = ref Stringmap.empty
let add_map s m = protect_maps := Stringmap.add s m !protect_maps
@@ -194,7 +194,7 @@ let dummy_goal env =
Evd.sigma = sigma}
let exec_tactic env n f args =
- let lid = list_tabulate(fun i -> id_of_string("x"^string_of_int i)) n in
+ let lid = List.tabulate(fun i -> id_of_string("x"^string_of_int i)) n in
let res = ref [||] in
let get_res ist =
let l = List.map (fun id -> List.assoc id ist.lfun) lid in
@@ -834,7 +834,7 @@ let ring_lookup (f:glob_tactic_expr) lH rl t gl =
TACTIC EXTEND ring_lookup
| [ "ring_lookup" tactic0(f) "[" constr_list(lH) "]" ne_constr_list(lrt) ] ->
- [ let (t,lr) = list_sep_last lrt in ring_lookup f lH lr t]
+ [ let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t]
END
@@ -1162,5 +1162,5 @@ let field_lookup (f:glob_tactic_expr) lH rl t gl =
TACTIC EXTEND field_lookup
| [ "field_lookup" tactic(f) "[" constr_list(lH) "]" ne_constr_list(lt) ] ->
- [ let (t,l) = list_sep_last lt in field_lookup f lH l t ]
+ [ let (t,l) = List.sep_last lt in field_lookup f lH l t ]
END
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 163675dfb..196f5a0e7 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -98,7 +98,7 @@ let msg_may_need_inversion () =
(* Utils *)
let make_anonymous_patvars n =
- list_make n (PatVar (Loc.ghost,Anonymous))
+ List.make n (PatVar (Loc.ghost,Anonymous))
(* We have x1:t1...xn:tn,xi':ti,y1..yk |- c and re-generalize
over xi:ti to get x1:t1...xn:tn,xi':ti,y1..yk |- c[xi:=xi'] *)
@@ -295,7 +295,7 @@ let try_find_ind env sigma typ realnames =
let names =
match realnames with
| Some names -> names
- | None -> list_make (List.length realargs) Anonymous in
+ | None -> List.make (List.length realargs) Anonymous in
IsInd (typ,ind,names)
let inh_coerce_to_ind evdref env ty tyi =
@@ -515,7 +515,7 @@ let mk_dep_patt_row (pats,_,eqn) =
let dependencies_in_pure_rhs nargs eqns =
if eqns = [] && not (Flags.is_program_mode ()) then
- list_make nargs false (* Only "_" patts *) else
+ List.make nargs false (* Only "_" patts *) else
let deps_rows = List.map mk_dep_patt_row eqns in
let deps_columns = matrix_transpose deps_rows in
List.map (List.exists ((=) true)) deps_columns
@@ -531,7 +531,7 @@ let rec dep_in_tomatch n = function
let dependencies_in_rhs nargs current tms eqns =
match kind_of_term current with
- | Rel n when dep_in_tomatch n tms -> list_make nargs true
+ | Rel n when dep_in_tomatch n tms -> List.make nargs true
| _ -> dependencies_in_pure_rhs nargs eqns
(* Computing the matrix of dependencies *)
@@ -551,7 +551,7 @@ let rec find_dependency_list tmblock = function
| (used,tdeps,d)::rest ->
let deps = find_dependency_list tmblock rest in
if used && List.exists (fun x -> dependent_decl x d) tmblock
- then list_add_set (List.length rest + 1) (list_union deps tdeps)
+ then List.add_set (List.length rest + 1) (List.union deps tdeps)
else deps
let find_dependencies is_dep_or_cstr_in_rhs (tm,(_,tmtypleaves),d) nextlist =
@@ -680,7 +680,7 @@ let merge_name get_name obj = function
let merge_names get_name = List.map2 (merge_name get_name)
let get_names env sign eqns =
- let names1 = list_make (List.length sign) Anonymous in
+ let names1 = List.make (List.length sign) Anonymous in
(* If any, we prefer names used in pats, from top to bottom *)
let names2,aliasname =
List.fold_right
@@ -691,7 +691,7 @@ let get_names env sign eqns =
(* Otherwise, we take names from the parameters of the constructor but
avoiding conflicts with user ids *)
let allvars =
- List.fold_left (fun l (_,_,eqn) -> list_union l eqn.rhs.avoid_ids)
+ List.fold_left (fun l (_,_,eqn) -> List.union l eqn.rhs.avoid_ids)
[] eqns in
let names3,_ =
List.fold_left2
@@ -723,7 +723,7 @@ let push_rels_eqn sign eqn =
rhs = {eqn.rhs with rhs_env = push_rel_context sign eqn.rhs.rhs_env} }
let push_rels_eqn_with_names sign eqn =
- let subpats = List.rev (list_firstn (List.length sign) eqn.patterns) in
+ let subpats = List.rev (List.firstn (List.length sign) eqn.patterns) in
let subpatnames = List.map alias_of_pat subpats in
let sign = recover_initial_subpattern_names subpatnames sign in
push_rels_eqn sign eqn
@@ -1134,7 +1134,7 @@ let build_branch current realargs deps (realnames,curname) pb arsign eqns const_
(* We adjust the terms to match in the context they will be once the *)
(* context [x1:T1,..,xn:Tn] will have been pushed on the current env *)
let typs' =
- list_map_i (fun i d -> (mkRel i,map_rel_declaration (lift i) d)) 1 typs in
+ List.map_i (fun i d -> (mkRel i,map_rel_declaration (lift i) d)) 1 typs in
let extenv = push_rel_context typs pb.env in
@@ -1462,7 +1462,7 @@ let abstract_tycon loc env evdref subst _tycon extenv t =
| Evar ev ->
let ty = get_type_of env sigma t in
let inst =
- list_map_i
+ List.map_i
(fun i _ ->
try list_assoc_in_triple i subst0 with Not_found -> mkRel i)
1 (rel_context env) in
@@ -1477,7 +1477,7 @@ let abstract_tycon loc env evdref subst _tycon extenv t =
let ty = lift (-k) (aux x (get_type_of env !evdref t)) in
let depvl = free_rels ty in
let inst =
- list_map_i
+ List.map_i
(fun i _ -> if List.mem i vl then u else mkRel i) 1
(rel_context extenv) in
let rel_filter =
@@ -1534,15 +1534,15 @@ let build_inversion_problem loc env sigma tms t =
| App (f,v) when isConstruct f ->
let cstr = destConstruct f in
let n = constructor_nrealargs env cstr in
- let l = list_lastn n (Array.to_list v) in
- let l,acc = list_fold_map' reveal_pattern l acc in
+ let l = List.lastn n (Array.to_list v) in
+ let l,acc = List.fold_map' reveal_pattern l acc in
PatCstr (Loc.ghost,cstr,l,Anonymous), acc
| _ -> make_patvar t acc in
let rec aux n env acc_sign tms acc =
match tms with
| [] -> [], acc_sign, acc
| (t, IsInd (_,IndType(indf,realargs),_)) :: tms ->
- let patl,acc = list_fold_map' reveal_pattern realargs acc in
+ let patl,acc = List.fold_map' reveal_pattern realargs acc in
let pat,acc = make_patvar t acc in
let indf' = lift_inductive_family n indf in
let sign = make_arity_signature env true indf' in
@@ -1570,14 +1570,14 @@ let build_inversion_problem loc env sigma tms t =
let n = List.length sign in
let decls =
- list_map_i (fun i d -> (mkRel i,map_rel_declaration (lift i) d)) 1 sign in
+ List.map_i (fun i d -> (mkRel i,map_rel_declaration (lift i) d)) 1 sign in
let pb_env = push_rel_context sign env in
let decls =
List.map (fun (c,d) -> (c,extract_inductive_data pb_env sigma d,d)) decls in
let decls = List.rev decls in
- let dep_sign = find_dependencies_signature (list_make n true) decls in
+ let dep_sign = find_dependencies_signature (List.make n true) decls in
let sub_tms =
List.map2 (fun deps (tm,(tmtyp,_),(na,b,t)) ->
@@ -1664,7 +1664,7 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign =
if nrealargs_ctxt <> List.length realnal then
anomaly "Ill-formed 'in' clause in cases";
List.rev realnal
- | None -> list_make nrealargs_ctxt Anonymous in
+ | None -> List.make nrealargs_ctxt Anonymous in
(na,None,build_dependent_inductive env0 indf')
::(List.map2 (fun x (_,c,t) ->(x,c,t)) realnal arsign) in
let rec buildrec n = function
@@ -2273,11 +2273,11 @@ let compile_program_cases loc style (typing_function, evdref) tycon env
let dep_sign =
find_dependencies_signature
- (list_make (List.length typs) true)
+ (List.make (List.length typs) true)
typs in
let typs' =
- list_map3
+ List.map3
(fun (tm,tmt) deps na ->
let deps = if not (isRel tm) then [] else deps in
((tm,tmt),deps,na))
@@ -2346,11 +2346,11 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e
let dep_sign =
find_dependencies_signature
- (list_make (List.length typs) true)
+ (List.make (List.length typs) true)
typs in
let typs' =
- list_map3
+ List.map3
(fun (tm,tmt) deps na ->
let deps = if not (isRel tm) then [] else deps in
((tm,tmt),deps,na))
diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml
index fd88049b2..099a2ab76 100644
--- a/pretyping/cbv.ml
+++ b/pretyping/cbv.ml
@@ -262,7 +262,7 @@ and cbv_stack_term info stack env t =
let eargs = Array.sub args nlams (nargs-nlams) in
cbv_stack_term info (APP(eargs,stk)) env' b
else
- let ctxt' = list_skipn nargs ctxt in
+ let ctxt' = List.skipn nargs ctxt in
LAM(nlams-nargs,ctxt', b, subs_cons(args,env))
(* a Fix applied enough -> IOTA *)
@@ -328,7 +328,7 @@ and cbv_norm_value info = function (* reduction under binders *)
map_constr_with_binders subs_lift (cbv_norm_term info) env t
| LAM (n,ctxt,b,env) ->
let nctxt =
- list_map_i (fun i (x,ty) ->
+ List.map_i (fun i (x,ty) ->
(x,cbv_norm_term info (subs_liftn i env) ty)) 0 ctxt in
compose_lam (List.rev nctxt) (cbv_norm_term info (subs_liftn n env) b)
| FIXP ((lij,(names,lty,bds)),env,args) ->
diff --git a/pretyping/classops.ml b/pretyping/classops.ml
index a1b7e5e9d..14476354b 100644
--- a/pretyping/classops.ml
+++ b/pretyping/classops.ml
@@ -313,7 +313,7 @@ let add_coercion_in_graph (ic,source,target) =
try_add_new_path1 (s,target) (p@[ic]);
Gmap.iter
(fun (u,v) q ->
- if u<>v & u = target && not (list_equal coe_info_typ_equal p q) then
+ if u<>v & u = target && not (List.equal coe_info_typ_equal p q) then
try_add_new_path1 (s,v) (p@[ic]@q))
old_inheritance_graph
end;
diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml
index 0b48654b6..0239a7e44 100644
--- a/pretyping/coercion.ml
+++ b/pretyping/coercion.ml
@@ -51,7 +51,7 @@ let apply_pattern_coercion loc pat p =
List.fold_left
(fun pat (co,n) ->
let f i = if i<n then Glob_term.PatVar (loc, Anonymous) else pat in
- Glob_term.PatCstr (loc, co, list_tabulate f (n+1), Anonymous))
+ Glob_term.PatCstr (loc, co, List.tabulate f (n+1), Anonymous))
pat p
(* raise Not_found if no coercion found *)
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 50410ad82..d30c1a11f 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -217,7 +217,7 @@ let lookup_index_as_renamed env t n =
let update_name na ((_,e),c) =
match na with
- | Name _ when force_wildcard () & noccurn (list_index na e) c ->
+ | Name _ when force_wildcard () & noccurn (List.index na e) c ->
Anonymous
| _ ->
na
@@ -240,14 +240,14 @@ let rec build_tree na isgoal e ci cl =
let mkpat n rhs pl = PatCstr(dl,(ci.ci_ind,n+1),pl,update_name na rhs) in
let cnl = ci.ci_cstr_ndecls in
List.flatten
- (list_tabulate (fun i -> contract_branch isgoal e (cnl.(i),mkpat i,cl.(i)))
+ (List.tabulate (fun i -> contract_branch isgoal e (cnl.(i),mkpat i,cl.(i)))
(Array.length cl))
and align_tree nal isgoal (e,c as rhs) = match nal with
| [] -> [[],rhs]
| na::nal ->
match kind_of_term c with
- | Case (ci,p,c,cl) when c = mkRel (list_index na (snd e))
+ | Case (ci,p,c,cl) when c = mkRel (List.index na (snd e))
& (* don't contract if p dependent *)
computable p (ci.ci_pp_info.ind_nargs) ->
let clauses = build_tree na isgoal e ci cl in
@@ -579,7 +579,7 @@ let rec subst_cases_pattern subst pat =
| PatVar _ -> pat
| PatCstr (loc,((kn,i),j),cpl,n) ->
let kn' = subst_ind subst kn
- and cpl' = list_smartmap (subst_cases_pattern subst) cpl in
+ and cpl' = List.smartmap (subst_cases_pattern subst) cpl in
if kn' == kn && cpl' == cpl then pat else
PatCstr (loc,((kn',i),j),cpl',n)
@@ -596,7 +596,7 @@ let rec subst_glob_constr subst raw =
| GApp (loc,r,rl) ->
let r' = subst_glob_constr subst r
- and rl' = list_smartmap (subst_glob_constr subst) rl in
+ and rl' = List.smartmap (subst_glob_constr subst) rl in
if r' == r && rl' == rl then raw else
GApp(loc,r',rl')
@@ -617,7 +617,7 @@ let rec subst_glob_constr subst raw =
| GCases (loc,sty,rtno,rl,branches) ->
let rtno' = Option.smartmap (subst_glob_constr subst) rtno
- and rl' = list_smartmap (fun (a,x as y) ->
+ and rl' = List.smartmap (fun (a,x as y) ->
let a' = subst_glob_constr subst a in
let (n,topt) = x in
let topt' = Option.smartmap
@@ -625,10 +625,10 @@ let rec subst_glob_constr subst raw =
let sp' = subst_ind subst sp in
if sp == sp' then t else (loc,(sp',i),y)) topt in
if a == a' && topt == topt' then y else (a',(n,topt'))) rl
- and branches' = list_smartmap
+ and branches' = List.smartmap
(fun (loc,idl,cpl,r as branch) ->
let cpl' =
- list_smartmap (subst_cases_pattern subst) cpl
+ List.smartmap (subst_cases_pattern subst) cpl
and r' = subst_glob_constr subst r in
if cpl' == cpl && r' == r then branch else
(loc,idl,cpl',r'))
@@ -656,7 +656,7 @@ let rec subst_glob_constr subst raw =
let ra1' = array_smartmap (subst_glob_constr subst) ra1
and ra2' = array_smartmap (subst_glob_constr subst) ra2 in
let bl' = array_smartmap
- (list_smartmap (fun (na,k,obd,ty as dcl) ->
+ (List.smartmap (fun (na,k,obd,ty as dcl) ->
let ty' = subst_glob_constr subst ty in
let obd' = Option.smartmap (subst_glob_constr subst) obd in
if ty'==ty & obd'==obd then dcl else (na,k,obd',ty')))
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index 074006872..9284e2c23 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -117,7 +117,7 @@ let check_conv_record (t1,sk1) (t2,sk2) =
| _ -> raise Not_found in
let us2,extra_args2 =
let l',s' = strip_app sk2_effective in
- let bef,aft = list_chop (List.length us) l' in
+ let bef,aft = List.chop (List.length us) l' in
(bef, append_stack_app_list aft s') in
c,bs,(params,params1),(us,us2),(extra_args1,extra_args2),c1,
(n,zip(t2,sk2))
@@ -537,7 +537,7 @@ let evar_eqappr_x ts env evd pbty appr1 appr2 =
(* We assume here |l1| <= |l2| *)
let first_order_unification ts env evd (ev1,l1) (term2,l2) =
- let (deb2,rest2) = list_chop (List.length l2-List.length l1) l2 in
+ let (deb2,rest2) = List.chop (List.length l2-List.length l1) l2 in
ise_and evd
(* First compare extra args for better failure message *)
[(fun i -> ise_list2 i (fun i -> evar_conv_x ts env i CONV) rest2 l1);
@@ -639,7 +639,7 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
| Some _ -> error "Selection of specific occurrences not supported"
| None ->
let evty = set_holes evdref cty subst in
- let instance = snd (list_filter2 (fun b c -> b) (filter,instance)) in
+ let instance = snd (List.filter2 (fun b c -> b) (filter,instance)) in
let evd,ev = new_evar_instance sign !evdref evty ~filter instance in
evdref := evd;
evsref := (fst (destEvar ev),evty)::!evsref;
@@ -761,7 +761,7 @@ let max_undefined_with_candidates evd =
| Some l -> (evk,ev_info,l)::evars) evd [] in
match l with
| [] -> None
- | a::l -> Some (list_last (a::l))
+ | a::l -> Some (List.last (a::l))
let rec solve_unconstrained_evars_with_canditates evd =
(* max_undefined is supposed to return the most recent, hence
diff --git a/pretyping/evarutil.ml b/pretyping/evarutil.ml
index 3256afd28..0243a5780 100644
--- a/pretyping/evarutil.ml
+++ b/pretyping/evarutil.ml
@@ -172,7 +172,7 @@ let collect_evars emap c =
else (* No recursion on the evar instantiation *) acc
| _ ->
fold_constr collrec acc c in
- list_uniquize (collrec [] c)
+ List.uniquize (collrec [] c)
let push_dependent_evars sigma emap =
Evd.fold_undefined (fun ev {evar_concl = ccl} (sigma',emap') ->
@@ -242,7 +242,7 @@ let apply_subfilter filter subfilter =
filter ([], List.rev subfilter))
let extract_subfilter initial_filter refined_filter =
- snd (list_filter2 (fun b1 b2 -> b1) (initial_filter,refined_filter))
+ snd (List.filter2 (fun b1 b2 -> b1) (initial_filter,refined_filter))
(**********************)
(* Creating new evars *)
@@ -320,7 +320,7 @@ let new_pure_evar evd sign ?(src=default_source) ?filter ?candidates typ =
let new_evar_instance sign evd typ ?src ?filter ?candidates instance =
assert (not !Flags.debug ||
- list_distinct (ids_of_named_context (named_context_of_val sign)));
+ List.distinct (ids_of_named_context (named_context_of_val sign)));
let evd,newevk = new_pure_evar evd sign ?src ?filter ?candidates typ in
(evd,mkEvar (newevk,Array.of_list instance))
@@ -333,7 +333,7 @@ let new_evar evd env ?src ?filter ?candidates typ =
let instance =
match filter with
| None -> instance
- | Some filter -> list_filter_with filter instance in
+ | Some filter -> List.filter_with filter instance in
new_evar_instance sign evd typ' ?src ?filter ?candidates instance
let new_type_evar ?src ?filter evd env =
@@ -369,7 +369,7 @@ let restrict_evar_key evd evk filter candidates =
let sign = evar_hyps evi in
let src = evi.evar_source in
let evd,newevk = new_pure_evar evd sign ccl ~src ~filter ?candidates in
- let ctxt = snd (list_filter2 (fun b c -> b) (filter,evar_context evi)) in
+ let ctxt = snd (List.filter2 (fun b c -> b) (filter,evar_context evi)) in
let id_inst = Array.of_list (List.map (fun (id,_,_) -> mkVar id) ctxt) in
Evd.define evk (mkEvar(newevk,id_inst)) evd,newevk
@@ -501,7 +501,7 @@ let extend_alias (_,b,_) (var_aliases,rel_aliases) =
let expand_alias_once aliases x =
match get_alias_chain_of aliases x with
| [] -> None
- | l -> Some (list_last l)
+ | l -> Some (List.last l)
let expansions_of_var aliases x =
match get_alias_chain_of aliases x with
@@ -744,7 +744,7 @@ let is_unification_pattern_evar env evd (evk,args) l t =
let args = remove_instance_local_defs evd evk (Array.to_list args) in
let n = List.length args in
match find_unification_pattern_args env (args @ l) t with
- | Some l -> Some (list_skipn n l)
+ | Some l -> Some (List.skipn n l)
| _ -> None
else
None
@@ -854,7 +854,7 @@ let define_evar_from_virtual_equation define_fun env evd t_in_env sign filter in
let t_in_env = whd_evar evd t_in_env in
let evd = define_fun env evd (destEvar evar_in_env) t_in_env in
let ids = List.map pi1 (named_context_of_val sign) in
- let inst_in_sign = List.map mkVar (list_filter_with filter ids) in
+ let inst_in_sign = List.map mkVar (List.filter_with filter ids) in
let evar_in_sign = mkEvar (fst (destEvar evar_in_env), Array.of_list inst_in_sign) in
(evd,whd_evar evd evar_in_sign)
@@ -881,7 +881,7 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env =
let sign1 = evar_hyps evi1 in
let filter1 = evar_filter evi1 in
let ids1 = List.map pi1 (named_context_of_val sign1) in
- let inst_in_sign = List.map mkVar (list_filter_with filter1 ids1) in
+ let inst_in_sign = List.map mkVar (List.filter_with filter1 ids1) in
let (sign2,filter2,inst2_in_env,inst2_in_sign,_,evd,_) =
List.fold_right (fun (na,b,t_in_env as d) (sign,filter,inst_in_env,inst_in_sign,env,evd,avoid) ->
let id = next_name_away na avoid in
@@ -1179,9 +1179,9 @@ let filter_candidates evd evk filter candidates =
match candidates,filter with
| None,_ | _, None -> candidates
| Some l, Some filter ->
- let ids = List.map pi1 (list_filter_with filter (evar_context evi)) in
+ let ids = List.map pi1 (List.filter_with filter (evar_context evi)) in
Some (List.filter (fun a ->
- list_subset (Idset.elements (collect_vars a)) ids) l)
+ List.subset (Idset.elements (collect_vars a)) ids) l)
let closure_of_filter evd evk filter =
let evi = Evd.find_undefined evd evk in
@@ -1407,7 +1407,7 @@ let solve_evar_evar ?(force=false) f g env evd (evk1,args1 as ev1) (evk2,args2 a
if are_canonical_instances args1 args2 env then
(* If instances are canonical, we solve the problem in linear time *)
let sign = evar_filtered_context (Evd.find evd evk2) in
- let id_inst = list_map_to_array (fun (id,_,_) -> mkVar id) sign in
+ let id_inst = List.map_to_array (fun (id,_,_) -> mkVar id) sign in
Evd.define evk2 (mkEvar(evk1,id_inst)) evd
else
let evd,ev1,ev2 =
@@ -1476,7 +1476,7 @@ let solve_candidates conv_algo env evd (evk,argsv as ev) rhs =
| None -> raise NoCandidates
| Some l ->
let l' =
- list_map_filter
+ List.map_filter
(filter_compatible_candidates conv_algo env evd evi args rhs) l in
match l' with
| [] -> error_cannot_unify env evd (mkEvar ev, rhs)
diff --git a/pretyping/evd.ml b/pretyping/evd.ml
index a73a74f45..7d95e743d 100644
--- a/pretyping/evd.ml
+++ b/pretyping/evd.ml
@@ -56,7 +56,7 @@ let evar_body evi = evi.evar_body
let evar_filter evi = evi.evar_filter
let evar_unfiltered_env evi = Global.env_of_context evi.evar_hyps
let evar_filtered_context evi =
- snd (list_filter2 (fun b c -> b) (evar_filter evi,evar_context evi))
+ snd (List.filter2 (fun b c -> b) (evar_filter evi,evar_context evi))
let evar_env evi =
List.fold_right push_named (evar_filtered_context evi)
(reset_context (Global.env()))
@@ -676,7 +676,7 @@ let rec list_assoc_in_triple x = function
let subst_defined_metas bl c =
let rec substrec c = match kind_of_term c with
- | Meta i -> substrec (list_assoc_snd_in_triple i bl)
+ | Meta i -> substrec (List.assoc_snd_in_triple i bl)
| _ -> map_constr substrec c
in try Some (substrec c) with Not_found -> None
@@ -797,9 +797,9 @@ let evar_dependency_closure n sigma =
if n=0 then l
else
let l' =
- list_map_append (fun (evk,_) ->
+ List.map_append (fun (evk,_) ->
try ExistentialMap.find evk graph with Not_found -> []) l in
- aux (n-1) (list_uniquize (Sort.list order (l@l'))) in
+ aux (n-1) (List.uniquize (Sort.list order (l@l'))) in
aux n (undefined_list sigma)
let pr_evar_map_t depth sigma =
diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml
index a090094aa..d20afaf40 100644
--- a/pretyping/glob_ops.ml
+++ b/pretyping/glob_ops.ml
@@ -36,7 +36,7 @@ let map_glob_decl_left_to_right f (na,k,obd,ty) =
let map_glob_constr_left_to_right f = function
| GApp (loc,g,args) ->
let comp1 = f g in
- let comp2 = Util.list_map_left f args in
+ let comp2 = Util.List.map_left f args in
GApp (loc,comp1,comp2)
| GLambda (loc,na,bk,ty,c) ->
let comp1 = f ty in
@@ -52,8 +52,8 @@ let map_glob_constr_left_to_right f = function
GLetIn (loc,na,comp1,comp2)
| GCases (loc,sty,rtntypopt,tml,pl) ->
let comp1 = Option.map f rtntypopt in
- let comp2 = Util.list_map_left (fun (tm,x) -> (f tm,x)) tml in
- let comp3 = Util.list_map_left (fun (loc,idl,p,c) -> (loc,idl,p,f c)) pl in
+ let comp2 = Util.List.map_left (fun (tm,x) -> (f tm,x)) tml in
+ let comp3 = Util.List.map_left (fun (loc,idl,p,c) -> (loc,idl,p,f c)) pl in
GCases (loc,sty,comp1,comp2,comp3)
| GLetTuple (loc,nal,(na,po),b,c) ->
let comp1 = Option.map f po in
@@ -66,7 +66,7 @@ let map_glob_constr_left_to_right f = function
let comp3 = f b2 in
GIf (loc,f c,(na,comp1),comp2,comp3)
| GRec (loc,fk,idl,bl,tyl,bv) ->
- let comp1 = Array.map (Util.list_map_left (map_glob_decl_left_to_right f)) bl in
+ let comp1 = Array.map (Util.List.map_left (map_glob_decl_left_to_right f)) bl in
let comp2 = Array.map f tyl in
let comp3 = Array.map f bv in
GRec (loc,fk,idl,comp1,comp2,comp3)
diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml
index 025f7f668..f11fb4613 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -134,7 +134,7 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let d = (n,Some b,t) in
mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::sign) c)
| Ind (_,_) ->
- let realargs = list_skipn nparams largs in
+ let realargs = List.skipn nparams largs in
let base = applist (lift i pk,realargs) in
if depK then
Reduction.beta_appvect
@@ -209,7 +209,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
let d = (n,Some b,t) in
mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::hyps) c)
| Ind _ ->
- let realargs = list_skipn nparrec largs
+ let realargs = List.skipn nparrec largs
and arg = appvect (mkRel (i+1), Termops.extended_rel_vect 0 hyps) in
applist(lift i fk,realargs@[arg])
| _ -> assert false
@@ -431,7 +431,7 @@ let mis_make_indrec env sigma listdepkind mib =
mis_make_case_com dep env sigma indi (mibi,mipi) kind
in
(* Body of mis_make_indrec *)
- list_tabulate make_one_rec nrec
+ List.tabulate make_one_rec nrec
(**********************************************************************)
(* This builds elimination predicate for Case tactic *)
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index 19126f01b..a026f53d4 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -91,7 +91,7 @@ let mis_nf_constructor_type (ind,mib,mip) j =
and nconstr = Array.length mip.mind_consnames in
let make_Ik k = mkInd ((fst ind),ntypes-k-1) in
if j > nconstr then error "Not enough constructors in the type.";
- substl (list_tabulate make_Ik ntypes) specif.(j-1)
+ substl (List.tabulate make_Ik ntypes) specif.(j-1)
(* Arity of constructors excluding parameters and local defs *)
@@ -219,7 +219,7 @@ let get_constructor (ind,mib,mip,params) j =
let typi = instantiate_params typi params mib.mind_params_ctxt in
let (args,ccl) = decompose_prod_assum typi in
let (_,allargs) = decompose_app ccl in
- let vargs = list_skipn (List.length params) allargs in
+ let vargs = List.skipn (List.length params) allargs in
{ cs_cstr = ith_constructor_of_inductive ind j;
cs_params = params;
cs_nargs = rel_context_length args;
@@ -258,11 +258,11 @@ let get_arity env (ind,params) =
let parsign = mib.mind_params_ctxt in
let nnonrecparams = mib.mind_nparams - mib.mind_nparams_rec in
if List.length params = rel_context_nhyps parsign - nnonrecparams then
- snd (list_chop nnonrecparams mib.mind_params_ctxt)
+ snd (List.chop nnonrecparams mib.mind_params_ctxt)
else
parsign in
let arproperlength = List.length mip.mind_arity_ctxt - List.length parsign in
- let arsign,_ = list_chop arproperlength mip.mind_arity_ctxt in
+ let arsign,_ = List.chop arproperlength mip.mind_arity_ctxt in
let subst = instantiate_context parsign params in
(substl_rel_context subst arsign, Inductive.inductive_sort_family mip)
@@ -324,7 +324,7 @@ let find_rectype env sigma c =
match kind_of_term t with
| Ind ind ->
let (mib,mip) = Inductive.lookup_mind_specif env ind in
- let (par,rargs) = list_chop mib.mind_nparams l in
+ let (par,rargs) = List.chop mib.mind_nparams l in
IndType((ind, par),rargs)
| _ -> raise Not_found
@@ -406,7 +406,7 @@ let type_case_branches_with_names env indspec p c =
let (ind,args) = indspec in
let (mib,mip as specif) = Inductive.lookup_mind_specif env ind in
let nparams = mib.mind_nparams in
- let (params,realargs) = list_chop nparams args in
+ let (params,realargs) = List.chop nparams args in
let lbrty = Inductive.build_branches_type ind specif params p in
(* Build case type *)
let conclty = Reduction.beta_appvect p (Array.of_list (realargs@[c])) in
diff --git a/pretyping/matching.ml b/pretyping/matching.ml
index bc9c832ae..37cbcc062 100644
--- a/pretyping/matching.ml
+++ b/pretyping/matching.ml
@@ -110,7 +110,7 @@ let dummy_constr = mkProp
let rec make_renaming ids = function
| (Name id,Name _,_)::stk ->
let renaming = make_renaming ids stk in
- (try mkRel (list_index id ids) :: renaming
+ (try mkRel (List.index id ids) :: renaming
with Not_found -> dummy_constr :: renaming)
| (_,_,_)::stk ->
dummy_constr :: make_renaming ids stk
@@ -152,7 +152,7 @@ let matches_core convert allow_partial_app allow_bound_rels pat c =
| _ -> error "Only bound indices allowed in second order pattern matching.")
args in
let frels = Intset.elements (free_rels cT) in
- if list_subset frels relargs then
+ if List.subset frels relargs then
constrain (n,([],build_lambda relargs stk cT)) subst
else
raise PatternMatchingFailure
diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml
index c3b03e209..e4ae1e4b8 100644
--- a/pretyping/patternops.ml
+++ b/pretyping/patternops.ml
@@ -146,13 +146,13 @@ let instantiate_pattern sigma lvar c =
let ctx,c = List.assoc id lvar in
try
let inst =
- List.map (fun id -> mkRel (list_index (Name id) vars)) ctx in
+ List.map (fun id -> mkRel (List.index (Name id) vars)) ctx in
let c = substl inst c in
snd (pattern_of_constr sigma c)
- with Not_found (* list_index failed *) ->
+ with Not_found (* List.index failed *) ->
let vars =
- list_map_filter (function Name id -> Some id | _ -> None) vars in
- error_instantiate_pattern id (list_subtract ctx vars)
+ List.map_filter (function Name id -> Some id | _ -> None) vars in
+ error_instantiate_pattern id (List.subtract ctx vars)
with Not_found (* List.assoc failed *) ->
x)
| (PFix _ | PCoFix _) -> error ("Non instantiable pattern.")
@@ -183,7 +183,7 @@ let rec subst_pattern subst pat =
if f' == f && args' == args then pat else
PApp (f',args')
| PSoApp (i,args) ->
- let args' = list_smartmap (subst_pattern subst) args in
+ let args' = List.smartmap (subst_pattern subst) args in
if args' == args then pat else
PSoApp (i,args')
| PLambda (name,c1,c2) ->
@@ -219,7 +219,7 @@ let rec subst_pattern subst pat =
let c' = subst_pattern subst c in
if c' == c then br else (i,n,c')
in
- let branches' = list_smartmap subst_branch branches in
+ let branches' = List.smartmap subst_branch branches in
if cip' == cip && typ' == typ && c' == c && branches' == branches
then pat
else PCase(cip', typ', c', branches')
@@ -241,7 +241,7 @@ let err loc pp = user_err_loc (loc,"pattern_of_glob_constr", pp)
let rec pat_of_raw metas vars = function
| GVar (_,id) ->
- (try PRel (list_index (Name id) vars)
+ (try PRel (List.index (Name id) vars)
with Not_found -> PVar id)
| GPatVar (_,(false,n)) ->
metas := n::!metas; PMeta (Some n)
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 511323695..0eff92b61 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -79,7 +79,7 @@ let search_guard loc env possible_indexes fixdefs =
let fix = ((indexes, 0),fixdefs) in
try check_fix env fix; raise (Found indexes)
with TypeError _ -> ())
- (list_combinations possible_indexes);
+ (List.combinations possible_indexes);
let errmsg = "Cannot guess decreasing argument of fix." in
if loc = Loc.ghost then error errmsg else
user_err_loc (loc,"search_guard", Pp.str errmsg)
@@ -363,7 +363,7 @@ let rec pretype (tycon : type_constraint) env evdref lvar = function
Array.to_list (Array.mapi
(fun i (n,_) -> match n with
| Some n -> [n]
- | None -> list_map_i (fun i _ -> i) 0 ctxtv.(i))
+ | None -> List.map_i (fun i _ -> i) 0 ctxtv.(i))
vn)
in
let fixdecls = (names,ftys,fdefs) in
diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml
index 7c128d245..96b57ae18 100644
--- a/pretyping/recordops.ml
+++ b/pretyping/recordops.ml
@@ -65,7 +65,7 @@ let subst_structure (subst,((kn,i),id,kl,projs as obj)) =
let projs' =
(* invariant: struc.s_PROJ is an evaluable reference. Thus we can take *)
(* the first component of subst_con. *)
- list_smartmap
+ List.smartmap
(Option.smartmap (fun kn -> fst (subst_con subst kn)))
projs
in
@@ -227,7 +227,7 @@ let compute_canonical_projections (con,ind) =
let args = snd (decompose_app t) in
let { s_EXPECTEDPARAM = p; s_PROJ = lpj; s_PROJKIND = kl } =
lookup_structure ind in
- let params, projs = list_chop p args in
+ let params, projs = List.chop p args in
let lpj = keep_true_projections lpj kl in
let lps = List.combine lpj projs in
let comp =
diff --git a/pretyping/redops.ml b/pretyping/redops.ml
index e10a64dc5..46f668476 100644
--- a/pretyping/redops.ml
+++ b/pretyping/redops.ml
@@ -18,13 +18,13 @@ let make_red_flag l =
if red.rDelta then
Errors.error
"Cannot set both constants to unfold and constants not to unfold";
- add_flag { red with rConst = Util.list_union red.rConst l } lf
+ add_flag { red with rConst = Util.List.union red.rConst l } lf
| FDeltaBut l :: lf ->
if red.rConst <> [] & not red.rDelta then
Errors.error
"Cannot set both constants to unfold and constants not to unfold";
add_flag
- { red with rConst = Util.list_union red.rConst l; rDelta = true }
+ { red with rConst = Util.List.union red.rConst l; rDelta = true }
lf
in
add_flag
diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml
index 0b5ad7b0b..fd9a312fc 100644
--- a/pretyping/reductionops.ml
+++ b/pretyping/reductionops.ml
@@ -58,7 +58,7 @@ let rec strip_app = function
let strip_n_app n s =
let apps,s' = strip_app s in
try
- let bef,aft = list_chop n apps in
+ let bef,aft = List.chop n apps in
match aft with
|h::[] -> Some (bef,h,s')
|h::t -> Some (bef,h,append_stack_app_list t s')
@@ -66,7 +66,7 @@ let strip_n_app n s =
with
|Failure _ -> None
let nfirsts_app_of_stack n s =
- let (args, _) = strip_app s in list_firstn n args
+ let (args, _) = strip_app s in List.firstn n args
let list_of_app_stack s =
let (out,s') = strip_app s in
Option.init (s' = []) out
@@ -88,7 +88,7 @@ let rec stack_assign s p c = match s with
if p >= q then
Zapp args :: stack_assign s (p-q) c
else
- (match list_chop p args with
+ (match List.chop p args with
(bef, _::aft) -> Zapp (bef@c::aft) :: s
| _ -> assert false)
| _ -> s
@@ -98,7 +98,7 @@ let rec stack_tail p s =
| Zapp args :: s ->
let q = List.length args in
if p >= q then stack_tail (p-q) s
- else Zapp (list_skipn p args) :: s
+ else Zapp (List.skipn p args) :: s
| _ -> failwith "stack_tail"
let rec stack_nth s p = match s with
| Zapp args :: s ->
@@ -241,13 +241,13 @@ let reducible_mind_case c = match kind_of_term c with
let contract_cofix (bodynum,(types,names,bodies as typedbodies)) =
let nbodies = Array.length bodies in
let make_Fi j = mkCoFix (nbodies-j-1,typedbodies) in
- substl (list_tabulate make_Fi nbodies) bodies.(bodynum)
+ substl (List.tabulate make_Fi nbodies) bodies.(bodynum)
let reduce_mind_case mia =
match kind_of_term mia.mconstr with
| Construct (ind_sp,i) ->
(* let ncargs = (fst mia.mci).(i-1) in*)
- let real_cargs = list_skipn mia.mci.ci_npar mia.mcargs in
+ let real_cargs = List.skipn mia.mci.ci_npar mia.mcargs in
applist (mia.mlf.(i-1),real_cargs)
| CoFix cofix ->
let cofix_def = contract_cofix cofix in
@@ -260,7 +260,7 @@ let reduce_mind_case mia =
let contract_fix ((recindices,bodynum),(types,names,bodies as typedbodies)) =
let nbodies = Array.length recindices in
let make_Fi j = mkFix ((recindices,nbodies-j-1),typedbodies) in
- substl (list_tabulate make_Fi nbodies) bodies.(bodynum)
+ substl (List.tabulate make_Fi nbodies) bodies.(bodynum)
let fix_recarg ((recindices,bodynum),_) stack =
assert (0 <= bodynum & bodynum < Array.length recindices);
@@ -338,7 +338,7 @@ let rec whd_state_gen flags ts env sigma =
if red_iota flags then
match strip_app stack with
|args, (Zcase(ci, _, lf)::s') ->
- whrec (lf.(c-1), append_stack_app_list (list_skipn ci.ci_npar args) s')
+ whrec (lf.(c-1), append_stack_app_list (List.skipn ci.ci_npar args) s')
|args, (Zfix (f,s')::s'') ->
let x' = applist(x,args) in
whrec (contract_fix f,append_stack_app_list s' (append_stack_app_list [x'] s''))
@@ -404,7 +404,7 @@ let local_whd_state_gen flags sigma =
if red_iota flags then
match strip_app stack with
|args, (Zcase(ci, _, lf)::s') ->
- whrec (lf.(c-1), append_stack_app_list (list_skipn ci.ci_npar args) s')
+ whrec (lf.(c-1), append_stack_app_list (List.skipn ci.ci_npar args) s')
|args, (Zfix (f,s')::s'') ->
let x' = applist(x,args) in
whrec (contract_fix f,append_stack_app_list s' (append_stack_app_list [x'] s''))
@@ -590,7 +590,7 @@ let whd_betaiota_preserving_vm_cast env sigma t =
| Construct (ind,c) -> begin
match strip_app stack with
|args, (Zcase(ci, _, lf)::s') ->
- whrec (lf.(c-1), append_stack_app_list (list_skipn ci.ci_npar args) s')
+ whrec (lf.(c-1), append_stack_app_list (List.skipn ci.ci_npar args) s')
|args, (Zfix (f,s')::s'') ->
let x' = applist(x,args) in
whrec (contract_fix f,append_stack_app_list s' (append_stack_app_list [x'] s''))
@@ -888,7 +888,7 @@ let whd_programs_stack env sigma =
| Construct (ind,c) -> begin
match strip_app stack with
|args, (Zcase(ci, _, lf)::s') ->
- whrec (lf.(c-1), append_stack_app_list (list_skipn ci.ci_npar args) s')
+ whrec (lf.(c-1), append_stack_app_list (List.skipn ci.ci_npar args) s')
|args, (Zfix (f,s')::s'') ->
let x' = applist(x,args) in
whrec (contract_fix f,append_stack_app_list s' (append_stack_app_list [x'] s''))
diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml
index 9791598fd..b897c01af 100644
--- a/pretyping/tacred.ml
+++ b/pretyping/tacred.ml
@@ -186,12 +186,12 @@ let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) =
raise Elimconst) args
in
let reversible_rels = List.map fst li in
- if not (list_distinct reversible_rels) then
+ if not (List.distinct reversible_rels) then
raise Elimconst;
- list_iter_i (fun i t_i ->
+ List.iter_i (fun i t_i ->
if not (List.mem_assoc (i+1) li) then
let fvs = List.map ((+) (i+1)) (Intset.elements (free_rels t_i)) in
- if list_intersect fvs reversible_rels <> [] then raise Elimconst)
+ if List.intersect fvs reversible_rels <> [] then raise Elimconst)
labs;
let k = lv.(i) in
if k < nargs then
@@ -322,13 +322,13 @@ let reference_eval sigma env = function
let x = Name (id_of_string "x")
let make_elim_fun (names,(nbfix,lv,n)) largs =
- let lu = list_firstn n largs in
+ let lu = List.firstn n largs in
let p = List.length lv in
let lyi = List.map fst lv in
let la =
- list_map_i (fun q aq ->
+ List.map_i (fun q aq ->
(* k from the comment is q+1 *)
- try mkRel (p+1-(list_index (n-q) lyi))
+ try mkRel (p+1-(List.index (n-q) lyi))
with Not_found -> aq)
0 (List.map (lift p) lu)
in
@@ -338,8 +338,8 @@ let make_elim_fun (names,(nbfix,lv,n)) largs =
| Some (minargs,ref) ->
let body = applistc (mkEvalRef ref) la in
let g =
- list_fold_left_i (fun q (* j = n+1-q *) c (ij,tij) ->
- let subst = List.map (lift (-q)) (list_firstn (n-ij) la) in
+ List.fold_left_i (fun q (* j = n+1-q *) c (ij,tij) ->
+ let subst = List.map (lift (-q)) (List.firstn (n-ij) la) in
let tij' = substl (List.rev subst) tij in
mkLambda (x,tij',c)) 1 body (List.rev lv)
in Some (minargs,g)
@@ -433,7 +433,7 @@ let reduce_fix whdfun sigma fix stack =
| None -> NotReducible
| Some (recargnum,recarg) ->
let (recarg'hd,_ as recarg') = whdfun sigma recarg in
- let stack' = list_assign stack recargnum (applist recarg') in
+ let stack' = List.assign stack recargnum (applist recarg') in
(match kind_of_term recarg'hd with
| Construct _ -> Reduced (contract_fix fix, stack')
| _ -> NotReducible)
@@ -442,7 +442,7 @@ let contract_fix_use_function env sigma f
((recindices,bodynum),(_names,_types,bodies as typedbodies)) =
let nbodies = Array.length recindices in
let make_Fi j = (mkFix((recindices,j),typedbodies), f j) in
- let lbodies = list_tabulate make_Fi nbodies in
+ let lbodies = List.tabulate make_Fi nbodies in
substl_checking_arity env (List.rev lbodies) (nf_beta sigma bodies.(bodynum))
let reduce_fix_use_function env sigma f whfun fix stack =
@@ -455,7 +455,7 @@ let reduce_fix_use_function env sigma f whfun fix stack =
(recarg, [])
else
whfun recarg in
- let stack' = list_assign stack recargnum (applist recarg') in
+ let stack' = List.assign stack recargnum (applist recarg') in
(match kind_of_term recarg'hd with
| Construct _ ->
Reduced (contract_fix_use_function env sigma f fix,stack')
@@ -465,14 +465,14 @@ let contract_cofix_use_function env sigma f
(bodynum,(_names,_,bodies as typedbodies)) =
let nbodies = Array.length bodies in
let make_Fi j = (mkCoFix(j,typedbodies), f j) in
- let subbodies = list_tabulate make_Fi nbodies in
+ let subbodies = List.tabulate make_Fi nbodies in
substl_checking_arity env (List.rev subbodies)
(nf_beta sigma bodies.(bodynum))
let reduce_mind_case_use_function func env sigma mia =
match kind_of_term mia.mconstr with
| Construct(ind_sp,i) ->
- let real_cargs = list_skipn mia.mci.ci_npar mia.mcargs in
+ let real_cargs = List.skipn mia.mci.ci_npar mia.mcargs in
applist (mia.mlf.(i-1), real_cargs)
| CoFix (bodynum,(names,_,_) as cofix) ->
let build_cofix_name =
@@ -661,7 +661,7 @@ let rec red_elim_const env sigma ref largs =
let arg = List.nth stack i in
let rarg = whd_construct_stack env sigma arg in
match kind_of_term (fst rarg) with
- | Construct _ -> list_assign stack i (applist rarg)
+ | Construct _ -> List.assign stack i (applist rarg)
| _ -> raise Redelimination)
largs l, n >= 0 && l = [] && nargs >= n,
n >= 0 && l <> [] && nargs >= n in
diff --git a/pretyping/termops.ml b/pretyping/termops.ml
index e4f481e58..765f1e4fa 100644
--- a/pretyping/termops.ml
+++ b/pretyping/termops.ml
@@ -286,10 +286,10 @@ let adjust_app_list_size f1 l1 f2 l2 =
let len1 = List.length l1 and len2 = List.length l2 in
if len1 = len2 then (f1,l1,f2,l2)
else if len1 < len2 then
- let extras,restl2 = list_chop (len2-len1) l2 in
+ let extras,restl2 = List.chop (len2-len1) l2 in
(f1, l1, applist (f2,extras), restl2)
else
- let extras,restl1 = list_chop (len1-len2) l1 in
+ let extras,restl1 = List.chop (len1-len2) l1 in
(applist (f1,extras), restl1, f2, l2)
let adjust_app_array_size f1 l1 f2 l2 =
@@ -550,7 +550,7 @@ let free_rels m =
let collect_metas c =
let rec collrec acc c =
match kind_of_term c with
- | Meta mv -> list_add_set mv acc
+ | Meta mv -> List.add_set mv acc
| _ -> fold_constr collrec acc c
in
List.rev (collrec [] c)
@@ -691,7 +691,7 @@ let replace_term = replace_term_gen eq_constr
except the ones in l *)
let error_invalid_occurrence l =
- let l = list_uniquize (List.sort Pervasives.compare l) in
+ let l = List.uniquize (List.sort Pervasives.compare l) in
errorlabstrm ""
(str ("Invalid occurrence " ^ plural (List.length l) "number" ^": ") ++
prlist_with_sep spc int l ++ str ".")
@@ -951,7 +951,7 @@ let align_prod_letin c a : rel_context * constr =
let (lc,_,_) = decompose_prod_letin c in
let (la,l,a) = decompose_prod_letin a in
if not (la >= lc) then invalid_arg "align_prod_letin";
- let (l1,l2) = Util.list_chop lc l in
+ let (l1,l2) = Util.List.chop lc l in
l2,it_mkProd_or_LetIn a l1
(* On reduit une serie d'eta-redex de tete ou rien du tout *)
@@ -1045,7 +1045,7 @@ let adjust_subst_to_rel_context sign l =
| _ -> anomaly "Instance and signature do not match"
in aux [] (List.rev sign) l
-let fold_named_context_both_sides f l ~init = list_fold_right_and_left f l init
+let fold_named_context_both_sides f l ~init = List.fold_right_and_left f l init
let rec mem_named_context id = function
| (id',_,_) :: _ when id=id' -> true
@@ -1098,7 +1098,7 @@ let context_chop k ctx =
(* Do not skip let-in's *)
let env_rel_context_chop k env =
let rels = rel_context env in
- let ctx1,ctx2 = list_chop k rels in
+ let ctx1,ctx2 = List.chop k rels in
push_rel_context ctx2 (reset_with_named_context (named_context_val env) env),
ctx1
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
index 8fd0f768e..426197af2 100644
--- a/pretyping/typeclasses.ml
+++ b/pretyping/typeclasses.ml
@@ -154,13 +154,13 @@ let subst_class (subst,cl) =
let do_subst_con c = fst (Mod_subst.subst_con subst c)
and do_subst c = Mod_subst.subst_mps subst c
and do_subst_gr gr = fst (subst_global subst gr) in
- let do_subst_ctx ctx = list_smartmap
+ let do_subst_ctx ctx = List.smartmap
(fun (na, b, t) -> (na, Option.smartmap do_subst b, do_subst t))
ctx in
let do_subst_context (grs,ctx) =
- list_smartmap (Option.smartmap (fun (gr,b) -> do_subst_gr gr, b)) grs,
+ List.smartmap (Option.smartmap (fun (gr,b) -> do_subst_gr gr, b)) grs,
do_subst_ctx ctx in
- let do_subst_projs projs = list_smartmap (fun (x, y, z) -> (x, y, Option.smartmap do_subst_con z)) projs in
+ let do_subst_projs projs = List.smartmap (fun (x, y, z) -> (x, y, Option.smartmap do_subst_con z)) projs in
{ cl_impl = do_subst_gr cl.cl_impl;
cl_context = do_subst_context cl.cl_context;
cl_props = do_subst_ctx cl.cl_props;
@@ -195,7 +195,7 @@ let discharge_class (_,cl) =
| Some (_, (tc, _)) -> Some (tc.cl_impl, true))
ctx'
in
- list_smartmap (Option.smartmap (fun (gr, b) -> Lib.discharge_global gr, b)) grs
+ List.smartmap (Option.smartmap (fun (gr, b) -> Lib.discharge_global gr, b)) grs
@ newgrs
in grs', discharge_rel_context subst 1 ctx @ ctx' in
let cl_impl' = Lib.discharge_global cl.cl_impl in
@@ -207,7 +207,7 @@ let discharge_class (_,cl) =
{ cl_impl = cl_impl';
cl_context = context;
cl_props = props;
- cl_projs = list_smartmap (fun (x, y, z) -> x, y, Option.smartmap Lib.discharge_con z) cl.cl_projs }
+ cl_projs = List.smartmap (fun (x, y, z) -> x, y, Option.smartmap Lib.discharge_con z) cl.cl_projs }
let rebuild_class cl =
try
@@ -246,7 +246,7 @@ let build_subclasses ~check env sigma glob pri =
| Some (rels, (tc, args)) ->
let instapp = Reductionops.whd_beta sigma (appvectc c (Termops.extended_rel_vect 0 rels)) in
let projargs = Array.of_list (args @ [instapp]) in
- let projs = list_map_filter
+ let projs = List.map_filter
(fun (n, b, proj) ->
match b with
| None -> None
@@ -408,12 +408,12 @@ let add_inductive_class ind =
let instance_constructor cl args =
let lenpars = List.length (List.filter (fun (na, b, t) -> b = None) (snd cl.cl_context)) in
- let pars = fst (list_chop lenpars args) in
+ let pars = fst (List.chop lenpars args) in
match cl.cl_impl with
| IndRef ind -> Some (applistc (mkConstruct (ind, 1)) args),
applistc (mkInd ind) pars
| ConstRef cst ->
- let term = if args = [] then None else Some (list_last args) in
+ let term = if args = [] then None else Some (List.last args) in
term, applistc (mkConst cst) pars
| _ -> assert false
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index df1833f84..b24992b8d 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -100,7 +100,7 @@ let e_is_correct_arity env evdref c pj ind specif params =
let e_type_case_branches env evdref (ind,largs) pj c =
let specif = lookup_mind_specif env ind in
let nparams = inductive_params specif in
- let (params,realargs) = list_chop nparams largs in
+ let (params,realargs) = List.chop nparams largs in
let p = pj.uj_val in
let univ = e_is_correct_arity env evdref c pj ind specif params in
let lc = build_branches_type ind specif params p in
diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml
index f99d8c109..190181c23 100644
--- a/printing/ppconstr.ml
+++ b/printing/ppconstr.ml
@@ -460,8 +460,8 @@ let pr pr sep inherited a =
pr spc ltop b),
lletin
| CAppExpl (_,(Some i,f),l) ->
- let l1,l2 = list_chop i l in
- let c,l1 = list_sep_last l1 in
+ let l1,l2 = List.chop i l in
+ let c,l1 = List.sep_last l1 in
let p = pr_proj (pr mt) pr_appexpl c f l1 in
if l2<>[] then
p ++ prlist (pr spc (lapp,L)) l2, lapp
@@ -473,8 +473,8 @@ let pr pr sep inherited a =
hov 0 (str ".." ++ pr spc (latom,E) t ++ spc () ++ str ".."), larg
| CAppExpl (_,(None,f),l) -> pr_appexpl (pr mt) f l, lapp
| CApp (_,(Some i,f),l) ->
- let l1,l2 = list_chop i l in
- let c,l1 = list_sep_last l1 in
+ let l1,l2 = List.chop i l in
+ let c,l1 = List.sep_last l1 in
assert (snd c = None);
let p = pr_proj (pr mt) pr_app (fst c) f l1 in
if l2<>[] then
diff --git a/printing/pptactic.ml b/printing/pptactic.ml
index 7118388cb..1a8f71397 100644
--- a/printing/pptactic.ml
+++ b/printing/pptactic.ml
@@ -592,7 +592,7 @@ let pr_fix_tac (id,n,c) =
let rec set_nth_name avoid n = function
(nal,ty)::bll ->
if n <= List.length nal then
- match list_chop (n-1) nal with
+ match List.chop (n-1) nal with
_, (_,Name id) :: _ -> id, (nal,ty)::bll
| bef, (loc,Anonymous) :: aft ->
let id = next_ident_away (id_of_string"y") avoid in
diff --git a/printing/prettyp.ml b/printing/prettyp.ml
index 86b2da2eb..7fbbc0a2e 100644
--- a/printing/prettyp.ml
+++ b/printing/prettyp.ml
@@ -93,7 +93,7 @@ let print_impargs_by_name max = function
let print_one_impargs_list l =
let imps = List.filter is_status_implicit l in
let maximps = List.filter Impargs.maximal_insertion_of imps in
- let nonmaximps = list_subtract imps maximps in
+ let nonmaximps = List.subtract imps maximps in
print_impargs_by_name false nonmaximps @
print_impargs_by_name true maximps
@@ -127,7 +127,7 @@ let need_expansion impl ref =
let ctx = (prod_assum typ) in
let nprods = List.length (List.filter (fun (_,b,_) -> b=None) ctx) in
impl <> [] & List.length impl >= nprods &
- let _,lastimpl = list_chop nprods impl in
+ let _,lastimpl = List.chop nprods impl in
List.filter is_status_implicit lastimpl <> []
let print_impargs ref =
diff --git a/printing/printer.ml b/printing/printer.ml
index 70a90b8ea..1ad9dba49 100644
--- a/printing/printer.ml
+++ b/printing/printer.ml
@@ -289,7 +289,7 @@ let pr_concl n sigma g =
(* display evar type: a context and a type *)
let pr_evgl_sign gl =
let ps = pr_named_context_of (evar_unfiltered_env gl) in
- let _,l = list_filter2 (fun b c -> not b) (evar_filter gl,evar_context gl) in
+ let _,l = List.filter2 (fun b c -> not b) (evar_filter gl,evar_context gl) in
let ids = List.rev (List.map pi1 l) in
let warn =
if ids = [] then mt () else
@@ -664,7 +664,7 @@ let print_one_inductive env mib ((_,i) as ind) =
brk(0,2) ++ print_constructors envpar mip.mind_consnames cstrtypes
let print_mutual_inductive env mind mib =
- let inds = list_tabulate (fun x -> (mind,x)) (Array.length mib.mind_packets)
+ let inds = List.tabulate (fun x -> (mind,x)) (Array.length mib.mind_packets)
in
hov 0 (
str (if mib.mind_finite then "Inductive " else "CoInductive ") ++
diff --git a/proofs/clenv.ml b/proofs/clenv.ml
index f3e414126..0dbaccc33 100644
--- a/proofs/clenv.ml
+++ b/proofs/clenv.ml
@@ -396,7 +396,7 @@ let clenv_independent clenv =
List.filter (fun mv -> not (Metaset.mem mv deps)) mvs
let check_bindings bl =
- match list_duplicates (List.map pi2 bl) with
+ match List.duplicates (List.map pi2 bl) with
| NamedHyp s :: _ ->
errorlabstrm ""
(str "The variable " ++ pr_id s ++
@@ -464,7 +464,7 @@ exception NoSuchBinding
let clenv_constrain_last_binding c clenv =
let all_mvs = collect_metas clenv.templval.rebus in
- let k = try list_last all_mvs with Failure _ -> raise NoSuchBinding in
+ let k = try List.last all_mvs with Failure _ -> raise NoSuchBinding in
clenv_assign_binding clenv k c
let error_not_right_number_missing_arguments n =
diff --git a/proofs/logic.ml b/proofs/logic.ml
index dcf1e4c73..53f5705a5 100644
--- a/proofs/logic.ml
+++ b/proofs/logic.ml
@@ -308,7 +308,7 @@ let collect_meta_variables c =
List.rev (collrec false [] c)
let check_meta_variables c =
- if not (list_distinct (collect_meta_variables c)) then
+ if not (List.distinct (collect_meta_variables c)) then
raise (RefinerError (NonLinearProof c))
let check_conv_leq_goal env sigma arg ty conclty =
@@ -536,7 +536,7 @@ let prim_refiner r sigma goal =
| _ -> error "Not enough products."
in
let (sp,_) = check_ind env n cl in
- let firsts,lasts = list_chop j rest in
+ let firsts,lasts = List.chop j rest in
let all = firsts@(f,n,cl)::lasts in
let rec mk_sign sign = function
| (f,n,ar)::oth ->
@@ -580,7 +580,7 @@ let prim_refiner r sigma goal =
error ("All methods must construct elements " ^
"in coinductive types.")
in
- let firsts,lasts = list_chop j others in
+ let firsts,lasts = List.chop j others in
let all = firsts@(f,cl)::lasts in
List.iter (fun (_,c) -> check_is_coind env c) all;
let rec mk_sign sign = function
diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml
index f7d9446b5..ab80841d8 100644
--- a/proofs/pfedit.ml
+++ b/proofs/pfedit.ml
@@ -126,7 +126,7 @@ let solve_nth ?(with_end_tac=false) gi tac =
Proof_global.run_tactic (Proofview.tclFOCUS gi gi tac)
with
| Proof_global.NoCurrentProof -> Errors.error "No focused proof"
- | Proofview.IndexOutOfRange | Failure "list_chop" ->
+ | Proofview.IndexOutOfRange | Failure "List.chop" ->
let msg = str "No such goal: " ++ int gi ++ str "." in
Errors.errorlabstrm "" msg
diff --git a/proofs/proofview.ml b/proofs/proofview.ml
index 0e2ac1dde..79b5ae731 100644
--- a/proofs/proofview.ml
+++ b/proofs/proofview.ml
@@ -112,7 +112,7 @@ let focus_sublist i j l =
let (left,sub_right) = list_goto (i-1) l in
let (sub, right) =
try
- Util.list_chop (j-i+1) sub_right
+ Util.List.chop (j-i+1) sub_right
with Failure "list_chop" ->
Errors.errorlabstrm "nth_unproven" (Pp.str"No such unproven subgoal")
in
diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml
index 75a2ae8c3..c16e02b3f 100644
--- a/proofs/redexpr.ml
+++ b/proofs/redexpr.ml
@@ -56,9 +56,9 @@ let cache_strategy (_,str) =
let subst_strategy (subs,(local,obj)) =
local,
- list_smartmap
+ List.smartmap
(fun (k,ql as entry) ->
- let ql' = list_smartmap (Mod_subst.subst_evaluable_reference subs) ql in
+ let ql' = List.smartmap (Mod_subst.subst_evaluable_reference subs) ql in
if ql==ql' then entry else (k,ql'))
obj
diff --git a/proofs/refiner.ml b/proofs/refiner.ml
index 910653ddb..0652f1986 100644
--- a/proofs/refiner.ml
+++ b/proofs/refiner.ml
@@ -109,7 +109,7 @@ let thens3parts_tac tacfi tac tacli (sigr,gs) =
let ng = List.length gs in
if ng<nf+nl then errorlabstrm "Refiner.thensn_tac" (str "Not enough subgoals.");
let gll =
- (list_map_i (fun i ->
+ (List.map_i (fun i ->
apply_sig_tac sigr (if i<nf then tacfi.(i) else if i>=ng-nl then tacli.(nl-ng+i) else tac))
0 gs) in
(sigr,List.flatten gll)
@@ -123,7 +123,7 @@ let thensl_tac tac taci = thens3parts_tac [||] tac taci
(* Apply [tac i] on the ith subgoal (no subgoals number check) *)
let thensi_tac tac (sigr,gs) =
let gll =
- list_map_i (fun i -> apply_sig_tac sigr (tac i)) 1 gs in
+ List.map_i (fun i -> apply_sig_tac sigr (tac i)) 1 gs in
(sigr, List.flatten gll)
let then_tac tac = thensf_tac [||] tac
diff --git a/tactics/auto.ml b/tactics/auto.ml
index 5068552d1..be1e8e701 100644
--- a/tactics/auto.ml
+++ b/tactics/auto.ml
@@ -280,7 +280,7 @@ let subst_path_atom subst p =
| PathAny -> p
| PathHints grs ->
let gr' gr = fst (subst_global subst gr) in
- let grs' = list_smartmap gr' grs in
+ let grs' = List.smartmap gr' grs in
if grs' == grs then p else PathHints grs'
let rec subst_hints_path subst hp =
@@ -393,7 +393,7 @@ module Hint_db = struct
let add_list l db = List.fold_left (fun db k -> add_one k db) db l
- let remove_sdl p sdl = list_smartfilter p sdl
+ let remove_sdl p sdl = List.smartfilter p sdl
let remove_he st p (sl1, sl2, dn as he) =
let sl1' = remove_sdl p sl1 and sl2' = remove_sdl p sl2 in
if sl1' == sl1 && sl2' == sl2 then he
@@ -402,7 +402,7 @@ module Hint_db = struct
let remove_list grs db =
let filter (_, h) = match h.name with PathHints [gr] -> not (List.mem gr grs) | _ -> true in
let hintmap = Constr_map.map (remove_he db.hintdb_state filter) db.hintdb_map in
- let hintnopat = list_smartfilter (fun (ge, sd) -> filter sd) db.hintdb_nopat in
+ let hintnopat = List.smartfilter (fun (ge, sd) -> filter sd) db.hintdb_nopat in
{ db with hintdb_map = hintmap; hintdb_nopat = hintnopat }
let remove_one gr db = remove_list [gr] db
@@ -691,14 +691,14 @@ let subst_autohint (subst,(local,name,hintlist as obj)) =
match hintlist with
| CreateDB _ -> obj
| AddTransparency (grs, b) ->
- let grs' = list_smartmap (subst_evaluable_reference subst) grs in
+ let grs' = List.smartmap (subst_evaluable_reference subst) grs in
if grs==grs' then obj else (local, name, AddTransparency (grs', b))
| AddHints hintlist ->
- let hintlist' = list_smartmap subst_hint hintlist in
+ let hintlist' = List.smartmap subst_hint hintlist in
if hintlist' == hintlist then obj else
(local,name,AddHints hintlist')
| RemoveHints grs ->
- let grs' = list_smartmap (fun x -> fst (subst_global subst x)) grs in
+ let grs' = List.smartmap (fun x -> fst (subst_global subst x)) grs in
if grs==grs' then obj else (local, name, RemoveHints grs')
| AddCut path ->
let path' = subst_hints_path subst path in
@@ -761,7 +761,7 @@ let add_extern pri pat tacast local dbname =
let tacmetas = [] in
match pat with
| Some (patmetas,pat) ->
- (match (list_subtract tacmetas patmetas) with
+ (match (List.subtract tacmetas patmetas) with
| i::_ ->
errorlabstrm "add_extern"
(str "The meta-variable ?" ++ Ppconstr.pr_patvar i ++ str" is not bound.")
@@ -859,7 +859,7 @@ let interp_hints h =
let constr_hints_of_ind qid =
let ind = global_inductive_with_alias qid in
Dumpglob.dump_reference (fst (qualid_of_reference qid)) "<>" (string_of_reference qid) "ind";
- list_tabulate (fun i -> let c = (ind,i+1) in
+ List.tabulate (fun i -> let c = (ind,i+1) in
None, true, PathHints [ConstructRef c], mkConstruct c)
(nconstructors ind) in
HintsResolveEntry (List.flatten (List.map constr_hints_of_ind lqid))
@@ -1054,10 +1054,10 @@ let unify_resolve_gen = function
(* Util *)
let expand_constructor_hints env lems =
- list_map_append (fun (sigma,lem) ->
+ List.map_append (fun (sigma,lem) ->
match kind_of_term lem with
| Ind ind ->
- list_tabulate (fun i -> mkConstruct (ind,i+1)) (nconstructors ind)
+ List.tabulate (fun i -> mkConstruct (ind,i+1)) (nconstructors ind)
| _ ->
[prepare_hint env (sigma,lem)]) lems
@@ -1067,7 +1067,7 @@ let expand_constructor_hints env lems =
let add_hint_lemmas eapply lems hint_db gl =
let lems = expand_constructor_hints (pf_env gl) lems in
let hintlist' =
- list_map_append (pf_apply make_resolves gl (eapply,true,false) None) lems in
+ List.map_append (pf_apply make_resolves gl (eapply,true,false) None) lems in
Hint_db.add_list hintlist' hint_db
let make_local_hint_db ?ts eapply lems gl =
@@ -1076,7 +1076,7 @@ let make_local_hint_db ?ts eapply lems gl =
| None -> Hint_db.transparent_state (searchtable_map "core")
| Some ts -> ts
in
- let hintlist = list_map_append (pf_apply make_resolve_hyp gl) sign in
+ let hintlist = List.map_append (pf_apply make_resolve_hyp gl) sign in
add_hint_lemmas eapply lems
(Hint_db.add_list hintlist (Hint_db.empty ts false)) gl
@@ -1271,7 +1271,7 @@ let rec trivial_fail_db dbg mod_delta db_list local_db gl =
and my_find_search_nodelta db_list local_db hdc concl =
List.map (fun hint -> (None,hint))
- (list_map_append (hintmap_of hdc concl) (local_db::db_list))
+ (List.map_append (hintmap_of hdc concl) (local_db::db_list))
and my_find_search mod_delta =
if mod_delta then my_find_search_delta
@@ -1281,7 +1281,7 @@ and my_find_search_delta db_list local_db hdc concl =
let flags = {auto_unif_flags with use_metas_eagerly_in_conv_on_closed_terms = true} in
let f = hintmap_of hdc concl in
if occur_existential concl then
- list_map_append
+ List.map_append
(fun db ->
if Hint_db.use_dn db then
let flags = flags_of_state (Hint_db.transparent_state db) in
@@ -1291,7 +1291,7 @@ and my_find_search_delta db_list local_db hdc concl =
List.map (fun x -> (Some flags,x)) (f db))
(local_db::db_list)
else
- list_map_append (fun db ->
+ List.map_append (fun db ->
if Hint_db.use_dn db then
let flags = flags_of_state (Hint_db.transparent_state db) in
List.map (fun x -> (Some flags, x)) (f db)
@@ -1346,7 +1346,7 @@ and trivial_resolve dbg mod_delta db_list local_db cl =
let make_db_list dbnames =
let use_core = not (List.mem "nocore" dbnames) in
- let dbnames = list_remove "nocore" dbnames in
+ let dbnames = List.remove "nocore" dbnames in
let dbnames = if use_core then "core"::dbnames else dbnames in
let lookup db =
try searchtable_map db with Not_found -> error_no_such_hint_database db
@@ -1361,7 +1361,7 @@ let trivial ?(debug=Off) lems dbnames gl =
let full_trivial ?(debug=Off) lems gl =
let dbnames = Hintdbmap.dom !searchtable in
- let dbnames = list_remove "v62" dbnames in
+ let dbnames = List.remove "v62" dbnames in
let db_list = List.map (fun x -> searchtable_map x) dbnames in
let d = mk_trivial_dbg debug in
tclTRY_dbg d
@@ -1454,7 +1454,7 @@ let default_auto = auto !default_search_depth [] []
let delta_full_auto ?(debug=Off) mod_delta n lems gl =
let dbnames = Hintdbmap.dom !searchtable in
- let dbnames = list_remove "v62" dbnames in
+ let dbnames = List.remove "v62" dbnames in
let db_list = List.map (fun x -> searchtable_map x) dbnames in
let d = mk_auto_dbg debug in
tclTRY_dbg d
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index d2d18c53b..dfa94102d 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -207,7 +207,7 @@ let auto_multi_rewrite_with ?(conds=Naive) tac_main lbas cl gl =
(* Functions necessary to the library object declaration *)
let cache_hintrewrite (_,(rbase,lrl)) =
let base = try find_base rbase with _ -> HintDN.empty in
- let max = try fst (Util.list_last (HintDN.find_all base)) with _ -> 0 in
+ let max = try fst (Util.List.last (HintDN.find_all base)) with _ -> 0 in
let lrl = HintDN.map (fun (i,h) -> (i + max, h)) lrl in
rewtab:=Stringmap.add rbase (HintDN.union lrl base) !rewtab
diff --git a/tactics/class_tactics.ml4 b/tactics/class_tactics.ml4
index 78df21a6d..1eecb1eb3 100644
--- a/tactics/class_tactics.ml4
+++ b/tactics/class_tactics.ml4
@@ -159,7 +159,7 @@ and e_my_find_search db_list local_db hdc complete concl =
let prods, concl = decompose_prod_assum concl in
let nprods = List.length prods in
let hintl =
- list_map_append
+ List.map_append
(fun db ->
if Hint_db.use_dn db then
let flags = flags_of_state (Hint_db.transparent_state db) in
@@ -251,7 +251,7 @@ let make_resolve_hyp env sigma st flags only_classes pri (id, _, cty) =
let hints =
if is_class then
let hints = build_subclasses ~check:false env sigma (VarRef id) None in
- (list_map_append
+ (List.map_append
(fun (pri, c) -> make_resolves env sigma
(true,false,Flags.is_verbose()) pri c)
hints)
@@ -376,7 +376,7 @@ let hints_tac hints =
(* Reorder with dependent subgoals. *)
(gls' @ List.map (fun (ev, x) -> Some ev, x) evgls, s')
in
- let gls' = list_map_i
+ let gls' = List.map_i
(fun j (evar, g) ->
let info =
{ info with auto_depth = j :: i :: info.auto_depth; auto_last_tac = pp;
@@ -472,7 +472,7 @@ let cut_of_hints h =
let make_autogoals ?(only_classes=true) ?(st=full_transparent_state) hints gs evm' =
let cut = cut_of_hints hints in
- { it = list_map_i (fun i g ->
+ { it = List.map_i (fun i g ->
let (gl, auto) = make_autogoal ~only_classes ~st cut (Some (fst g)) {it = snd g; sigma = evm'} in
(gl, { auto with auto_depth = [i]})) 1 gs; sigma = evm' }
@@ -779,7 +779,7 @@ END
let typeclasses_eauto ?(only_classes=false) ?(st=full_transparent_state) dbs gl =
try
- let dbs = list_map_filter (fun db -> try Some (Auto.searchtable_map db) with _ -> None) dbs in
+ let dbs = List.map_filter (fun db -> try Some (Auto.searchtable_map db) with _ -> None) dbs in
let st = match dbs with x :: _ -> Hint_db.transparent_state x | _ -> st in
eauto ?limit:!typeclasses_depth ~only_classes ~st dbs gl
with Not_found -> tclFAIL 0 (str" typeclasses eauto failed on: " ++ Printer.pr_goal gl) gl
diff --git a/tactics/eauto.ml4 b/tactics/eauto.ml4
index 28840669f..fd5fbe25a 100644
--- a/tactics/eauto.ml4
+++ b/tactics/eauto.ml4
@@ -119,11 +119,11 @@ and e_my_find_search db_list local_db hdc concl =
let hdc = head_of_constr_reference hdc in
let hintl =
if occur_existential concl then
- list_map_append (fun db ->
+ List.map_append (fun db ->
let flags = {auto_unif_flags with modulo_delta = Hint_db.transparent_state db} in
List.map (fun x -> flags, x) (Hint_db.map_all hdc db)) (local_db::db_list)
else
- list_map_append (fun db ->
+ List.map_append (fun db ->
let flags = {auto_unif_flags with modulo_delta = Hint_db.transparent_state db} in
List.map (fun x -> flags, x) (Hint_db.map_auto (hdc,concl) db)) (local_db::db_list)
in
@@ -261,7 +261,7 @@ module SearchProblem = struct
{ depth = pred s.depth; tacres = res;
dblist = s.dblist; last_tactic = pp; prev = ps;
localdb =
- list_addn (nbgl'-nbgl) (List.hd s.localdb) s.localdb })
+ List.addn (nbgl'-nbgl) (List.hd s.localdb) s.localdb })
l
in
List.sort compare (assumption_tacs @ intro_tac @ rec_tacs)
@@ -364,7 +364,7 @@ let eauto ?(debug=Off) np lems dbnames =
let full_eauto ?(debug=Off) n lems gl =
let dbnames = current_db_names () in
- let dbnames = list_remove "v62" dbnames in
+ let dbnames = List.remove "v62" dbnames in
let db_list = List.map searchtable_map dbnames in
tclTRY (e_search_auto debug n lems db_list) gl
diff --git a/tactics/elim.ml b/tactics/elim.ml
index a2ec6dfa5..32acc5af5 100644
--- a/tactics/elim.ml
+++ b/tactics/elim.ml
@@ -37,10 +37,10 @@ let introCaseAssumsThen tac ba =
let n1 = List.length case_thin_sign in
let n2 = List.length ba.branchnames in
let (l1,l2),l3 =
- if n1 < n2 then list_chop n1 ba.branchnames, []
+ if n1 < n2 then List.chop n1 ba.branchnames, []
else
(ba.branchnames, []),
- if n1 > n2 then snd (list_chop n2 case_thin_sign) else [] in
+ if n1 > n2 then snd (List.chop n2 case_thin_sign) else [] in
let introCaseAssums =
tclTHEN (intros_pattern MoveLast l1) (intros_clearing l3) in
(tclTHEN introCaseAssums (case_on_ba (tac l2) ba))
diff --git a/tactics/eqdecide.ml4 b/tactics/eqdecide.ml4
index 7bc372ca9..f26eb1024 100644
--- a/tactics/eqdecide.ml4
+++ b/tactics/eqdecide.ml4
@@ -129,7 +129,7 @@ let solveEqBranch rectype g =
let (eqonleft,op,lhs,rhs,_) = match_eqdec (pf_concl g) in
let (mib,mip) = Global.lookup_inductive rectype in
let nparams = mib.mind_nparams in
- let getargs l = list_skipn nparams (snd (decompose_app l)) in
+ let getargs l = List.skipn nparams (snd (decompose_app l)) in
let rargs = getargs rhs
and largs = getargs lhs in
List.fold_right2
diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml
index f744b015a..33eb7c618 100644
--- a/tactics/eqschemes.ml
+++ b/tactics/eqschemes.ml
@@ -63,7 +63,7 @@ let default_id_of_sort = function InProp | InSet -> hid | InType -> xid
let fresh env id = next_global_ident_away id []
let build_dependent_inductive ind (mib,mip) =
- let realargs,_ = list_chop mip.mind_nrealargs_ctxt mip.mind_arity_ctxt in
+ let realargs,_ = List.chop mip.mind_nrealargs_ctxt mip.mind_arity_ctxt in
applist
(mkInd ind,
extended_rel_list mip.mind_nrealargs_ctxt mib.mind_params_ctxt
@@ -96,20 +96,20 @@ let get_sym_eq_data env ind =
let (mib,mip as specif) = lookup_mind_specif env ind in
if Array.length mib.mind_packets <> 1 or Array.length mip.mind_nf_lc <> 1 then
error "Not an inductive type with a single constructor.";
- let realsign,_ = list_chop mip.mind_nrealargs_ctxt mip.mind_arity_ctxt in
+ let realsign,_ = List.chop mip.mind_nrealargs_ctxt mip.mind_arity_ctxt in
if List.exists (fun (_,b,_) -> b <> None) realsign then
error "Inductive equalities with local definitions in arity not supported.";
let constrsign,ccl = decompose_prod_assum mip.mind_nf_lc.(0) in
let _,constrargs = decompose_app ccl in
if rel_context_length constrsign<>rel_context_length mib.mind_params_ctxt then
error "Constructor must have no arguments"; (* This can be relaxed... *)
- let params,constrargs = list_chop mib.mind_nparams constrargs in
+ let params,constrargs = List.chop mib.mind_nparams constrargs in
if mip.mind_nrealargs > mib.mind_nparams then
error "Constructors arguments must repeat the parameters.";
- let _,params2 = list_chop (mib.mind_nparams-mip.mind_nrealargs) params in
+ let _,params2 = List.chop (mib.mind_nparams-mip.mind_nrealargs) params in
let paramsctxt1,_ =
- list_chop (mib.mind_nparams-mip.mind_nrealargs) mib.mind_params_ctxt in
- if not (list_equal eq_constr params2 constrargs) then
+ List.chop (mib.mind_nparams-mip.mind_nrealargs) mib.mind_params_ctxt in
+ if not (List.equal eq_constr params2 constrargs) then
error "Constructors arguments must repeat the parameters.";
(* nrealargs_ctxt and nrealargs are the same here *)
(specif,mip.mind_nrealargs,realsign,mib.mind_params_ctxt,paramsctxt1)
@@ -128,14 +128,14 @@ let get_non_sym_eq_data env ind =
let (mib,mip as specif) = lookup_mind_specif env ind in
if Array.length mib.mind_packets <> 1 or Array.length mip.mind_nf_lc <> 1 then
error "Not an inductive type with a single constructor.";
- let realsign,_ = list_chop mip.mind_nrealargs_ctxt mip.mind_arity_ctxt in
+ let realsign,_ = List.chop mip.mind_nrealargs_ctxt mip.mind_arity_ctxt in
if List.exists (fun (_,b,_) -> b <> None) realsign then
error "Inductive equalities with local definitions in arity not supported";
let constrsign,ccl = decompose_prod_assum mip.mind_nf_lc.(0) in
let _,constrargs = decompose_app ccl in
if rel_context_length constrsign<>rel_context_length mib.mind_params_ctxt then
error "Constructor must have no arguments";
- let _,constrargs = list_chop mib.mind_nparams constrargs in
+ let _,constrargs = List.chop mib.mind_nparams constrargs in
(specif,constrargs,realsign,mip.mind_nrealargs)
(**********************************************************************)
@@ -679,7 +679,7 @@ let build_congr env (eq,refl) ind =
if mip.mind_nrealargs <> 1 then
error "Expect an inductive type with one predicate parameter.";
let i = 1 in
- let realsign,_ = list_chop mip.mind_nrealargs_ctxt mip.mind_arity_ctxt in
+ let realsign,_ = List.chop mip.mind_nrealargs_ctxt mip.mind_arity_ctxt in
if List.exists (fun (_,b,_) -> b <> None) realsign then
error "Inductive equalities with local definitions in arity not supported.";
let env_with_arity = push_rel_context mip.mind_arity_ctxt env in
diff --git a/tactics/equality.ml b/tactics/equality.ml
index b4cb48285..4d67fef00 100644
--- a/tactics/equality.ml
+++ b/tactics/equality.ml
@@ -397,7 +397,7 @@ let general_multi_rewrite l2r with_evars ?tac c cl =
(* If the term to rewrite uses an hypothesis H, don't rewrite in H *)
let ids =
let ids_in_c = Environ.global_vars_set (Global.env()) (fst c) in
- Idset.fold (fun id l -> list_remove id l) ids_in_c (pf_ids_of_hyps gl)
+ Idset.fold (fun id l -> List.remove id l) ids_in_c (pf_ids_of_hyps gl)
in do_hyps_atleastonce ids gl
in
if cl.concl_occs = NoOccurrences then do_hyps else
@@ -532,15 +532,15 @@ let find_positions env sigma t1 t2 =
| Construct sp1, Construct sp2
when List.length args1 = mis_constructor_nargs_env env sp1
->
- let sorts = list_intersect sorts (allowed_sorts env (fst sp1)) in
+ let sorts = List.intersect sorts (allowed_sorts env (fst sp1)) in
(* both sides are fully applied constructors, so either we descend,
or we can discriminate here. *)
if is_conv env sigma hd1 hd2 then
let nrealargs = constructor_nrealargs env sp1 in
- let rargs1 = list_lastn nrealargs args1 in
- let rargs2 = list_lastn nrealargs args2 in
+ let rargs1 = List.lastn nrealargs args1 in
+ let rargs2 = List.lastn nrealargs args2 in
List.flatten
- (list_map2_i (fun i -> findrec sorts ((sp1,i)::posn))
+ (List.map2_i (fun i -> findrec sorts ((sp1,i)::posn))
0 rargs1 rargs2)
else if List.mem InType sorts then (* see build_discriminator *)
raise (DiscrFound (List.rev posn,sp1,sp2))
@@ -1495,7 +1495,7 @@ let subst_all ?(flags=default_subst_tactic_flags ()) gl =
with PatternMatchingFailure -> failwith "caught"
in
let ids = map_succeed test (pf_hyps_types gl) in
- let ids = list_uniquize ids in
+ let ids = List.uniquize ids in
subst_gen flags.rewrite_dependent_proof ids gl
(* Rewrite the first assumption for which the condition faildir does not fail
diff --git a/tactics/hipattern.ml4 b/tactics/hipattern.ml4
index a6dcb9d58..bb35cedea 100644
--- a/tactics/hipattern.ml4
+++ b/tactics/hipattern.ml4
@@ -319,7 +319,7 @@ let match_with_nodep_ind t =
if array_for_all nodep_constr mip.mind_nf_lc then
let params=
if mip.mind_nrealargs=0 then args else
- fst (list_chop mib.mind_nparams args) in
+ fst (List.chop mib.mind_nparams args) in
Some (hdapp,params,mip.mind_nrealargs)
else
None
diff --git a/tactics/inv.ml b/tactics/inv.ml
index cb1ffb385..189c73a16 100644
--- a/tactics/inv.ml
+++ b/tactics/inv.ml
@@ -111,7 +111,7 @@ let make_inv_predicate env sigma indf realargs id status concl =
let nhyps = rel_context_length hyps in
let env' = push_rel_context hyps env in
let realargs' = List.map (lift nhyps) realargs in
- let pairs = list_map_i (compute_eqn env' sigma nhyps) 0 realargs' in
+ let pairs = List.map_i (compute_eqn env' sigma nhyps) 0 realargs' in
(* Now the arity is pushed, and we need to construct the pairs
* ai,mkRel(n-i+1) *)
(* Now, we can recurse down this list, for each ai,(mkRel k) whether to
@@ -469,7 +469,7 @@ let raw_inversion inv_kind id status names gl =
(fun id ->
(tclTHEN
(apply_term (mkVar id)
- (list_tabulate (fun _ -> Evarutil.mk_new_meta()) neqns))
+ (List.tabulate (fun _ -> Evarutil.mk_new_meta()) neqns))
reflexivity))])
gl
diff --git a/tactics/leminv.ml b/tactics/leminv.ml
index 5242d0f3e..3031734fb 100644
--- a/tactics/leminv.ml
+++ b/tactics/leminv.ml
@@ -187,7 +187,7 @@ let inversion_scheme env sigma t sort dep_option inv_op =
compute_first_inversion_scheme env sigma ind sort dep_option
in
assert
- (list_subset
+ (List.subset
(global_vars env invGoal)
(ids_of_named_context (named_context invEnv)));
(*
diff --git a/tactics/refine.ml b/tactics/refine.ml
index fc2da8d0a..1cb4f01e2 100644
--- a/tactics/refine.ml
+++ b/tactics/refine.ml
@@ -348,7 +348,7 @@ let rec tcc_aux subst (TH (c,mm,sgp) as _th) gl =
| _ -> error "Recursive functions must have names."
in
let fixes = array_map3 (fun f n c -> (out_name f,succ n,c)) fi ni ai in
- let firsts,lasts = list_chop j (Array.to_list fixes) in
+ let firsts,lasts = List.chop j (Array.to_list fixes) in
tclTHENS
(tclTHEN
(ensure_products (succ ni.(j)))
@@ -365,7 +365,7 @@ let rec tcc_aux subst (TH (c,mm,sgp) as _th) gl =
| _ -> error "Recursive functions must have names."
in
let cofixes = array_map2 (fun f c -> (out_name f,c)) fi ai in
- let firsts,lasts = list_chop j (Array.to_list cofixes) in
+ let firsts,lasts = List.chop j (Array.to_list cofixes) in
tclTHENS
(mutual_cofix (out_name fi.(j)) (firsts@List.tl lasts) j)
(List.map (function
diff --git a/tactics/rewrite.ml4 b/tactics/rewrite.ml4
index f9aa1f025..c36ab2f83 100644
--- a/tactics/rewrite.ml4
+++ b/tactics/rewrite.ml4
@@ -725,7 +725,7 @@ let fold_match ?(force=false) env sigma c =
in
let app =
let ind, args = Inductive.find_rectype env cty in
- let pars, args = list_chop ci.ci_npar args in
+ let pars, args = List.chop ci.ci_npar args in
let meths = List.map (fun br -> br) (Array.to_list brs) in
applist (mkConst sk, pars @ [pred] @ meths @ args @ [c])
in
diff --git a/tactics/tacinterp.ml b/tactics/tacinterp.ml
index b0997c067..4ce382df2 100644
--- a/tactics/tacinterp.ml
+++ b/tactics/tacinterp.ml
@@ -92,7 +92,7 @@ let catch_error call_trace tac g =
| LtacLocated _ as e -> raise e
| Loc.Exc_located (_,LtacLocated _) as e -> raise e
| e ->
- let (nrep,loc',c),tail = list_sep_last call_trace in
+ let (nrep,loc',c),tail = List.sep_last call_trace in
let loc,e' = match e with Loc.Exc_located(loc,e) -> loc,e | _ ->dloc,e in
if tail = [] then
let loc = if loc = dloc then loc' else loc in
@@ -928,7 +928,7 @@ and intern_match_rule onlytac ist = function
let {ltacvars=(lfun,l2); gsigma=sigma; genv=env} = ist in
let lfun',metas1,hyps = intern_match_goal_hyps ist lfun rl in
let ido,metas2,pat = intern_pattern ist lfun mp in
- let metas = list_uniquize (metas1@metas2) in
+ let metas = List.uniquize (metas1@metas2) in
let ist' = { ist with ltacvars = (metas@(Option.List.cons ido lfun'),l2) } in
Pat (hyps,pat,intern_tactic onlytac ist' tc) :: (intern_match_rule onlytac ist tl)
| [] -> []
@@ -1329,7 +1329,7 @@ let interp_constr_in_compound_list inj_fun dest_fun interp_fun ist env sigma l =
(*all of dest_fun, List.assoc, constr_list_of_VList may raise Not_found*)
let sigma, c = interp_fun ist env sigma x in
sigma, [c] in
- let sigma, l = list_fold_map try_expand_ltac_var sigma l in
+ let sigma, l = List.fold_map try_expand_ltac_var sigma l in
sigma, List.flatten l
let interp_constr_list ist env sigma c =
@@ -1541,7 +1541,7 @@ let interp_bindings ist env sigma = function
let sigma, l = interp_open_constr_list ist env sigma l in
sigma, ImplicitBindings l
| ExplicitBindings l ->
- let sigma, l = list_fold_map (interp_binding ist env) sigma l in
+ let sigma, l = List.fold_map (interp_binding ist env) sigma l in
sigma, ExplicitBindings l
let interp_constr_with_bindings ist env sigma (c,bl) =
@@ -1556,8 +1556,8 @@ let interp_open_constr_with_bindings ist env sigma (c,bl) =
let loc_of_bindings = function
| NoBindings -> Loc.ghost
-| ImplicitBindings l -> loc_of_glob_constr (fst (list_last l))
-| ExplicitBindings l -> pi1 (list_last l)
+| ImplicitBindings l -> loc_of_glob_constr (fst (List.last l))
+| ExplicitBindings l -> pi1 (List.last l)
let interp_open_constr_with_bindings_loc ist env sigma ((c,_),bl as cb) =
let loc1 = loc_of_glob_constr c in
@@ -1982,7 +1982,7 @@ and eval_with_fail ist is_lazy goal tac =
(* Interprets the clauses of a recursive LetIn *)
and interp_letrec ist gl llc u =
let lref = ref ist.lfun in
- let lve = list_map_left (fun ((_,id),b) -> (id,VRec (lref,TacArg (dloc,b)))) llc in
+ let lve = List.map_left (fun ((_,id),b) -> (id,VRec (lref,TacArg (dloc,b)))) llc in
lref := lve@ist.lfun;
let ist = { ist with lfun = lve@ist.lfun } in
val_interp ist gl u
@@ -2069,7 +2069,7 @@ and apply_hyps_context ist env lz goal mt lctxt lgmatch mhyps hyps =
db_matched_hyp ist.debug (pf_env goal) hyp_match hypname;
try
let id_match = pi1 hyp_match in
- let nextlhyps = list_remove_assoc_in_triple id_match lhyps_rest in
+ let nextlhyps = List.remove_assoc_in_triple id_match lhyps_rest in
apply_hyps_context_rec (lfun@lids) lm nextlhyps tl
with e when is_match_catchable e ->
match_next_pattern find_next' in
@@ -2335,7 +2335,7 @@ and interp_atomic ist gl tac =
(h_vm_cast_no_check c_interp)
| TacApply (a,ev,cb,cl) ->
let sigma, l =
- list_fold_map (interp_open_constr_with_bindings_loc ist env) sigma cb
+ List.fold_map (interp_open_constr_with_bindings_loc ist env) sigma cb
in
let tac = match cl with
| None -> h_apply a ev
@@ -2435,7 +2435,7 @@ and interp_atomic ist gl tac =
h_simple_induction_destruct isrec (interp_quantified_hypothesis ist h)
| TacInductionDestruct (isrec,ev,(l,el,cls)) ->
let sigma, l =
- list_fold_map (fun sigma (c,(ipato,ipats)) ->
+ List.fold_map (fun sigma (c,(ipato,ipats)) ->
let c = interp_induction_arg ist gl c in
(sigma,(c,
(Option.map (interp_intro_pattern ist gl) ipato,
@@ -2492,7 +2492,7 @@ and interp_atomic ist gl tac =
let sigma, bl = interp_bindings ist env sigma bl in
tclWITHHOLES ev (h_right ev) sigma bl
| TacSplit (ev,_,bll) ->
- let sigma, bll = list_fold_map (interp_bindings ist env) sigma bll in
+ let sigma, bll = List.fold_map (interp_bindings ist env) sigma bll in
tclWITHHOLES ev (h_split ev) sigma bll
| TacAnyConstructor (ev,t) ->
abstract_tactic (TacAnyConstructor (ev,t))
@@ -3142,7 +3142,7 @@ let add_tacdef local isrec tacl =
let rfun = List.map (fun (ident, b, _) -> make_absolute_name ident b) tacl in
let ist =
{(make_empty_glob_sign()) with ltacrecvars =
- if isrec then list_map_filter
+ if isrec then List.map_filter
(function (Some id, qid) -> Some (id, qid) | (None, _) -> None) rfun
else []} in
let gtacl =
diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml
index aaa75a4e2..f88530eec 100644
--- a/tactics/tacticals.ml
+++ b/tactics/tacticals.ml
@@ -87,7 +87,7 @@ let lastHypId gl = nthHypId 1 gl
let lastHyp gl = nthHyp 1 gl
let nLastDecls n gl =
- try list_firstn n (pf_hyps gl)
+ try List.firstn n (pf_hyps gl)
with Failure _ -> error "Not enough hypotheses in the goal."
let nLastHypsId n gl = List.map pi1 (nLastDecls n gl)
@@ -108,7 +108,7 @@ let onNLastHypsId n tac = onHyps (nLastHypsId n) tac
let onNLastHyps n tac = onHyps (nLastHyps n) tac
let afterHyp id gl =
- fst (list_split_when (fun (hyp,_,_) -> hyp = id) (pf_hyps gl))
+ fst (List.split_when (fun (hyp,_,_) -> hyp = id) (pf_hyps gl))
(***************************************)
(* Clause Tacticals *)
@@ -174,7 +174,7 @@ let fix_empty_or_and_pattern nv l =
names and "[ ]" for no clause at all *)
(* 2- More generally, we admit "[ ]" for any disjunctive pattern of
arbitrary length *)
- if l = [[]] then list_make nv [] else l
+ if l = [[]] then List.make nv [] else l
let check_or_and_pattern_size loc names n =
if List.length names <> n then
@@ -335,7 +335,7 @@ let make_elim_branch_assumptions ba gl =
| (_, _) -> anomaly "make_elim_branch_assumptions"
in
makerec ([],[],[],[],[]) ba.branchsign
- (try list_firstn ba.nassums (pf_hyps gl)
+ (try List.firstn ba.nassums (pf_hyps gl)
with Failure _ -> anomaly "make_elim_branch_assumptions")
let elim_on_ba tac ba gl = tac (make_elim_branch_assumptions ba gl) gl
@@ -359,7 +359,7 @@ let make_case_branch_assumptions ba gl =
| (_, _) -> anomaly "make_case_branch_assumptions"
in
makerec ([],[],[],[]) ba.branchsign
- (try list_firstn ba.nassums (pf_hyps gl)
+ (try List.firstn ba.nassums (pf_hyps gl)
with Failure _ -> anomaly "make_case_branch_assumptions")
let case_on_ba tac ba gl = tac (make_case_branch_assumptions ba gl) gl
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 1ccb3ffdb..b1a9c6732 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -857,7 +857,7 @@ let clenv_fchain_in id ?(flags=elim_flags) mv elimclause hypclause =
let elimination_in_clause_scheme with_evars ?(flags=elim_flags) id i elimclause indclause gl =
let indmv = destMeta (nth_arg i elimclause.templval.rebus) in
let hypmv =
- try match list_remove indmv (clenv_independent elimclause) with
+ try match List.remove indmv (clenv_independent elimclause) with
| [a] -> a
| _ -> failwith ""
with Failure _ -> errorlabstrm "elimination_clause"
@@ -929,7 +929,7 @@ let descend_in_conjunctions tac exit c gl =
let elim = pf_apply build_case_analysis_scheme gl ind false sort in
NotADefinedRecordUseScheme elim in
tclFIRST
- (list_tabulate (fun i gl ->
+ (List.tabulate (fun i gl ->
match make_projection (project gl) params cstr sign elim i n c with
| None -> tclFAIL 0 (mt()) gl
| Some (p,pt) ->
@@ -1025,7 +1025,7 @@ let progress_with_clause flags innerclause clause =
if ordered_metas = [] then error "Statement without assumptions.";
let f mv =
find_matching_clause (clenv_fchain mv ~flags clause) innerclause in
- try list_try_find f ordered_metas
+ try List.try_find f ordered_metas
with Failure _ -> error "Unable to unify."
let apply_in_once_main flags innerclause (d,lbind) gl =
@@ -1164,7 +1164,7 @@ let specialize mopt (c,lbind) g =
let nargs = List.length tstack in
let tstack = match mopt with
| Some m ->
- if m < nargs then list_firstn m tstack else tstack
+ if m < nargs then List.firstn m tstack else tstack
| None ->
let rec chk = function
| [] -> []
@@ -1565,8 +1565,8 @@ let generalize_dep ?(with_let=false) c gl =
let generalize_gen_let lconstr gl =
let newcl =
- list_fold_right_i (generalize_goal gl) 0 lconstr (pf_concl gl) in
- apply_type newcl (list_map_filter (fun ((_,c,b),_) ->
+ List.fold_right_i (generalize_goal gl) 0 lconstr (pf_concl gl) in
+ apply_type newcl (List.map_filter (fun ((_,c,b),_) ->
if b = None then Some c else None) lconstr) gl
let generalize_gen lconstr =
@@ -1753,7 +1753,7 @@ let letin_abstract id c (test,out) (occs,check_occs) gl =
| Some occ ->
subst1 (mkVar id) (subst_closed_term_occ_modulo occ test None (pf_concl gl)) in
let lastlhyp =
- if depdecls = [] then MoveLast else MoveAfter(pi1(list_last depdecls)) in
+ if depdecls = [] then MoveLast else MoveAfter(pi1(List.last depdecls)) in
(depdecls,lastlhyp,ccl,out test)
let letin_tac_gen with_eq name (sigmac,c) test ty occs gl =
@@ -1998,7 +1998,7 @@ let atomize_param_of_ind (indref,nparams,_) hyp0 gl =
let typ0 = pf_apply reduce_to_quantified_ref gl indref tmptyp0 in
let prods, indtyp = decompose_prod typ0 in
let argl = snd (decompose_app indtyp) in
- let params = list_firstn nparams argl in
+ let params = List.firstn nparams argl in
(* le gl est important pour ne pas préévaluer *)
let rec atomize_one i avoid gl =
if i<>nparams then
@@ -2031,7 +2031,7 @@ let atomize_param_of_ind (indref,nparams,_) hyp0 gl =
let find_atomic_param_of_ind nparams indtyp =
let argl = snd (decompose_app indtyp) in
let argv = Array.of_list argl in
- let params = list_firstn nparams argl in
+ let params = List.firstn nparams argl in
let indvars = ref Idset.empty in
for i = nparams to (Array.length argv)-1 do
match kind_of_term argv.(i) with
@@ -2749,8 +2749,8 @@ let compute_scheme_signature scheme names_info ind_type_guess =
(* Check again conclusion *)
let ccl_arg_ok = is_pred (p + scheme.nargs + 1) f = IndArg in
let ind_is_ok =
- list_equal eq_constr
- (list_lastn scheme.nargs indargs)
+ List.equal eq_constr
+ (List.lastn scheme.nargs indargs)
(extended_rel_list 0 scheme.args) in
if not (ccl_arg_ok & ind_is_ok) then
error_ind_scheme "the conclusion of"
@@ -2895,11 +2895,11 @@ let recolle_clenv nparams lid elimclause gl =
let nidargs = List.length lidargs in
(* parameters correspond to first elts of lid. *)
let clauses_params =
- list_map_i (fun i id -> mkVar id , pf_get_hyp_typ gl id , lindmv.(i))
+ List.map_i (fun i id -> mkVar id , pf_get_hyp_typ gl id , lindmv.(i))
0 lidparams in
(* arguments correspond to last elts of lid. *)
let clauses_args =
- list_map_i
+ List.map_i
(fun i id -> mkVar id , pf_get_hyp_typ gl id , lindmv.(nmv-nidargs+i))
0 lidargs in
let clauses = clauses_params@clauses_args in
@@ -3087,7 +3087,7 @@ let clear_unselected_context id inhyps cls gl =
(* erase if not selected and dependent on id or selected hyps *)
let test id = occur_var_in_decl (pf_env gl) id d in
if List.exists test (id::inhyps) then Some id' else None in
- let ids = list_map_filter to_erase (pf_hyps gl) in
+ let ids = List.map_filter to_erase (pf_hyps gl) in
thin ids gl
| None -> tclIDTAC gl
diff --git a/tactics/tauto.ml4 b/tactics/tauto.ml4
index 969669668..fdfc2b783 100644
--- a/tactics/tauto.ml4
+++ b/tactics/tauto.ml4
@@ -173,7 +173,7 @@ let flatten_contravariant_disj flags ist =
typ with
| Some (_,args) ->
let hyp = valueIn (VConstr ([],hyp)) in
- iter_tac (list_map_i (fun i arg ->
+ iter_tac (List.map_i (fun i arg ->
let typ = valueIn (VConstr ([],mkArrow arg c)) in
let i = Tacexpr.Integer i in
<:tactic<
diff --git a/tools/coq_makefile.ml b/tools/coq_makefile.ml
index 7e400a332..b6268f233 100644
--- a/tools/coq_makefile.ml
+++ b/tools/coq_makefile.ml
@@ -158,7 +158,7 @@ let vars_to_put_by_root var_x_files_l (inc_i,inc_r) =
if inc_r = [] then
[".","$(INSTALLDEFAULTROOT)",[]]
else
- Util.list_fold_left_i (fun i out (pdir,ldir,abspdir) ->
+ Util.List.fold_left_i (fun i out (pdir,ldir,abspdir) ->
let vars_r = var_filter (List.exists (is_prefix abspdir)) (fun x -> x^string_of_int i) in
let pdir' = physical_dir_of_logical_dir ldir in
(pdir,pdir',vars_r)::out) 1 [] inc_r
diff --git a/toplevel/command.ml b/toplevel/command.ml
index 3ccfe21d8..c54dc8120 100644
--- a/toplevel/command.ml
+++ b/toplevel/command.ml
@@ -229,12 +229,12 @@ let minductive_message = function
let check_all_names_different indl =
let ind_names = List.map (fun ind -> ind.ind_name) indl in
- let cstr_names = list_map_append (fun ind -> List.map fst ind.ind_lc) indl in
- let l = list_duplicates ind_names in
+ let cstr_names = List.map_append (fun ind -> List.map fst ind.ind_lc) indl in
+ let l = List.duplicates ind_names in
if l <> [] then raise (InductiveError (SameNamesTypes (List.hd l)));
- let l = list_duplicates cstr_names in
+ let l = List.duplicates cstr_names in
if l <> [] then raise (InductiveError (SameNamesConstructors (List.hd l)));
- let l = list_intersect ind_names cstr_names in
+ let l = List.intersect ind_names cstr_names in
if l <> [] then raise (InductiveError (SameNamesOverlap l))
let mk_mltype_data evdref env assums arity indname =
@@ -286,7 +286,7 @@ let interp_mutual_inductive (paramsl,indl) notations finite =
(* Temporary declaration of notations and scopes *)
List.iter (Metasyntax.set_notation_for_interpretation impls) notations;
(* Interpret the constructor types *)
- list_map3 (interp_cstrs evdref env_ar_params impls) mldatas arities indl)
+ List.map3 (interp_cstrs evdref env_ar_params impls) mldatas arities indl)
() in
(* Instantiate evars and check all are resolved *)
@@ -303,7 +303,7 @@ let interp_mutual_inductive (paramsl,indl) notations finite =
constructors;
(* Build the inductive entries *)
- let entries = list_map3 (fun ind arity (cnames,ctypes,cimpls) -> {
+ let entries = List.map3 (fun ind arity (cnames,ctypes,cimpls) -> {
mind_entry_typename = ind.ind_name;
mind_entry_arity = arity;
mind_entry_consnames = cnames;
@@ -368,11 +368,11 @@ let declare_mutual_inductive_with_eliminations isrecord mie impls =
let names = List.map (fun e -> e.mind_entry_typename) mie.mind_entry_inds in
let (_,kn) = declare_mind isrecord mie in
let mind = Global.mind_of_delta_kn kn in
- list_iter_i (fun i (indimpls, constrimpls) ->
+ List.iter_i (fun i (indimpls, constrimpls) ->
let ind = (mind,i) in
Autoinstance.search_declaration (IndRef ind);
maybe_declare_manual_implicits false (IndRef ind) indimpls;
- list_iter_i
+ List.iter_i
(fun j impls ->
(* Autoinstance.search_declaration (ConstructRef (ind,j));*)
maybe_declare_manual_implicits false (ConstructRef (ind, succ j)) impls)
@@ -423,7 +423,7 @@ let rec partial_order = function
let rec browse res xge' = function
| [] ->
let res = List.map (function
- | (z, Inr zge) when List.mem x zge -> (z, Inr (list_union zge xge'))
+ | (z, Inr zge) when List.mem x zge -> (z, Inr (List.union zge xge'))
| r -> r) res in
(x,Inr xge')::res
| y::xge ->
@@ -438,13 +438,13 @@ let rec partial_order = function
if t = y then (z, Inl x) else (z, Inl t)
| (z, Inr zge) ->
if List.mem y zge then
- (z, Inr (list_add_set x (list_remove y zge)))
+ (z, Inr (List.add_set x (List.remove y zge)))
else
(z, Inr zge)) res in
- browse ((y,Inl x)::res) xge' (list_union xge (list_remove x yge))
+ browse ((y,Inl x)::res) xge' (List.union xge (List.remove x yge))
else
- browse res (list_add_set y (list_union xge' yge)) xge
- with Not_found -> browse res (list_add_set y xge') xge
+ browse res (List.add_set y (List.union xge' yge)) xge
+ with Not_found -> browse res (List.add_set y xge') xge
in link y
in browse (partial_order rest) [] xge
@@ -734,12 +734,12 @@ let interp_recursive isfix fixl notations =
(* Interp arities allowing for unresolved types *)
let evdref = ref Evd.empty in
let fixctxs, fiximppairs, fixannots =
- list_split3 (List.map (interp_fix_context evdref env isfix) fixl) in
+ List.split3 (List.map (interp_fix_context evdref env isfix) fixl) in
let fixctximpenvs, fixctximps = List.split fiximppairs in
- let fixccls,fixcclimps = List.split (list_map3 (interp_fix_ccl evdref) fixctximpenvs fixctxs fixl) in
+ let fixccls,fixcclimps = List.split (List.map3 (interp_fix_ccl evdref) fixctximpenvs fixctxs fixl) in
let fixtypes = List.map2 build_fix_type fixctxs fixccls in
let fixtypes = List.map (nf_evar !evdref) fixtypes in
- let fiximps = list_map3
+ let fiximps = List.map3
(fun ctximps cclimps (_,ctx) -> ctximps@(Impargs.lift_implicits (List.length ctx) cclimps))
fixctximps fixcclimps fixctxs in
let rec_sign =
@@ -764,7 +764,7 @@ let interp_recursive isfix fixl notations =
let fixdefs =
Metasyntax.with_syntax_protection (fun () ->
List.iter (Metasyntax.set_notation_for_interpretation impls) notations;
- list_map4
+ List.map4
(fun fixctximpenv -> interp_fix_body evdref env_rec (Idmap.fold Idmap.add fixctximpenv impls))
fixctximpenvs fixctxs fixl fixccls)
() in
@@ -776,7 +776,7 @@ let interp_recursive isfix fixl notations =
let fixctxnames = List.map (fun (_,ctx) -> List.map pi1 ctx) fixctxs in
(* Build the fix declaration block *)
- (env,rec_sign,evd), (fixnames,fixdefs,fixtypes), list_combine3 fixctxnames fiximps fixannots
+ (env,rec_sign,evd), (fixnames,fixdefs,fixtypes), List.combine3 fixctxnames fiximps fixannots
let check_recursive isfix ((env,rec_sign,evd),(fixnames,fixdefs,fixtypes),info) =
let evd = Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals ~fail:true env evd in
@@ -795,7 +795,7 @@ let declare_fixpoint ((fixnames,fixdefs,fixtypes),fiximps) indexes ntns =
if List.mem None fixdefs then
(* Some bodies to define by proof *)
let thms =
- list_map3 (fun id t (len,imps,_) -> (id,(t,(len,imps)))) fixnames fixtypes fiximps in
+ List.map3 (fun id t (len,imps,_) -> (id,(t,(len,imps)))) fixnames fixtypes fiximps in
let init_tac =
Some (List.map (Option.cata Tacmach.refine_no_check Tacticals.tclIDTAC)
fixdefs) in
@@ -808,8 +808,8 @@ let declare_fixpoint ((fixnames,fixdefs,fixtypes),fiximps) indexes ntns =
let indexes = search_guard Loc.ghost (Global.env()) indexes fixdecls in
let fiximps = List.map (fun (n,r,p) -> r) fiximps in
let fixdecls =
- list_map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 fixnames in
- ignore (list_map4 (declare_fix Fixpoint) fixnames fixdecls fixtypes fiximps);
+ List.map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 fixnames in
+ ignore (List.map4 (declare_fix Fixpoint) fixnames fixdecls fixtypes fiximps);
(* Declare the recursive definitions *)
fixpoint_message (Some indexes) fixnames;
end;
@@ -820,7 +820,7 @@ let declare_cofixpoint ((fixnames,fixdefs,fixtypes),fiximps) ntns =
if List.mem None fixdefs then
(* Some bodies to define by proof *)
let thms =
- list_map3 (fun id t (len,imps,_) -> (id,(t,(len,imps)))) fixnames fixtypes fiximps in
+ List.map3 (fun id t (len,imps,_) -> (id,(t,(len,imps)))) fixnames fixtypes fiximps in
let init_tac =
Some (List.map (Option.cata Tacmach.refine_no_check Tacticals.tclIDTAC)
fixdefs) in
@@ -830,9 +830,9 @@ let declare_cofixpoint ((fixnames,fixdefs,fixtypes),fiximps) ntns =
(* We shortcut the proof process *)
let fixdefs = List.map Option.get fixdefs in
let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in
- let fixdecls = list_map_i (fun i _ -> mkCoFix (i,fixdecls)) 0 fixnames in
+ let fixdecls = List.map_i (fun i _ -> mkCoFix (i,fixdecls)) 0 fixnames in
let fiximps = List.map (fun (len,imps,idx) -> imps) fiximps in
- ignore (list_map4 (declare_fix CoFixpoint) fixnames fixdecls fixtypes fiximps);
+ ignore (List.map4 (declare_fix CoFixpoint) fixnames fixdecls fixtypes fiximps);
(* Declare the recursive definitions *)
cofixpoint_message fixnames
end;
@@ -911,7 +911,7 @@ let do_program_recursive fixkind fixl ntns =
let (fixnames,fixdefs,fixtypes) = fix in
let fiximps = List.map pi2 info in
let fixdefs = List.map Option.get fixdefs in
- let defs = list_map4 collect_evars fixnames fixdefs fixtypes fiximps in
+ let defs = List.map4 collect_evars fixnames fixdefs fixtypes fiximps in
if isfix then begin
let possible_indexes = List.map compute_possible_guardness_evidences info in
let fixdecls = Array.of_list (List.map (fun x -> Name x) fixnames),
@@ -920,7 +920,7 @@ let do_program_recursive fixkind fixl ntns =
in
let indexes =
Pretyping.search_guard Loc.ghost (Global.env ()) possible_indexes fixdecls in
- list_iter_i (fun i _ -> Inductive.check_fix env ((indexes,i),fixdecls)) fixl
+ List.iter_i (fun i _ -> Inductive.check_fix env ((indexes,i),fixdecls)) fixl
end;
Obligations.add_mutual_definitions defs ntns fixkind
diff --git a/toplevel/discharge.ml b/toplevel/discharge.ml
index 5ed45535a..92b0d6536 100644
--- a/toplevel/discharge.ml
+++ b/toplevel/discharge.ml
@@ -38,7 +38,7 @@ let abstract_inductive hyps nparams inds =
let ntyp = List.length inds in
let nhyp = named_context_length hyps in
let args = instance_from_named_context (List.rev hyps) in
- let subs = list_tabulate (fun k -> lift nhyp (mkApp(mkRel (k+1),args))) ntyp in
+ let subs = List.tabulate (fun k -> lift nhyp (mkApp(mkRel (k+1),args))) ntyp in
let inds' =
List.map
(function (tname,arity,cnames,lc) ->
diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml
index 3cdeb0be9..53c4325e1 100644
--- a/toplevel/himsg.ml
+++ b/toplevel/himsg.ml
@@ -1014,7 +1014,7 @@ let explain_ltac_call_trace (nrep,last,trace,loc) =
| Proof_type.LtacConstrInterp (c,(vars,unboundvars)) ->
let filter =
function (id,None) -> None | (id,Some id') -> Some(id,([],mkVar id')) in
- let unboundvars = list_map_filter filter unboundvars in
+ let unboundvars = List.map_filter filter unboundvars in
quote (pr_glob_constr_env (Global.env()) c) ++
(if unboundvars <> [] or vars <> [] then
strbrk " (with " ++
@@ -1027,7 +1027,7 @@ let explain_ltac_call_trace (nrep,last,trace,loc) =
else if n>2 then str " (repeated "++int n++str" times)"
else mt()) in
if calls <> [] then
- let kind_of_last_call = match list_last calls with
+ let kind_of_last_call = match List.last calls with
| (_,Proof_type.LtacConstrInterp _) -> ", last term evaluation failed."
| _ -> ", last call failed." in
hov 0 (str "In nested Ltac calls to " ++
diff --git a/toplevel/indschemes.ml b/toplevel/indschemes.ml
index b99f83e5c..b3ea8438a 100644
--- a/toplevel/indschemes.ml
+++ b/toplevel/indschemes.ml
@@ -176,7 +176,7 @@ let beq_scheme_msg mind =
(* TODO: mutual inductive case *)
str "Boolean equality on " ++
pr_enum (fun ind -> quote (Printer.pr_inductive (Global.env()) ind))
- (list_tabulate (fun i -> (mind,i)) (Array.length mib.mind_packets))
+ (List.tabulate (fun i -> (mind,i)) (Array.length mib.mind_packets))
let declare_beq_scheme_with l kn =
try_declare_scheme (beq_scheme_msg kn) declare_beq_scheme_gen UserVerbose l kn
@@ -219,7 +219,7 @@ let declare_one_induction_scheme ind =
let from_prop = kind = InProp in
let kelim = elim_sorts (mib,mip) in
let elims =
- list_map_filter (fun (sort,kind) ->
+ List.map_filter (fun (sort,kind) ->
if List.mem sort kelim then Some kind else None)
(if from_prop then kinds_from_prop else kinds_from_type) in
List.iter (fun kind -> ignore (define_individual_scheme kind KernelVerbose None ind))
@@ -365,10 +365,10 @@ let get_common_underlying_mutual_inductive = function
| (_,ind')::_ ->
raise (RecursionSchemeError (NotMutualInScheme (ind,ind')))
| [] ->
- if not (list_distinct (List.map snd (List.map snd all))) then
+ if not (List.distinct (List.map snd (List.map snd all))) then
error "A type occurs twice";
mind,
- list_map_filter
+ List.map_filter
(function (Some id,(_,i)) -> Some (i,snd id) | (None,_) -> None) all
let do_scheme l =
@@ -394,7 +394,7 @@ let list_split_rev_at index l =
let rec aux i acc = function
hd :: tl when i = index -> acc, tl
| hd :: tl -> aux (succ i) (hd :: acc) tl
- | [] -> failwith "list_split_when: Invalid argument"
+ | [] -> failwith "List.split_when: Invalid argument"
in aux 0 [] l
let fold_left' f = function
diff --git a/toplevel/lemmas.ml b/toplevel/lemmas.ml
index 0c9cf877d..25a8e9208 100644
--- a/toplevel/lemmas.ml
+++ b/toplevel/lemmas.ml
@@ -84,7 +84,7 @@ let find_mutually_recursive_statements thms =
(fun env c -> fst (whd_betadeltaiota_stack env Evd.empty c))
(Global.env()) hyps in
let ind_hyps =
- List.flatten (list_map_i (fun i (_,b,t) ->
+ List.flatten (List.map_i (fun i (_,b,t) ->
match kind_of_term t with
| Ind (kn,_ as ind) when
let mind = Global.lookup_mind kn in
@@ -108,14 +108,14 @@ let find_mutually_recursive_statements thms =
(* Check if all conclusions are coinductive in the same type *)
(* (degenerated cartesian product since there is at most one coind ccl) *)
let same_indccl =
- list_cartesians_filter (fun hyp oks ->
+ List.cartesians_filter (fun hyp oks ->
if List.for_all (of_same_mutind hyp) oks
then Some (hyp::oks) else None) [] ind_ccls in
let ordered_same_indccl =
- List.filter (list_for_all_i (fun i ((kn,j),_,_) -> i=j) 0) same_indccl in
+ List.filter (List.for_all_i (fun i ((kn,j),_,_) -> i=j) 0) same_indccl in
(* Check if some hypotheses are inductive in the same type *)
let common_same_indhyp =
- list_cartesians_filter (fun hyp oks ->
+ List.cartesians_filter (fun hyp oks ->
if List.for_all (of_same_mutind hyp) oks
then Some (hyp::oks) else None) [] inds_hyps in
let ordered_inds,finite,guard =
@@ -129,11 +129,11 @@ let find_mutually_recursive_statements thms =
indccl, true, []
| [], _::_ ->
if same_indccl <> [] &&
- list_distinct (List.map pi1 (List.hd same_indccl)) then
+ List.distinct (List.map pi1 (List.hd same_indccl)) then
if_verbose msg_info (strbrk "Coinductive statements do not follow the order of definition, assuming the proof to be by induction."); flush_all ();
let possible_guards = List.map (List.map pi3) inds_hyps in
(* assume the largest indices as possible *)
- list_last common_same_indhyp, false, possible_guards
+ List.last common_same_indhyp, false, possible_guards
| _, [] ->
error
("Cannot find common (mutual) inductive premises or coinductive" ^
@@ -265,7 +265,7 @@ let rec_tac_initializer finite guard thms snl =
else
(* nl is dummy: it will be recomputed at Qed-time *)
let nl = match snl with
- | None -> List.map succ (List.map list_last guard)
+ | None -> List.map succ (List.map List.last guard)
| Some nl -> nl
in match List.map2 (fun (id,(t,_)) n -> (id,n,t)) thms nl with
| (id,n,_)::l -> Hiddentac.h_mutual_fix true id n l
@@ -302,7 +302,7 @@ let start_proof_with_initialization kind recguard thms snl hook =
if other_thms = [] then [] else
(* there are several theorems defined mutually *)
let body,opaq = retrieve_first_recthm ref in
- list_map_i (save_remaining_recthms kind body opaq) 1 other_thms in
+ List.map_i (save_remaining_recthms kind body opaq) 1 other_thms in
let thms_data = (strength,ref,imps)::other_thms_data in
List.iter (fun (strength,ref,imps) ->
maybe_declare_manual_implicits false ref imps;
diff --git a/toplevel/metasyntax.ml b/toplevel/metasyntax.ml
index 4eebfc8f7..534a0a7dd 100644
--- a/toplevel/metasyntax.ml
+++ b/toplevel/metasyntax.ml
@@ -372,7 +372,7 @@ let analyze_notation_tokens l =
let l = raw_analyze_notation_tokens l in
let vars = get_notation_vars l in
let recvars,l = interp_list_parser [] l in
- recvars, list_subtract vars (List.map snd recvars), l
+ recvars, List.subtract vars (List.map snd recvars), l
let error_not_same_scope x y =
error ("Variables "^string_of_id x^" and "^string_of_id y^
@@ -448,7 +448,7 @@ let make_hunks etyps symbols from =
let vars,typs = List.split etyps in
let rec make ws = function
| NonTerminal m :: prods ->
- let i = list_index m vars in
+ let i = List.index m vars in
let _,prec = precedence_of_entry_type from (List.nth typs (i-1)) in
let u = UnpMetaVar (i,prec) in
if prods <> [] && is_non_terminal (List.hd prods) then
@@ -494,14 +494,14 @@ let make_hunks etyps symbols from =
add_break n (make NoBreak prods)
| SProdList (m,sl) :: prods ->
- let i = list_index m vars in
+ let i = List.index m vars in
let typ = List.nth typs (i-1) in
let _,prec = precedence_of_entry_type from typ in
let sl' =
(* If no separator: add a break *)
if sl = [] then add_break 1 []
(* We add NonTerminal for simulation but remove it afterwards *)
- else snd (list_sep_last (make NoBreak (sl@[NonTerminal m]))) in
+ else snd (List.sep_last (make NoBreak (sl@[NonTerminal m]))) in
let hunk = match typ with
| ETConstr _ -> UnpListMetaVar (i,prec,sl')
| ETBinder isopen ->
@@ -560,7 +560,7 @@ let hunks_of_format (from,(vars,typs)) symfmt =
when s = drop_simple_quotes s' ->
let symbs, l = aux (symbs,fmt) in symbs, UnpTerminal s :: l
| NonTerminal s :: symbs, UnpTerminal s' :: fmt when s = id_of_string s' ->
- let i = list_index s vars in
+ let i = List.index s vars in
let _,prec = precedence_of_entry_type from (List.nth typs (i-1)) in
let symbs, l = aux (symbs,fmt) in symbs, UnpMetaVar (i,prec) :: l
| symbs, UnpBox (a,b) :: fmt ->
@@ -570,7 +570,7 @@ let hunks_of_format (from,(vars,typs)) symfmt =
| symbs, (UnpCut _ as u) :: fmt ->
let symbs, l = aux (symbs,fmt) in symbs, u :: l
| SProdList (m,sl) :: symbs, fmt ->
- let i = list_index m vars in
+ let i = List.index m vars in
let typ = List.nth typs (i-1) in
let _,prec = precedence_of_entry_type from typ in
let slfmt,fmt = read_recursive_format sl fmt in
@@ -889,7 +889,7 @@ let find_precedence lev etyps symbols =
error "A left-recursive notation must have an explicit level."
else [],Option.get lev)
| Terminal _ ::l when
- (match list_last symbols with Terminal _ -> true |_ -> false)
+ (match List.last symbols with Terminal _ -> true |_ -> false)
->
if lev = None then
([msg_info,strbrk "Setting notation at level 0."], 0)
@@ -1181,7 +1181,7 @@ let cache_scope_command o =
let subst_scope_command (subst,(scope,o as x)) = match o with
| ScopeClasses cl ->
- let cl' = list_map_filter (subst_scope_class subst) cl in
+ let cl' = List.map_filter (subst_scope_class subst) cl in
let cl' =
if List.length cl = List.length cl' && List.for_all2 (==) cl cl' then cl
else cl' in
diff --git a/toplevel/obligations.ml b/toplevel/obligations.ml
index 2a0f3c167..b9783a9cb 100644
--- a/toplevel/obligations.ml
+++ b/toplevel/obligations.ml
@@ -31,7 +31,7 @@ let trace s =
let succfix (depth, fixrels) =
(succ depth, List.map succ fixrels)
-let mkMetas n = list_tabulate (fun _ -> Evarutil.mk_new_meta ()) n
+let mkMetas n = List.tabulate (fun _ -> Evarutil.mk_new_meta ()) n
let check_evars env evm =
List.iter
@@ -80,7 +80,7 @@ let subst_evar_constr evs n idf t =
*)
let args =
let n = match chop with None -> 0 | Some c -> c in
- let (l, r) = list_chop n (List.rev (Array.to_list args)) in
+ let (l, r) = List.chop n (List.rev (Array.to_list args)) in
List.rev r
in
let args =
@@ -108,7 +108,7 @@ let subst_evar_constr evs n idf t =
(** Substitute variable references in t using De Bruijn indices,
where n binders were passed through. *)
let subst_vars acc n t =
- let var_index id = Util.list_index id acc in
+ let var_index id = Util.List.index id acc in
let rec substrec depth c = match kind_of_term c with
| Var v -> (try mkRel (depth + (var_index v)) with Not_found -> c)
| _ -> map_constr_with_binders succ substrec depth c
@@ -146,7 +146,7 @@ open Tacticals
let trunc_named_context n ctx =
let len = List.length ctx in
- list_firstn (len - n) ctx
+ List.firstn (len - n) ctx
let rec chop_product n t =
if n = 0 then Some t
@@ -547,13 +547,13 @@ let compute_possible_guardness_evidences (n,_) fixbody fixtype =
fixpoints ?) *)
let m = Term.nb_prod fixtype in
let ctx = fst (decompose_prod_n_assum m fixtype) in
- list_map_i (fun i _ -> i) 0 ctx
+ List.map_i (fun i _ -> i) 0 ctx
let declare_mutual_definition l =
let len = List.length l in
let first = List.hd l in
let fixdefs, fixtypes, fiximps =
- list_split3
+ List.split3
(List.map (fun x ->
let subs, typ = (subst_body true x) in
let term = snd (Reductionops.splay_lam_n (Global.env ()) Evd.empty len subs) in
@@ -571,14 +571,14 @@ let declare_mutual_definition l =
match fixkind with
| IsFixpoint wfl ->
let possible_indexes =
- list_map3 compute_possible_guardness_evidences wfl fixdefs fixtypes in
+ List.map3 compute_possible_guardness_evidences wfl fixdefs fixtypes in
let indexes = Pretyping.search_guard Loc.ghost (Global.env ()) possible_indexes fixdecls in
- Some indexes, list_map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 l
+ Some indexes, List.map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 l
| IsCoFixpoint ->
- None, list_map_i (fun i _ -> mkCoFix (i,fixdecls)) 0 l
+ None, List.map_i (fun i _ -> mkCoFix (i,fixdecls)) 0 l
in
(* Declare the recursive definitions *)
- let kns = list_map4 (!declare_fix_ref kind) fixnames fixdecls fixtypes fiximps in
+ let kns = List.map4 (!declare_fix_ref kind) fixnames fixdecls fixtypes fiximps in
(* Declare notations *)
List.iter Metasyntax.add_notation_interpretation first.prg_notations;
Declare.recursive_message (fixkind<>IsCoFixpoint) indexes fixnames;
diff --git a/toplevel/record.ml b/toplevel/record.ml
index 978329e13..487add316 100644
--- a/toplevel/record.ml
+++ b/toplevel/record.ml
@@ -158,7 +158,7 @@ let subst_projection fid l c =
c''
let instantiate_possibly_recursive_type indsp paramdecls fields =
- let subst = list_map_i (fun i _ -> mkRel i) 1 paramdecls in
+ let subst = List.map_i (fun i _ -> mkRel i) 1 paramdecls in
Termops.substl_rel_context (subst@[mkInd indsp]) fields
(* We build projections *)
@@ -173,7 +173,7 @@ let declare_projections indsp ?(kind=StructureComponent) ?name coers fieldimpls
let fields = instantiate_possibly_recursive_type indsp paramdecls fields in
let lifted_fields = Termops.lift_rel_context 1 fields in
let (_,kinds,sp_projs,_) =
- list_fold_left3
+ List.fold_left3
(fun (nfi,kinds,sp_projs,subst) coe (fi,optci,ti) impls ->
let (sp_projs,subst) =
match fi with
@@ -234,7 +234,7 @@ let structure_signature ctx =
| (_,_,typ)::tl ->
let ev = Evarutil.new_untyped_evar() in
let evm = Evd.add evm ev (Evd.make_evar Environ.empty_named_context_val typ) in
- let new_tl = Util.list_map_i
+ let new_tl = Util.List.map_i
(fun pos (n,c,t) -> n,c,
Termops.replace_term (mkRel pos) (mkEvar(ev,[||])) t) 1 tl in
deps_to_evar evm new_tl in
@@ -280,7 +280,7 @@ let declare_structure finite infer id idbuild paramimpls params arity fieldimpls
rsp
let implicits_of_context ctx =
- list_map_i (fun i name ->
+ List.map_i (fun i name ->
let explname =
match name with
| Name n -> Some n
@@ -346,7 +346,7 @@ let declare_class finite def infer id idbuild paramimpls params arity fieldimpls
Option.map (fun b -> if b then Backward, pri else Forward, pri) coe)
coers priorities
in
- IndRef ind, (list_map3 (fun (id, _, _) b y -> (id, b, y))
+ IndRef ind, (List.map3 (fun (id, _, _) b y -> (id, b, y))
(List.rev fields) coers (Recordops.lookup_projections ind))
in
let ctx_context =
@@ -362,7 +362,7 @@ let declare_class finite def infer id idbuild paramimpls params arity fieldimpls
cl_props = fields;
cl_projs = projs }
in
-(* list_iter3 (fun p sub pri -> *)
+(* List.iter3 (fun p sub pri -> *)
(* if sub then match p with (_, _, Some p) -> declare_instance_cst true p pri | _ -> ()) *)
(* k.cl_projs coers priorities; *)
add_class k; impl
@@ -389,7 +389,7 @@ let definition_structure (kind,finite,infer,(is_coe,(loc,idstruc)),ps,cfs,idbuil
| Vernacexpr.DefExpr ((_,Name id),_,_) -> id::acc
| _ -> acc in
let allnames = idstruc::(List.fold_left extract_name [] fs) in
- if not (list_distinct allnames) then error "Two objects have the same name";
+ if not (List.distinct allnames) then error "Two objects have the same name";
if not (kind = Class false) && List.exists ((<>) None) priorities then
error "Priorities only allowed for type class substructures";
(* Now, younger decl in params and fields is on top *)
diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml
index 18cef702c..dc451fe05 100644
--- a/toplevel/vernacentries.ml
+++ b/toplevel/vernacentries.ml
@@ -153,9 +153,9 @@ let show_intro all =
msg_notice (hov 0 (prlist_with_sep spc pr_id lid))
else
try
- let n = list_last l in
+ let n = List.last l in
msg_notice (pr_id (List.hd (Tactics.find_intro_names [n] gl)))
- with Failure "list_last" -> ()
+ with Failure "List.last" -> ()
(** Prepare a "match" template for a given inductive type.
For each branch of the match, we list the constructor name
@@ -174,7 +174,7 @@ let make_cases s =
Util.array_fold_right2
(fun consname typ l ->
let al = List.rev (fst (Term.decompose_prod typ)) in
- let al = Util.list_skipn np al in
+ let al = Util.List.skipn np al in
let rec rename avoid = function
| [] -> []
| (n,_)::l ->
@@ -217,8 +217,8 @@ let print_modules () =
and loaded = Library.loaded_libraries () in
(* we intersect over opened to preserve the order of opened since *)
(* non-commutative operations (e.g. visibility) are done at import time *)
- let loaded_opened = list_intersect opened loaded
- and only_loaded = list_subtract loaded opened in
+ let loaded_opened = List.intersect opened loaded
+ and only_loaded = List.subtract loaded opened in
str"Loaded and imported library files: " ++
pr_vertical_list pr_dirpath loaded_opened ++ fnl () ++
str"Loaded and not imported library files: " ++
@@ -289,7 +289,7 @@ let print_namespace ns =
| MPfile dir -> Names.repr_dirpath dir
| MPdot (mp,lbl) -> (Names.id_of_label lbl)::(list_of_modulepath mp)
in
- snd (Util.list_chop n (List.rev (list_of_modulepath mp)))
+ snd (Util.List.chop n (List.rev (list_of_modulepath mp)))
in
let print_list pr l = prlist_with_sep (fun () -> str".") pr l in
let print_kn kn =
@@ -817,7 +817,7 @@ let vernac_set_end_tac tac =
let vernac_set_used_variables l =
let l = List.map snd l in
- if not (list_distinct l) then error "Used variables list contains duplicates";
+ if not (List.distinct l) then error "Used variables list contains duplicates";
let vars = Environ.named_context (Global.env ()) in
List.iter (fun id ->
if not (List.exists (fun (id',_,_) -> id = id') vars) then
@@ -905,7 +905,7 @@ let vernac_declare_arguments local r l nargs flags =
let scopes = List.map (List.map (fun (_,_, s, _,_) -> s)) l in
if List.exists ((<>) names) rest then
error "All arguments lists must declare the same names.";
- if not (Util.list_distinct (List.filter ((<>) Anonymous) names)) then
+ if not (Util.List.distinct (List.filter ((<>) Anonymous) names)) then
error "Arguments names must be distinct.";
let sr = smart_global r in
let inf_names =
@@ -929,7 +929,7 @@ let vernac_declare_arguments local r l nargs flags =
if List.exists (List.exists ((<>) None)) rest then
error "Notation scopes can be given only once";
if not extra_scope_flag then l, scopes else
- let l, _ = List.split (List.map (list_chop (List.length inf_names)) l) in
+ let l, _ = List.split (List.map (List.chop (List.length inf_names)) l) in
l, scopes in
(* we interpret _ as the inferred names *)
let l = if l = [[]] then l else
@@ -943,8 +943,8 @@ let vernac_declare_arguments local r l nargs flags =
with Not_found -> false in
let some_renaming_specified, implicits =
if l = [[]] then false, [[]] else
- Util.list_fold_map (fun sr il ->
- let sr', impl = Util.list_fold_map (fun b -> function
+ Util.List.fold_map (fun sr il ->
+ let sr', impl = Util.List.fold_map (fun b -> function
| (Anonymous, _,_, true, max), Name id -> assert false
| (Name x, _,_, true, _), Anonymous ->
error ("Argument "^string_of_id x^" cannot be declared implicit.")
@@ -953,7 +953,7 @@ let vernac_declare_arguments local r l nargs flags =
| (Name iid, _,_, _, _), Name id -> b || iid <> id, None
| _ -> b, None)
sr (List.combine il inf_names) in
- sr || sr', Util.list_map_filter (fun x -> x) impl)
+ sr || sr', Util.List.map_filter (fun x -> x) impl)
some_renaming_specified l in
if some_renaming_specified then
if not (List.mem `Rename flags) then
@@ -969,8 +969,8 @@ let vernac_declare_arguments local r l nargs flags =
with _ -> Some (Notation.find_delimiters_scope o k)) scopes in
let some_scopes_specified = List.exists ((<>) None) scopes in
let rargs =
- Util.list_map_filter (function (n, true) -> Some n | _ -> None)
- (Util.list_map_i (fun i (_, b, _,_,_) -> i, b) 0 l) in
+ Util.List.map_filter (function (n, true) -> Some n | _ -> None)
+ (Util.List.map_i (fun i (_, b, _,_,_) -> i, b) 0 l) in
if some_scopes_specified || List.mem `ClearScopes flags then
vernac_arguments_scope local r scopes;
if not some_implicits_specified && List.mem `DefaultImplicits flags then