1 (* Title: Pure/library.ML
3 Author: Lawrence C Paulson, Cambridge University Computer Laboratory
4 Copyright 1992 University of Cambridge
6 Basic library: functions, options, pairs, booleans, lists, integers,
7 strings, lists as sets, association lists, generic tables, balanced
8 trees, orders, I/O and diagnostics, timing, misc.
11 infix |> ~~ \ \\ ins ins_string ins_int orf andf prefix upto downto
12 mem mem_int mem_string union union_int union_string inter inter_int
13 inter_string subset subset_int subset_string;
20 val curry: ('a * 'b -> 'c) -> 'a -> 'b -> 'c
21 val uncurry: ('a -> 'b -> 'c) -> 'a * 'b -> 'c
24 val |> : 'a * ('a -> 'b) -> 'b
25 val apl: 'a * ('a * 'b -> 'c) -> 'b -> 'c
26 val apr: ('a * 'b -> 'c) * 'b -> 'a -> 'c
27 val funpow: int -> ('a -> 'a) -> 'a -> 'a
28 val oo: ('a -> 'b) * ('c -> 'd -> 'a) -> 'c -> 'd -> 'b
29 val ooo: ('a -> 'b) * ('c -> 'd -> 'e -> 'a) -> 'c -> 'd -> 'e -> 'b
33 val stamp: unit -> stamp
36 datatype 'a option = None | Some of 'a
38 val the: 'a option -> 'a
39 val if_none: 'a option -> 'a -> 'a
40 val is_some: 'a option -> bool
41 val is_none: 'a option -> bool
42 val apsome: ('a -> 'b) -> 'a option -> 'b option
43 val can: ('a -> 'b) -> 'a -> bool
44 val try: ('a -> 'b) -> 'a -> 'b option
47 val pair: 'a -> 'b -> 'a * 'b
48 val rpair: 'a -> 'b -> 'b * 'a
49 val fst: 'a * 'b -> 'a
50 val snd: 'a * 'b -> 'b
51 val eq_fst: (''a * 'b) * (''a * 'c) -> bool
52 val eq_snd: ('a * ''b) * ('c * ''b) -> bool
53 val swap: 'a * 'b -> 'b * 'a
54 val apfst: ('a -> 'b) -> 'a * 'c -> 'b * 'c
55 val apsnd: ('a -> 'b) -> 'c * 'a -> 'c * 'b
56 val pairself: ('a -> 'b) -> 'a * 'a -> 'b * 'b
59 val equal: ''a -> ''a -> bool
60 val not_equal: ''a -> ''a -> bool
61 val orf: ('a -> bool) * ('a -> bool) -> 'a -> bool
62 val andf: ('a -> bool) * ('a -> bool) -> 'a -> bool
63 val exists: ('a -> bool) -> 'a list -> bool
64 val forall: ('a -> bool) -> 'a list -> bool
65 val set: bool ref -> bool
66 val reset: bool ref -> bool
67 val toggle: bool ref -> bool
68 val setmp: 'a ref -> 'a -> ('b -> 'c) -> 'b -> 'c
71 exception LIST of string
72 val null: 'a list -> bool
74 val tl: 'a list -> 'a list
75 val cons: 'a -> 'a list -> 'a list
76 val single: 'a -> 'a list
77 val append: 'a list -> 'a list -> 'a list
78 val apply: ('a -> 'a) list -> 'a -> 'a
79 val foldl: ('a * 'b -> 'a) -> 'a * 'b list -> 'a
80 val foldr: ('a * 'b -> 'b) -> 'a list * 'b -> 'b
81 val foldr1: ('a * 'a -> 'a) -> 'a list -> 'a
82 val foldl_map: ('a * 'b -> 'a * 'c) -> 'a * 'b list -> 'a * 'c list
83 val length: 'a list -> int
84 val take: int * 'a list -> 'a list
85 val drop: int * 'a list -> 'a list
86 val dropwhile: ('a -> bool) -> 'a list -> 'a list
87 val nth_elem: int * 'a list -> 'a
88 val last_elem: 'a list -> 'a
89 val split_last: 'a list -> 'a list * 'a
90 val nth_update: 'a -> int * 'a list -> 'a list
91 val find_index: ('a -> bool) -> 'a list -> int
92 val find_index_eq: ''a -> ''a list -> int
93 val find_first: ('a -> bool) -> 'a list -> 'a option
94 val get_first: ('a -> 'b option) -> 'a list -> 'b option
95 val flat: 'a list list -> 'a list
96 val seq: ('a -> unit) -> 'a list -> unit
97 val separate: 'a -> 'a list -> 'a list
98 val replicate: int -> 'a -> 'a list
99 val multiply: 'a list * 'a list list -> 'a list list
100 val filter: ('a -> bool) -> 'a list -> 'a list
101 val filter_out: ('a -> bool) -> 'a list -> 'a list
102 val mapfilter: ('a -> 'b option) -> 'a list -> 'b list
103 val map2: ('a * 'b -> 'c) -> 'a list * 'b list -> 'c list
104 val exists2: ('a * 'b -> bool) -> 'a list * 'b list -> bool
105 val forall2: ('a * 'b -> bool) -> 'a list * 'b list -> bool
106 val seq2: ('a * 'b -> unit) -> 'a list * 'b list -> unit
107 val ~~ : 'a list * 'b list -> ('a * 'b) list
108 val split_list: ('a * 'b) list -> 'a list * 'b list
109 val prefix: ''a list * ''a list -> bool
110 val take_prefix: ('a -> bool) -> 'a list -> 'a list * 'a list
111 val take_suffix: ('a -> bool) -> 'a list -> 'a list * 'a list
114 val inc: int ref -> int
115 val dec: int ref -> int
116 val upto: int * int -> int list
117 val downto: int * int -> int list
118 val downto0: int list * int -> bool
119 val radixpand: int * int -> int list
120 val radixstring: int * string * int -> string
121 val string_of_int: int -> string
122 val string_of_indexname: string * int -> string
125 val enclose: string -> string -> string -> string
126 val quote: string -> string
127 val space_implode: string -> string list -> string
128 val commas: string list -> string
129 val commas_quote: string list -> string
130 val cat_lines: string list -> string
131 val space_explode: string -> string -> string list
132 val split_lines: string -> string list
133 val suffix: string -> string -> string
134 val unsuffix: string -> string -> string
137 val mem: ''a * ''a list -> bool
138 val mem_int: int * int list -> bool
139 val mem_string: string * string list -> bool
140 val gen_mem: ('a * 'b -> bool) -> 'a * 'b list -> bool
141 val ins: ''a * ''a list -> ''a list
142 val ins_int: int * int list -> int list
143 val ins_string: string * string list -> string list
144 val gen_ins: ('a * 'a -> bool) -> 'a * 'a list -> 'a list
145 val union: ''a list * ''a list -> ''a list
146 val union_int: int list * int list -> int list
147 val union_string: string list * string list -> string list
148 val gen_union: ('a * 'a -> bool) -> 'a list * 'a list -> 'a list
149 val inter: ''a list * ''a list -> ''a list
150 val inter_int: int list * int list -> int list
151 val inter_string: string list * string list -> string list
152 val subset: ''a list * ''a list -> bool
153 val subset_int: int list * int list -> bool
154 val subset_string: string list * string list -> bool
155 val eq_set: ''a list * ''a list -> bool
156 val eq_set_string: string list * string list -> bool
157 val gen_subset: ('a * 'b -> bool) -> 'a list * 'b list -> bool
158 val \ : ''a list * ''a -> ''a list
159 val \\ : ''a list * ''a list -> ''a list
160 val gen_rem: ('a * 'b -> bool) -> 'a list * 'b -> 'a list
161 val gen_rems: ('a * 'b -> bool) -> 'a list * 'b list -> 'a list
162 val gen_distinct: ('a * 'a -> bool) -> 'a list -> 'a list
163 val distinct: ''a list -> ''a list
164 val findrep: ''a list -> ''a list
165 val gen_duplicates: ('a * 'a -> bool) -> 'a list -> 'a list
166 val duplicates: ''a list -> ''a list
168 (*association lists*)
169 val assoc: (''a * 'b) list * ''a -> 'b option
170 val assoc_int: (int * 'a) list * int -> 'a option
171 val assoc_string: (string * 'a) list * string -> 'a option
172 val assoc_string_int: ((string * int) * 'a) list * (string * int) -> 'a option
173 val assocs: (''a * 'b list) list -> ''a -> 'b list
174 val assoc2: (''a * (''b * 'c) list) list * (''a * ''b) -> 'c option
175 val gen_assoc: ('a * 'b -> bool) -> ('b * 'c) list * 'a -> 'c option
176 val overwrite: (''a * 'b) list * (''a * 'b) -> (''a * 'b) list
177 val gen_overwrite: ('a * 'a -> bool) -> ('a * 'b) list * ('a * 'b) -> ('a * 'b) list
180 val generic_extend: ('a * 'a -> bool)
181 -> ('b -> 'a list) -> ('a list -> 'b) -> 'b -> 'a list -> 'b
182 val generic_merge: ('a * 'a -> bool) -> ('b -> 'a list) -> ('a list -> 'b) -> 'b -> 'b -> 'b
183 val extend_list: ''a list -> ''a list -> ''a list
184 val merge_lists: ''a list -> ''a list -> ''a list
185 val merge_alists: (''a * 'b) list -> (''a * 'b) list -> (''a * 'b) list
186 val merge_rev_lists: ''a list -> ''a list -> ''a list
190 val fold_bal: ('a * 'a -> 'a) -> 'a list -> 'a
191 val access_bal: ('a -> 'a) * ('a -> 'a) * 'a -> int -> int -> 'a
192 val accesses_bal: ('a -> 'a) * ('a -> 'a) * 'a -> int -> 'a list
195 datatype order = EQUAL | GREATER | LESS
196 val rev_order: order -> order
197 val make_ord: ('a * 'a -> bool) -> 'a * 'a -> order
198 val int_ord: int * int -> order
199 val string_ord: string * string -> order
200 val prod_ord: ('a * 'b -> order) -> ('c * 'd -> order) -> ('a * 'c) * ('b * 'd) -> order
201 val dict_ord: ('a * 'b -> order) -> 'a list * 'b list -> order
202 val list_ord: ('a * 'b -> order) -> 'a list * 'b list -> order
203 val sort: ('a * 'a -> order) -> 'a list -> 'a list
204 val sort_strings: string list -> string list
205 val sort_wrt: ('a -> string) -> 'a list -> 'a list
207 (*I/O and diagnostics*)
208 val cd: string -> unit
209 val pwd: unit -> string
210 val prs_fn: (string -> unit) ref
211 val warning_fn: (string -> unit) ref
212 val error_fn: (string -> unit) ref
213 val prs: string -> unit
214 val writeln: string -> unit
215 val warning: string -> unit
217 val error_msg: string -> unit
218 val error: string -> 'a
219 val sys_error: string -> 'a
220 val assert: bool -> string -> unit
221 val deny: bool -> string -> unit
222 val assert_all: ('a -> bool) -> 'a list -> ('a -> string) -> unit
223 datatype 'a error = Error of string | OK of 'a
224 val get_error: 'a error -> string option
225 val get_ok: 'a error -> 'a option
226 val handle_error: ('a -> 'b) -> 'a -> 'b error
227 exception ERROR_MESSAGE of string
228 val transform_error: ('a -> 'b) -> 'a -> 'b
229 val transform_failure: (exn -> exn) -> ('a -> 'b) -> 'a -> 'b
232 val cond_timeit: bool -> (unit -> 'a) -> 'a
233 val timeit: (unit -> 'a) -> 'a
234 val timeap: ('a -> 'b) -> 'a -> 'b
237 val make_keylist: ('a -> 'b) -> 'a list -> ('a * 'b) list
238 val keyfilter: ('a * ''b) list -> ''b -> 'a list
239 val partition: ('a -> bool) -> 'a list -> 'a list * 'a list
240 val partition_eq: ('a * 'a -> bool) -> 'a list -> 'a list list
241 val partition_list: (int -> 'a -> bool) -> int -> int -> 'a list -> 'a list list
242 val transitive_closure: (string * string list) list -> (string * string list) list
243 val init_gensym: unit -> unit
244 val gensym: string -> string
245 val bump_int_list: string list -> string list
246 val bump_list: string list * string -> string list
247 val bump_string: string -> string
248 val scanwords: (string -> bool) -> string list -> string list
249 datatype 'a mtree = Join of 'a * 'a mtree list
252 structure Library: LIBRARY =
258 (*handy combinators*)
259 fun curry f x y = f (x, y);
260 fun uncurry f (x, y) = f x y;
267 (*application of (infix) operator to its left or right argument*)
268 fun apl (x, f) y = f (x, y);
269 fun apr (f, y) x = f (x, y);
271 (*function exponentiation: f(...(f x)...) with n applications of f*)
273 let fun rep (0, x) = x
274 | rep (n, x) = rep (n - 1, f x)
277 (*concatenation: 2 and 3 args*)
278 fun (f oo g) x y = f (g x y);
279 fun (f ooo g) x y z = f (g x y z);
284 type stamp = unit ref;
285 val stamp: unit -> stamp = ref;
291 datatype 'a option = None | Some of 'a;
296 | the None = raise OPTION;
299 fun if_none None y = y
300 | if_none (Some x) _ = x;
302 fun is_some (Some _) = true
303 | is_some None = false;
305 fun is_none (Some _) = false
306 | is_none None = true;
308 fun apsome f (Some x) = Some (f x)
309 | apsome _ None = None;
311 (*handle partial functions*)
312 fun can f x = (f x; true) handle _ => false;
313 fun try f x = Some (f x) handle _ => None;
319 fun pair x y = (x, y);
320 fun rpair x y = (y, x);
325 fun eq_fst ((x1, _), (x2, _)) = x1 = x2;
326 fun eq_snd ((_, y1), (_, y2)) = y1 = y2;
328 fun swap (x, y) = (y, x);
330 (*apply function to components*)
331 fun apfst f (x, y) = (f x, y);
332 fun apsnd f (x, y) = (x, f y);
333 fun pairself f (x, y) = (f x, f y);
341 fun equal x y = x = y;
342 fun not_equal x y = x <> y;
345 (* operators for combining predicates *)
347 fun (p orf q) = fn x => p x orelse q x;
348 fun (p andf q) = fn x => p x andalso q x;
351 (* predicates on lists *)
353 (*exists pred [x1, ..., xn] ===> pred x1 orelse ... orelse pred xn*)
354 fun exists (pred: 'a -> bool) : 'a list -> bool =
355 let fun boolf [] = false
356 | boolf (x :: xs) = pred x orelse boolf xs
359 (*forall pred [x1, ..., xn] ===> pred x1 andalso ... andalso pred xn*)
360 fun forall (pred: 'a -> bool) : 'a list -> bool =
361 let fun boolf [] = true
362 | boolf (x :: xs) = pred x andalso boolf xs
368 fun set flag = (flag := true; true);
369 fun reset flag = (flag := false; false);
370 fun toggle flag = (flag := not (! flag); ! flag);
372 (*temporarily set flag, handling errors*)
373 fun setmp flag value f x =
375 val orig_value = ! flag;
376 fun return y = (flag := orig_value; y);
379 return (f x handle exn => (return (); raise exn))
386 exception LIST of string;
389 | null (_ :: _) = false;
391 fun hd [] = raise LIST "hd"
394 fun tl [] = raise LIST "tl"
397 fun cons x xs = x :: xs;
400 fun append xs ys = xs @ ys;
403 | apply (f :: fs) x = apply fs (f x);
408 (*the following versions of fold are designed to fit nicely with infixes*)
410 (* (op @) (e, [x1, ..., xn]) ===> ((e @ x1) @ x2) ... @ xn
411 for operators that associate to the left (TAIL RECURSIVE)*)
412 fun foldl (f: 'a * 'b -> 'a) : 'a * 'b list -> 'a =
413 let fun itl (e, []) = e
414 | itl (e, a::l) = itl (f(e, a), l)
417 (* (op @) ([x1, ..., xn], e) ===> x1 @ (x2 ... @ (xn @ e))
418 for operators that associate to the right (not tail recursive)*)
421 | itr (a::l) = f(a, itr l)
424 (* (op @) [x1, ..., xn] ===> x1 @ (x2 ... @ (x[n-1] @ xn))
425 for n > 0, operators that associate to the right (not tail recursive)*)
428 | itr (x::l) = f(x, itr l)
431 fun foldl_map _ (x, []) = (x, [])
432 | foldl_map f (x, y :: ys) =
434 val (x', y') = f (x, y);
435 val (x'', ys') = foldl_map f (x', ys);
436 in (x'', y' :: ys') end;
439 (* basic list functions *)
441 (*length of a list, should unquestionably be a standard function*)
442 local fun length1 (n, []) = n (*TAIL RECURSIVE*)
443 | length1 (n, x :: xs) = length1 (n + 1, xs)
444 in fun length l = length1 (0, l) end;
446 (*take the first n elements from a list*)
447 fun take (n, []) = []
448 | take (n, x :: xs) =
449 if n > 0 then x :: take (n - 1, xs) else [];
451 (*drop the first n elements from a list*)
452 fun drop (n, []) = []
453 | drop (n, x :: xs) =
454 if n > 0 then drop (n - 1, xs) else x :: xs;
456 fun dropwhile P [] = []
457 | dropwhile P (ys as x::xs) = if P x then dropwhile P xs else ys;
459 (*return nth element of a list, where 0 designates the first element;
460 raise EXCEPTION if list too short*)
463 [] => raise LIST "nth_elem"
466 (*last element of a list*)
467 fun last_elem [] = raise LIST "last_elem"
469 | last_elem (_ :: xs) = last_elem xs;
471 (*rear decomposition*)
472 fun split_last [] = raise LIST "split_last"
473 | split_last [x] = ([], x)
474 | split_last (x :: xs) = apfst (cons x) (split_last xs);
476 (*update nth element*)
477 fun nth_update x (n, xs) =
479 val prfx = take (n, xs);
480 val sffx = drop (n, xs);
483 [] => raise LIST "nth_update"
484 | _ :: sffx' => prfx @ (x :: sffx'))
487 (*find the position of an element in a list*)
488 fun find_index pred =
489 let fun find _ [] = ~1
490 | find n (x :: xs) = if pred x then n else find (n + 1) xs;
493 fun find_index_eq x = find_index (equal x);
495 (*find first element satisfying predicate*)
496 fun find_first _ [] = None
497 | find_first pred (x :: xs) =
498 if pred x then Some x else find_first pred xs;
500 (*get first element by lookup function*)
501 fun get_first _ [] = None
502 | get_first f (x :: xs) =
504 None => get_first f xs
507 (*flatten a list of lists to a list*)
508 fun flat (ls: 'c list list) : 'c list = foldr (op @) (ls, []);
510 (*like Lisp's MAPC -- seq proc [x1, ..., xn] evaluates
511 (proc x1; ...; proc xn) for side effects*)
512 fun seq (proc: 'a -> unit) : 'a list -> unit =
514 | seqf (x :: xs) = (proc x; seqf xs)
517 (*separate s [x1, x2, ..., xn] ===> [x1, s, x2, s, ..., s, xn]*)
518 fun separate s (x :: (xs as _ :: _)) = x :: s :: separate s xs
519 | separate _ xs = xs;
521 (*make the list [x, x, ..., x] of length n*)
522 fun replicate n (x: 'a) : 'a list =
523 let fun rep (0, xs) = xs
524 | rep (n, xs) = rep (n - 1, x :: xs)
526 if n < 0 then raise LIST "replicate"
530 (*multiply [a, b, c, ...] * [xs, ys, zs, ...]*)
531 fun multiply ([], _) = []
532 | multiply (x :: xs, yss) = map (cons x) yss @ multiply (xs, yss);
537 (*copy the list preserving elements that satisfy the predicate*)
538 fun filter (pred: 'a->bool) : 'a list -> 'a list =
540 | filt (x :: xs) = if pred x then x :: filt xs else filt xs
543 fun filter_out f = filter (not o f);
545 fun mapfilter (f: 'a -> 'b option) ([]: 'a list) = [] : 'b list
546 | mapfilter f (x :: xs) =
548 None => mapfilter f xs
549 | Some y => y :: mapfilter f xs);
554 fun map2 _ ([], []) = []
555 | map2 f (x :: xs, y :: ys) = (f (x, y) :: map2 f (xs, ys))
556 | map2 _ _ = raise LIST "map2";
558 fun exists2 _ ([], []) = false
559 | exists2 pred (x :: xs, y :: ys) = pred (x, y) orelse exists2 pred (xs, ys)
560 | exists2 _ _ = raise LIST "exists2";
562 fun forall2 _ ([], []) = true
563 | forall2 pred (x :: xs, y :: ys) = pred (x, y) andalso forall2 pred (xs, ys)
564 | forall2 _ _ = raise LIST "forall2";
566 fun seq2 _ ([], []) = ()
567 | seq2 f (x :: xs, y :: ys) = (f (x, y); seq2 f (xs, ys))
568 | seq2 _ _ = raise LIST "seq2";
570 (*combine two lists forming a list of pairs:
571 [x1, ..., xn] ~~ [y1, ..., yn] ===> [(x1, y1), ..., (xn, yn)]*)
573 | (x :: xs) ~~ (y :: ys) = (x, y) :: (xs ~~ ys)
574 | _ ~~ _ = raise LIST "~~";
576 (*inverse of ~~; the old 'split':
577 [(x1, y1), ..., (xn, yn)] ===> ([x1, ..., xn], [y1, ..., yn])*)
578 fun split_list (l: ('a * 'b) list) = (map #1 l, map #2 l);
581 (* prefixes, suffixes *)
583 fun [] prefix _ = true
584 | (x :: xs) prefix (y :: ys) = x = y andalso (xs prefix ys)
585 | _ prefix _ = false;
587 (* [x1, ..., xi, ..., xn] ---> ([x1, ..., x(i-1)], [xi, ..., xn])
588 where xi is the first element that does not satisfy the predicate*)
589 fun take_prefix (pred : 'a -> bool) (xs: 'a list) : 'a list * 'a list =
590 let fun take (rxs, []) = (rev rxs, [])
591 | take (rxs, x :: xs) =
592 if pred x then take(x :: rxs, xs) else (rev rxs, x :: xs)
595 (* [x1, ..., xi, ..., xn] ---> ([x1, ..., xi], [x(i+1), ..., xn])
596 where xi is the last element that does not satisfy the predicate*)
597 fun take_suffix _ [] = ([], [])
598 | take_suffix pred (x :: xs) =
599 (case take_suffix pred xs of
600 ([], sffx) => if pred x then ([], x :: sffx) else ([x], sffx)
601 | (prfx, sffx) => (x :: prfx, sffx));
607 fun inc i = (i := ! i + 1; ! i);
608 fun dec i = (i := ! i - 1; ! i);
611 (* lists of integers *)
613 (*make the list [from, from + 1, ..., to]*)
615 if from > to then [] else from :: ((from + 1) upto to);
617 (*make the list [from, from - 1, ..., to]*)
618 fun (from downto to) =
619 if from < to then [] else from :: ((from - 1) downto to);
621 (*predicate: downto0 (is, n) <=> is = [n, n - 1, ..., 0]*)
622 fun downto0 (i :: is, n) = i = n andalso downto0 (is, n - 1)
623 | downto0 ([], n) = n = ~1;
626 (* convert integers to strings *)
628 (*expand the number in the given base;
629 example: radixpand (2, 8) gives [1, 0, 0, 0]*)
630 fun radixpand (base, num) : int list =
632 fun radix (n, tail) =
633 if n < base then n :: tail
634 else radix (n div base, (n mod base) :: tail)
635 in radix (num, []) end;
637 (*expands a number into a string of characters starting from "zerochar";
638 example: radixstring (2, "0", 8) gives "1000"*)
639 fun radixstring (base, zerochar, num) =
640 let val offset = ord zerochar;
641 fun chrof n = chr (offset + n)
642 in implode (map chrof (radixpand (base, num))) end;
645 val string_of_int = Int.toString;
647 fun string_of_indexname (a,0) = a
648 | string_of_indexname (a,i) = a ^ "_" ^ Int.toString i;
654 (*enclose in brackets*)
655 fun enclose lpar rpar str = lpar ^ str ^ rpar;
657 (*simple quoting (does not escape special chars)*)
658 val quote = enclose "\"" "\"";
660 (*space_implode "..." (explode "hello") = "h...e...l...l...o"*)
661 fun space_implode a bs = implode (separate a bs);
663 val commas = space_implode ", ";
664 val commas_quote = commas o map quote;
666 (*concatenate messages, one per line, into a string*)
667 val cat_lines = space_implode "\n";
669 (*space_explode "." "h.e..l.lo" = ["h", "e", "", "l", "lo"]*)
670 fun space_explode _ "" = []
671 | space_explode sep str =
674 (case take_prefix (not_equal sep) chs of
675 (cs, []) => [implode cs]
676 | (cs, _ :: cs') => implode cs :: expl cs');
677 in expl (explode str) end;
679 val split_lines = space_explode "\n";
682 fun suffix sfx s = s ^ sfx;
688 val prfx_len = size s - size sfx;
690 if prfx_len >= 0 andalso implode (drop (prfx_len, cs)) = sfx then
691 implode (take (prfx_len, cs))
692 else raise LIST "unsuffix"
697 (** lists as sets **)
699 (*membership in a list*)
701 | x mem (y :: ys) = x = y orelse x mem ys;
703 (*membership in a list, optimized version for ints*)
704 fun (x:int) mem_int [] = false
705 | x mem_int (y :: ys) = x = y orelse x mem_int ys;
707 (*membership in a list, optimized version for strings*)
708 fun (x:string) mem_string [] = false
709 | x mem_string (y :: ys) = x = y orelse x mem_string ys;
711 (*generalized membership test*)
712 fun gen_mem eq (x, []) = false
713 | gen_mem eq (x, y :: ys) = eq (x, y) orelse gen_mem eq (x, ys);
716 (*insertion into list if not already there*)
717 fun (x ins xs) = if x mem xs then xs else x :: xs;
719 (*insertion into list, optimized version for ints*)
720 fun (x ins_int xs) = if x mem_int xs then xs else x :: xs;
722 (*insertion into list, optimized version for strings*)
723 fun (x ins_string xs) = if x mem_string xs then xs else x :: xs;
725 (*generalized insertion*)
726 fun gen_ins eq (x, xs) = if gen_mem eq (x, xs) then xs else x :: xs;
729 (*union of sets represented as lists: no repetitions*)
732 | (x :: xs) union ys = xs union (x ins ys);
734 (*union of sets, optimized version for ints*)
735 fun (xs:int list) union_int [] = xs
736 | [] union_int ys = ys
737 | (x :: xs) union_int ys = xs union_int (x ins_int ys);
739 (*union of sets, optimized version for strings*)
740 fun (xs:string list) union_string [] = xs
741 | [] union_string ys = ys
742 | (x :: xs) union_string ys = xs union_string (x ins_string ys);
744 (*generalized union*)
745 fun gen_union eq (xs, []) = xs
746 | gen_union eq ([], ys) = ys
747 | gen_union eq (x :: xs, ys) = gen_union eq (xs, gen_ins eq (x, ys));
752 | (x :: xs) inter ys =
753 if x mem ys then x :: (xs inter ys) else xs inter ys;
755 (*intersection, optimized version for ints*)
756 fun ([]:int list) inter_int ys = []
757 | (x :: xs) inter_int ys =
758 if x mem_int ys then x :: (xs inter_int ys) else xs inter_int ys;
760 (*intersection, optimized version for strings *)
761 fun ([]:string list) inter_string ys = []
762 | (x :: xs) inter_string ys =
763 if x mem_string ys then x :: (xs inter_string ys) else xs inter_string ys;
767 fun [] subset ys = true
768 | (x :: xs) subset ys = x mem ys andalso xs subset ys;
770 (*subset, optimized version for ints*)
771 fun ([]:int list) subset_int ys = true
772 | (x :: xs) subset_int ys = x mem_int ys andalso xs subset_int ys;
774 (*subset, optimized version for strings*)
775 fun ([]:string list) subset_string ys = true
776 | (x :: xs) subset_string ys = x mem_string ys andalso xs subset_string ys;
779 fun eq_set (xs, ys) =
780 xs = ys orelse (xs subset ys andalso ys subset xs);
782 (*set equality for strings*)
783 fun eq_set_string ((xs:string list), ys) =
784 xs = ys orelse (xs subset_string ys andalso ys subset_string xs);
786 fun gen_subset eq (xs, ys) = forall (fn x => gen_mem eq (x, ys)) xs;
789 (*removing an element from a list WITHOUT duplicates*)
790 fun (y :: ys) \ x = if x = y then ys else y :: (ys \ x)
793 fun ys \\ xs = foldl (op \) (ys,xs);
795 (*removing an element from a list -- possibly WITH duplicates*)
796 fun gen_rem eq (xs, y) = filter_out (fn x => eq (x, y)) xs;
798 fun gen_rems eq = foldl (gen_rem eq);
801 (*makes a list of the distinct members of the input; preserves order, takes
802 first of equal elements*)
803 fun gen_distinct eq lst =
805 val memb = gen_mem eq;
807 fun dist (rev_seen, []) = rev rev_seen
808 | dist (rev_seen, x :: xs) =
809 if memb (x, rev_seen) then dist (rev_seen, xs)
810 else dist (x :: rev_seen, xs);
815 fun distinct l = gen_distinct (op =) l;
818 (*returns the tail beginning with the first repeated element, or []*)
820 | findrep (x :: xs) = if x mem xs then x :: xs else findrep xs;
823 (*returns a list containing all repeated elements exactly once; preserves
824 order, takes first of equal elements*)
825 fun gen_duplicates eq lst =
827 val memb = gen_mem eq;
829 fun dups (rev_dups, []) = rev rev_dups
830 | dups (rev_dups, x :: xs) =
831 if memb (x, rev_dups) orelse not (memb (x, xs)) then
833 else dups (x :: rev_dups, xs);
838 fun duplicates l = gen_duplicates (op =) l;
842 (** association lists **)
844 (*association list lookup*)
845 fun assoc ([], key) = None
846 | assoc ((keyi, xi) :: pairs, key) =
847 if key = keyi then Some xi else assoc (pairs, key);
849 (*association list lookup, optimized version for ints*)
850 fun assoc_int ([], (key:int)) = None
851 | assoc_int ((keyi, xi) :: pairs, key) =
852 if key = keyi then Some xi else assoc_int (pairs, key);
854 (*association list lookup, optimized version for strings*)
855 fun assoc_string ([], (key:string)) = None
856 | assoc_string ((keyi, xi) :: pairs, key) =
857 if key = keyi then Some xi else assoc_string (pairs, key);
859 (*association list lookup, optimized version for string*ints*)
860 fun assoc_string_int ([], (key:string*int)) = None
861 | assoc_string_int ((keyi, xi) :: pairs, key) =
862 if key = keyi then Some xi else assoc_string_int (pairs, key);
865 (case assoc (ps, x) of
869 (*two-fold association list lookup*)
870 fun assoc2 (aal, (key1, key2)) =
871 (case assoc (aal, key1) of
872 Some al => assoc (al, key2)
875 (*generalized association list lookup*)
876 fun gen_assoc eq ([], key) = None
877 | gen_assoc eq ((keyi, xi) :: pairs, key) =
878 if eq (key, keyi) then Some xi else gen_assoc eq (pairs, key);
880 (*association list update*)
881 fun overwrite (al, p as (key, _)) =
882 let fun over ((q as (keyi, _)) :: pairs) =
883 if keyi = key then p :: pairs else q :: (over pairs)
887 fun gen_overwrite eq (al, p as (key, _)) =
888 let fun over ((q as (keyi, _)) :: pairs) =
889 if eq (keyi, key) then p :: pairs else q :: (over pairs)
895 (** generic tables **)
897 (*Tables are supposed to be 'efficient' encodings of lists of elements distinct
898 wrt. an equality "eq". The extend and merge operations below are optimized
899 for long-term space efficiency.*)
901 (*append (new) elements to a table*)
902 fun generic_extend _ _ _ tab [] = tab
903 | generic_extend eq dest_tab mk_tab tab1 lst2 =
905 val lst1 = dest_tab tab1;
906 val new_lst2 = gen_rems eq (lst2, lst1);
908 if null new_lst2 then tab1
909 else mk_tab (lst1 @ new_lst2)
912 (*append (new) elements of 2nd table to 1st table*)
913 fun generic_merge eq dest_tab mk_tab tab1 tab2 =
915 val lst1 = dest_tab tab1;
916 val lst2 = dest_tab tab2;
917 val new_lst2 = gen_rems eq (lst2, lst1);
919 if null new_lst2 then tab1
920 else if gen_subset eq (lst1, lst2) then tab2
921 else mk_tab (lst1 @ new_lst2)
926 fun extend_list tab = generic_extend (op =) I I tab;
927 fun merge_lists tab = generic_merge (op =) I I tab;
928 fun merge_alists tab = generic_merge eq_fst I I tab;
930 fun merge_rev_lists xs [] = xs
931 | merge_rev_lists [] ys = ys
932 | merge_rev_lists xs (y :: ys) =
933 (if y mem xs then I else cons y) (merge_rev_lists xs ys);
937 (** balanced trees **)
939 exception Balance; (*indicates non-positive argument to balancing fun*)
941 (*balanced folding; avoids deep nesting*)
942 fun fold_bal f [x] = x
943 | fold_bal f [] = raise Balance
945 let val k = length xs div 2
946 in f (fold_bal f (take(k, xs)),
947 fold_bal f (drop(k, xs)))
950 (*construct something of the form f(...g(...(x)...)) for balanced access*)
951 fun access_bal (f, g, x) n i =
952 let fun acc n i = (*1<=i<=n*)
955 in if i<=n2 then f (acc n2 i)
956 else g (acc (n-n2) (i-n2))
958 in if 1<=i andalso i<=n then acc n i else raise Balance end;
960 (*construct ALL such accesses; could try harder to share recursive calls!*)
961 fun accesses_bal (f, g, x) n =
966 in if n-n2=n2 then map f acc2 @ map g acc2
967 else map f acc2 @ map g (acc (n-n2)) end
968 in if 1<=n then acc n else raise Balance end;
974 datatype order = LESS | EQUAL | GREATER;
976 fun rev_order LESS = GREATER
977 | rev_order EQUAL = EQUAL
978 | rev_order GREATER = LESS;
980 (*assume rel is a linear strict order*)
981 fun make_ord rel (x, y) =
982 if rel (x, y) then LESS
983 else if rel (y, x) then GREATER
986 fun int_ord (i, j: int) =
988 else if i = j then EQUAL
991 fun string_ord (a, b: string) =
993 else if a = b then EQUAL
996 (*lexicographic product*)
997 fun prod_ord a_ord b_ord ((x, y), (x', y')) =
998 (case a_ord (x, x') of EQUAL => b_ord (y, y') | ord => ord);
1000 (*dictionary order -- in general NOT well-founded!*)
1001 fun dict_ord _ ([], []) = EQUAL
1002 | dict_ord _ ([], _ :: _) = LESS
1003 | dict_ord _ (_ :: _, []) = GREATER
1004 | dict_ord elem_ord (x :: xs, y :: ys) =
1005 (case elem_ord (x, y) of EQUAL => dict_ord elem_ord (xs, ys) | ord => ord);
1007 (*lexicographic product of lists*)
1008 fun list_ord elem_ord (xs, ys) =
1009 prod_ord int_ord (dict_ord elem_ord) ((length xs, xs), (length ys, ys));
1014 (*quicksort (stable, i.e. does not reorder equal elements)*)
1018 let val len = length xs in
1021 let val (lts, eqs, gts) = part (nth_elem (len div 2, xs)) xs in
1022 qsort lts @ eqs @ qsort gts
1025 and part _ [] = ([], [], [])
1026 | part pivot (x :: xs) = add (ord (x, pivot)) x (part pivot xs)
1027 and add LESS x (lts, eqs, gts) = (x :: lts, eqs, gts)
1028 | add EQUAL x (lts, eqs, gts) = (lts, x :: eqs, gts)
1029 | add GREATER x (lts, eqs, gts) = (lts, eqs, x :: gts);
1033 val sort_strings = sort string_ord;
1034 fun sort_wrt sel xs = sort (string_ord o pairself sel) xs;
1038 (** input / output and diagnostics **)
1040 val cd = OS.FileSys.chDir;
1041 val pwd = OS.FileSys.getDir;
1046 (TextIO.output (TextIO.stdOut, s); TextIO.flushOut TextIO.stdOut);
1048 fun prefix_lines prfx txt =
1049 txt |> split_lines |> map (fn s => prfx ^ s ^ "\n") |> implode;
1052 (*hooks for output channels: normal, warning, error*)
1053 val prs_fn = ref (fn s => out s);
1054 val warning_fn = ref (fn s => out (prefix_lines "### " s));
1055 val error_fn = ref (fn s => out (prefix_lines "*** " s));
1059 fun prs s = !prs_fn s;
1060 fun writeln s = prs (s ^ "\n");
1062 fun warning s = !warning_fn s;
1064 (*print error message and abort to top level*)
1066 fun error_msg s = !error_fn s;
1067 fun error s = (error_msg s; raise ERROR);
1068 fun sys_error msg = error ("## SYSTEM ERROR ##\n" ^ msg);
1070 fun assert p msg = if p then () else error msg;
1071 fun deny p msg = if p then error msg else ();
1073 (*Assert pred for every member of l, generating a message if pred fails*)
1074 fun assert_all pred l msg_fn =
1076 | asl (x::xs) = if pred x then asl xs else error (msg_fn x)
1080 (* handle errors capturing messages *)
1086 fun get_error (Error msg) = Some msg
1087 | get_error _ = None;
1089 fun get_ok (OK x) = Some x
1092 datatype 'a result =
1096 fun handle_error f x =
1098 val buffer = ref ([]: string list);
1099 fun capture s = buffer := ! buffer @ [s];
1100 fun err_msg () = if not (null (! buffer)) then error_msg (cat_lines (! buffer)) else ();
1102 (case Result (setmp error_fn capture f x) handle exn => Exn exn of
1103 Result y => (err_msg (); OK y)
1104 | Exn ERROR => Error (cat_lines (! buffer))
1105 | Exn exn => (err_msg (); raise exn))
1109 (* transform ERROR into ERROR_MESSAGE *)
1111 exception ERROR_MESSAGE of string;
1113 fun transform_error f x =
1114 (case handle_error f x of
1116 | Error msg => raise ERROR_MESSAGE msg);
1119 (* transform any exception, including ERROR *)
1121 fun transform_failure exn f x =
1122 transform_error f x handle e => raise exn e;
1128 (*a conditional timing function: applies f to () and, if the flag is true,
1129 prints its runtime*)
1130 fun cond_timeit flag f =
1132 let val start = startTiming()
1135 writeln (endTiming start); result
1139 (*unconditional timing function*)
1140 fun timeit x = cond_timeit true x;
1142 (*timed application function*)
1143 fun timeap f x = timeit (fn () => f x);
1149 (*use the keyfun to make a list of (x, key) pairs*)
1150 fun make_keylist (keyfun: 'a->'b) : 'a list -> ('a * 'b) list =
1151 let fun keypair x = (x, keyfun x)
1154 (*given a list of (x, key) pairs and a searchkey
1155 return the list of xs from each pair whose key equals searchkey*)
1156 fun keyfilter [] searchkey = []
1157 | keyfilter ((x, key) :: pairs) searchkey =
1158 if key = searchkey then x :: keyfilter pairs searchkey
1159 else keyfilter pairs searchkey;
1162 (*Partition list into elements that satisfy predicate and those that don't.
1163 Preserves order of elements in both lists.*)
1164 fun partition (pred: 'a->bool) (ys: 'a list) : ('a list * 'a list) =
1165 let fun part ([], answer) = answer
1166 | part (x::xs, (ys, ns)) = if pred(x)
1167 then part (xs, (x::ys, ns))
1168 else part (xs, (ys, x::ns))
1169 in part (rev ys, ([], [])) end;
1172 fun partition_eq (eq:'a * 'a -> bool) =
1173 let fun part [] = []
1174 | part (x::ys) = let val (xs, xs') = partition (apl(x, eq)) ys
1175 in (x::xs)::(part xs') end
1179 (*Partition a list into buckets [ bi, b(i+1), ..., bj ]
1180 putting x in bk if p(k)(x) holds. Preserve order of elements if possible.*)
1181 fun partition_list p i j =
1184 (case xs of [] => []
1185 | _ => raise LIST "partition_list")
1187 let val (ns, rest) = partition (p k) xs;
1188 in ns :: part(k+1)rest end
1192 (* transitive closure (not Warshall's algorithm) *)
1194 fun transitive_closure [] = []
1195 | transitive_closure ((x, ys)::ps) =
1196 let val qs = transitive_closure ps
1197 val zs = foldl (fn (zs, y) => assocs qs y union_string zs) (ys, ys)
1198 fun step(u, us) = (u, if x mem_string us then zs union_string us
1200 in (x, zs) :: map step qs end;
1203 (* generating identifiers *)
1205 (** Freshly generated identifiers; supplied prefix MUST start with a letter **)
1207 (*Maps 0-63 to A-Z, a-z, 0-9 or _ or ' for generating random identifiers*)
1208 fun char i = if i<26 then chr (ord "A" + i)
1209 else if i<52 then chr (ord "a" + i - 26)
1210 else if i<62 then chr (ord"0" + i - 52)
1211 else if i=62 then "_"
1214 val charVec = Vector.tabulate (64, char);
1218 in implode (map (fn i => Vector.sub(charVec,i)) (radixpand (64,n))) end;
1224 fun init_gensym() = (seedr := 0);
1226 fun gensym pre = pre ^ (#1(newid (!seedr), inc seedr));
1231 (*Identifies those character codes legal in identifiers.
1232 chould use Basis Library character functions if Poly/ML provided characters*)
1233 fun idCode k = (ord "a" <= k andalso k < ord "z") orelse
1234 (ord "A" <= k andalso k < ord "Z") orelse
1235 (ord "0" <= k andalso k < ord "9");
1237 val idCodeVec = Vector.tabulate (256, idCode);
1241 (*Increment a list of letters like a reversed base 26 number.
1242 If head is "z", bumps chars in tail.
1243 Digits are incremented as if they were integers.
1244 "_" and "'" are not changed.
1245 For making variants of identifiers.*)
1247 fun bump_int_list(c::cs) =
1248 if c="9" then "0" :: bump_int_list cs
1250 if "0" <= c andalso c < "9" then chr(ord(c)+1) :: cs
1252 | bump_int_list([]) = error("bump_int_list: not an identifier");
1254 fun bump_list([], d) = [d]
1255 | bump_list(["'"], d) = [d, "'"]
1256 | bump_list("z"::cs, _) = "a" :: bump_list(cs, "a")
1257 | bump_list("Z"::cs, _) = "A" :: bump_list(cs, "A")
1258 | bump_list("9"::cs, _) = "0" :: bump_int_list cs
1259 | bump_list(c::cs, _) =
1261 in if Vector.sub(idCodeVec,k) then chr(k+1) :: cs
1263 if c="'" orelse c="_" then c :: bump_list(cs, "")
1264 else error("bump_list: not legal in identifier: " ^
1265 implode(rev(c::cs)))
1270 fun bump_string s : string = implode (rev (bump_list(rev(explode s), "")));
1273 (* lexical scanning *)
1275 (*scan a list of characters into "words" composed of "letters" (recognized by
1276 is_let) and separated by any number of non-"letters"*)
1277 fun scanwords is_let cs =
1278 let fun scan1 [] = []
1280 let val (lets, rest) = take_prefix is_let cs
1281 in implode lets :: scanwords is_let rest end;
1282 in scan1 (#2 (take_prefix (not o is_let) cs)) end;
1286 (* Variable-branching trees: for proof terms etc. *)
1287 datatype 'a mtree = Join of 'a * 'a mtree list;