| author | hoelzl | 
| Tue, 05 Jan 2016 13:35:06 +0100 | |
| changeset 62055 | 755fda743c49 | 
| parent 59012 | f4e9bd04e1d5 | 
| child 62503 | 19afb533028e | 
| permissions | -rw-r--r-- | 
| 6118 | 1  | 
(* Title: Pure/General/table.ML  | 
| 15014 | 2  | 
Author: Markus Wenzel and Stefan Berghofer, TU Muenchen  | 
| 5015 | 3  | 
|
| 16342 | 4  | 
Generic tables. Efficient purely functional implementation using  | 
5  | 
balanced 2-3 trees.  | 
|
| 5015 | 6  | 
*)  | 
7  | 
||
8  | 
signature KEY =  | 
|
9  | 
sig  | 
|
10  | 
type key  | 
|
11  | 
val ord: key * key -> order  | 
|
12  | 
end;  | 
|
13  | 
||
14  | 
signature TABLE =  | 
|
15  | 
sig  | 
|
16  | 
type key  | 
|
17  | 
type 'a table  | 
|
18  | 
exception DUP of key  | 
|
| 19031 | 19  | 
exception SAME  | 
| 15014 | 20  | 
exception UNDEF of key  | 
| 5015 | 21  | 
val empty: 'a table  | 
22  | 
val is_empty: 'a table -> bool  | 
|
| 39020 | 23  | 
val map: (key -> 'a -> 'b) -> 'a table -> 'b table  | 
| 16446 | 24  | 
val fold: (key * 'b -> 'a -> 'a) -> 'b table -> 'a -> 'a  | 
| 28334 | 25  | 
val fold_rev: (key * 'b -> 'a -> 'a) -> 'b table -> 'a -> 'a  | 
| 5015 | 26  | 
val dest: 'a table -> (key * 'a) list  | 
| 5681 | 27  | 
val keys: 'a table -> key list  | 
| 52049 | 28  | 
val min: 'a table -> (key * 'a) option  | 
29  | 
val max: 'a table -> (key * 'a) option  | 
|
| 
35012
 
c3e3ac3ca091
removed unused "boundary" of Table/Graph.get_first;
 
wenzelm 
parents: 
33162 
diff
changeset
 | 
30  | 
val get_first: (key * 'a -> 'b option) -> 'a table -> 'b option  | 
| 27508 | 31  | 
val exists: (key * 'a -> bool) -> 'a table -> bool  | 
32  | 
val forall: (key * 'a -> bool) -> 'a table -> bool  | 
|
| 
43792
 
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
 
wenzelm 
parents: 
39020 
diff
changeset
 | 
33  | 
val lookup_key: 'a table -> key -> (key * 'a) option  | 
| 31616 | 34  | 
val lookup: 'a table -> key -> 'a option  | 
| 16887 | 35  | 
val defined: 'a table -> key -> bool  | 
| 31209 | 36  | 
val update: key * 'a -> 'a table -> 'a table  | 
37  | 
val update_new: key * 'a -> 'a table -> 'a table (*exception DUP*)  | 
|
| 17179 | 38  | 
val default: key * 'a -> 'a table -> 'a table  | 
| 56051 | 39  | 
  val map_entry: key -> ('a -> 'a) (*exception SAME*) -> 'a table -> 'a table
 | 
| 31209 | 40  | 
  val map_default: key * 'a -> ('a -> 'a) -> 'a table -> 'a table
 | 
| 
23655
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
21065 
diff
changeset
 | 
41  | 
val make: (key * 'a) list -> 'a table (*exception DUP*)  | 
| 56051 | 42  | 
val join: (key -> 'a * 'a -> 'a) (*exception SAME*) ->  | 
| 
23655
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
21065 
diff
changeset
 | 
43  | 
'a table * 'a table -> 'a table (*exception DUP*)  | 
| 
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
21065 
diff
changeset
 | 
44  | 
  val merge: ('a * 'a -> bool) -> 'a table * 'a table -> 'a table      (*exception DUP*)
 | 
| 15665 | 45  | 
val delete: key -> 'a table -> 'a table (*exception UNDEF*)  | 
| 15761 | 46  | 
val delete_safe: key -> 'a table -> 'a table  | 
| 16466 | 47  | 
  val member: ('b * 'a -> bool) -> 'a table -> key * 'b -> bool
 | 
| 15761 | 48  | 
  val insert: ('a * 'a -> bool) -> key * 'a -> 'a table -> 'a table    (*exception DUP*)
 | 
| 16139 | 49  | 
  val remove: ('b * 'a -> bool) -> key * 'b -> 'a table -> 'a table
 | 
| 18946 | 50  | 
val lookup_list: 'a list table -> key -> 'a list  | 
| 31209 | 51  | 
val cons_list: key * 'a -> 'a list table -> 'a list table  | 
| 19506 | 52  | 
  val insert_list: ('a * 'a -> bool) -> key * 'a -> 'a list table -> 'a list table
 | 
| 18946 | 53  | 
  val remove_list: ('b * 'a -> bool) -> key * 'b -> 'a list table -> 'a list table
 | 
| 25391 | 54  | 
  val update_list: ('a * 'a -> bool) -> key * 'a -> 'a list table -> 'a list table
 | 
| 18946 | 55  | 
val make_list: (key * 'a) list -> 'a list table  | 
56  | 
val dest_list: 'a list table -> (key * 'a) list  | 
|
| 33162 | 57  | 
  val merge_list: ('a * 'a -> bool) -> 'a list table * 'a list table -> 'a list table
 | 
| 50234 | 58  | 
type set = unit table  | 
| 
59012
 
f4e9bd04e1d5
clarified Table.make_set: duplicate arguments are allowed, like Table.make_list or Scala Set() formation;
 
wenzelm 
parents: 
56051 
diff
changeset
 | 
59  | 
val make_set: key list -> set  | 
| 5015 | 60  | 
end;  | 
61  | 
||
| 
31971
 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 
wenzelm 
parents: 
31621 
diff
changeset
 | 
62  | 
functor Table(Key: KEY): TABLE =  | 
| 5015 | 63  | 
struct  | 
64  | 
||
65  | 
||
66  | 
(* datatype table *)  | 
|
67  | 
||
68  | 
type key = Key.key;  | 
|
69  | 
||
70  | 
datatype 'a table =  | 
|
71  | 
Empty |  | 
|
72  | 
Branch2 of 'a table * (key * 'a) * 'a table |  | 
|
73  | 
Branch3 of 'a table * (key * 'a) * 'a table * (key * 'a) * 'a table;  | 
|
74  | 
||
75  | 
exception DUP of key;  | 
|
76  | 
||
77  | 
||
| 5681 | 78  | 
(* empty *)  | 
79  | 
||
| 5015 | 80  | 
val empty = Empty;  | 
81  | 
||
82  | 
fun is_empty Empty = true  | 
|
83  | 
| is_empty _ = false;  | 
|
84  | 
||
| 5681 | 85  | 
|
86  | 
(* map and fold combinators *)  | 
|
87  | 
||
| 16657 | 88  | 
fun map_table f =  | 
89  | 
let  | 
|
90  | 
fun map Empty = Empty  | 
|
91  | 
| map (Branch2 (left, (k, x), right)) =  | 
|
92  | 
Branch2 (map left, (k, f k x), map right)  | 
|
93  | 
| map (Branch3 (left, (k1, x1), mid, (k2, x2), right)) =  | 
|
94  | 
Branch3 (map left, (k1, f k1 x1), map mid, (k2, f k2 x2), map right);  | 
|
95  | 
in map end;  | 
|
| 5681 | 96  | 
|
| 16657 | 97  | 
fun fold_table f =  | 
98  | 
let  | 
|
99  | 
fun fold Empty x = x  | 
|
100  | 
| fold (Branch2 (left, p, right)) x =  | 
|
101  | 
fold right (f p (fold left x))  | 
|
102  | 
| fold (Branch3 (left, p1, mid, p2, right)) x =  | 
|
103  | 
fold right (f p2 (fold mid (f p1 (fold left x))));  | 
|
104  | 
in fold end;  | 
|
| 5681 | 105  | 
|
| 28334 | 106  | 
fun fold_rev_table f =  | 
107  | 
let  | 
|
108  | 
fun fold Empty x = x  | 
|
109  | 
| fold (Branch2 (left, p, right)) x =  | 
|
110  | 
fold left (f p (fold right x))  | 
|
111  | 
| fold (Branch3 (left, p1, mid, p2, right)) x =  | 
|
112  | 
fold left (f p1 (fold mid (f p2 (fold right x))));  | 
|
113  | 
in fold end;  | 
|
114  | 
||
115  | 
fun dest tab = fold_rev_table cons tab [];  | 
|
116  | 
fun keys tab = fold_rev_table (cons o #1) tab [];  | 
|
| 16192 | 117  | 
|
| 27508 | 118  | 
|
| 52049 | 119  | 
(* min/max entries *)  | 
| 5015 | 120  | 
|
| 52049 | 121  | 
fun min Empty = NONE  | 
122  | 
| min (Branch2 (Empty, p, _)) = SOME p  | 
|
123  | 
| min (Branch3 (Empty, p, _, _, _)) = SOME p  | 
|
124  | 
| min (Branch2 (left, _, _)) = min left  | 
|
125  | 
| min (Branch3 (left, _, _, _, _)) = min left;  | 
|
| 8409 | 126  | 
|
| 52049 | 127  | 
fun max Empty = NONE  | 
128  | 
| max (Branch2 (_, p, Empty)) = SOME p  | 
|
129  | 
| max (Branch3 (_, _, _, p, Empty)) = SOME p  | 
|
130  | 
| max (Branch2 (_, _, right)) = max right  | 
|
131  | 
| max (Branch3 (_, _, _, _, right)) = max right;  | 
|
| 15665 | 132  | 
|
| 5015 | 133  | 
|
| 31616 | 134  | 
(* get_first *)  | 
135  | 
||
| 
35012
 
c3e3ac3ca091
removed unused "boundary" of Table/Graph.get_first;
 
wenzelm 
parents: 
33162 
diff
changeset
 | 
136  | 
fun get_first f =  | 
| 31616 | 137  | 
let  | 
| 
35012
 
c3e3ac3ca091
removed unused "boundary" of Table/Graph.get_first;
 
wenzelm 
parents: 
33162 
diff
changeset
 | 
138  | 
fun get Empty = NONE  | 
| 
31618
 
2e4430b84303
improved get_first: check boundary before entering subtrees;
 
wenzelm 
parents: 
31616 
diff
changeset
 | 
139  | 
| get (Branch2 (left, (k, x), right)) =  | 
| 
35012
 
c3e3ac3ca091
removed unused "boundary" of Table/Graph.get_first;
 
wenzelm 
parents: 
33162 
diff
changeset
 | 
140  | 
(case get left of  | 
| 31616 | 141  | 
NONE =>  | 
| 
35012
 
c3e3ac3ca091
removed unused "boundary" of Table/Graph.get_first;
 
wenzelm 
parents: 
33162 
diff
changeset
 | 
142  | 
(case f (k, x) of  | 
| 31616 | 143  | 
NONE => get right  | 
144  | 
| some => some)  | 
|
145  | 
| some => some)  | 
|
| 
31618
 
2e4430b84303
improved get_first: check boundary before entering subtrees;
 
wenzelm 
parents: 
31616 
diff
changeset
 | 
146  | 
| get (Branch3 (left, (k1, x1), mid, (k2, x2), right)) =  | 
| 
35012
 
c3e3ac3ca091
removed unused "boundary" of Table/Graph.get_first;
 
wenzelm 
parents: 
33162 
diff
changeset
 | 
147  | 
(case get left of  | 
| 31616 | 148  | 
NONE =>  | 
| 
35012
 
c3e3ac3ca091
removed unused "boundary" of Table/Graph.get_first;
 
wenzelm 
parents: 
33162 
diff
changeset
 | 
149  | 
(case f (k1, x1) of  | 
| 31616 | 150  | 
NONE =>  | 
| 
35012
 
c3e3ac3ca091
removed unused "boundary" of Table/Graph.get_first;
 
wenzelm 
parents: 
33162 
diff
changeset
 | 
151  | 
(case get mid of  | 
| 31616 | 152  | 
NONE =>  | 
| 
35012
 
c3e3ac3ca091
removed unused "boundary" of Table/Graph.get_first;
 
wenzelm 
parents: 
33162 
diff
changeset
 | 
153  | 
(case f (k2, x2) of  | 
| 31616 | 154  | 
NONE => get right  | 
155  | 
| some => some)  | 
|
156  | 
| some => some)  | 
|
157  | 
| some => some)  | 
|
158  | 
| some => some);  | 
|
| 
35012
 
c3e3ac3ca091
removed unused "boundary" of Table/Graph.get_first;
 
wenzelm 
parents: 
33162 
diff
changeset
 | 
159  | 
in get end;  | 
| 31616 | 160  | 
|
| 
35012
 
c3e3ac3ca091
removed unused "boundary" of Table/Graph.get_first;
 
wenzelm 
parents: 
33162 
diff
changeset
 | 
161  | 
fun exists pred = is_some o get_first (fn entry => if pred entry then SOME () else NONE);  | 
| 31616 | 162  | 
fun forall pred = not o exists (not o pred);  | 
163  | 
||
164  | 
||
| 5015 | 165  | 
(* lookup *)  | 
166  | 
||
| 
43792
 
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
 
wenzelm 
parents: 
39020 
diff
changeset
 | 
167  | 
fun lookup_key tab key =  | 
| 19031 | 168  | 
let  | 
169  | 
fun look Empty = NONE  | 
|
170  | 
| look (Branch2 (left, (k, x), right)) =  | 
|
171  | 
(case Key.ord (key, k) of  | 
|
172  | 
LESS => look left  | 
|
| 
43792
 
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
 
wenzelm 
parents: 
39020 
diff
changeset
 | 
173  | 
| EQUAL => SOME (k, x)  | 
| 19031 | 174  | 
| GREATER => look right)  | 
175  | 
| look (Branch3 (left, (k1, x1), mid, (k2, x2), right)) =  | 
|
176  | 
(case Key.ord (key, k1) of  | 
|
177  | 
LESS => look left  | 
|
| 
43792
 
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
 
wenzelm 
parents: 
39020 
diff
changeset
 | 
178  | 
| EQUAL => SOME (k1, x1)  | 
| 19031 | 179  | 
| GREATER =>  | 
180  | 
(case Key.ord (key, k2) of  | 
|
181  | 
LESS => look mid  | 
|
| 
43792
 
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
 
wenzelm 
parents: 
39020 
diff
changeset
 | 
182  | 
| EQUAL => SOME (k2, x2)  | 
| 19031 | 183  | 
| GREATER => look right));  | 
184  | 
in look tab end;  | 
|
| 5015 | 185  | 
|
| 
43792
 
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
 
wenzelm 
parents: 
39020 
diff
changeset
 | 
186  | 
fun lookup tab key = Option.map #2 (lookup_key tab key);  | 
| 
 
d5803c3d537a
Table.lookup_key and Graph.get_entry allow to retrieve the original key, which is not necessarily identical to the given one;
 
wenzelm 
parents: 
39020 
diff
changeset
 | 
187  | 
|
| 19031 | 188  | 
fun defined tab key =  | 
189  | 
let  | 
|
190  | 
fun def Empty = false  | 
|
191  | 
| def (Branch2 (left, (k, x), right)) =  | 
|
192  | 
(case Key.ord (key, k) of  | 
|
193  | 
LESS => def left  | 
|
194  | 
| EQUAL => true  | 
|
195  | 
| GREATER => def right)  | 
|
196  | 
| def (Branch3 (left, (k1, x1), mid, (k2, x2), right)) =  | 
|
197  | 
(case Key.ord (key, k1) of  | 
|
198  | 
LESS => def left  | 
|
199  | 
| EQUAL => true  | 
|
200  | 
| GREATER =>  | 
|
201  | 
(case Key.ord (key, k2) of  | 
|
202  | 
LESS => def mid  | 
|
203  | 
| EQUAL => true  | 
|
204  | 
| GREATER => def right));  | 
|
205  | 
in def tab end;  | 
|
| 16887 | 206  | 
|
| 5015 | 207  | 
|
| 19031 | 208  | 
(* modify *)  | 
| 5015 | 209  | 
|
210  | 
datatype 'a growth =  | 
|
211  | 
Stay of 'a table |  | 
|
212  | 
Sprout of 'a table * (key * 'a) * 'a table;  | 
|
213  | 
||
| 15761 | 214  | 
exception SAME;  | 
215  | 
||
| 15665 | 216  | 
fun modify key f tab =  | 
217  | 
let  | 
|
218  | 
fun modfy Empty = Sprout (Empty, (key, f NONE), Empty)  | 
|
219  | 
| modfy (Branch2 (left, p as (k, x), right)) =  | 
|
220  | 
(case Key.ord (key, k) of  | 
|
221  | 
LESS =>  | 
|
222  | 
(case modfy left of  | 
|
223  | 
Stay left' => Stay (Branch2 (left', p, right))  | 
|
224  | 
| Sprout (left1, q, left2) => Stay (Branch3 (left1, q, left2, p, right)))  | 
|
225  | 
| EQUAL => Stay (Branch2 (left, (k, f (SOME x)), right))  | 
|
226  | 
| GREATER =>  | 
|
227  | 
(case modfy right of  | 
|
228  | 
Stay right' => Stay (Branch2 (left, p, right'))  | 
|
229  | 
| Sprout (right1, q, right2) =>  | 
|
230  | 
Stay (Branch3 (left, p, right1, q, right2))))  | 
|
231  | 
| modfy (Branch3 (left, p1 as (k1, x1), mid, p2 as (k2, x2), right)) =  | 
|
232  | 
(case Key.ord (key, k1) of  | 
|
| 5015 | 233  | 
LESS =>  | 
| 15665 | 234  | 
(case modfy left of  | 
235  | 
Stay left' => Stay (Branch3 (left', p1, mid, p2, right))  | 
|
236  | 
| Sprout (left1, q, left2) =>  | 
|
237  | 
Sprout (Branch2 (left1, q, left2), p1, Branch2 (mid, p2, right)))  | 
|
238  | 
| EQUAL => Stay (Branch3 (left, (k1, f (SOME x1)), mid, p2, right))  | 
|
| 5015 | 239  | 
| GREATER =>  | 
| 15665 | 240  | 
(case Key.ord (key, k2) of  | 
241  | 
LESS =>  | 
|
242  | 
(case modfy mid of  | 
|
243  | 
Stay mid' => Stay (Branch3 (left, p1, mid', p2, right))  | 
|
244  | 
| Sprout (mid1, q, mid2) =>  | 
|
245  | 
Sprout (Branch2 (left, p1, mid1), q, Branch2 (mid2, p2, right)))  | 
|
246  | 
| EQUAL => Stay (Branch3 (left, p1, mid, (k2, f (SOME x2)), right))  | 
|
247  | 
| GREATER =>  | 
|
248  | 
(case modfy right of  | 
|
249  | 
Stay right' => Stay (Branch3 (left, p1, mid, p2, right'))  | 
|
250  | 
| Sprout (right1, q, right2) =>  | 
|
251  | 
Sprout (Branch2 (left, p1, mid), p2, Branch2 (right1, q, right2)))));  | 
|
| 5015 | 252  | 
|
| 15665 | 253  | 
in  | 
254  | 
(case modfy tab of  | 
|
255  | 
Stay tab' => tab'  | 
|
256  | 
| Sprout br => Branch2 br)  | 
|
257  | 
handle SAME => tab  | 
|
258  | 
end;  | 
|
| 5015 | 259  | 
|
| 17412 | 260  | 
fun update (key, x) tab = modify key (fn _ => x) tab;  | 
261  | 
fun update_new (key, x) tab = modify key (fn NONE => x | SOME _ => raise DUP key) tab;  | 
|
| 17709 | 262  | 
fun default (key, x) tab = modify key (fn NONE => x | SOME _ => raise SAME) tab;  | 
| 15761 | 263  | 
fun map_entry key f = modify key (fn NONE => raise SAME | SOME x => f x);  | 
| 27783 | 264  | 
fun map_default (key, x) f = modify key (fn NONE => f x | SOME y => f y);  | 
| 5015 | 265  | 
|
266  | 
||
| 15014 | 267  | 
(* delete *)  | 
268  | 
||
| 15665 | 269  | 
exception UNDEF of key;  | 
270  | 
||
271  | 
local  | 
|
272  | 
||
273  | 
fun compare NONE (k2, _) = LESS  | 
|
274  | 
| compare (SOME k1) (k2, _) = Key.ord (k1, k2);  | 
|
| 15014 | 275  | 
|
276  | 
fun if_eq EQUAL x y = x  | 
|
277  | 
| if_eq _ x y = y;  | 
|
278  | 
||
| 15531 | 279  | 
fun del (SOME k) Empty = raise UNDEF k  | 
280  | 
| del NONE (Branch2 (Empty, p, Empty)) = (p, (true, Empty))  | 
|
281  | 
| del NONE (Branch3 (Empty, p, Empty, q, Empty)) =  | 
|
| 15014 | 282  | 
(p, (false, Branch2 (Empty, q, Empty)))  | 
| 15665 | 283  | 
| del k (Branch2 (Empty, p, Empty)) = (case compare k p of  | 
| 16002 | 284  | 
EQUAL => (p, (true, Empty)) | _ => raise UNDEF (the k))  | 
| 15665 | 285  | 
| del k (Branch3 (Empty, p, Empty, q, Empty)) = (case compare k p of  | 
| 15014 | 286  | 
EQUAL => (p, (false, Branch2 (Empty, q, Empty)))  | 
| 15665 | 287  | 
| _ => (case compare k q of  | 
| 15014 | 288  | 
EQUAL => (q, (false, Branch2 (Empty, p, Empty)))  | 
| 16002 | 289  | 
| _ => raise UNDEF (the k)))  | 
| 15665 | 290  | 
| del k (Branch2 (l, p, r)) = (case compare k p of  | 
| 15014 | 291  | 
LESS => (case del k l of  | 
292  | 
(p', (false, l')) => (p', (false, Branch2 (l', p, r)))  | 
|
293  | 
| (p', (true, l')) => (p', case r of  | 
|
294  | 
Branch2 (rl, rp, rr) =>  | 
|
295  | 
(true, Branch3 (l', p, rl, rp, rr))  | 
|
296  | 
| Branch3 (rl, rp, rm, rq, rr) => (false, Branch2  | 
|
297  | 
(Branch2 (l', p, rl), rp, Branch2 (rm, rq, rr)))))  | 
|
| 15531 | 298  | 
| ord => (case del (if_eq ord NONE k) r of  | 
| 15014 | 299  | 
(p', (false, r')) => (p', (false, Branch2 (l, if_eq ord p' p, r')))  | 
300  | 
| (p', (true, r')) => (p', case l of  | 
|
301  | 
Branch2 (ll, lp, lr) =>  | 
|
302  | 
(true, Branch3 (ll, lp, lr, if_eq ord p' p, r'))  | 
|
303  | 
| Branch3 (ll, lp, lm, lq, lr) => (false, Branch2  | 
|
304  | 
(Branch2 (ll, lp, lm), lq, Branch2 (lr, if_eq ord p' p, r'))))))  | 
|
| 15665 | 305  | 
| del k (Branch3 (l, p, m, q, r)) = (case compare k q of  | 
306  | 
LESS => (case compare k p of  | 
|
| 15014 | 307  | 
LESS => (case del k l of  | 
308  | 
(p', (false, l')) => (p', (false, Branch3 (l', p, m, q, r)))  | 
|
309  | 
| (p', (true, l')) => (p', (false, case (m, r) of  | 
|
310  | 
(Branch2 (ml, mp, mr), Branch2 _) =>  | 
|
311  | 
Branch2 (Branch3 (l', p, ml, mp, mr), q, r)  | 
|
312  | 
| (Branch3 (ml, mp, mm, mq, mr), _) =>  | 
|
313  | 
Branch3 (Branch2 (l', p, ml), mp, Branch2 (mm, mq, mr), q, r)  | 
|
314  | 
| (Branch2 (ml, mp, mr), Branch3 (rl, rp, rm, rq, rr)) =>  | 
|
315  | 
Branch3 (Branch2 (l', p, ml), mp, Branch2 (mr, q, rl), rp,  | 
|
316  | 
Branch2 (rm, rq, rr)))))  | 
|
| 15531 | 317  | 
| ord => (case del (if_eq ord NONE k) m of  | 
| 15014 | 318  | 
(p', (false, m')) =>  | 
319  | 
(p', (false, Branch3 (l, if_eq ord p' p, m', q, r)))  | 
|
320  | 
| (p', (true, m')) => (p', (false, case (l, r) of  | 
|
321  | 
(Branch2 (ll, lp, lr), Branch2 _) =>  | 
|
322  | 
Branch2 (Branch3 (ll, lp, lr, if_eq ord p' p, m'), q, r)  | 
|
323  | 
| (Branch3 (ll, lp, lm, lq, lr), _) =>  | 
|
324  | 
Branch3 (Branch2 (ll, lp, lm), lq,  | 
|
325  | 
Branch2 (lr, if_eq ord p' p, m'), q, r)  | 
|
326  | 
| (_, Branch3 (rl, rp, rm, rq, rr)) =>  | 
|
327  | 
Branch3 (l, if_eq ord p' p, Branch2 (m', q, rl), rp,  | 
|
328  | 
Branch2 (rm, rq, rr))))))  | 
|
| 15531 | 329  | 
| ord => (case del (if_eq ord NONE k) r of  | 
| 15014 | 330  | 
(q', (false, r')) =>  | 
331  | 
(q', (false, Branch3 (l, p, m, if_eq ord q' q, r')))  | 
|
332  | 
| (q', (true, r')) => (q', (false, case (l, m) of  | 
|
333  | 
(Branch2 _, Branch2 (ml, mp, mr)) =>  | 
|
334  | 
Branch2 (l, p, Branch3 (ml, mp, mr, if_eq ord q' q, r'))  | 
|
335  | 
| (_, Branch3 (ml, mp, mm, mq, mr)) =>  | 
|
336  | 
Branch3 (l, p, Branch2 (ml, mp, mm), mq,  | 
|
337  | 
Branch2 (mr, if_eq ord q' q, r'))  | 
|
338  | 
| (Branch3 (ll, lp, lm, lq, lr), Branch2 (ml, mp, mr)) =>  | 
|
339  | 
Branch3 (Branch2 (ll, lp, lm), lq, Branch2 (lr, p, ml), mp,  | 
|
340  | 
Branch2 (mr, if_eq ord q' q, r'))))));  | 
|
341  | 
||
| 15665 | 342  | 
in  | 
343  | 
||
| 15761 | 344  | 
fun delete key tab = snd (snd (del (SOME key) tab));  | 
| 
44336
 
59ff5a93eef4
tuned Table.delete_safe: avoid potentially expensive attempt of delete;
 
wenzelm 
parents: 
43792 
diff
changeset
 | 
345  | 
fun delete_safe key tab = if defined tab key then delete key tab else tab;  | 
| 15014 | 346  | 
|
| 15665 | 347  | 
end;  | 
348  | 
||
| 15014 | 349  | 
|
| 19031 | 350  | 
(* membership operations *)  | 
| 16466 | 351  | 
|
352  | 
fun member eq tab (key, x) =  | 
|
| 17412 | 353  | 
(case lookup tab key of  | 
| 16466 | 354  | 
NONE => false  | 
355  | 
| SOME y => eq (x, y));  | 
|
| 15761 | 356  | 
|
357  | 
fun insert eq (key, x) =  | 
|
358  | 
modify key (fn NONE => x | SOME y => if eq (x, y) then raise SAME else raise DUP key);  | 
|
359  | 
||
360  | 
fun remove eq (key, x) tab =  | 
|
| 17412 | 361  | 
(case lookup tab key of  | 
| 15761 | 362  | 
NONE => tab  | 
363  | 
| SOME y => if eq (x, y) then delete key tab else tab);  | 
|
364  | 
||
365  | 
||
| 19031 | 366  | 
(* simultaneous modifications *)  | 
367  | 
||
| 29004 | 368  | 
fun make entries = fold update_new entries empty;  | 
| 5015 | 369  | 
|
| 12287 | 370  | 
fun join f (table1, table2) =  | 
| 
55727
 
7e330ae052bb
optimize special case according to Library.merge (see also 8fbc355100f2, 520872460b7b);
 
wenzelm 
parents: 
52049 
diff
changeset
 | 
371  | 
let  | 
| 
 
7e330ae052bb
optimize special case according to Library.merge (see also 8fbc355100f2, 520872460b7b);
 
wenzelm 
parents: 
52049 
diff
changeset
 | 
372  | 
fun add (key, y) tab = modify key (fn NONE => y | SOME x => f key (x, y)) tab;  | 
| 
 
7e330ae052bb
optimize special case according to Library.merge (see also 8fbc355100f2, 520872460b7b);
 
wenzelm 
parents: 
52049 
diff
changeset
 | 
373  | 
in  | 
| 
 
7e330ae052bb
optimize special case according to Library.merge (see also 8fbc355100f2, 520872460b7b);
 
wenzelm 
parents: 
52049 
diff
changeset
 | 
374  | 
if pointer_eq (table1, table2) then table1  | 
| 
 
7e330ae052bb
optimize special case according to Library.merge (see also 8fbc355100f2, 520872460b7b);
 
wenzelm 
parents: 
52049 
diff
changeset
 | 
375  | 
else if is_empty table1 then table2  | 
| 
 
7e330ae052bb
optimize special case according to Library.merge (see also 8fbc355100f2, 520872460b7b);
 
wenzelm 
parents: 
52049 
diff
changeset
 | 
376  | 
else fold_table add table2 table1  | 
| 
 
7e330ae052bb
optimize special case according to Library.merge (see also 8fbc355100f2, 520872460b7b);
 
wenzelm 
parents: 
52049 
diff
changeset
 | 
377  | 
end;  | 
| 12287 | 378  | 
|
| 19031 | 379  | 
fun merge eq = join (fn key => fn xy => if eq xy then raise SAME else raise DUP key);  | 
| 5015 | 380  | 
|
381  | 
||
| 19031 | 382  | 
(* list tables *)  | 
| 15761 | 383  | 
|
| 18946 | 384  | 
fun lookup_list tab key = these (lookup tab key);  | 
| 25391 | 385  | 
|
386  | 
fun cons_list (key, x) tab = modify key (fn NONE => [x] | SOME xs => x :: xs) tab;  | 
|
| 5015 | 387  | 
|
| 20124 | 388  | 
fun insert_list eq (key, x) =  | 
389  | 
modify key (fn NONE => [x] | SOME xs => if Library.member eq xs x then raise SAME else x :: xs);  | 
|
| 25391 | 390  | 
|
| 18946 | 391  | 
fun remove_list eq (key, x) tab =  | 
| 15761 | 392  | 
map_entry key (fn xs => (case Library.remove eq x xs of [] => raise UNDEF key | ys => ys)) tab  | 
393  | 
handle UNDEF _ => delete key tab;  | 
|
| 5015 | 394  | 
|
| 25391 | 395  | 
fun update_list eq (key, x) =  | 
396  | 
modify key (fn NONE => [x] | SOME [] => [x] | SOME (xs as y :: _) =>  | 
|
397  | 
if eq (x, y) then raise SAME else Library.update eq x xs);  | 
|
398  | 
||
399  | 
fun make_list args = fold_rev cons_list args empty;  | 
|
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19073 
diff
changeset
 | 
400  | 
fun dest_list tab = maps (fn (key, xs) => map (pair key) xs) (dest tab);  | 
| 19031 | 401  | 
fun merge_list eq = join (fn _ => Library.merge eq);  | 
| 5015 | 402  | 
|
403  | 
||
| 50234 | 404  | 
(* unit tables *)  | 
405  | 
||
406  | 
type set = unit table;  | 
|
407  | 
||
| 
59012
 
f4e9bd04e1d5
clarified Table.make_set: duplicate arguments are allowed, like Table.make_list or Scala Set() formation;
 
wenzelm 
parents: 
56051 
diff
changeset
 | 
408  | 
fun make_set entries = fold (fn x => update (x, ())) entries empty;  | 
| 50234 | 409  | 
|
410  | 
||
| 
47980
 
c81801f881b3
simplified Poly/ML setup -- 5.3.0 is now the common base-line;
 
wenzelm 
parents: 
44336 
diff
changeset
 | 
411  | 
(* ML pretty-printing *)  | 
| 
38635
 
f76ad0771f67
added ML toplevel pretty-printing for tables, using dummy for anything other than Poly/ML 5.3.0 (or later);
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
412  | 
|
| 
 
f76ad0771f67
added ML toplevel pretty-printing for tables, using dummy for anything other than Poly/ML 5.3.0 (or later);
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
413  | 
val _ =  | 
| 
 
f76ad0771f67
added ML toplevel pretty-printing for tables, using dummy for anything other than Poly/ML 5.3.0 (or later);
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
414  | 
PolyML.addPrettyPrinter (fn depth => fn pretty => fn tab =>  | 
| 
 
f76ad0771f67
added ML toplevel pretty-printing for tables, using dummy for anything other than Poly/ML 5.3.0 (or later);
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
415  | 
ml_pretty  | 
| 
 
f76ad0771f67
added ML toplevel pretty-printing for tables, using dummy for anything other than Poly/ML 5.3.0 (or later);
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
416  | 
      (ML_Pretty.enum "," "{" "}"
 | 
| 
 
f76ad0771f67
added ML toplevel pretty-printing for tables, using dummy for anything other than Poly/ML 5.3.0 (or later);
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
417  | 
(ML_Pretty.pair (pretty_ml o PolyML.prettyRepresentation) (pretty_ml o pretty))  | 
| 
 
f76ad0771f67
added ML toplevel pretty-printing for tables, using dummy for anything other than Poly/ML 5.3.0 (or later);
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
418  | 
(dest tab, depth)));  | 
| 
 
f76ad0771f67
added ML toplevel pretty-printing for tables, using dummy for anything other than Poly/ML 5.3.0 (or later);
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
419  | 
|
| 
 
f76ad0771f67
added ML toplevel pretty-printing for tables, using dummy for anything other than Poly/ML 5.3.0 (or later);
 
wenzelm 
parents: 
35012 
diff
changeset
 | 
420  | 
|
| 5681 | 421  | 
(*final declarations of this structure!*)  | 
| 39020 | 422  | 
val map = map_table;  | 
| 16446 | 423  | 
val fold = fold_table;  | 
| 28334 | 424  | 
val fold_rev = fold_rev_table;  | 
| 5015 | 425  | 
|
426  | 
end;  | 
|
427  | 
||
| 
31971
 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 
wenzelm 
parents: 
31621 
diff
changeset
 | 
428  | 
structure Inttab = Table(type key = int val ord = int_ord);  | 
| 
 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 
wenzelm 
parents: 
31621 
diff
changeset
 | 
429  | 
structure Symtab = Table(type key = string val ord = fast_string_ord);  | 
| 
 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 
wenzelm 
parents: 
31621 
diff
changeset
 | 
430  | 
structure Symreltab = Table(type key = string * string  | 
| 
30647
 
ef8f46c3158a
added Symreltab (binary relations of symbols) instance of TableFun
 
haftmann 
parents: 
30290 
diff
changeset
 | 
431  | 
val ord = prod_ord fast_string_ord fast_string_ord);  | 
| 
 
ef8f46c3158a
added Symreltab (binary relations of symbols) instance of TableFun
 
haftmann 
parents: 
30290 
diff
changeset
 | 
432  |