clasohm@1460: (* Title: net clasohm@0: ID: $Id$ clasohm@1460: 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: clasohm@0: From the book clasohm@0: E. Charniak, C. K. Riesbeck, D. V. McDermott. clasohm@0: Artificial Intelligence Programming. clasohm@0: (Lawrence Erlbaum Associates, 1980). [Chapter 14] nipkow@225: nipkow@228: 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: clasohm@0: signature NET = clasohm@0: sig clasohm@0: type key clasohm@0: type 'a net clasohm@0: exception DELETE and INSERT clasohm@0: val delete: (key list * 'a) * 'a net * ('a*'a -> bool) -> 'a net clasohm@0: val delete_term: (term * 'a) * 'a net * ('a*'a -> bool) -> 'a net clasohm@0: val empty: 'a net clasohm@0: val insert: (key list * 'a) * 'a net * ('a*'a -> bool) -> 'a net clasohm@0: val insert_term: (term * 'a) * 'a net * ('a*'a -> bool) -> 'a net clasohm@0: val lookup: 'a net * key list -> 'a list clasohm@0: val match_term: 'a net -> term -> 'a list clasohm@0: val key_of_term: term -> key list clasohm@0: val unify_term: 'a net -> term -> 'a list clasohm@0: end; clasohm@0: clasohm@0: paulson@1500: structure Net : NET = clasohm@0: struct clasohm@0: clasohm@0: datatype key = CombK | VarK | AtomK of string; clasohm@0: nipkow@225: (*Bound variables*) clasohm@0: fun string_of_bound i = "*B*" ^ chr i; 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: *) clasohm@0: fun add_key_of_terms (t, cs) = clasohm@0: let fun rands (f$t, cs) = CombK :: rands (f, add_key_of_terms(t, cs)) clasohm@1460: | rands (Const(c,_), cs) = AtomK c :: cs clasohm@1460: | rands (Free(c,_), cs) = AtomK c :: cs clasohm@1460: | rands (Bound i, cs) = AtomK (string_of_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: Keys in the association list (alist) are stored in ascending order. 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 clasohm@1460: | Net of {comb: 'a net, clasohm@1460: var: 'a net, clasohm@1460: alist: (string * 'a net) list}; clasohm@0: clasohm@0: val empty = Leaf[]; clasohm@0: val emptynet = Net{comb=empty, var=empty, alist=[]}; clasohm@0: clasohm@0: clasohm@0: (*** Insertion into a discrimination net ***) clasohm@0: clasohm@1460: 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. clasohm@0: eq is the equality test for items. clasohm@0: The empty list of keys generates a Leaf node, others a Net node. clasohm@0: *) clasohm@0: fun insert ((keys,x), net, eq) = clasohm@0: let fun ins1 ([], Leaf xs) = clasohm@0: if gen_mem eq (x,xs) then raise INSERT else Leaf(x::xs) clasohm@0: | ins1 (keys, Leaf[]) = ins1 (keys, emptynet) (*expand empty...*) clasohm@0: | ins1 (CombK :: keys, Net{comb,var,alist}) = clasohm@1460: Net{comb=ins1(keys,comb), var=var, alist=alist} clasohm@1460: | ins1 (VarK :: keys, Net{comb,var,alist}) = clasohm@1460: Net{comb=comb, var=ins1(keys,var), alist=alist} clasohm@1460: | ins1 (AtomK a :: keys, Net{comb,var,alist}) = clasohm@1460: let fun newpair net = (a, ins1(keys,net)) clasohm@1460: fun inslist [] = [newpair empty] clasohm@1460: | inslist((b: string, netb) :: alist) = clasohm@1460: if a=b then newpair netb :: alist clasohm@1460: else if ab*) (b,netb) :: inslist alist clasohm@1460: in Net{comb=comb, var=var, alist= inslist alist} end clasohm@0: in ins1 (keys,net) end; clasohm@0: clasohm@0: fun insert_term ((t,x), net, eq) = insert((key_of_term t, x), net, eq); clasohm@0: clasohm@0: (*** Deletion from a discrimination net ***) clasohm@0: clasohm@1460: exception DELETE; (*missing item in the net*) clasohm@0: clasohm@0: (*Create a new Net node if it would be nonempty*) clasohm@0: fun newnet {comb=Leaf[], var=Leaf[], alist=[]} = empty clasohm@0: | newnet {comb,var,alist} = Net{comb=comb, var=var, alist=alist}; clasohm@0: clasohm@0: (*add new (b,net) pair to the alist provided net is nonempty*) clasohm@0: fun conspair((b, Leaf[]), alist) = alist clasohm@0: | conspair((b, net), alist) = (b, net) :: alist; 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. *) clasohm@0: fun delete ((keys, x), net, eq) = clasohm@0: let fun del1 ([], Leaf xs) = clasohm@0: if gen_mem eq (x,xs) then Leaf (gen_rem eq (xs,x)) clasohm@0: else raise DELETE clasohm@1460: | del1 (keys, Leaf[]) = raise DELETE clasohm@1460: | del1 (CombK :: keys, Net{comb,var,alist}) = clasohm@1460: newnet{comb=del1(keys,comb), var=var, alist=alist} clasohm@1460: | del1 (VarK :: keys, Net{comb,var,alist}) = clasohm@1460: newnet{comb=comb, var=del1(keys,var), alist=alist} clasohm@1460: | del1 (AtomK a :: keys, Net{comb,var,alist}) = clasohm@1460: let fun newpair net = (a, del1(keys,net)) clasohm@1460: fun dellist [] = raise DELETE clasohm@1460: | dellist((b: string, netb) :: alist) = clasohm@1460: if a=b then conspair(newpair netb, alist) clasohm@1460: else if ab*) (b,netb) :: dellist alist clasohm@1460: in newnet{comb=comb, var=var, alist= dellist alist} end clasohm@0: in del1 (keys,net) end; clasohm@0: clasohm@0: fun delete_term ((t,x), net, eq) = delete((key_of_term t, x), net, eq); clasohm@0: clasohm@0: (*** Retrieval functions for discrimination nets ***) clasohm@0: clasohm@0: exception OASSOC; clasohm@0: clasohm@0: (*Ordered association list lookup*) clasohm@0: fun oassoc ([], a: string) = raise OASSOC clasohm@0: | oassoc ((b,x)::pairs, a) = clasohm@0: if b []; clasohm@0: clasohm@0: clasohm@0: (*Skipping a term in a net. Recursively skip 2 levels if a combination*) clasohm@0: fun net_skip (Leaf _, nets) = nets clasohm@0: | net_skip (Net{comb,var,alist}, nets) = clasohm@0: foldr net_skip clasohm@0: (net_skip (comb,[]), clasohm@1460: foldr (fn ((_,net), nets) => net::nets) (alist, var::nets)); clasohm@0: clasohm@0: (** Matching and Unification**) clasohm@0: clasohm@0: (*conses the linked net, if present, to nets*) clasohm@0: fun look1 (alist, a) nets = clasohm@0: oassoc(alist,a) :: nets handle OASSOC => nets; clasohm@0: clasohm@0: (*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. nipkow@225: Abs or Var in object: if "unif", regarded as wildcard, nipkow@225: else matches only a variable in net. nipkow@225: *) clasohm@0: fun matching unif t (net,nets) = clasohm@0: let fun rands _ (Leaf _, nets) = nets clasohm@1460: | rands t (Net{comb,alist,...}, nets) = clasohm@1460: case t of clasohm@1460: f$t => foldr (matching unif t) (rands f (comb,[]), nets) clasohm@1460: | Const(c,_) => look1 (alist, c) nets clasohm@1460: | Free(c,_) => look1 (alist, c) nets clasohm@1460: | Bound i => look1 (alist, string_of_bound i) nets clasohm@1460: | _ => nets clasohm@0: in clasohm@0: case net of clasohm@1460: Leaf _ => nets clasohm@0: | Net{var,...} => paulson@2836: let val etat = Pattern.eta_contract_atom t paulson@2836: in case (head_of etat) of paulson@2836: Var _ => if unif then net_skip (net,nets) paulson@2836: 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.*) paulson@2836: | Abs _ => if unif then net_skip (net,nets) paulson@2836: else var::nets (*only a Var can match*) paulson@2836: | _ => rands etat (net, var::nets) (*var could match also*) paulson@2836: end clasohm@0: end; clasohm@0: paulson@2672: fun extract_leaves l = List.concat (map (fn Leaf(xs) => xs) l); clasohm@0: nipkow@225: (*return items whose key could match t, WHICH MUST BE BETA-ETA NORMAL*) clasohm@0: fun match_term net t = clasohm@0: extract_leaves (matching false t (net,[])); clasohm@0: clasohm@0: (*return items whose key could unify with t*) clasohm@0: fun unify_term net t = clasohm@0: extract_leaves (matching true t (net,[])); clasohm@0: clasohm@0: end;