| 6118 |      1 | (*  Title:      Pure/General/table.ML
 | 
| 5015 |      2 |     ID:         $Id$
 | 
| 15014 |      3 |     Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
 | 
| 5015 |      4 | 
 | 
| 16342 |      5 | Generic tables.  Efficient purely functional implementation using
 | 
|  |      6 | balanced 2-3 trees.
 | 
| 5015 |      7 | *)
 | 
|  |      8 | 
 | 
|  |      9 | signature KEY =
 | 
|  |     10 | sig
 | 
|  |     11 |   type key
 | 
|  |     12 |   val ord: key * key -> order
 | 
|  |     13 | end;
 | 
|  |     14 | 
 | 
|  |     15 | signature TABLE =
 | 
|  |     16 | sig
 | 
|  |     17 |   type key
 | 
|  |     18 |   type 'a table
 | 
|  |     19 |   exception DUP of key
 | 
|  |     20 |   exception DUPS of key list
 | 
| 15014 |     21 |   exception UNDEF of key
 | 
| 5015 |     22 |   val empty: 'a table
 | 
|  |     23 |   val is_empty: 'a table -> bool
 | 
| 5681 |     24 |   val map: ('a -> 'b) -> 'a table -> 'b table
 | 
| 16446 |     25 |   val map': (key -> 'a -> 'b) -> 'a table -> 'b table
 | 
|  |     26 |   val fold: (key * 'b -> 'a -> 'a) -> 'b table -> 'a -> 'a
 | 
| 5681 |     27 |   val foldl: ('a * (key * 'b) -> 'a) -> 'a * 'b table -> 'a
 | 
| 5015 |     28 |   val dest: 'a table -> (key * 'a) list
 | 
| 5681 |     29 |   val keys: 'a table -> key list
 | 
| 8409 |     30 |   val min_key: 'a table -> key option
 | 
| 15160 |     31 |   val max_key: 'a table -> key option
 | 
| 7061 |     32 |   val exists: (key * 'a -> bool) -> 'a table -> bool
 | 
| 16192 |     33 |   val forall: (key * 'a -> bool) -> 'a table -> bool
 | 
| 5015 |     34 |   val lookup: 'a table * key -> 'a option
 | 
|  |     35 |   val update: (key * 'a) * 'a table -> 'a table
 | 
| 12287 |     36 |   val update_new: (key * 'a) * 'a table -> 'a table                    (*exception DUP*)
 | 
| 15665 |     37 |   val map_entry: key -> ('a -> 'a) -> 'a table -> 'a table
 | 
| 12287 |     38 |   val make: (key * 'a) list -> 'a table                                (*exception DUPS*)
 | 
|  |     39 |   val extend: 'a table * (key * 'a) list -> 'a table                   (*exception DUPS*)
 | 
| 16446 |     40 |   val join: (key -> 'a * 'a -> 'a option) ->
 | 
|  |     41 |     'a table * 'a table -> 'a table                                    (*exception DUPS*)
 | 
| 12287 |     42 |   val merge: ('a * 'a -> bool) -> 'a table * 'a table -> 'a table      (*exception DUPS*)
 | 
| 15665 |     43 |   val delete: key -> 'a table -> 'a table                              (*exception UNDEF*)
 | 
| 15761 |     44 |   val delete_safe: key -> 'a table -> 'a table
 | 
| 16466 |     45 |   val member: ('b * 'a -> bool) -> 'a table -> key * 'b -> bool
 | 
| 15761 |     46 |   val insert: ('a * 'a -> bool) -> key * 'a -> 'a table -> 'a table    (*exception DUP*)
 | 
| 16139 |     47 |   val remove: ('b * 'a -> bool) -> key * 'b -> 'a table -> 'a table
 | 
| 5015 |     48 |   val lookup_multi: 'a list table * key -> 'a list
 | 
| 8606 |     49 |   val update_multi: (key * 'a) * 'a list table -> 'a list table
 | 
| 16139 |     50 |   val remove_multi: ('b * 'a -> bool) -> key * 'b -> 'a list table -> 'a list table
 | 
| 5015 |     51 |   val make_multi: (key * 'a) list -> 'a list table
 | 
|  |     52 |   val dest_multi: 'a list table -> (key * 'a) list
 | 
| 12287 |     53 |   val merge_multi: ('a * 'a -> bool) ->
 | 
| 15761 |     54 |     'a list table * 'a list table -> 'a list table                     (*exception DUPS*)
 | 
| 12287 |     55 |   val merge_multi': ('a * 'a -> bool) ->
 | 
| 15761 |     56 |     'a list table * 'a list table -> 'a list table                     (*exception DUPS*)
 | 
| 5015 |     57 | end;
 | 
|  |     58 | 
 | 
|  |     59 | functor TableFun(Key: KEY): TABLE =
 | 
|  |     60 | struct
 | 
|  |     61 | 
 | 
|  |     62 | 
 | 
|  |     63 | (* datatype table *)
 | 
|  |     64 | 
 | 
|  |     65 | type key = Key.key;
 | 
|  |     66 | 
 | 
|  |     67 | datatype 'a table =
 | 
|  |     68 |   Empty |
 | 
|  |     69 |   Branch2 of 'a table * (key * 'a) * 'a table |
 | 
|  |     70 |   Branch3 of 'a table * (key * 'a) * 'a table * (key * 'a) * 'a table;
 | 
|  |     71 | 
 | 
|  |     72 | exception DUP of key;
 | 
|  |     73 | exception DUPS of key list;
 | 
|  |     74 | 
 | 
|  |     75 | 
 | 
| 5681 |     76 | (* empty *)
 | 
|  |     77 | 
 | 
| 5015 |     78 | val empty = Empty;
 | 
|  |     79 | 
 | 
|  |     80 | fun is_empty Empty = true
 | 
|  |     81 |   | is_empty _ = false;
 | 
|  |     82 | 
 | 
| 5681 |     83 | 
 | 
|  |     84 | (* map and fold combinators *)
 | 
|  |     85 | 
 | 
|  |     86 | fun map_table _ Empty = Empty
 | 
|  |     87 |   | map_table f (Branch2 (left, (k, x), right)) =
 | 
| 16446 |     88 |       Branch2 (map_table f left, (k, f k x), map_table f right)
 | 
| 5681 |     89 |   | map_table f (Branch3 (left, (k1, x1), mid, (k2, x2), right)) =
 | 
| 16446 |     90 |       Branch3 (map_table f left, (k1, f k1 x1),
 | 
|  |     91 |         map_table f mid, (k2, f k2 x2), map_table f right);
 | 
| 5681 |     92 | 
 | 
| 16446 |     93 | fun fold_table f Empty x = x
 | 
|  |     94 |   | fold_table f (Branch2 (left, p, right)) x =
 | 
|  |     95 |       fold_table f right (f p (fold_table f left x))
 | 
|  |     96 |   | fold_table f (Branch3 (left, p1, mid, p2, right)) x =
 | 
|  |     97 |       fold_table f right (f p2 (fold_table f mid (f p1 (fold_table f left x))));
 | 
| 5681 |     98 | 
 | 
| 16446 |     99 | fun dest tab = rev (fold_table cons tab []);
 | 
|  |    100 | fun keys tab = rev (fold_table (cons o #1) tab []);
 | 
| 16192 |    101 | 
 | 
|  |    102 | local exception TRUE in
 | 
|  |    103 | 
 | 
|  |    104 | fun exists pred tab =
 | 
| 16446 |    105 |   (fold_table (fn p => fn () => if pred p then raise TRUE else ()) tab (); false)
 | 
| 16192 |    106 |     handle TRUE => true;
 | 
|  |    107 | 
 | 
|  |    108 | fun forall pred = not o exists (not o pred);
 | 
|  |    109 | 
 | 
|  |    110 | end;
 | 
| 5015 |    111 | 
 | 
| 15531 |    112 | fun min_key Empty = NONE
 | 
| 16002 |    113 |   | min_key (Branch2 (left, (k, _), _)) = SOME (if_none (min_key left) k)
 | 
|  |    114 |   | min_key (Branch3 (left, (k, _), _, _, _)) = SOME (if_none (min_key left) k);
 | 
| 8409 |    115 | 
 | 
| 15531 |    116 | fun max_key Empty = NONE
 | 
| 16002 |    117 |   | max_key (Branch2 (_, (k, _), right)) = SOME (if_none (max_key right) k)
 | 
|  |    118 |   | max_key (Branch3 (_, _, _, (k,_), right)) = SOME (if_none (max_key right) k);
 | 
| 15665 |    119 | 
 | 
| 5015 |    120 | 
 | 
|  |    121 | (* lookup *)
 | 
|  |    122 | 
 | 
| 15531 |    123 | fun lookup (Empty, _) = NONE
 | 
| 5015 |    124 |   | lookup (Branch2 (left, (k, x), right), key) =
 | 
|  |    125 |       (case Key.ord (key, k) of
 | 
|  |    126 |         LESS => lookup (left, key)
 | 
| 15531 |    127 |       | EQUAL => SOME x
 | 
| 5015 |    128 |       | GREATER => lookup (right, key))
 | 
|  |    129 |   | lookup (Branch3 (left, (k1, x1), mid, (k2, x2), right), key) =
 | 
|  |    130 |       (case Key.ord (key, k1) of
 | 
|  |    131 |         LESS => lookup (left, key)
 | 
| 15531 |    132 |       | EQUAL => SOME x1
 | 
| 5015 |    133 |       | GREATER =>
 | 
|  |    134 |           (case Key.ord (key, k2) of
 | 
|  |    135 |             LESS => lookup (mid, key)
 | 
| 15531 |    136 |           | EQUAL => SOME x2
 | 
| 5015 |    137 |           | GREATER => lookup (right, key)));
 | 
|  |    138 | 
 | 
|  |    139 | 
 | 
| 15665 |    140 | (* updates *)
 | 
| 5015 |    141 | 
 | 
|  |    142 | datatype 'a growth =
 | 
|  |    143 |   Stay of 'a table |
 | 
|  |    144 |   Sprout of 'a table * (key * 'a) * 'a table;
 | 
|  |    145 | 
 | 
| 15761 |    146 | exception SAME;
 | 
|  |    147 | 
 | 
| 15665 |    148 | fun modify key f tab =
 | 
|  |    149 |   let
 | 
|  |    150 |     fun modfy Empty = Sprout (Empty, (key, f NONE), Empty)
 | 
|  |    151 |       | modfy (Branch2 (left, p as (k, x), right)) =
 | 
|  |    152 |           (case Key.ord (key, k) of
 | 
|  |    153 |             LESS =>
 | 
|  |    154 |               (case modfy left of
 | 
|  |    155 |                 Stay left' => Stay (Branch2 (left', p, right))
 | 
|  |    156 |               | Sprout (left1, q, left2) => Stay (Branch3 (left1, q, left2, p, right)))
 | 
|  |    157 |           | EQUAL => Stay (Branch2 (left, (k, f (SOME x)), right))
 | 
|  |    158 |           | GREATER =>
 | 
|  |    159 |               (case modfy right of
 | 
|  |    160 |                 Stay right' => Stay (Branch2 (left, p, right'))
 | 
|  |    161 |               | Sprout (right1, q, right2) =>
 | 
|  |    162 |                   Stay (Branch3 (left, p, right1, q, right2))))
 | 
|  |    163 |       | modfy (Branch3 (left, p1 as (k1, x1), mid, p2 as (k2, x2), right)) =
 | 
|  |    164 |           (case Key.ord (key, k1) of
 | 
| 5015 |    165 |             LESS =>
 | 
| 15665 |    166 |               (case modfy left of
 | 
|  |    167 |                 Stay left' => Stay (Branch3 (left', p1, mid, p2, right))
 | 
|  |    168 |               | Sprout (left1, q, left2) =>
 | 
|  |    169 |                   Sprout (Branch2 (left1, q, left2), p1, Branch2 (mid, p2, right)))
 | 
|  |    170 |           | EQUAL => Stay (Branch3 (left, (k1, f (SOME x1)), mid, p2, right))
 | 
| 5015 |    171 |           | GREATER =>
 | 
| 15665 |    172 |               (case Key.ord (key, k2) of
 | 
|  |    173 |                 LESS =>
 | 
|  |    174 |                   (case modfy mid of
 | 
|  |    175 |                     Stay mid' => Stay (Branch3 (left, p1, mid', p2, right))
 | 
|  |    176 |                   | Sprout (mid1, q, mid2) =>
 | 
|  |    177 |                       Sprout (Branch2 (left, p1, mid1), q, Branch2 (mid2, p2, right)))
 | 
|  |    178 |               | EQUAL => Stay (Branch3 (left, p1, mid, (k2, f (SOME x2)), right))
 | 
|  |    179 |               | GREATER =>
 | 
|  |    180 |                   (case modfy right of
 | 
|  |    181 |                     Stay right' => Stay (Branch3 (left, p1, mid, p2, right'))
 | 
|  |    182 |                   | Sprout (right1, q, right2) =>
 | 
|  |    183 |                       Sprout (Branch2 (left, p1, mid), p2, Branch2 (right1, q, right2)))));
 | 
| 5015 |    184 | 
 | 
| 15665 |    185 |   in
 | 
|  |    186 |     (case modfy tab of
 | 
|  |    187 |       Stay tab' => tab'
 | 
|  |    188 |     | Sprout br => Branch2 br)
 | 
|  |    189 |     handle SAME => tab
 | 
|  |    190 |   end;
 | 
| 5015 |    191 | 
 | 
| 15761 |    192 | fun update ((key, x), tab) = modify key (fn _ => x) tab;
 | 
|  |    193 | fun update_new ((key, x), tab) = modify key (fn NONE => x | SOME _ => raise DUP key) tab;
 | 
|  |    194 | fun map_entry key f = modify key (fn NONE => raise SAME | SOME x => f x);
 | 
| 5015 |    195 | 
 | 
|  |    196 | 
 | 
| 12287 |    197 | (* extend and make *)
 | 
| 5015 |    198 | 
 | 
| 15761 |    199 | fun extend (table, args) =
 | 
| 12287 |    200 |   let
 | 
| 15761 |    201 |     fun add (key, x) (tab, dups) =
 | 
| 16002 |    202 |       if is_some (lookup (tab, key)) then (tab, key :: dups)
 | 
| 15761 |    203 |       else (update ((key, x), tab), dups);
 | 
| 12287 |    204 |   in
 | 
| 15761 |    205 |     (case fold add args (table, []) of
 | 
| 12287 |    206 |       (table', []) => table'
 | 
|  |    207 |     | (_, dups) => raise DUPS (rev dups))
 | 
|  |    208 |   end;
 | 
|  |    209 | 
 | 
|  |    210 | fun make pairs = extend (empty, pairs);
 | 
|  |    211 | 
 | 
| 5015 |    212 | 
 | 
| 15014 |    213 | (* delete *)
 | 
|  |    214 | 
 | 
| 15665 |    215 | exception UNDEF of key;
 | 
|  |    216 | 
 | 
|  |    217 | local
 | 
|  |    218 | 
 | 
|  |    219 | fun compare NONE (k2, _) = LESS
 | 
|  |    220 |   | compare (SOME k1) (k2, _) = Key.ord (k1, k2);
 | 
| 15014 |    221 | 
 | 
|  |    222 | fun if_eq EQUAL x y = x
 | 
|  |    223 |   | if_eq _ x y = y;
 | 
|  |    224 | 
 | 
| 15531 |    225 | fun del (SOME k) Empty = raise UNDEF k
 | 
|  |    226 |   | del NONE (Branch2 (Empty, p, Empty)) = (p, (true, Empty))
 | 
|  |    227 |   | del NONE (Branch3 (Empty, p, Empty, q, Empty)) =
 | 
| 15014 |    228 |       (p, (false, Branch2 (Empty, q, Empty)))
 | 
| 15665 |    229 |   | del k (Branch2 (Empty, p, Empty)) = (case compare k p of
 | 
| 16002 |    230 |       EQUAL => (p, (true, Empty)) | _ => raise UNDEF (the k))
 | 
| 15665 |    231 |   | del k (Branch3 (Empty, p, Empty, q, Empty)) = (case compare k p of
 | 
| 15014 |    232 |       EQUAL => (p, (false, Branch2 (Empty, q, Empty)))
 | 
| 15665 |    233 |     | _ => (case compare k q of
 | 
| 15014 |    234 |         EQUAL => (q, (false, Branch2 (Empty, p, Empty)))
 | 
| 16002 |    235 |       | _ => raise UNDEF (the k)))
 | 
| 15665 |    236 |   | del k (Branch2 (l, p, r)) = (case compare k p of
 | 
| 15014 |    237 |       LESS => (case del k l of
 | 
|  |    238 |         (p', (false, l')) => (p', (false, Branch2 (l', p, r)))
 | 
|  |    239 |       | (p', (true, l')) => (p', case r of
 | 
|  |    240 |           Branch2 (rl, rp, rr) =>
 | 
|  |    241 |             (true, Branch3 (l', p, rl, rp, rr))
 | 
|  |    242 |         | Branch3 (rl, rp, rm, rq, rr) => (false, Branch2
 | 
|  |    243 |             (Branch2 (l', p, rl), rp, Branch2 (rm, rq, rr)))))
 | 
| 15531 |    244 |     | ord => (case del (if_eq ord NONE k) r of
 | 
| 15014 |    245 |         (p', (false, r')) => (p', (false, Branch2 (l, if_eq ord p' p, r')))
 | 
|  |    246 |       | (p', (true, r')) => (p', case l of
 | 
|  |    247 |           Branch2 (ll, lp, lr) =>
 | 
|  |    248 |             (true, Branch3 (ll, lp, lr, if_eq ord p' p, r'))
 | 
|  |    249 |         | Branch3 (ll, lp, lm, lq, lr) => (false, Branch2
 | 
|  |    250 |             (Branch2 (ll, lp, lm), lq, Branch2 (lr, if_eq ord p' p, r'))))))
 | 
| 15665 |    251 |   | del k (Branch3 (l, p, m, q, r)) = (case compare k q of
 | 
|  |    252 |       LESS => (case compare k p of
 | 
| 15014 |    253 |         LESS => (case del k l of
 | 
|  |    254 |           (p', (false, l')) => (p', (false, Branch3 (l', p, m, q, r)))
 | 
|  |    255 |         | (p', (true, l')) => (p', (false, case (m, r) of
 | 
|  |    256 |             (Branch2 (ml, mp, mr), Branch2 _) =>
 | 
|  |    257 |               Branch2 (Branch3 (l', p, ml, mp, mr), q, r)
 | 
|  |    258 |           | (Branch3 (ml, mp, mm, mq, mr), _) =>
 | 
|  |    259 |               Branch3 (Branch2 (l', p, ml), mp, Branch2 (mm, mq, mr), q, r)
 | 
|  |    260 |           | (Branch2 (ml, mp, mr), Branch3 (rl, rp, rm, rq, rr)) =>
 | 
|  |    261 |               Branch3 (Branch2 (l', p, ml), mp, Branch2 (mr, q, rl), rp,
 | 
|  |    262 |                 Branch2 (rm, rq, rr)))))
 | 
| 15531 |    263 |       | ord => (case del (if_eq ord NONE k) m of
 | 
| 15014 |    264 |           (p', (false, m')) =>
 | 
|  |    265 |             (p', (false, Branch3 (l, if_eq ord p' p, m', q, r)))
 | 
|  |    266 |         | (p', (true, m')) => (p', (false, case (l, r) of
 | 
|  |    267 |             (Branch2 (ll, lp, lr), Branch2 _) =>
 | 
|  |    268 |               Branch2 (Branch3 (ll, lp, lr, if_eq ord p' p, m'), q, r)
 | 
|  |    269 |           | (Branch3 (ll, lp, lm, lq, lr), _) =>
 | 
|  |    270 |               Branch3 (Branch2 (ll, lp, lm), lq,
 | 
|  |    271 |                 Branch2 (lr, if_eq ord p' p, m'), q, r)
 | 
|  |    272 |           | (_, Branch3 (rl, rp, rm, rq, rr)) =>
 | 
|  |    273 |               Branch3 (l, if_eq ord p' p, Branch2 (m', q, rl), rp,
 | 
|  |    274 |                 Branch2 (rm, rq, rr))))))
 | 
| 15531 |    275 |     | ord => (case del (if_eq ord NONE k) r of
 | 
| 15014 |    276 |         (q', (false, r')) =>
 | 
|  |    277 |           (q', (false, Branch3 (l, p, m, if_eq ord q' q, r')))
 | 
|  |    278 |       | (q', (true, r')) => (q', (false, case (l, m) of
 | 
|  |    279 |           (Branch2 _, Branch2 (ml, mp, mr)) =>
 | 
|  |    280 |             Branch2 (l, p, Branch3 (ml, mp, mr, if_eq ord q' q, r'))
 | 
|  |    281 |         | (_, Branch3 (ml, mp, mm, mq, mr)) =>
 | 
|  |    282 |             Branch3 (l, p, Branch2 (ml, mp, mm), mq,
 | 
|  |    283 |               Branch2 (mr, if_eq ord q' q, r'))
 | 
|  |    284 |         | (Branch3 (ll, lp, lm, lq, lr), Branch2 (ml, mp, mr)) =>
 | 
|  |    285 |             Branch3 (Branch2 (ll, lp, lm), lq, Branch2 (lr, p, ml), mp,
 | 
|  |    286 |               Branch2 (mr, if_eq ord q' q, r'))))));
 | 
|  |    287 | 
 | 
| 15665 |    288 | in
 | 
|  |    289 | 
 | 
| 15761 |    290 | fun delete key tab = snd (snd (del (SOME key) tab));
 | 
|  |    291 | fun delete_safe key tab = delete key tab handle UNDEF _ => tab;
 | 
| 15014 |    292 | 
 | 
| 15665 |    293 | end;
 | 
|  |    294 | 
 | 
| 15014 |    295 | 
 | 
| 16466 |    296 | (* member, insert and remove *)
 | 
|  |    297 | 
 | 
|  |    298 | fun member eq tab (key, x) =
 | 
|  |    299 |   (case lookup (tab, key) of
 | 
|  |    300 |     NONE => false
 | 
|  |    301 |   | SOME y => eq (x, y));
 | 
| 15761 |    302 | 
 | 
|  |    303 | fun insert eq (key, x) =
 | 
|  |    304 |   modify key (fn NONE => x | SOME y => if eq (x, y) then raise SAME else raise DUP key);
 | 
|  |    305 | 
 | 
|  |    306 | fun remove eq (key, x) tab =
 | 
|  |    307 |   (case lookup (tab, key) of
 | 
|  |    308 |     NONE => tab
 | 
|  |    309 |   | SOME y => if eq (x, y) then delete key tab else tab);
 | 
|  |    310 | 
 | 
|  |    311 | 
 | 
| 12287 |    312 | (* join and merge *)
 | 
| 5015 |    313 | 
 | 
| 12287 |    314 | fun join f (table1, table2) =
 | 
|  |    315 |   let
 | 
| 16446 |    316 |     fun add (key, x) (tab, dups) =
 | 
| 12287 |    317 |       (case lookup (tab, key) of
 | 
| 15531 |    318 |         NONE => (update ((key, x), tab), dups)
 | 
|  |    319 |       | SOME y =>
 | 
| 16446 |    320 |           (case f key (y, x) of
 | 
| 15531 |    321 |             SOME z => (update ((key, z), tab), dups)
 | 
|  |    322 |           | NONE => (tab, key :: dups)));
 | 
| 12287 |    323 |   in
 | 
| 16446 |    324 |     (case fold_table add table2 (table1, []) of
 | 
| 12287 |    325 |       (table', []) => table'
 | 
|  |    326 |     | (_, dups) => raise DUPS (rev dups))
 | 
|  |    327 |   end;
 | 
|  |    328 | 
 | 
| 16446 |    329 | fun merge eq tabs = join (fn _ => fn (y, x) => if eq (y, x) then SOME y else NONE) tabs;
 | 
| 5015 |    330 | 
 | 
|  |    331 | 
 | 
| 15761 |    332 | (* tables with multiple entries per key *)
 | 
|  |    333 | 
 | 
| 16002 |    334 | fun lookup_multi arg = if_none (lookup arg) [];
 | 
| 15761 |    335 | 
 | 
|  |    336 | fun update_multi ((key, x), tab) = modify key (fn NONE => [x] | SOME xs => x :: xs) tab;
 | 
| 5015 |    337 | 
 | 
| 15761 |    338 | fun remove_multi eq (key, x) tab =
 | 
|  |    339 |   map_entry key (fn xs => (case Library.remove eq x xs of [] => raise UNDEF key | ys => ys)) tab
 | 
|  |    340 |   handle UNDEF _ => delete key tab;
 | 
| 5015 |    341 | 
 | 
| 15761 |    342 | fun make_multi args = foldr update_multi empty args;
 | 
|  |    343 | fun dest_multi tab = List.concat (map (fn (key, xs) => map (pair key) xs) (dest tab));
 | 
| 16446 |    344 | fun merge_multi eq = join (fn _ => fn (xs, xs') => SOME (gen_merge_lists eq xs xs'));
 | 
|  |    345 | fun merge_multi' eq = join (fn _ => fn (xs, xs') => SOME (gen_merge_lists' eq xs xs'));
 | 
| 5015 |    346 | 
 | 
|  |    347 | 
 | 
| 5681 |    348 | (*final declarations of this structure!*)
 | 
| 16446 |    349 | fun map f = map_table (K f);
 | 
|  |    350 | val map' = map_table;
 | 
|  |    351 | val fold = fold_table;
 | 
|  |    352 | fun foldl f (x, tab) = fold (fn p => fn x' => f (x', p)) tab x;
 | 
| 5015 |    353 | 
 | 
|  |    354 | end;
 | 
|  |    355 | 
 | 
| 16342 |    356 | structure Inttab = TableFun(type key = int val ord = int_ord);
 | 
| 5015 |    357 | structure Symtab = TableFun(type key = string val ord = string_ord);
 |