| author | wenzelm | 
| Mon, 07 Jan 2013 10:17:11 +0100 | |
| changeset 50757 | 37091451ba1a | 
| parent 45404 | 69ec395ef6ca | 
| child 55741 | b969263fcf02 | 
| permissions | -rw-r--r-- | 
| 12319 | 1  | 
(* Title: Pure/net.ML  | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
| 0 | 3  | 
Copyright 1993 University of Cambridge  | 
4  | 
||
5  | 
Discrimination nets: a data structure for indexing items  | 
|
6  | 
||
| 12319 | 7  | 
From the book  | 
8  | 
E. Charniak, C. K. Riesbeck, D. V. McDermott.  | 
|
| 0 | 9  | 
Artificial Intelligence Programming.  | 
10  | 
(Lawrence Erlbaum Associates, 1980). [Chapter 14]  | 
|
| 
225
 
76f60e6400e8
optimized net for matching of abstractions to speed up simplifier
 
nipkow 
parents: 
0 
diff
changeset
 | 
11  | 
|
| 12319 | 12  | 
match_term no longer treats abstractions as wildcards; instead they match  | 
| 228 | 13  | 
only wildcards in patterns. Requires operands to be beta-eta-normal.  | 
| 0 | 14  | 
*)  | 
15  | 
||
| 12319 | 16  | 
signature NET =  | 
| 16808 | 17  | 
sig  | 
| 0 | 18  | 
type key  | 
| 16808 | 19  | 
val key_of_term: term -> key list  | 
| 37523 | 20  | 
val encode_type: typ -> term  | 
| 0 | 21  | 
type 'a net  | 
22  | 
val empty: 'a net  | 
|
| 16808 | 23  | 
exception INSERT  | 
24  | 
  val insert: ('a * 'a -> bool) -> key list * 'a -> 'a net -> 'a net
 | 
|
25  | 
  val insert_term: ('a * 'a -> bool) -> term * 'a -> 'a net -> 'a net
 | 
|
| 33371 | 26  | 
  val insert_safe: ('a * 'a -> bool) -> key list * 'a -> 'a net -> 'a net
 | 
27  | 
  val insert_term_safe: ('a * 'a -> bool) -> term * 'a -> 'a net -> 'a net
 | 
|
| 16808 | 28  | 
exception DELETE  | 
29  | 
  val delete: ('b * 'a -> bool) -> key list * 'b -> 'a net -> 'a net
 | 
|
30  | 
  val delete_term: ('b * 'a -> bool) -> term * 'b -> 'a net -> 'a net
 | 
|
| 33371 | 31  | 
  val delete_safe: ('b * 'a -> bool) -> key list * 'b -> 'a net -> 'a net
 | 
32  | 
  val delete_term_safe: ('b * 'a -> bool) -> term * 'b -> 'a net -> 'a net
 | 
|
| 16808 | 33  | 
val lookup: 'a net -> key list -> 'a list  | 
| 0 | 34  | 
val match_term: 'a net -> term -> 'a list  | 
35  | 
val unify_term: 'a net -> term -> 'a list  | 
|
| 16808 | 36  | 
val entries: 'a net -> 'a list  | 
37  | 
  val subtract: ('b * 'a -> bool) -> 'a net -> 'b net -> 'b list
 | 
|
38  | 
  val merge: ('a * 'a -> bool) -> 'a net * 'a net -> 'a net
 | 
|
| 20011 | 39  | 
val content: 'a net -> 'a list  | 
| 16808 | 40  | 
end;  | 
| 0 | 41  | 
|
| 16808 | 42  | 
structure Net: NET =  | 
| 0 | 43  | 
struct  | 
44  | 
||
45  | 
datatype key = CombK | VarK | AtomK of string;  | 
|
46  | 
||
| 228 | 47  | 
(*Keys are preorder lists of symbols -- Combinations, Vars, Atoms.  | 
| 
225
 
76f60e6400e8
optimized net for matching of abstractions to speed up simplifier
 
nipkow 
parents: 
0 
diff
changeset
 | 
48  | 
Any term whose head is a Var is regarded entirely as a Var.  | 
| 228 | 49  | 
Abstractions are also regarded as Vars; this covers eta-conversion  | 
50  | 
and "near" eta-conversions such as %x.?P(?f(x)).  | 
|
| 0 | 51  | 
*)  | 
| 12319 | 52  | 
fun add_key_of_terms (t, cs) =  | 
| 0 | 53  | 
let fun rands (f$t, cs) = CombK :: rands (f, add_key_of_terms(t, cs))  | 
| 12319 | 54  | 
| rands (Const(c,_), cs) = AtomK c :: cs  | 
55  | 
| rands (Free(c,_), cs) = AtomK c :: cs  | 
|
| 20080 | 56  | 
| rands (Bound i, cs) = AtomK (Name.bound i) :: cs  | 
| 45404 | 57  | 
in case head_of t of  | 
| 
225
 
76f60e6400e8
optimized net for matching of abstractions to speed up simplifier
 
nipkow 
parents: 
0 
diff
changeset
 | 
58  | 
Var _ => VarK :: cs  | 
| 228 | 59  | 
| Abs _ => VarK :: cs  | 
| 
225
 
76f60e6400e8
optimized net for matching of abstractions to speed up simplifier
 
nipkow 
parents: 
0 
diff
changeset
 | 
60  | 
| _ => rands(t,cs)  | 
| 0 | 61  | 
end;  | 
62  | 
||
| 
225
 
76f60e6400e8
optimized net for matching of abstractions to speed up simplifier
 
nipkow 
parents: 
0 
diff
changeset
 | 
63  | 
(*convert a term to a list of keys*)  | 
| 0 | 64  | 
fun key_of_term t = add_key_of_terms (t, []);  | 
65  | 
||
| 37523 | 66  | 
(*encode_type -- for indexing purposes*)  | 
67  | 
fun encode_type (Type (c, Ts)) = Term.list_comb (Const (c, dummyT), map encode_type Ts)  | 
|
68  | 
| encode_type (TFree (a, _)) = Free (a, dummyT)  | 
|
69  | 
| encode_type (TVar (a, _)) = Var (a, dummyT);  | 
|
70  | 
||
| 0 | 71  | 
|
72  | 
(*Trees indexed by key lists: each arc is labelled by a key.  | 
|
73  | 
Each node contains a list of items, and arcs to children.  | 
|
74  | 
The empty key addresses the entire net.  | 
|
75  | 
Lookup functions preserve order in items stored at same level.  | 
|
76  | 
*)  | 
|
77  | 
datatype 'a net = Leaf of 'a list  | 
|
| 12319 | 78  | 
                | Net of {comb: 'a net,
 | 
79  | 
var: 'a net,  | 
|
| 16708 | 80  | 
atoms: 'a net Symtab.table};  | 
| 0 | 81  | 
|
82  | 
val empty = Leaf[];  | 
|
| 16708 | 83  | 
fun is_empty (Leaf []) = true | is_empty _ = false;  | 
84  | 
val emptynet = Net{comb=empty, var=empty, atoms=Symtab.empty};
 | 
|
| 0 | 85  | 
|
86  | 
||
87  | 
(*** Insertion into a discrimination net ***)  | 
|
88  | 
||
| 12319 | 89  | 
exception INSERT; (*duplicate item in the net*)  | 
| 0 | 90  | 
|
91  | 
||
92  | 
(*Adds item x to the list at the node addressed by the keys.  | 
|
93  | 
Creates node if not already present.  | 
|
| 12319 | 94  | 
eq is the equality test for items.  | 
| 0 | 95  | 
The empty list of keys generates a Leaf node, others a Net node.  | 
96  | 
*)  | 
|
| 16808 | 97  | 
fun insert eq (keys,x) net =  | 
| 12319 | 98  | 
let fun ins1 ([], Leaf xs) =  | 
| 16686 | 99  | 
if member eq xs x then raise INSERT else Leaf(x::xs)  | 
| 0 | 100  | 
| ins1 (keys, Leaf[]) = ins1 (keys, emptynet) (*expand empty...*)  | 
| 16708 | 101  | 
        | ins1 (CombK :: keys, Net{comb,var,atoms}) =
 | 
102  | 
            Net{comb=ins1(keys,comb), var=var, atoms=atoms}
 | 
|
103  | 
        | ins1 (VarK :: keys, Net{comb,var,atoms}) =
 | 
|
104  | 
            Net{comb=comb, var=ins1(keys,var), atoms=atoms}
 | 
|
105  | 
        | ins1 (AtomK a :: keys, Net{comb,var,atoms}) =
 | 
|
| 45404 | 106  | 
let val atoms' = Symtab.map_default (a, empty) (fn net' => ins1 (keys, net')) atoms;  | 
| 16708 | 107  | 
            in  Net{comb=comb, var=var, atoms=atoms'}  end
 | 
| 0 | 108  | 
in ins1 (keys,net) end;  | 
109  | 
||
| 33371 | 110  | 
fun insert_term eq (t, x) = insert eq (key_of_term t, x);  | 
111  | 
||
| 16808 | 112  | 
fun insert_safe eq entry net = insert eq entry net handle INSERT => net;  | 
| 33371 | 113  | 
fun insert_term_safe eq entry net = insert_term eq entry net handle INSERT => net;  | 
| 16808 | 114  | 
|
| 0 | 115  | 
|
116  | 
(*** Deletion from a discrimination net ***)  | 
|
117  | 
||
| 12319 | 118  | 
exception DELETE; (*missing item in the net*)  | 
| 0 | 119  | 
|
120  | 
(*Create a new Net node if it would be nonempty*)  | 
|
| 16708 | 121  | 
fun newnet (args as {comb,var,atoms}) =
 | 
122  | 
if is_empty comb andalso is_empty var andalso Symtab.is_empty atoms  | 
|
123  | 
then empty else Net args;  | 
|
| 0 | 124  | 
|
125  | 
(*Deletes item x from the list at the node addressed by the keys.  | 
|
126  | 
Raises DELETE if absent. Collapses the net if possible.  | 
|
127  | 
eq is the equality test for items. *)  | 
|
| 16808 | 128  | 
fun delete eq (keys, x) net =  | 
| 0 | 129  | 
let fun del1 ([], Leaf xs) =  | 
| 16686 | 130  | 
if member eq xs x then Leaf (remove eq x xs)  | 
| 0 | 131  | 
else raise DELETE  | 
| 12319 | 132  | 
| del1 (keys, Leaf[]) = raise DELETE  | 
| 16708 | 133  | 
        | del1 (CombK :: keys, Net{comb,var,atoms}) =
 | 
134  | 
            newnet{comb=del1(keys,comb), var=var, atoms=atoms}
 | 
|
135  | 
        | del1 (VarK :: keys, Net{comb,var,atoms}) =
 | 
|
136  | 
            newnet{comb=comb, var=del1(keys,var), atoms=atoms}
 | 
|
137  | 
        | del1 (AtomK a :: keys, Net{comb,var,atoms}) =
 | 
|
138  | 
let val atoms' =  | 
|
| 17412 | 139  | 
(case Symtab.lookup atoms a of  | 
| 16708 | 140  | 
NONE => raise DELETE  | 
141  | 
| SOME net' =>  | 
|
142  | 
(case del1 (keys, net') of  | 
|
143  | 
Leaf [] => Symtab.delete a atoms  | 
|
| 17412 | 144  | 
| net'' => Symtab.update (a, net'') atoms))  | 
| 16708 | 145  | 
            in  newnet{comb=comb, var=var, atoms=atoms'}  end
 | 
| 0 | 146  | 
in del1 (keys,net) end;  | 
147  | 
||
| 16808 | 148  | 
fun delete_term eq (t, x) = delete eq (key_of_term t, x);  | 
| 0 | 149  | 
|
| 33371 | 150  | 
fun delete_safe eq entry net = delete eq entry net handle DELETE => net;  | 
151  | 
fun delete_term_safe eq entry net = delete_term eq entry net handle DELETE => net;  | 
|
152  | 
||
| 16677 | 153  | 
|
| 0 | 154  | 
(*** Retrieval functions for discrimination nets ***)  | 
155  | 
||
| 16708 | 156  | 
exception ABSENT;  | 
| 0 | 157  | 
|
| 16708 | 158  | 
fun the_atom atoms a =  | 
| 17412 | 159  | 
(case Symtab.lookup atoms a of  | 
| 16708 | 160  | 
NONE => raise ABSENT  | 
161  | 
| SOME net => net);  | 
|
| 0 | 162  | 
|
163  | 
(*Return the list of items at the given node, [] if no such node*)  | 
|
| 16808 | 164  | 
fun lookup (Leaf xs) [] = xs  | 
165  | 
| lookup (Leaf _) (_ :: _) = [] (*non-empty keys and empty net*)  | 
|
166  | 
  | lookup (Net {comb, var, atoms}) (CombK :: keys) = lookup comb keys
 | 
|
167  | 
  | lookup (Net {comb, var, atoms}) (VarK :: keys) = lookup var keys
 | 
|
168  | 
  | lookup (Net {comb, var, atoms}) (AtomK a :: keys) =
 | 
|
169  | 
lookup (the_atom atoms a) keys handle ABSENT => [];  | 
|
| 0 | 170  | 
|
171  | 
||
172  | 
(*Skipping a term in a net. Recursively skip 2 levels if a combination*)  | 
|
| 23178 | 173  | 
fun net_skip (Leaf _) nets = nets  | 
174  | 
  | net_skip (Net{comb,var,atoms}) nets =
 | 
|
175  | 
fold_rev net_skip (net_skip comb []) (Symtab.fold (cons o #2) atoms (var::nets));  | 
|
| 0 | 176  | 
|
| 16808 | 177  | 
|
178  | 
(** Matching and Unification **)  | 
|
| 0 | 179  | 
|
180  | 
(*conses the linked net, if present, to nets*)  | 
|
| 16708 | 181  | 
fun look1 (atoms, a) nets =  | 
182  | 
the_atom atoms a :: nets handle ABSENT => nets;  | 
|
| 0 | 183  | 
|
| 12319 | 184  | 
(*Return the nodes accessible from the term (cons them before nets)  | 
| 0 | 185  | 
"unif" signifies retrieval for unification rather than matching.  | 
186  | 
Var in net matches any term.  | 
|
| 12319 | 187  | 
Abs or Var in object: if "unif", regarded as wildcard,  | 
| 
225
 
76f60e6400e8
optimized net for matching of abstractions to speed up simplifier
 
nipkow 
parents: 
0 
diff
changeset
 | 
188  | 
else matches only a variable in net.  | 
| 
 
76f60e6400e8
optimized net for matching of abstractions to speed up simplifier
 
nipkow 
parents: 
0 
diff
changeset
 | 
189  | 
*)  | 
| 23178 | 190  | 
fun matching unif t net nets =  | 
| 0 | 191  | 
let fun rands _ (Leaf _, nets) = nets  | 
| 16708 | 192  | 
        | rands t (Net{comb,atoms,...}, nets) =
 | 
| 12319 | 193  | 
case t of  | 
| 23178 | 194  | 
f$t => fold_rev (matching unif t) (rands f (comb,[])) nets  | 
| 16708 | 195  | 
| Const(c,_) => look1 (atoms, c) nets  | 
196  | 
| Free(c,_) => look1 (atoms, c) nets  | 
|
| 20080 | 197  | 
| Bound i => look1 (atoms, Name.bound i) nets  | 
| 12319 | 198  | 
| _ => nets  | 
199  | 
in  | 
|
| 0 | 200  | 
case net of  | 
| 12319 | 201  | 
Leaf _ => nets  | 
| 0 | 202  | 
       | Net{var,...} =>
 | 
| 12319 | 203  | 
case head_of t of  | 
| 23178 | 204  | 
Var _ => if unif then net_skip net nets  | 
| 12319 | 205  | 
else var::nets (*only matches Var in net*)  | 
| 2836 | 206  | 
(*If "unif" then a var instantiation in the abstraction could allow  | 
207  | 
an eta-reduction, so regard the abstraction as a wildcard.*)  | 
|
| 23178 | 208  | 
| Abs _ => if unif then net_skip net nets  | 
| 12319 | 209  | 
else var::nets (*only a Var can match*)  | 
210  | 
| _ => rands t (net, var::nets) (*var could match also*)  | 
|
| 0 | 211  | 
end;  | 
212  | 
||
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
18939 
diff
changeset
 | 
213  | 
fun extract_leaves l = maps (fn Leaf xs => xs) l;  | 
| 0 | 214  | 
|
| 
225
 
76f60e6400e8
optimized net for matching of abstractions to speed up simplifier
 
nipkow 
parents: 
0 
diff
changeset
 | 
215  | 
(*return items whose key could match t, WHICH MUST BE BETA-ETA NORMAL*)  | 
| 12319 | 216  | 
fun match_term net t =  | 
| 23178 | 217  | 
extract_leaves (matching false t net []);  | 
| 0 | 218  | 
|
219  | 
(*return items whose key could unify with t*)  | 
|
| 12319 | 220  | 
fun unify_term net t =  | 
| 23178 | 221  | 
extract_leaves (matching true t net []);  | 
| 0 | 222  | 
|
| 3548 | 223  | 
|
| 16808 | 224  | 
(** operations on nets **)  | 
225  | 
||
226  | 
(*subtraction: collect entries of second net that are NOT present in first net*)  | 
|
227  | 
fun subtract eq net1 net2 =  | 
|
228  | 
let  | 
|
229  | 
fun subtr (Net _) (Leaf ys) = append ys  | 
|
230  | 
| subtr (Leaf xs) (Leaf ys) =  | 
|
231  | 
fold_rev (fn y => if member eq xs y then I else cons y) ys  | 
|
232  | 
| subtr (Leaf _) (net as Net _) = subtr emptynet net  | 
|
233  | 
      | subtr (Net {comb = comb1, var = var1, atoms = atoms1})
 | 
|
234  | 
            (Net {comb = comb2, var = var2, atoms = atoms2}) =
 | 
|
| 16842 | 235  | 
subtr comb1 comb2  | 
236  | 
#> subtr var1 var2  | 
|
237  | 
#> Symtab.fold (fn (a, net) =>  | 
|
| 18939 | 238  | 
subtr (the_default emptynet (Symtab.lookup atoms1 a)) net) atoms2  | 
| 16808 | 239  | 
in subtr net1 net2 [] end;  | 
240  | 
||
241  | 
fun entries net = subtract (K false) empty net;  | 
|
242  | 
||
243  | 
||
244  | 
(* merge *)  | 
|
| 3548 | 245  | 
|
246  | 
fun cons_fst x (xs, y) = (x :: xs, y);  | 
|
247  | 
||
248  | 
fun dest (Leaf xs) = map (pair []) xs  | 
|
| 16708 | 249  | 
  | dest (Net {comb, var, atoms}) =
 | 
| 3560 | 250  | 
map (cons_fst CombK) (dest comb) @  | 
251  | 
map (cons_fst VarK) (dest var) @  | 
|
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
18939 
diff
changeset
 | 
252  | 
maps (fn (a, net) => map (cons_fst (AtomK a)) (dest net)) (Symtab.dest atoms);  | 
| 3548 | 253  | 
|
| 16808 | 254  | 
fun merge eq (net1, net2) = fold (insert_safe eq) (dest net2) net1;  | 
| 3548 | 255  | 
|
| 20011 | 256  | 
fun content net = map #2 (dest net);  | 
257  | 
||
| 0 | 258  | 
end;  |