wenzelm@12319: (* Title: Pure/net.ML wenzelm@12319: Author: Lawrence C Paulson, Cambridge University Computer Laboratory clasohm@0: Copyright 1993 University of Cambridge clasohm@0: clasohm@0: Discrimination nets: a data structure for indexing items clasohm@0: wenzelm@12319: From the book wenzelm@12319: E. Charniak, C. K. Riesbeck, D. V. McDermott. clasohm@0: Artificial Intelligence Programming. clasohm@0: (Lawrence Erlbaum Associates, 1980). [Chapter 14] nipkow@225: wenzelm@12319: match_term no longer treats abstractions as wildcards; instead they match nipkow@228: only wildcards in patterns. Requires operands to be beta-eta-normal. clasohm@0: *) clasohm@0: wenzelm@12319: signature NET = wenzelm@16808: sig clasohm@0: type key wenzelm@16808: val key_of_term: term -> key list clasohm@0: type 'a net clasohm@0: val empty: 'a net wenzelm@16808: exception INSERT wenzelm@16808: val insert: ('a * 'a -> bool) -> key list * 'a -> 'a net -> 'a net wenzelm@16808: val insert_term: ('a * 'a -> bool) -> term * 'a -> 'a net -> 'a net wenzelm@16808: exception DELETE wenzelm@16808: val delete: ('b * 'a -> bool) -> key list * 'b -> 'a net -> 'a net wenzelm@16808: val delete_term: ('b * 'a -> bool) -> term * 'b -> 'a net -> 'a net wenzelm@16808: val lookup: 'a net -> key list -> 'a list clasohm@0: val match_term: 'a net -> term -> 'a list clasohm@0: val unify_term: 'a net -> term -> 'a list wenzelm@16808: val entries: 'a net -> 'a list wenzelm@16808: val subtract: ('b * 'a -> bool) -> 'a net -> 'b net -> 'b list wenzelm@16808: val merge: ('a * 'a -> bool) -> 'a net * 'a net -> 'a net wenzelm@20011: val content: 'a net -> 'a list wenzelm@16808: end; clasohm@0: wenzelm@16808: structure Net: NET = clasohm@0: struct clasohm@0: clasohm@0: datatype key = CombK | VarK | AtomK of string; clasohm@0: nipkow@228: (*Keys are preorder lists of symbols -- Combinations, Vars, Atoms. nipkow@225: Any term whose head is a Var is regarded entirely as a Var. nipkow@228: Abstractions are also regarded as Vars; this covers eta-conversion nipkow@228: and "near" eta-conversions such as %x.?P(?f(x)). clasohm@0: *) wenzelm@12319: fun add_key_of_terms (t, cs) = clasohm@0: let fun rands (f$t, cs) = CombK :: rands (f, add_key_of_terms(t, cs)) wenzelm@12319: | rands (Const(c,_), cs) = AtomK c :: cs wenzelm@12319: | rands (Free(c,_), cs) = AtomK c :: cs wenzelm@20080: | rands (Bound i, cs) = AtomK (Name.bound i) :: cs clasohm@0: in case (head_of t) of nipkow@225: Var _ => VarK :: cs nipkow@228: | Abs _ => VarK :: cs nipkow@225: | _ => rands(t,cs) clasohm@0: end; clasohm@0: nipkow@225: (*convert a term to a list of keys*) clasohm@0: fun key_of_term t = add_key_of_terms (t, []); clasohm@0: clasohm@0: clasohm@0: (*Trees indexed by key lists: each arc is labelled by a key. clasohm@0: Each node contains a list of items, and arcs to children. clasohm@0: The empty key addresses the entire net. clasohm@0: Lookup functions preserve order in items stored at same level. clasohm@0: *) clasohm@0: datatype 'a net = Leaf of 'a list wenzelm@12319: | Net of {comb: 'a net, wenzelm@12319: var: 'a net, wenzelm@16708: atoms: 'a net Symtab.table}; clasohm@0: clasohm@0: val empty = Leaf[]; wenzelm@16708: fun is_empty (Leaf []) = true | is_empty _ = false; wenzelm@16708: val emptynet = Net{comb=empty, var=empty, atoms=Symtab.empty}; clasohm@0: clasohm@0: clasohm@0: (*** Insertion into a discrimination net ***) clasohm@0: wenzelm@12319: exception INSERT; (*duplicate item in the net*) clasohm@0: clasohm@0: clasohm@0: (*Adds item x to the list at the node addressed by the keys. clasohm@0: Creates node if not already present. wenzelm@12319: eq is the equality test for items. clasohm@0: The empty list of keys generates a Leaf node, others a Net node. clasohm@0: *) wenzelm@16808: fun insert eq (keys,x) net = wenzelm@12319: let fun ins1 ([], Leaf xs) = wenzelm@16686: if member eq xs x then raise INSERT else Leaf(x::xs) clasohm@0: | ins1 (keys, Leaf[]) = ins1 (keys, emptynet) (*expand empty...*) wenzelm@16708: | ins1 (CombK :: keys, Net{comb,var,atoms}) = wenzelm@16708: Net{comb=ins1(keys,comb), var=var, atoms=atoms} wenzelm@16708: | ins1 (VarK :: keys, Net{comb,var,atoms}) = wenzelm@16708: Net{comb=comb, var=ins1(keys,var), atoms=atoms} wenzelm@16708: | ins1 (AtomK a :: keys, Net{comb,var,atoms}) = wenzelm@16708: let wenzelm@18939: val net' = the_default empty (Symtab.lookup atoms a); wenzelm@17412: val atoms' = Symtab.update (a, ins1 (keys, net')) atoms; wenzelm@16708: in Net{comb=comb, var=var, atoms=atoms'} end clasohm@0: in ins1 (keys,net) end; clasohm@0: wenzelm@16808: fun insert_safe eq entry net = insert eq entry net handle INSERT => net; wenzelm@16808: fun insert_term eq (t, x) = insert eq (key_of_term t, x); wenzelm@16808: clasohm@0: clasohm@0: (*** Deletion from a discrimination net ***) clasohm@0: wenzelm@12319: exception DELETE; (*missing item in the net*) clasohm@0: clasohm@0: (*Create a new Net node if it would be nonempty*) wenzelm@16708: fun newnet (args as {comb,var,atoms}) = wenzelm@16708: if is_empty comb andalso is_empty var andalso Symtab.is_empty atoms wenzelm@16708: then empty else Net args; clasohm@0: clasohm@0: (*Deletes item x from the list at the node addressed by the keys. clasohm@0: Raises DELETE if absent. Collapses the net if possible. clasohm@0: eq is the equality test for items. *) wenzelm@16808: fun delete eq (keys, x) net = clasohm@0: let fun del1 ([], Leaf xs) = wenzelm@16686: if member eq xs x then Leaf (remove eq x xs) clasohm@0: else raise DELETE wenzelm@12319: | del1 (keys, Leaf[]) = raise DELETE wenzelm@16708: | del1 (CombK :: keys, Net{comb,var,atoms}) = wenzelm@16708: newnet{comb=del1(keys,comb), var=var, atoms=atoms} wenzelm@16708: | del1 (VarK :: keys, Net{comb,var,atoms}) = wenzelm@16708: newnet{comb=comb, var=del1(keys,var), atoms=atoms} wenzelm@16708: | del1 (AtomK a :: keys, Net{comb,var,atoms}) = wenzelm@16708: let val atoms' = wenzelm@17412: (case Symtab.lookup atoms a of wenzelm@16708: NONE => raise DELETE wenzelm@16708: | SOME net' => wenzelm@16708: (case del1 (keys, net') of wenzelm@16708: Leaf [] => Symtab.delete a atoms wenzelm@17412: | net'' => Symtab.update (a, net'') atoms)) wenzelm@16708: in newnet{comb=comb, var=var, atoms=atoms'} end clasohm@0: in del1 (keys,net) end; clasohm@0: wenzelm@16808: fun delete_term eq (t, x) = delete eq (key_of_term t, x); clasohm@0: wenzelm@16677: clasohm@0: (*** Retrieval functions for discrimination nets ***) clasohm@0: wenzelm@16708: exception ABSENT; clasohm@0: wenzelm@16708: fun the_atom atoms a = wenzelm@17412: (case Symtab.lookup atoms a of wenzelm@16708: NONE => raise ABSENT wenzelm@16708: | SOME net => net); clasohm@0: clasohm@0: (*Return the list of items at the given node, [] if no such node*) wenzelm@16808: fun lookup (Leaf xs) [] = xs wenzelm@16808: | lookup (Leaf _) (_ :: _) = [] (*non-empty keys and empty net*) wenzelm@16808: | lookup (Net {comb, var, atoms}) (CombK :: keys) = lookup comb keys wenzelm@16808: | lookup (Net {comb, var, atoms}) (VarK :: keys) = lookup var keys wenzelm@16808: | lookup (Net {comb, var, atoms}) (AtomK a :: keys) = wenzelm@16808: lookup (the_atom atoms a) keys handle ABSENT => []; clasohm@0: clasohm@0: clasohm@0: (*Skipping a term in a net. Recursively skip 2 levels if a combination*) wenzelm@23178: fun net_skip (Leaf _) nets = nets wenzelm@23178: | net_skip (Net{comb,var,atoms}) nets = wenzelm@23178: fold_rev net_skip (net_skip comb []) (Symtab.fold (cons o #2) atoms (var::nets)); clasohm@0: wenzelm@16808: wenzelm@16808: (** Matching and Unification **) clasohm@0: clasohm@0: (*conses the linked net, if present, to nets*) wenzelm@16708: fun look1 (atoms, a) nets = wenzelm@16708: the_atom atoms a :: nets handle ABSENT => nets; clasohm@0: wenzelm@12319: (*Return the nodes accessible from the term (cons them before nets) clasohm@0: "unif" signifies retrieval for unification rather than matching. clasohm@0: Var in net matches any term. wenzelm@12319: Abs or Var in object: if "unif", regarded as wildcard, nipkow@225: else matches only a variable in net. nipkow@225: *) wenzelm@23178: fun matching unif t net nets = clasohm@0: let fun rands _ (Leaf _, nets) = nets wenzelm@16708: | rands t (Net{comb,atoms,...}, nets) = wenzelm@12319: case t of wenzelm@23178: f$t => fold_rev (matching unif t) (rands f (comb,[])) nets wenzelm@16708: | Const(c,_) => look1 (atoms, c) nets wenzelm@16708: | Free(c,_) => look1 (atoms, c) nets wenzelm@20080: | Bound i => look1 (atoms, Name.bound i) nets wenzelm@12319: | _ => nets wenzelm@12319: in clasohm@0: case net of wenzelm@12319: Leaf _ => nets clasohm@0: | Net{var,...} => wenzelm@12319: case head_of t of wenzelm@23178: Var _ => if unif then net_skip net nets wenzelm@12319: else var::nets (*only matches Var in net*) paulson@2836: (*If "unif" then a var instantiation in the abstraction could allow paulson@2836: an eta-reduction, so regard the abstraction as a wildcard.*) wenzelm@23178: | Abs _ => if unif then net_skip net nets wenzelm@12319: else var::nets (*only a Var can match*) wenzelm@12319: | _ => rands t (net, var::nets) (*var could match also*) clasohm@0: end; clasohm@0: wenzelm@19482: fun extract_leaves l = maps (fn Leaf xs => xs) l; clasohm@0: nipkow@225: (*return items whose key could match t, WHICH MUST BE BETA-ETA NORMAL*) wenzelm@12319: fun match_term net t = wenzelm@23178: extract_leaves (matching false t net []); clasohm@0: clasohm@0: (*return items whose key could unify with t*) wenzelm@12319: fun unify_term net t = wenzelm@23178: extract_leaves (matching true t net []); clasohm@0: wenzelm@3548: wenzelm@16808: (** operations on nets **) wenzelm@16808: wenzelm@16808: (*subtraction: collect entries of second net that are NOT present in first net*) wenzelm@16808: fun subtract eq net1 net2 = wenzelm@16808: let wenzelm@16808: fun subtr (Net _) (Leaf ys) = append ys wenzelm@16808: | subtr (Leaf xs) (Leaf ys) = wenzelm@16808: fold_rev (fn y => if member eq xs y then I else cons y) ys wenzelm@16808: | subtr (Leaf _) (net as Net _) = subtr emptynet net wenzelm@16808: | subtr (Net {comb = comb1, var = var1, atoms = atoms1}) wenzelm@16808: (Net {comb = comb2, var = var2, atoms = atoms2}) = wenzelm@16842: subtr comb1 comb2 wenzelm@16842: #> subtr var1 var2 wenzelm@16842: #> Symtab.fold (fn (a, net) => wenzelm@18939: subtr (the_default emptynet (Symtab.lookup atoms1 a)) net) atoms2 wenzelm@16808: in subtr net1 net2 [] end; wenzelm@16808: wenzelm@16808: fun entries net = subtract (K false) empty net; wenzelm@16808: wenzelm@16808: wenzelm@16808: (* merge *) wenzelm@3548: wenzelm@3548: fun cons_fst x (xs, y) = (x :: xs, y); wenzelm@3548: wenzelm@3548: fun dest (Leaf xs) = map (pair []) xs wenzelm@16708: | dest (Net {comb, var, atoms}) = wenzelm@3560: map (cons_fst CombK) (dest comb) @ wenzelm@3560: map (cons_fst VarK) (dest var) @ wenzelm@19482: maps (fn (a, net) => map (cons_fst (AtomK a)) (dest net)) (Symtab.dest atoms); wenzelm@3548: wenzelm@16808: fun merge eq (net1, net2) = fold (insert_safe eq) (dest net2) net1; wenzelm@3548: wenzelm@20011: fun content net = map #2 (dest net); wenzelm@20011: clasohm@0: end;