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;
18 val curry: ('a * 'b -> 'c) -> 'a -> 'b -> 'c
19 val uncurry: ('a -> 'b -> 'c) -> 'a * 'b -> 'c
22 val |> : 'a * ('a -> 'b) -> 'b
23 val apl: 'a * ('a * 'b -> 'c) -> 'b -> 'c
24 val apr: ('a * 'b -> 'c) * 'b -> 'a -> 'c
25 val funpow: int -> ('a -> 'a) -> 'a -> 'a
29 val stamp: unit -> stamp
32 datatype 'a option = None | Some of 'a
34 val the: 'a option -> 'a
35 val if_none: 'a option -> 'a -> 'a
36 val is_some: 'a option -> bool
37 val is_none: 'a option -> bool
38 val apsome: ('a -> 'b) -> 'a option -> 'b option
39 val can: ('a -> 'b) -> 'a -> bool
40 val try: ('a -> 'b) -> 'a -> 'b option
43 val pair: 'a -> 'b -> 'a * 'b
44 val rpair: 'a -> 'b -> 'b * 'a
45 val fst: 'a * 'b -> 'a
46 val snd: 'a * 'b -> 'b
47 val eq_fst: (''a * 'b) * (''a * 'c) -> bool
48 val eq_snd: ('a * ''b) * ('c * ''b) -> bool
49 val swap: 'a * 'b -> 'b * 'a
50 val apfst: ('a -> 'b) -> 'a * 'c -> 'b * 'c
51 val apsnd: ('a -> 'b) -> 'c * 'a -> 'c * 'b
52 val pairself: ('a -> 'b) -> 'a * 'a -> 'b * 'b
55 val equal: ''a -> ''a -> bool
56 val not_equal: ''a -> ''a -> bool
57 val orf: ('a -> bool) * ('a -> bool) -> 'a -> bool
58 val andf: ('a -> bool) * ('a -> bool) -> 'a -> bool
59 val exists: ('a -> bool) -> 'a list -> bool
60 val forall: ('a -> bool) -> 'a list -> bool
61 val set: bool ref -> bool
62 val reset: bool ref -> bool
63 val toggle: bool ref -> bool
64 val setmp: 'a ref -> 'a -> ('b -> 'c) -> 'b -> 'c
67 exception LIST of string
68 val null: 'a list -> bool
70 val tl: 'a list -> 'a list
71 val cons: 'a -> 'a list -> 'a list
72 val single: 'a -> 'a list
73 val append: 'a list -> 'a list -> 'a list
74 val foldl: ('a * 'b -> 'a) -> 'a * 'b list -> 'a
75 val foldr: ('a * 'b -> 'b) -> 'a list * 'b -> 'b
76 val foldr1: ('a * 'a -> 'a) -> 'a list -> 'a
77 val foldl_map: ('a * 'b -> 'a * 'c) -> 'a * 'b list -> 'a * 'c list
78 val length: 'a list -> int
79 val take: int * 'a list -> 'a list
80 val drop: int * 'a list -> 'a list
81 val dropwhile: ('a -> bool) -> 'a list -> 'a list
82 val nth_elem: int * 'a list -> 'a
83 val last_elem: 'a list -> 'a
84 val split_last: 'a list -> 'a list * 'a
85 val nth_update: 'a -> int * 'a list -> 'a list
86 val find_index: ('a -> bool) -> 'a list -> int
87 val find_index_eq: ''a -> ''a list -> int
88 val find_first: ('a -> bool) -> 'a list -> 'a option
89 val get_first: ('a -> 'b option) -> 'a list -> 'b option
90 val flat: 'a list list -> 'a list
91 val seq: ('a -> unit) -> 'a list -> unit
92 val separate: 'a -> 'a list -> 'a list
93 val replicate: int -> 'a -> 'a list
94 val multiply: 'a list * 'a list list -> 'a list list
95 val filter: ('a -> bool) -> 'a list -> 'a list
96 val filter_out: ('a -> bool) -> 'a list -> 'a list
97 val mapfilter: ('a -> 'b option) -> 'a list -> 'b list
98 val map2: ('a * 'b -> 'c) -> 'a list * 'b list -> 'c list
99 val exists2: ('a * 'b -> bool) -> 'a list * 'b list -> bool
100 val forall2: ('a * 'b -> bool) -> 'a list * 'b list -> bool
101 val seq2: ('a * 'b -> unit) -> 'a list * 'b list -> unit
102 val ~~ : 'a list * 'b list -> ('a * 'b) list
103 val split_list: ('a * 'b) list -> 'a list * 'b list
104 val prefix: ''a list * ''a list -> bool
105 val take_prefix: ('a -> bool) -> 'a list -> 'a list * 'a list
106 val take_suffix: ('a -> bool) -> 'a list -> 'a list * 'a list
109 val inc: int ref -> int
110 val dec: int ref -> int
111 val upto: int * int -> int list
112 val downto: int * int -> int list
113 val downto0: int list * int -> bool
114 val radixpand: int * int -> int list
115 val radixstring: int * string * int -> string
116 val string_of_int: int -> string
117 val string_of_indexname: string * int -> string
120 val enclose: string -> string -> string -> string
121 val quote: string -> string
122 val space_implode: string -> string list -> string
123 val commas: string list -> string
124 val commas_quote: string list -> string
125 val cat_lines: string list -> string
126 val space_explode: string -> string -> string list
127 val split_lines: string -> string list
128 val suffix: string -> string -> string
129 val unsuffix: string -> string -> string
132 val mem: ''a * ''a list -> bool
133 val mem_int: int * int list -> bool
134 val mem_string: string * string list -> bool
135 val gen_mem: ('a * 'b -> bool) -> 'a * 'b list -> bool
136 val ins: ''a * ''a list -> ''a list
137 val ins_int: int * int list -> int list
138 val ins_string: string * string list -> string list
139 val gen_ins: ('a * 'a -> bool) -> 'a * 'a list -> 'a list
140 val union: ''a list * ''a list -> ''a list
141 val union_int: int list * int list -> int list
142 val union_string: string list * string list -> string list
143 val gen_union: ('a * 'a -> bool) -> 'a list * 'a list -> 'a list
144 val inter: ''a list * ''a list -> ''a list
145 val inter_int: int list * int list -> int list
146 val inter_string: string list * string list -> string list
147 val subset: ''a list * ''a list -> bool
148 val subset_int: int list * int list -> bool
149 val subset_string: string list * string list -> bool
150 val eq_set: ''a list * ''a list -> bool
151 val eq_set_string: string list * string list -> bool
152 val gen_subset: ('a * 'b -> bool) -> 'a list * 'b list -> bool
153 val \ : ''a list * ''a -> ''a list
154 val \\ : ''a list * ''a list -> ''a list
155 val gen_rem: ('a * 'b -> bool) -> 'a list * 'b -> 'a list
156 val gen_rems: ('a * 'b -> bool) -> 'a list * 'b list -> 'a list
157 val gen_distinct: ('a * 'a -> bool) -> 'a list -> 'a list
158 val distinct: ''a list -> ''a list
159 val findrep: ''a list -> ''a list
160 val gen_duplicates: ('a * 'a -> bool) -> 'a list -> 'a list
161 val duplicates: ''a list -> ''a list
163 (*association lists*)
164 val assoc: (''a * 'b) list * ''a -> 'b option
165 val assoc_int: (int * 'a) list * int -> 'a option
166 val assoc_string: (string * 'a) list * string -> 'a option
167 val assoc_string_int: ((string * int) * 'a) list * (string * int) -> 'a option
168 val assocs: (''a * 'b list) list -> ''a -> 'b list
169 val assoc2: (''a * (''b * 'c) list) list * (''a * ''b) -> 'c option
170 val gen_assoc: ('a * 'b -> bool) -> ('b * 'c) list * 'a -> 'c option
171 val overwrite: (''a * 'b) list * (''a * 'b) -> (''a * 'b) list
172 val gen_overwrite: ('a * 'a -> bool) -> ('a * 'b) list * ('a * 'b) -> ('a * 'b) list
175 val generic_extend: ('a * 'a -> bool)
176 -> ('b -> 'a list) -> ('a list -> 'b) -> 'b -> 'a list -> 'b
177 val generic_merge: ('a * 'a -> bool) -> ('b -> 'a list) -> ('a list -> 'b) -> 'b -> 'b -> 'b
178 val extend_list: ''a list -> ''a list -> ''a list
179 val merge_lists: ''a list -> ''a list -> ''a list
180 val merge_alists: (''a * 'b) list -> (''a * 'b) list -> (''a * 'b) list
181 val merge_rev_lists: ''a list -> ''a list -> ''a list
185 val fold_bal: ('a * 'a -> 'a) -> 'a list -> 'a
186 val access_bal: ('a -> 'a) * ('a -> 'a) * 'a -> int -> int -> 'a
187 val accesses_bal: ('a -> 'a) * ('a -> 'a) * 'a -> int -> 'a list
190 datatype order = EQUAL | GREATER | LESS
191 val rev_order: order -> order
192 val make_ord: ('a * 'a -> bool) -> 'a * 'a -> order
193 val int_ord: int * int -> order
194 val string_ord: string * string -> order
195 val prod_ord: ('a * 'b -> order) -> ('c * 'd -> order) -> ('a * 'c) * ('b * 'd) -> order
196 val dict_ord: ('a * 'b -> order) -> 'a list * 'b list -> order
197 val list_ord: ('a * 'b -> order) -> 'a list * 'b list -> order
198 val sort: ('a * 'a -> order) -> 'a list -> 'a list
199 val sort_strings: string list -> string list
200 val sort_wrt: ('a -> string) -> 'a list -> 'a list
202 (*I/O and diagnostics*)
203 val cd: string -> unit
204 val pwd: unit -> string
205 val prs_fn: (string -> unit) ref
206 val warning_fn: (string -> unit) ref
207 val error_fn: (string -> unit) ref
208 val prs: string -> unit
209 val writeln: string -> unit
210 val warning: string -> unit
212 val error_msg: string -> unit
213 val error: string -> 'a
214 val sys_error: string -> 'a
215 val assert: bool -> string -> unit
216 val deny: bool -> string -> unit
217 val assert_all: ('a -> bool) -> 'a list -> ('a -> string) -> unit
218 datatype 'a error = Error of string | OK of 'a
219 val get_error: 'a error -> string option
220 val get_ok: 'a error -> 'a option
221 val handle_error: ('a -> 'b) -> 'a -> 'b error
222 exception ERROR_MESSAGE of string
223 val transform_error: ('a -> 'b) -> 'a -> 'b
226 val cond_timeit: bool -> (unit -> 'a) -> 'a
227 val timeit: (unit -> 'a) -> 'a
228 val timeap: ('a -> 'b) -> 'a -> 'b
231 val make_keylist: ('a -> 'b) -> 'a list -> ('a * 'b) list
232 val keyfilter: ('a * ''b) list -> ''b -> 'a list
233 val partition: ('a -> bool) -> 'a list -> 'a list * 'a list
234 val partition_eq: ('a * 'a -> bool) -> 'a list -> 'a list list
235 val partition_list: (int -> 'a -> bool) -> int -> int -> 'a list -> 'a list list
236 val transitive_closure: (string * string list) list -> (string * string list) list
237 val init_gensym: unit -> unit
238 val gensym: string -> string
239 val bump_int_list: string list -> string list
240 val bump_list: string list * string -> string list
241 val bump_string: string -> string
242 val scanwords: (string -> bool) -> string list -> string list
243 datatype 'a mtree = Join of 'a * 'a mtree list
246 structure Library: LIBRARY =
252 (*handy combinators*)
253 fun curry f x y = f (x, y);
254 fun uncurry f (x, y) = f x y;
261 (*application of (infix) operator to its left or right argument*)
262 fun apl (x, f) y = f (x, y);
263 fun apr (f, y) x = f (x, y);
265 (*function exponentiation: f(...(f x)...) with n applications of f*)
267 let fun rep (0, x) = x
268 | rep (n, x) = rep (n - 1, f x)
275 type stamp = unit ref;
276 val stamp: unit -> stamp = ref;
282 datatype 'a option = None | Some of 'a;
287 | the None = raise OPTION;
290 fun if_none None y = y
291 | if_none (Some x) _ = x;
293 fun is_some (Some _) = true
294 | is_some None = false;
296 fun is_none (Some _) = false
297 | is_none None = true;
299 fun apsome f (Some x) = Some (f x)
300 | apsome _ None = None;
302 (*handle partial functions*)
303 fun can f x = (f x; true) handle _ => false;
304 fun try f x = Some (f x) handle _ => None;
310 fun pair x y = (x, y);
311 fun rpair x y = (y, x);
316 fun eq_fst ((x1, _), (x2, _)) = x1 = x2;
317 fun eq_snd ((_, y1), (_, y2)) = y1 = y2;
319 fun swap (x, y) = (y, x);
321 (*apply function to components*)
322 fun apfst f (x, y) = (f x, y);
323 fun apsnd f (x, y) = (x, f y);
324 fun pairself f (x, y) = (f x, f y);
332 fun equal x y = x = y;
333 fun not_equal x y = x <> y;
336 (* operators for combining predicates *)
338 fun (p orf q) = fn x => p x orelse q x;
339 fun (p andf q) = fn x => p x andalso q x;
342 (* predicates on lists *)
344 (*exists pred [x1, ..., xn] ===> pred x1 orelse ... orelse pred xn*)
345 fun exists (pred: 'a -> bool) : 'a list -> bool =
346 let fun boolf [] = false
347 | boolf (x :: xs) = pred x orelse boolf xs
350 (*forall pred [x1, ..., xn] ===> pred x1 andalso ... andalso pred xn*)
351 fun forall (pred: 'a -> bool) : 'a list -> bool =
352 let fun boolf [] = true
353 | boolf (x :: xs) = pred x andalso boolf xs
359 fun set flag = (flag := true; true);
360 fun reset flag = (flag := false; false);
361 fun toggle flag = (flag := not (! flag); ! flag);
363 (*temporarily set flag, handling errors*)
364 fun setmp flag value f x =
366 val orig_value = ! flag;
367 fun return y = (flag := orig_value; y);
370 return (f x handle exn => (return (); raise exn))
377 exception LIST of string;
380 | null (_ :: _) = false;
382 fun hd [] = raise LIST "hd"
385 fun tl [] = raise LIST "tl"
388 fun cons x xs = x :: xs;
391 fun append xs ys = xs @ ys;
396 (*the following versions of fold are designed to fit nicely with infixes*)
398 (* (op @) (e, [x1, ..., xn]) ===> ((e @ x1) @ x2) ... @ xn
399 for operators that associate to the left (TAIL RECURSIVE)*)
400 fun foldl (f: 'a * 'b -> 'a) : 'a * 'b list -> 'a =
401 let fun itl (e, []) = e
402 | itl (e, a::l) = itl (f(e, a), l)
405 (* (op @) ([x1, ..., xn], e) ===> x1 @ (x2 ... @ (xn @ e))
406 for operators that associate to the right (not tail recursive)*)
409 | itr (a::l) = f(a, itr l)
412 (* (op @) [x1, ..., xn] ===> x1 @ (x2 ... @ (x[n-1] @ xn))
413 for n > 0, operators that associate to the right (not tail recursive)*)
416 | itr (x::l) = f(x, itr l)
419 fun foldl_map _ (x, []) = (x, [])
420 | foldl_map f (x, y :: ys) =
422 val (x', y') = f (x, y);
423 val (x'', ys') = foldl_map f (x', ys);
424 in (x'', y' :: ys') end;
427 (* basic list functions *)
429 (*length of a list, should unquestionably be a standard function*)
430 local fun length1 (n, []) = n (*TAIL RECURSIVE*)
431 | length1 (n, x :: xs) = length1 (n + 1, xs)
432 in fun length l = length1 (0, l) end;
434 (*take the first n elements from a list*)
435 fun take (n, []) = []
436 | take (n, x :: xs) =
437 if n > 0 then x :: take (n - 1, xs) else [];
439 (*drop the first n elements from a list*)
440 fun drop (n, []) = []
441 | drop (n, x :: xs) =
442 if n > 0 then drop (n - 1, xs) else x :: xs;
444 fun dropwhile P [] = []
445 | dropwhile P (ys as x::xs) = if P x then dropwhile P xs else ys;
447 (*return nth element of a list, where 0 designates the first element;
448 raise EXCEPTION if list too short*)
451 [] => raise LIST "nth_elem"
454 (*last element of a list*)
455 fun last_elem [] = raise LIST "last_elem"
457 | last_elem (_ :: xs) = last_elem xs;
459 (*rear decomposition*)
460 fun split_last [] = raise LIST "split_last"
461 | split_last [x] = ([], x)
462 | split_last (x :: xs) = apfst (cons x) (split_last xs);
464 (*update nth element*)
465 fun nth_update x (n, xs) =
467 val prfx = take (n, xs);
468 val sffx = drop (n, xs);
471 [] => raise LIST "nth_update"
472 | _ :: sffx' => prfx @ (x :: sffx'))
475 (*find the position of an element in a list*)
476 fun find_index pred =
477 let fun find _ [] = ~1
478 | find n (x :: xs) = if pred x then n else find (n + 1) xs;
481 fun find_index_eq x = find_index (equal x);
483 (*find first element satisfying predicate*)
484 fun find_first _ [] = None
485 | find_first pred (x :: xs) =
486 if pred x then Some x else find_first pred xs;
488 (*get first element by lookup function*)
489 fun get_first _ [] = None
490 | get_first f (x :: xs) =
492 None => get_first f xs
495 (*flatten a list of lists to a list*)
496 fun flat (ls: 'c list list) : 'c list = foldr (op @) (ls, []);
498 (*like Lisp's MAPC -- seq proc [x1, ..., xn] evaluates
499 (proc x1; ...; proc xn) for side effects*)
500 fun seq (proc: 'a -> unit) : 'a list -> unit =
502 | seqf (x :: xs) = (proc x; seqf xs)
505 (*separate s [x1, x2, ..., xn] ===> [x1, s, x2, s, ..., s, xn]*)
506 fun separate s (x :: (xs as _ :: _)) = x :: s :: separate s xs
507 | separate _ xs = xs;
509 (*make the list [x, x, ..., x] of length n*)
510 fun replicate n (x: 'a) : 'a list =
511 let fun rep (0, xs) = xs
512 | rep (n, xs) = rep (n - 1, x :: xs)
514 if n < 0 then raise LIST "replicate"
518 (*multiply [a, b, c, ...] * [xs, ys, zs, ...]*)
519 fun multiply ([], _) = []
520 | multiply (x :: xs, yss) = map (cons x) yss @ multiply (xs, yss);
525 (*copy the list preserving elements that satisfy the predicate*)
526 fun filter (pred: 'a->bool) : 'a list -> 'a list =
528 | filt (x :: xs) = if pred x then x :: filt xs else filt xs
531 fun filter_out f = filter (not o f);
533 fun mapfilter (f: 'a -> 'b option) ([]: 'a list) = [] : 'b list
534 | mapfilter f (x :: xs) =
536 None => mapfilter f xs
537 | Some y => y :: mapfilter f xs);
542 fun map2 _ ([], []) = []
543 | map2 f (x :: xs, y :: ys) = (f (x, y) :: map2 f (xs, ys))
544 | map2 _ _ = raise LIST "map2";
546 fun exists2 _ ([], []) = false
547 | exists2 pred (x :: xs, y :: ys) = pred (x, y) orelse exists2 pred (xs, ys)
548 | exists2 _ _ = raise LIST "exists2";
550 fun forall2 _ ([], []) = true
551 | forall2 pred (x :: xs, y :: ys) = pred (x, y) andalso forall2 pred (xs, ys)
552 | forall2 _ _ = raise LIST "forall2";
554 fun seq2 _ ([], []) = ()
555 | seq2 f (x :: xs, y :: ys) = (f (x, y); seq2 f (xs, ys))
556 | seq2 _ _ = raise LIST "seq2";
558 (*combine two lists forming a list of pairs:
559 [x1, ..., xn] ~~ [y1, ..., yn] ===> [(x1, y1), ..., (xn, yn)]*)
561 | (x :: xs) ~~ (y :: ys) = (x, y) :: (xs ~~ ys)
562 | _ ~~ _ = raise LIST "~~";
564 (*inverse of ~~; the old 'split':
565 [(x1, y1), ..., (xn, yn)] ===> ([x1, ..., xn], [y1, ..., yn])*)
566 fun split_list (l: ('a * 'b) list) = (map #1 l, map #2 l);
569 (* prefixes, suffixes *)
571 fun [] prefix _ = true
572 | (x :: xs) prefix (y :: ys) = x = y andalso (xs prefix ys)
573 | _ prefix _ = false;
575 (* [x1, ..., xi, ..., xn] ---> ([x1, ..., x(i-1)], [xi, ..., xn])
576 where xi is the first element that does not satisfy the predicate*)
577 fun take_prefix (pred : 'a -> bool) (xs: 'a list) : 'a list * 'a list =
578 let fun take (rxs, []) = (rev rxs, [])
579 | take (rxs, x :: xs) =
580 if pred x then take(x :: rxs, xs) else (rev rxs, x :: xs)
583 (* [x1, ..., xi, ..., xn] ---> ([x1, ..., xi], [x(i+1), ..., xn])
584 where xi is the last element that does not satisfy the predicate*)
585 fun take_suffix _ [] = ([], [])
586 | take_suffix pred (x :: xs) =
587 (case take_suffix pred xs of
588 ([], sffx) => if pred x then ([], x :: sffx) else ([x], sffx)
589 | (prfx, sffx) => (x :: prfx, sffx));
595 fun inc i = (i := ! i + 1; ! i);
596 fun dec i = (i := ! i - 1; ! i);
599 (* lists of integers *)
601 (*make the list [from, from + 1, ..., to]*)
603 if from > to then [] else from :: ((from + 1) upto to);
605 (*make the list [from, from - 1, ..., to]*)
606 fun (from downto to) =
607 if from < to then [] else from :: ((from - 1) downto to);
609 (*predicate: downto0 (is, n) <=> is = [n, n - 1, ..., 0]*)
610 fun downto0 (i :: is, n) = i = n andalso downto0 (is, n - 1)
611 | downto0 ([], n) = n = ~1;
614 (* convert integers to strings *)
616 (*expand the number in the given base;
617 example: radixpand (2, 8) gives [1, 0, 0, 0]*)
618 fun radixpand (base, num) : int list =
620 fun radix (n, tail) =
621 if n < base then n :: tail
622 else radix (n div base, (n mod base) :: tail)
623 in radix (num, []) end;
625 (*expands a number into a string of characters starting from "zerochar";
626 example: radixstring (2, "0", 8) gives "1000"*)
627 fun radixstring (base, zerochar, num) =
628 let val offset = ord zerochar;
629 fun chrof n = chr (offset + n)
630 in implode (map chrof (radixpand (base, num))) end;
633 val string_of_int = Int.toString;
635 fun string_of_indexname (a,0) = a
636 | string_of_indexname (a,i) = a ^ "_" ^ Int.toString i;
642 (*enclose in brackets*)
643 fun enclose lpar rpar str = lpar ^ str ^ rpar;
645 (*simple quoting (does not escape special chars)*)
646 val quote = enclose "\"" "\"";
648 (*space_implode "..." (explode "hello") = "h...e...l...l...o"*)
649 fun space_implode a bs = implode (separate a bs);
651 val commas = space_implode ", ";
652 val commas_quote = commas o map quote;
654 (*concatenate messages, one per line, into a string*)
655 val cat_lines = space_implode "\n";
657 (*space_explode "." "h.e..l.lo" = ["h", "e", "", "l", "lo"]*)
658 fun space_explode _ "" = []
659 | space_explode sep str =
662 (case take_prefix (not_equal sep) chs of
663 (cs, []) => [implode cs]
664 | (cs, _ :: cs') => implode cs :: expl cs');
665 in expl (explode str) end;
667 val split_lines = space_explode "\n";
670 fun suffix sfx s = s ^ sfx;
676 val prfx_len = size s - size sfx;
678 if prfx_len >= 0 andalso implode (drop (prfx_len, cs)) = sfx then
679 implode (take (prfx_len, cs))
680 else raise LIST "unsuffix"
685 (** lists as sets **)
687 (*membership in a list*)
689 | x mem (y :: ys) = x = y orelse x mem ys;
691 (*membership in a list, optimized version for ints*)
692 fun (x:int) mem_int [] = false
693 | x mem_int (y :: ys) = x = y orelse x mem_int ys;
695 (*membership in a list, optimized version for strings*)
696 fun (x:string) mem_string [] = false
697 | x mem_string (y :: ys) = x = y orelse x mem_string ys;
699 (*generalized membership test*)
700 fun gen_mem eq (x, []) = false
701 | gen_mem eq (x, y :: ys) = eq (x, y) orelse gen_mem eq (x, ys);
704 (*insertion into list if not already there*)
705 fun (x ins xs) = if x mem xs then xs else x :: xs;
707 (*insertion into list, optimized version for ints*)
708 fun (x ins_int xs) = if x mem_int xs then xs else x :: xs;
710 (*insertion into list, optimized version for strings*)
711 fun (x ins_string xs) = if x mem_string xs then xs else x :: xs;
713 (*generalized insertion*)
714 fun gen_ins eq (x, xs) = if gen_mem eq (x, xs) then xs else x :: xs;
717 (*union of sets represented as lists: no repetitions*)
720 | (x :: xs) union ys = xs union (x ins ys);
722 (*union of sets, optimized version for ints*)
723 fun (xs:int list) union_int [] = xs
724 | [] union_int ys = ys
725 | (x :: xs) union_int ys = xs union_int (x ins_int ys);
727 (*union of sets, optimized version for strings*)
728 fun (xs:string list) union_string [] = xs
729 | [] union_string ys = ys
730 | (x :: xs) union_string ys = xs union_string (x ins_string ys);
732 (*generalized union*)
733 fun gen_union eq (xs, []) = xs
734 | gen_union eq ([], ys) = ys
735 | gen_union eq (x :: xs, ys) = gen_union eq (xs, gen_ins eq (x, ys));
740 | (x :: xs) inter ys =
741 if x mem ys then x :: (xs inter ys) else xs inter ys;
743 (*intersection, optimized version for ints*)
744 fun ([]:int list) inter_int ys = []
745 | (x :: xs) inter_int ys =
746 if x mem_int ys then x :: (xs inter_int ys) else xs inter_int ys;
748 (*intersection, optimized version for strings *)
749 fun ([]:string list) inter_string ys = []
750 | (x :: xs) inter_string ys =
751 if x mem_string ys then x :: (xs inter_string ys) else xs inter_string ys;
755 fun [] subset ys = true
756 | (x :: xs) subset ys = x mem ys andalso xs subset ys;
758 (*subset, optimized version for ints*)
759 fun ([]:int list) subset_int ys = true
760 | (x :: xs) subset_int ys = x mem_int ys andalso xs subset_int ys;
762 (*subset, optimized version for strings*)
763 fun ([]:string list) subset_string ys = true
764 | (x :: xs) subset_string ys = x mem_string ys andalso xs subset_string ys;
767 fun eq_set (xs, ys) =
768 xs = ys orelse (xs subset ys andalso ys subset xs);
770 (*set equality for strings*)
771 fun eq_set_string ((xs:string list), ys) =
772 xs = ys orelse (xs subset_string ys andalso ys subset_string xs);
774 fun gen_subset eq (xs, ys) = forall (fn x => gen_mem eq (x, ys)) xs;
777 (*removing an element from a list WITHOUT duplicates*)
778 fun (y :: ys) \ x = if x = y then ys else y :: (ys \ x)
781 fun ys \\ xs = foldl (op \) (ys,xs);
783 (*removing an element from a list -- possibly WITH duplicates*)
784 fun gen_rem eq (xs, y) = filter_out (fn x => eq (x, y)) xs;
786 fun gen_rems eq = foldl (gen_rem eq);
789 (*makes a list of the distinct members of the input; preserves order, takes
790 first of equal elements*)
791 fun gen_distinct eq lst =
793 val memb = gen_mem eq;
795 fun dist (rev_seen, []) = rev rev_seen
796 | dist (rev_seen, x :: xs) =
797 if memb (x, rev_seen) then dist (rev_seen, xs)
798 else dist (x :: rev_seen, xs);
803 fun distinct l = gen_distinct (op =) l;
806 (*returns the tail beginning with the first repeated element, or []*)
808 | findrep (x :: xs) = if x mem xs then x :: xs else findrep xs;
811 (*returns a list containing all repeated elements exactly once; preserves
812 order, takes first of equal elements*)
813 fun gen_duplicates eq lst =
815 val memb = gen_mem eq;
817 fun dups (rev_dups, []) = rev rev_dups
818 | dups (rev_dups, x :: xs) =
819 if memb (x, rev_dups) orelse not (memb (x, xs)) then
821 else dups (x :: rev_dups, xs);
826 fun duplicates l = gen_duplicates (op =) l;
830 (** association lists **)
832 (*association list lookup*)
833 fun assoc ([], key) = None
834 | assoc ((keyi, xi) :: pairs, key) =
835 if key = keyi then Some xi else assoc (pairs, key);
837 (*association list lookup, optimized version for ints*)
838 fun assoc_int ([], (key:int)) = None
839 | assoc_int ((keyi, xi) :: pairs, key) =
840 if key = keyi then Some xi else assoc_int (pairs, key);
842 (*association list lookup, optimized version for strings*)
843 fun assoc_string ([], (key:string)) = None
844 | assoc_string ((keyi, xi) :: pairs, key) =
845 if key = keyi then Some xi else assoc_string (pairs, key);
847 (*association list lookup, optimized version for string*ints*)
848 fun assoc_string_int ([], (key:string*int)) = None
849 | assoc_string_int ((keyi, xi) :: pairs, key) =
850 if key = keyi then Some xi else assoc_string_int (pairs, key);
853 (case assoc (ps, x) of
857 (*two-fold association list lookup*)
858 fun assoc2 (aal, (key1, key2)) =
859 (case assoc (aal, key1) of
860 Some al => assoc (al, key2)
863 (*generalized association list lookup*)
864 fun gen_assoc eq ([], key) = None
865 | gen_assoc eq ((keyi, xi) :: pairs, key) =
866 if eq (key, keyi) then Some xi else gen_assoc eq (pairs, key);
868 (*association list update*)
869 fun overwrite (al, p as (key, _)) =
870 let fun over ((q as (keyi, _)) :: pairs) =
871 if keyi = key then p :: pairs else q :: (over pairs)
875 fun gen_overwrite eq (al, p as (key, _)) =
876 let fun over ((q as (keyi, _)) :: pairs) =
877 if eq (keyi, key) then p :: pairs else q :: (over pairs)
883 (** generic tables **)
885 (*Tables are supposed to be 'efficient' encodings of lists of elements distinct
886 wrt. an equality "eq". The extend and merge operations below are optimized
887 for long-term space efficiency.*)
889 (*append (new) elements to a table*)
890 fun generic_extend _ _ _ tab [] = tab
891 | generic_extend eq dest_tab mk_tab tab1 lst2 =
893 val lst1 = dest_tab tab1;
894 val new_lst2 = gen_rems eq (lst2, lst1);
896 if null new_lst2 then tab1
897 else mk_tab (lst1 @ new_lst2)
900 (*append (new) elements of 2nd table to 1st table*)
901 fun generic_merge eq dest_tab mk_tab tab1 tab2 =
903 val lst1 = dest_tab tab1;
904 val lst2 = dest_tab tab2;
905 val new_lst2 = gen_rems eq (lst2, lst1);
907 if null new_lst2 then tab1
908 else if gen_subset eq (lst1, lst2) then tab2
909 else mk_tab (lst1 @ new_lst2)
914 fun extend_list tab = generic_extend (op =) I I tab;
915 fun merge_lists tab = generic_merge (op =) I I tab;
916 fun merge_alists tab = generic_merge eq_fst I I tab;
918 fun merge_rev_lists xs [] = xs
919 | merge_rev_lists [] ys = ys
920 | merge_rev_lists xs (y :: ys) =
921 (if y mem xs then I else cons y) (merge_rev_lists xs ys);
925 (** balanced trees **)
927 exception Balance; (*indicates non-positive argument to balancing fun*)
929 (*balanced folding; avoids deep nesting*)
930 fun fold_bal f [x] = x
931 | fold_bal f [] = raise Balance
933 let val k = length xs div 2
934 in f (fold_bal f (take(k, xs)),
935 fold_bal f (drop(k, xs)))
938 (*construct something of the form f(...g(...(x)...)) for balanced access*)
939 fun access_bal (f, g, x) n i =
940 let fun acc n i = (*1<=i<=n*)
943 in if i<=n2 then f (acc n2 i)
944 else g (acc (n-n2) (i-n2))
946 in if 1<=i andalso i<=n then acc n i else raise Balance end;
948 (*construct ALL such accesses; could try harder to share recursive calls!*)
949 fun accesses_bal (f, g, x) n =
954 in if n-n2=n2 then map f acc2 @ map g acc2
955 else map f acc2 @ map g (acc (n-n2)) end
956 in if 1<=n then acc n else raise Balance end;
962 datatype order = LESS | EQUAL | GREATER;
964 fun rev_order LESS = GREATER
965 | rev_order EQUAL = EQUAL
966 | rev_order GREATER = LESS;
968 (*assume rel is a linear strict order*)
969 fun make_ord rel (x, y) =
970 if rel (x, y) then LESS
971 else if rel (y, x) then GREATER
974 fun int_ord (i, j: int) =
976 else if i = j then EQUAL
979 fun string_ord (a, b: string) =
981 else if a = b then EQUAL
984 (*lexicographic product*)
985 fun prod_ord a_ord b_ord ((x, y), (x', y')) =
986 (case a_ord (x, x') of EQUAL => b_ord (y, y') | ord => ord);
988 (*dictionary order -- in general NOT well-founded!*)
989 fun dict_ord _ ([], []) = EQUAL
990 | dict_ord _ ([], _ :: _) = LESS
991 | dict_ord _ (_ :: _, []) = GREATER
992 | dict_ord elem_ord (x :: xs, y :: ys) =
993 (case elem_ord (x, y) of EQUAL => dict_ord elem_ord (xs, ys) | ord => ord);
995 (*lexicographic product of lists*)
996 fun list_ord elem_ord (xs, ys) =
997 prod_ord int_ord (dict_ord elem_ord) ((length xs, xs), (length ys, ys));
1002 (*quicksort (stable, i.e. does not reorder equal elements)*)
1006 let val len = length xs in
1009 let val (lts, eqs, gts) = part (nth_elem (len div 2, xs)) xs in
1010 qsort lts @ eqs @ qsort gts
1013 and part _ [] = ([], [], [])
1014 | part pivot (x :: xs) = add (ord (x, pivot)) x (part pivot xs)
1015 and add LESS x (lts, eqs, gts) = (x :: lts, eqs, gts)
1016 | add EQUAL x (lts, eqs, gts) = (lts, x :: eqs, gts)
1017 | add GREATER x (lts, eqs, gts) = (lts, eqs, x :: gts);
1021 val sort_strings = sort string_ord;
1022 fun sort_wrt sel xs = sort (string_ord o pairself sel) xs;
1026 (** input / output and diagnostics **)
1028 val cd = OS.FileSys.chDir;
1029 val pwd = OS.FileSys.getDir;
1034 (TextIO.output (TextIO.stdOut, s); TextIO.flushOut TextIO.stdOut);
1036 fun prefix_lines prfx txt =
1037 txt |> split_lines |> map (fn s => prfx ^ s ^ "\n") |> implode;
1040 (*hooks for output channels: normal, warning, error*)
1041 val prs_fn = ref (fn s => out s);
1042 val warning_fn = ref (fn s => out (prefix_lines "### " s));
1043 val error_fn = ref (fn s => out (prefix_lines "*** " s));
1047 fun prs s = !prs_fn s;
1048 fun writeln s = prs (s ^ "\n");
1050 fun warning s = !warning_fn s;
1052 (*print error message and abort to top level*)
1054 fun error_msg s = !error_fn s;
1055 fun error s = (error_msg s; raise ERROR);
1056 fun sys_error msg = error ("## SYSTEM ERROR ##\n" ^ msg);
1058 fun assert p msg = if p then () else error msg;
1059 fun deny p msg = if p then error msg else ();
1061 (*Assert pred for every member of l, generating a message if pred fails*)
1062 fun assert_all pred l msg_fn =
1064 | asl (x::xs) = if pred x then asl xs else error (msg_fn x)
1068 (* handle errors capturing messages *)
1074 fun get_error (Error msg) = Some msg
1075 | get_error _ = None;
1077 fun get_ok (OK x) = Some x
1080 datatype 'a result =
1084 fun handle_error f x =
1086 val buffer = ref ([]: string list);
1087 fun capture s = buffer := ! buffer @ [s];
1088 fun err_msg () = if not (null (! buffer)) then error_msg (cat_lines (! buffer)) else ();
1090 (case Result (setmp error_fn capture f x) handle exn => Exn exn of
1091 Result y => (err_msg (); OK y)
1092 | Exn ERROR => Error (cat_lines (! buffer))
1093 | Exn exn => (err_msg (); raise exn))
1097 (* transform ERROR into ERROR_MESSAGE *)
1099 exception ERROR_MESSAGE of string;
1101 fun transform_error f x =
1102 (case handle_error f x of
1104 | Error msg => raise ERROR_MESSAGE msg);
1110 (*a conditional timing function: applies f to () and, if the flag is true,
1111 prints its runtime*)
1112 fun cond_timeit flag f =
1114 let val start = startTiming()
1117 writeln (endTiming start); result
1121 (*unconditional timing function*)
1122 fun timeit x = cond_timeit true x;
1124 (*timed application function*)
1125 fun timeap f x = timeit (fn () => f x);
1131 (*use the keyfun to make a list of (x, key) pairs*)
1132 fun make_keylist (keyfun: 'a->'b) : 'a list -> ('a * 'b) list =
1133 let fun keypair x = (x, keyfun x)
1136 (*given a list of (x, key) pairs and a searchkey
1137 return the list of xs from each pair whose key equals searchkey*)
1138 fun keyfilter [] searchkey = []
1139 | keyfilter ((x, key) :: pairs) searchkey =
1140 if key = searchkey then x :: keyfilter pairs searchkey
1141 else keyfilter pairs searchkey;
1144 (*Partition list into elements that satisfy predicate and those that don't.
1145 Preserves order of elements in both lists.*)
1146 fun partition (pred: 'a->bool) (ys: 'a list) : ('a list * 'a list) =
1147 let fun part ([], answer) = answer
1148 | part (x::xs, (ys, ns)) = if pred(x)
1149 then part (xs, (x::ys, ns))
1150 else part (xs, (ys, x::ns))
1151 in part (rev ys, ([], [])) end;
1154 fun partition_eq (eq:'a * 'a -> bool) =
1155 let fun part [] = []
1156 | part (x::ys) = let val (xs, xs') = partition (apl(x, eq)) ys
1157 in (x::xs)::(part xs') end
1161 (*Partition a list into buckets [ bi, b(i+1), ..., bj ]
1162 putting x in bk if p(k)(x) holds. Preserve order of elements if possible.*)
1163 fun partition_list p i j =
1166 (case xs of [] => []
1167 | _ => raise LIST "partition_list")
1169 let val (ns, rest) = partition (p k) xs;
1170 in ns :: part(k+1)rest end
1174 (* transitive closure (not Warshall's algorithm) *)
1176 fun transitive_closure [] = []
1177 | transitive_closure ((x, ys)::ps) =
1178 let val qs = transitive_closure ps
1179 val zs = foldl (fn (zs, y) => assocs qs y union_string zs) (ys, ys)
1180 fun step(u, us) = (u, if x mem_string us then zs union_string us
1182 in (x, zs) :: map step qs end;
1185 (* generating identifiers *)
1187 (** Freshly generated identifiers; supplied prefix MUST start with a letter **)
1189 (*Maps 0-63 to A-Z, a-z, 0-9 or _ or ' for generating random identifiers*)
1190 fun char i = if i<26 then chr (ord "A" + i)
1191 else if i<52 then chr (ord "a" + i - 26)
1192 else if i<62 then chr (ord"0" + i - 52)
1193 else if i=62 then "_"
1196 val charVec = Vector.tabulate (64, char);
1200 in implode (map (fn i => Vector.sub(charVec,i)) (radixpand (64,n))) end;
1206 fun init_gensym() = (seedr := 0);
1208 fun gensym pre = pre ^ (#1(newid (!seedr), inc seedr));
1213 (*Identifies those character codes legal in identifiers.
1214 chould use Basis Library character functions if Poly/ML provided characters*)
1215 fun idCode k = (ord "a" <= k andalso k < ord "z") orelse
1216 (ord "A" <= k andalso k < ord "Z") orelse
1217 (ord "0" <= k andalso k < ord "9");
1219 val idCodeVec = Vector.tabulate (256, idCode);
1223 (*Increment a list of letters like a reversed base 26 number.
1224 If head is "z", bumps chars in tail.
1225 Digits are incremented as if they were integers.
1226 "_" and "'" are not changed.
1227 For making variants of identifiers.*)
1229 fun bump_int_list(c::cs) =
1230 if c="9" then "0" :: bump_int_list cs
1232 if "0" <= c andalso c < "9" then chr(ord(c)+1) :: cs
1234 | bump_int_list([]) = error("bump_int_list: not an identifier");
1236 fun bump_list([], d) = [d]
1237 | bump_list(["'"], d) = [d, "'"]
1238 | bump_list("z"::cs, _) = "a" :: bump_list(cs, "a")
1239 | bump_list("Z"::cs, _) = "A" :: bump_list(cs, "A")
1240 | bump_list("9"::cs, _) = "0" :: bump_int_list cs
1241 | bump_list(c::cs, _) =
1243 in if Vector.sub(idCodeVec,k) then chr(k+1) :: cs
1245 if c="'" orelse c="_" then c :: bump_list(cs, "")
1246 else error("bump_list: not legal in identifier: " ^
1247 implode(rev(c::cs)))
1252 fun bump_string s : string = implode (rev (bump_list(rev(explode s), "")));
1255 (* lexical scanning *)
1257 (*scan a list of characters into "words" composed of "letters" (recognized by
1258 is_let) and separated by any number of non-"letters"*)
1259 fun scanwords is_let cs =
1260 let fun scan1 [] = []
1262 let val (lets, rest) = take_prefix is_let cs
1263 in implode lets :: scanwords is_let rest end;
1264 in scan1 (#2 (take_prefix (not o is_let) cs)) end;
1268 (* Variable-branching trees: for proof terms etc. *)
1269 datatype 'a mtree = Join of 'a * 'a mtree list;