# HG changeset patch # User nipkow # Date 758531347 -3600 # Node ID 76f60e6400e89f4587bd0e1749bfd194a35c19fd # Parent d762f9421933d9a1e0354798fe7e3910145f01f4 optimized net for matching of abstractions to speed up simplifier diff -r d762f9421933 -r 76f60e6400e8 src/Pure/net.ML --- a/src/Pure/net.ML Tue Jan 11 12:58:19 1994 +0100 +++ b/src/Pure/net.ML Fri Jan 14 08:09:07 1994 +0100 @@ -9,6 +9,9 @@ E. Charniak, C. K. Riesbeck, D. V. McDermott. Artificial Intelligence Programming. (Lawrence Erlbaum Associates, 1980). [Chapter 14] + +match_term no longer treats abstractions as wildcards but as the constant +*Abs*. Requires operands to be beta-eta-normal. *) signature NET = @@ -33,25 +36,26 @@ datatype key = CombK | VarK | AtomK of string; -(*Only 'loose' bound variables could arise, since Abs nodes are skipped*) +(*Bound variables*) fun string_of_bound i = "*B*" ^ chr i; (*Keys are preorder lists of symbols -- constants, Vars, bound vars, ... - Any term whose head is a Var is regarded entirely as a Var; - abstractions are also regarded as Vars (to cover eta-conversion) + Any term whose head is a Var is regarded entirely as a Var. + Abstractions are also regarded as Vars. This covers eta-conversion + and "near" eta-conversions such as %x.P(?f(x)). *) fun add_key_of_terms (t, cs) = let fun rands (f$t, cs) = CombK :: rands (f, add_key_of_terms(t, cs)) | rands (Const(c,_), cs) = AtomK c :: cs | rands (Free(c,_), cs) = AtomK c :: cs - | rands (Bound i, cs) = AtomK (string_of_bound i) :: cs + | rands (Bound i, cs) = AtomK (string_of_bound i) :: cs in case (head_of t) of - Var _ => VarK :: cs + Var _ => VarK :: cs | Abs (_,_,t) => VarK :: cs - | _ => rands(t,cs) + | _ => rands(t,cs) end; -(*convert a term to a key -- a list of keys*) +(*convert a term to a list of keys*) fun key_of_term t = add_key_of_terms (t, []); @@ -173,16 +177,18 @@ (*Return the nodes accessible from the term (cons them before nets) "unif" signifies retrieval for unification rather than matching. Var in net matches any term. - Abs in object (and Var if "unif") regarded as wildcard. - If not "unif", Var in object only matches a variable in net.*) + Abs or Var in object: if "unif", regarded as wildcard, + else matches only a variable in net. +*) fun matching unif t (net,nets) = let fun rands _ (Leaf _, nets) = nets | rands t (Net{comb,alist,...}, nets) = case t of - (f$t) => foldr (matching unif t) (rands f (comb,[]), nets) - | (Const(c,_)) => look1 (alist, c) nets - | (Free(c,_)) => look1 (alist, c) nets - | (Bound i) => look1 (alist, string_of_bound i) nets + f$t => foldr (matching unif t) (rands f (comb,[]), nets) + | Const(c,_) => look1 (alist, c) nets + | Free(c,_) => look1 (alist, c) nets + | Bound i => look1 (alist, string_of_bound i) nets + | Abs _ => look1 (alist, "*Abs*") nets in case net of Leaf _ => nets @@ -190,13 +196,16 @@ case (head_of t) of Var _ => if unif then net_skip (net,nets) else var::nets (*only matches Var in net*) - | Abs(_,_,u) => net_skip (net,nets) (*could match anything*) +(*If "unif" then a var instantiation in the abstraction could allow + an eta-reduction, so regard the abstraction as a wildcard.*) + | Abs(_,_,u) => if unif then net_skip (net,nets) + else rands t (net, var::nets) | _ => rands t (net, var::nets) (*var could match also*) end; val extract_leaves = flat o map (fn Leaf(xs) => xs); -(*return items whose key could match t*) +(*return items whose key could match t, WHICH MUST BE BETA-ETA NORMAL*) fun match_term net t = extract_leaves (matching false t (net,[])); diff -r d762f9421933 -r 76f60e6400e8 src/Pure/thm.ML --- a/src/Pure/thm.ML Tue Jan 11 12:58:19 1994 +0100 +++ b/src/Pure/thm.ML Fri Jan 14 08:09:07 1994 +0100 @@ -906,7 +906,8 @@ (*Conversion to apply the meta simpset to a term*) fun rewritec (prover,signt) (mss as Mss{net,...}) (hypst,t) = - let fun rew (t, {thm as Thm{sign,hyps,maxidx,prop,...}, lhs}) = + let val t = Pattern.eta_contract t; + fun rew {thm as Thm{sign,hyps,maxidx,prop,...}, lhs} = let val unit = if Sign.subsig(sign,signt) then () else (writeln"Warning: rewrite rule from different theory"; raise Pattern.MATCH) @@ -923,19 +924,14 @@ | Some(thm2) => check_conv(thm2,prop')) end - fun rews t = - let fun rews1 [] = None - | rews1 (rrule::rrules) = - let val opt = rew(t,rrule) handle Pattern.MATCH => None - in case opt of None => rews1 rrules | some => some end; - in rews1 end - - fun eta_rews([]) = None - | eta_rews(rrules) = rews (Pattern.eta_contract t) rrules + fun rews [] = None + | rews (rrule::rrules) = + let val opt = rew rrule handle Pattern.MATCH => None + in case opt of None => rews rrules | some => some end; in case t of Abs(_,_,body) $ u => Some(hypst,subst_bounds([u], body)) - | _ => eta_rews(Net.match_term net t) + | _ => rews(Net.match_term net t) end; (*Conversion to apply a congruence rule to a term*)