src/Pure/unify.ML
changeset 1460 5a6f2aabd538
parent 1458 fd510875fb71
child 1505 14f4c55bbe9a
--- a/src/Pure/unify.ML	Mon Jan 29 13:58:15 1996 +0100
+++ b/src/Pure/unify.ML	Mon Jan 29 14:16:13 1996 +0100
@@ -1,6 +1,6 @@
-(*  Title:      unify
+(*  Title: 	unify
     ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   Cambridge University 1992
 
 Higher-Order Unification
@@ -28,16 +28,16 @@
   val combound : (term*int*int) -> term
   val rlist_abs: (string*typ)list * term -> term   
   val smash_unifiers : Sign.sg * Envir.env * (term*term)list
-        -> (Envir.env Sequence.seq)
+	-> (Envir.env Sequence.seq)
   val unifiers: Sign.sg * Envir.env * ((term*term)list)
-        -> (Envir.env * (term * term)list) Sequence.seq
+	-> (Envir.env * (term * term)list) Sequence.seq
 end;
 
 functor UnifyFun (structure Sign: SIGN and Envir: ENVIR and Sequence: SEQUENCE
                   and Pattern:PATTERN
                   sharing type Sign.sg = Pattern.sg
                   and     type Envir.env = Pattern.env)
-        : UNIFY = 
+	: UNIFY = 
 struct
 
 structure Sign = Sign;
@@ -47,11 +47,11 @@
 
 (*Unification options*)
 
-val trace_bound = ref 10        (*tracing starts above this depth, 0 for full*)
-and search_bound = ref 20       (*unification quits above this depth*)
-and trace_simp = ref false      (*print dpairs before calling SIMPL*)
-and trace_types = ref false     (*announce potential incompleteness
-                                  of type unification*)
+val trace_bound = ref 10	(*tracing starts above this depth, 0 for full*)
+and search_bound = ref 20	(*unification quits above this depth*)
+and trace_simp = ref false	(*print dpairs before calling SIMPL*)
+and trace_types = ref false	(*announce potential incompleteness
+				  of type unification*)
 
 val sgr = ref(Sign.proto_pure);
 
@@ -62,14 +62,14 @@
 fun body_type(Envir.Envir{iTs,...}) = 
 let fun bT(Type("fun",[_,T])) = bT T
       | bT(T as TVar(ixn,_)) = (case assoc(iTs,ixn) of
-                None => T | Some(T') => bT T')
+		None => T | Some(T') => bT T')
       | bT T = T
 in bT end;
 
 fun binder_types(Envir.Envir{iTs,...}) = 
 let fun bTs(Type("fun",[T,U])) = T :: bTs U
       | bTs(T as TVar(ixn,_)) = (case assoc(iTs,ixn) of
-                None => [] | Some(T') => bTs T')
+		None => [] | Some(T') => bTs T')
       | bTs _ = []
 in bTs end;
 
@@ -82,7 +82,7 @@
   inefficiency, avoid partial application:  if an atom is applied to
   any arguments at all, apply it to its full number of arguments.
   For
-    rbinder = [(x1,T),...,(xm,Tm)]              (user's var names preserved!)
+    rbinder = [(x1,T),...,(xm,Tm)]		(user's var names preserved!)
     args  =   [arg1,...,argn]
   the value of 
       (xm,...,x1)(head(arg1,...,argn))  remains invariant.
@@ -92,18 +92,18 @@
 in
   fun head_norm (env,t) : term =
     let fun hnorm (Var (v,T)) = 
-              (case Envir.lookup (env,v) of
-                  Some u => head_norm (env, u)
-                | None   => raise SAME)
-          | hnorm (Abs(a,T,body)) =  Abs(a, T, hnorm body)
-          | hnorm (Abs(_,_,body) $ t) =
-              head_norm (env, subst_bounds([t], body))
-          | hnorm (f $ t) =
-              (case hnorm f of
-                 Abs(_,_,body) =>
-                   head_norm (env, subst_bounds([t], body))
-               | nf => nf $ t)
-          | hnorm _ =  raise SAME
+	      (case Envir.lookup (env,v) of
+		  Some u => head_norm (env, u)
+		| None   => raise SAME)
+	  | hnorm (Abs(a,T,body)) =  Abs(a, T, hnorm body)
+	  | hnorm (Abs(_,_,body) $ t) =
+	      head_norm (env, subst_bounds([t], body))
+	  | hnorm (f $ t) =
+	      (case hnorm f of
+		 Abs(_,_,body) =>
+		   head_norm (env, subst_bounds([t], body))
+	       | nf => nf $ t)
+	  | hnorm _ =  raise SAME
     in  hnorm t  handle SAME=> t  end
 end;
 
@@ -113,18 +113,18 @@
 fun fastype (Envir.Envir{iTs,...}) =
 let val funerr = "fastype: expected function type";
     fun fast(rbinder, f$u) =
-        (case (fast (rbinder, f)) of
-           Type("fun",[_,T]) => T
-         | TVar(ixn,_) =>
-                (case assoc(iTs,ixn) of
-                   Some(Type("fun",[_,T])) => T
-                 | _ => raise TERM(funerr, [f$u]))
-         | _ => raise TERM(funerr, [f$u]))
+	(case (fast (rbinder, f)) of
+	   Type("fun",[_,T]) => T
+	 | TVar(ixn,_) =>
+		(case assoc(iTs,ixn) of
+		   Some(Type("fun",[_,T])) => T
+		 | _ => raise TERM(funerr, [f$u]))
+	 | _ => raise TERM(funerr, [f$u]))
       | fast (rbinder, Const (_,T)) = T
       | fast (rbinder, Free (_,T)) = T
       | fast (rbinder, Bound i) =
-        (#2 (nth_elem (i,rbinder))
-         handle LIST _=> raise TERM("fastype: Bound", [Bound i]))
+	(#2 (nth_elem (i,rbinder))
+  	 handle LIST _=> raise TERM("fastype: Bound", [Bound i]))
       | fast (rbinder, Var (_,T)) = T 
       | fast (rbinder, Abs (_,T,u)) =  T --> fast (("",T) :: rbinder, u)
 in fast end;
@@ -133,14 +133,14 @@
 (*Eta normal form*)
 fun eta_norm(env as Envir.Envir{iTs,...}) =
   let fun etif (Type("fun",[T,U]), t) =
-            Abs("", T, etif(U, incr_boundvars 1 t $ Bound 0))
-        | etif (TVar(ixn,_),t) = 
-            (case assoc(iTs,ixn) of
-                  None => t | Some(T) => etif(T,t))
-        | etif (_,t) = t;
+	    Abs("", T, etif(U, incr_boundvars 1 t $ Bound 0))
+	| etif (TVar(ixn,_),t) = 
+	    (case assoc(iTs,ixn) of
+		  None => t | Some(T) => etif(T,t))
+	| etif (_,t) = t;
       fun eta_nm (rbinder, Abs(a,T,body)) =
-            Abs(a, T, eta_nm ((a,T)::rbinder, body))
-        | eta_nm (rbinder, t) = etif(fastype env (rbinder,t), t)
+	    Abs(a, T, eta_nm ((a,T)::rbinder, body))
+	| eta_nm (rbinder, t) = etif(fastype env (rbinder,t), t)
   in eta_nm end;
 
 
@@ -150,22 +150,22 @@
   "seen" is list of variables passed thru, is a memo variable for sharing.
   This version searches for nonrigid occurrence, returns true if found. *)
 fun occurs_terms (seen: (indexname list) ref,
-                  env: Envir.env, v: indexname, ts: term list): bool =
+ 		  env: Envir.env, v: indexname, ts: term list): bool =
   let fun occurs [] = false
-        | occurs (t::ts) =  occur t  orelse  occurs ts
+	| occurs (t::ts) =  occur t  orelse  occurs ts
       and occur (Const _)  = false
-        | occur (Bound _)  = false
-        | occur (Free _)  = false
-        | occur (Var (w,_))  = 
-            if w mem !seen then false
-            else if v=w then true
-              (*no need to lookup: v has no assignment*)
-            else (seen := w:: !seen;  
-                  case  Envir.lookup(env,w)  of
-                      None    => false
-                    | Some t => occur t)
-        | occur (Abs(_,_,body)) = occur body
-        | occur (f$t) = occur t  orelse   occur f
+	| occur (Bound _)  = false
+	| occur (Free _)  = false
+	| occur (Var (w,_))  = 
+	    if w mem !seen then false
+	    else if v=w then true
+	      (*no need to lookup: v has no assignment*)
+	    else (seen := w:: !seen;  
+	          case  Envir.lookup(env,w)  of
+		      None    => false
+		    | Some t => occur t)
+	| occur (Abs(_,_,body)) = occur body
+	| occur (f$t) = occur t  orelse   occur f
   in  occurs ts  end;
 
 
@@ -174,7 +174,7 @@
 fun head_of_in (env,t) : term = case t of
     f$_ => head_of_in(env,f)
   | Var (v,_) => (case  Envir.lookup(env,v)  of  
-                        Some u => head_of_in(env,u)  |  None   => t)
+			Some u => head_of_in(env,u)  |  None   => t)
   | _ => t;
 
 
@@ -205,9 +205,9 @@
 *)
 fun rigid_occurs_term (seen: (indexname list)ref, env, v: indexname, t) = 
   let fun nonrigid t = if occurs_terms(seen,env,v,[t]) then Nonrigid 
-                       else NoOcc
+		       else NoOcc
       fun occurs [] = NoOcc
-        | occurs (t::ts) =
+	| occurs (t::ts) =
             (case occur t of
                Rigid => Rigid
              | oc =>  (case occurs ts of NoOcc => oc  |  oc2 => oc2))
@@ -217,28 +217,28 @@
              | oc =>  (case occomb f of NoOcc => oc  |  oc2 => oc2))
         | occomb t = occur t
       and occur (Const _)  = NoOcc
-        | occur (Bound _)  = NoOcc
-        | occur (Free _)  = NoOcc
-        | occur (Var (w,_))  = 
-            if w mem !seen then NoOcc
-            else if v=w then Rigid
-            else (seen := w:: !seen;  
-                  case  Envir.lookup(env,w)  of
-                      None    => NoOcc
-                    | Some t => occur t)
-        | occur (Abs(_,_,body)) = occur body
-        | occur (t as f$_) =  (*switch to nonrigid search?*)
-           (case head_of_in (env,f) of
-              Var (w,_) => (*w is not assigned*)
-                if v=w then Rigid  
-                else  nonrigid t
-            | Abs(_,_,body) => nonrigid t (*not in normal form*)
-            | _ => occomb t)
+	| occur (Bound _)  = NoOcc
+	| occur (Free _)  = NoOcc
+	| occur (Var (w,_))  = 
+	    if w mem !seen then NoOcc
+	    else if v=w then Rigid
+	    else (seen := w:: !seen;  
+	          case  Envir.lookup(env,w)  of
+		      None    => NoOcc
+		    | Some t => occur t)
+	| occur (Abs(_,_,body)) = occur body
+	| occur (t as f$_) =  (*switch to nonrigid search?*)
+	   (case head_of_in (env,f) of
+	      Var (w,_) => (*w is not assigned*)
+		if v=w then Rigid  
+		else  nonrigid t
+	    | Abs(_,_,body) => nonrigid t (*not in normal form*)
+	    | _ => occomb t)
   in  occur t  end;
 
 
-exception CANTUNIFY;    (*Signals non-unifiability.  Does not signal errors!*)
-exception ASSIGN;       (*Raised if not an assignment*)
+exception CANTUNIFY;	(*Signals non-unifiability.  Does not signal errors!*)
+exception ASSIGN;	(*Raised if not an assignment*)
 
 
 fun unify_types(T,U, env as Envir.Envir{asol,iTs,maxidx}) =
@@ -261,8 +261,8 @@
   Result is var a for use in SIMPL. *)
 fun get_eta_var ([], _, Var vT)  =  vT
   | get_eta_var (_::rbinder, n, f $ Bound i) =
-        if  n=i  then  get_eta_var (rbinder, n+1, f) 
-                 else  raise ASSIGN
+	if  n=i  then  get_eta_var (rbinder, n+1, f) 
+		 else  raise ASSIGN
   | get_eta_var _ = raise ASSIGN;
 
 
@@ -277,11 +277,11 @@
 fun assignment (env, rbinder, t, u) =
     let val (v,T) = get_eta_var(rbinder,0,t)
     in  case rigid_occurs_term (ref[], env, v, u) of
-              NoOcc => let val env = unify_types(body_type env T,
-                                                 fastype env (rbinder,u),env)
-                in Envir.update ((v, rlist_abs(rbinder,u)), env) end
-            | Nonrigid =>  raise ASSIGN
-            | Rigid =>  raise CANTUNIFY
+	      NoOcc => let val env = unify_types(body_type env T,
+						 fastype env (rbinder,u),env)
+		in Envir.update ((v, rlist_abs(rbinder,u)), env) end
+	    | Nonrigid =>  raise ASSIGN
+	    | Rigid =>  raise CANTUNIFY
     end;
 
 
@@ -291,9 +291,9 @@
     if not, raises TERM, probably indicating type mismatch.
   Uses variable a (unless the null string) to preserve user's naming.*) 
 fun new_dpair (rbinder, Abs(a,T,body1), Abs(b,U,body2), env) =
-        let val env' = unify_types(T,U,env)
-            val c = if a="" then b else a
-        in new_dpair((c,T) :: rbinder, body1, body2, env') end
+	let val env' = unify_types(T,U,env)
+	    val c = if a="" then b else a
+	in new_dpair((c,T) :: rbinder, body1, body2, env') end
     | new_dpair (_, Abs _, _, _) = raise TERM ("new_dpair", [])
     | new_dpair (_, _, Abs _, _) = raise TERM ("new_dpair", [])
     | new_dpair (rbinder, t1, t2, env) = ((rbinder, t1, t2), env);
@@ -301,8 +301,8 @@
 
 fun head_norm_dpair (env, (rbinder,t,u)) : dpair * Envir.env =
      new_dpair (rbinder,
-                eta_norm env (rbinder, head_norm(env,t)),
-                eta_norm env (rbinder, head_norm(env,u)), env);
+		eta_norm env (rbinder, head_norm(env,t)),
+	  	eta_norm env (rbinder, head_norm(env,u)), env);
 
 
 
@@ -313,34 +313,34 @@
     do so caused numerous problems with no compensating advantage.
 *)
 fun SIMPL0 (dp0, (env,flexflex,flexrigid))
-        : Envir.env * dpair list * dpair list =
+	: Envir.env * dpair list * dpair list =
     let val (dp as (rbinder,t,u), env) = head_norm_dpair(env,dp0);
-            fun SIMRANDS(f$t, g$u, env) =
-                        SIMPL0((rbinder,t,u), SIMRANDS(f,g,env))
-              | SIMRANDS (t as _$_, _, _) =
-                raise TERM ("SIMPL: operands mismatch", [t,u])
-              | SIMRANDS (t, u as _$_, _) =
-                raise TERM ("SIMPL: operands mismatch", [t,u])
-              | SIMRANDS(_,_,env) = (env,flexflex,flexrigid);
+	    fun SIMRANDS(f$t, g$u, env) =
+			SIMPL0((rbinder,t,u), SIMRANDS(f,g,env))
+	      | SIMRANDS (t as _$_, _, _) =
+		raise TERM ("SIMPL: operands mismatch", [t,u])
+	      | SIMRANDS (t, u as _$_, _) =
+		raise TERM ("SIMPL: operands mismatch", [t,u])
+	      | SIMRANDS(_,_,env) = (env,flexflex,flexrigid);
     in case (head_of t, head_of u) of
        (Var(_,T), Var(_,U)) =>
-            let val T' = body_type env T and U' = body_type env U;
-                val env = unify_types(T',U',env)
-            in (env, dp::flexflex, flexrigid) end
+	    let val T' = body_type env T and U' = body_type env U;
+		val env = unify_types(T',U',env)
+	    in (env, dp::flexflex, flexrigid) end
      | (Var _, _) =>
-            ((assignment (env,rbinder,t,u), flexflex, flexrigid)
-             handle ASSIGN => (env, flexflex, dp::flexrigid))
+	    ((assignment (env,rbinder,t,u), flexflex, flexrigid)
+	     handle ASSIGN => (env, flexflex, dp::flexrigid))
      | (_, Var _) =>
-            ((assignment (env,rbinder,u,t), flexflex, flexrigid)
-             handle ASSIGN => (env, flexflex, (rbinder,u,t)::flexrigid))
+	    ((assignment (env,rbinder,u,t), flexflex, flexrigid)
+	     handle ASSIGN => (env, flexflex, (rbinder,u,t)::flexrigid))
      | (Const(a,T), Const(b,U)) =>
-            if a=b then SIMRANDS(t,u, unify_types(T,U,env))
-            else raise CANTUNIFY
+	    if a=b then SIMRANDS(t,u, unify_types(T,U,env))
+	    else raise CANTUNIFY
      | (Bound i,    Bound j)    =>
-            if i=j  then SIMRANDS(t,u,env) else raise CANTUNIFY
+	    if i=j  then SIMRANDS(t,u,env) else raise CANTUNIFY
      | (Free(a,T),  Free(b,U))  =>
-            if a=b then SIMRANDS(t,u, unify_types(T,U,env))
-            else raise CANTUNIFY
+	    if a=b then SIMRANDS(t,u, unify_types(T,U,env))
+	    else raise CANTUNIFY
      | _ => raise CANTUNIFY
     end;
 
@@ -356,8 +356,8 @@
   Clever would be to re-do just the affected dpairs*)
 fun SIMPL (env,dpairs) : Envir.env * dpair list * dpair list =
     let val all as (env',flexflex,flexrigid) =
-            foldr SIMPL0 (dpairs, (env,[],[]));
-        val dps = flexrigid@flexflex
+	    foldr SIMPL0 (dpairs, (env,[],[]));
+	val dps = flexrigid@flexflex
     in if exists (fn ((_,t,u)) => changed(env',t) orelse changed(env',u)) dps
        then SIMPL(env',dps) else all
     end;
@@ -376,7 +376,7 @@
 fun make_args name (binder: typ list, env, []) = (env, [])   (*frequent case*)
   | make_args name (binder: typ list, env, Ts) : Envir.env * term list =
        let fun funtype T = binder--->T;
-           val (env', vars) = Envir.genvars name (env, map funtype Ts)
+	   val (env', vars) = Envir.genvars name (env, map funtype Ts)
        in  (env',  map (fn var=> combound(var, 0, length binder)) vars)  end;
 
 
@@ -398,60 +398,60 @@
   The order for trying projections is crucial in ?b'(?a)   
   NB "vname" is only used in the call to make_args!!   *)
 fun matchcopy vname = let fun mc(rbinder, targs, u, ed as (env,dpairs)) 
-        : (term * (Envir.env * dpair list))Sequence.seq =
+	: (term * (Envir.env * dpair list))Sequence.seq =
 let (*Produce copies of uarg and cons them in front of uargs*)
     fun copycons uarg (uargs, (env, dpairs)) =
-        Sequence.maps(fn (uarg', ed') => (uarg'::uargs, ed'))
-            (mc (rbinder, targs,eta_norm env (rbinder,head_norm(env,uarg)),
-                 (env, dpairs)));
-        (*Produce sequence of all possible ways of copying the arg list*)
+	Sequence.maps(fn (uarg', ed') => (uarg'::uargs, ed'))
+	    (mc (rbinder, targs,eta_norm env (rbinder,head_norm(env,uarg)),
+		 (env, dpairs)));
+	(*Produce sequence of all possible ways of copying the arg list*)
     fun copyargs [] = Sequence.cons( ([],ed), Sequence.null)
       | copyargs (uarg::uargs) =
-            Sequence.flats (Sequence.maps (copycons uarg) (copyargs uargs));
+	    Sequence.flats (Sequence.maps (copycons uarg) (copyargs uargs));
     val (uhead,uargs) = strip_comb u;
     val base = body_type env (fastype env (rbinder,uhead));
     fun joinargs (uargs',ed') = (list_comb(uhead,uargs'), ed');
     (*attempt projection on argument with given typ*)
     val Ts = map (curry (fastype env) rbinder) targs;
     fun projenv (head, (Us,bary), targ, tail) = 
-        let val env = if !trace_types then test_unify_types(base,bary,env)
-                      else unify_types(base,bary,env)
-        in Sequence.seqof (fn () =>  
-            let val (env',args) = make_args vname (Ts,env,Us);
-                (*higher-order projection: plug in targs for bound vars*)
-                fun plugin arg = list_comb(head_of arg, targs);
-                val dp = (rbinder, list_comb(targ, map plugin args), u);
-                val (env2,frigid,fflex) = SIMPL (env', dp::dpairs)
-                    (*may raise exception CANTUNIFY*)
-            in  Some ((list_comb(head,args), (env2, frigid@fflex)),
-                        tail)
-            end  handle CANTUNIFY => Sequence.pull tail)
-        end handle CANTUNIFY => tail;
+	let val env = if !trace_types then test_unify_types(base,bary,env)
+		      else unify_types(base,bary,env)
+	in Sequence.seqof (fn () =>  
+	    let val (env',args) = make_args vname (Ts,env,Us);
+		(*higher-order projection: plug in targs for bound vars*)
+		fun plugin arg = list_comb(head_of arg, targs);
+		val dp = (rbinder, list_comb(targ, map plugin args), u);
+		val (env2,frigid,fflex) = SIMPL (env', dp::dpairs)
+		    (*may raise exception CANTUNIFY*)
+	    in  Some ((list_comb(head,args), (env2, frigid@fflex)),
+			tail)
+	    end  handle CANTUNIFY => Sequence.pull tail)
+	end handle CANTUNIFY => tail;
     (*make a list of projections*)
     fun make_projs (T::Ts, targ::targs) =
-              (Bound(length Ts), T, targ) :: make_projs (Ts,targs)
+	      (Bound(length Ts), T, targ) :: make_projs (Ts,targs)
       | make_projs ([],[]) = []
       | make_projs _ = raise TERM ("make_projs", u::targs);
     (*try projections and imitation*)
     fun matchfun ((bvar,T,targ)::projs) =
-               (projenv(bvar, strip_type env T, targ, matchfun projs))
+	       (projenv(bvar, strip_type env T, targ, matchfun projs))
       | matchfun [] = (*imitation last of all*)
-              (case uhead of
-                 Const _ => Sequence.maps joinargs (copyargs uargs)
-               | Free _  => Sequence.maps joinargs (copyargs uargs)
-               | _ => Sequence.null)  (*if Var, would be a loop!*)
+	      (case uhead of
+		 Const _ => Sequence.maps joinargs (copyargs uargs)
+	       | Free _  => Sequence.maps joinargs (copyargs uargs)
+	       | _ => Sequence.null)  (*if Var, would be a loop!*)
 in case uhead of
-        Abs(a, T, body) =>
-            Sequence.maps(fn (body', ed') => (Abs (a,T,body'), ed')) 
-                (mc ((a,T)::rbinder,
-                        (map (incr_boundvars 1) targs) @ [Bound 0], body, ed))
+	Abs(a, T, body) =>
+	    Sequence.maps(fn (body', ed') => (Abs (a,T,body'), ed')) 
+		(mc ((a,T)::rbinder,
+			(map (incr_boundvars 1) targs) @ [Bound 0], body, ed))
       | Var (w,uary) => 
-            (*a flex-flex dpair: make variable for t*)
-            let val (env', newhd) = Envir.genvar (#1 w) (env, Ts---> base)
-                val tabs = combound(newhd, 0, length Ts)
-                val tsub = list_comb(newhd,targs)
-            in  Sequence.single (tabs, (env', (rbinder,tsub,u):: dpairs)) 
-            end
+	    (*a flex-flex dpair: make variable for t*)
+	    let val (env', newhd) = Envir.genvar (#1 w) (env, Ts---> base)
+		val tabs = combound(newhd, 0, length Ts)
+		val tsub = list_comb(newhd,targs)
+	    in  Sequence.single (tabs, (env', (rbinder,tsub,u):: dpairs)) 
+	    end
       | _ =>  matchfun(rev(make_projs(Ts, targs)))
 end
 in mc end;
@@ -459,14 +459,14 @@
 
 (*Call matchcopy to produce assignments to the variable in the dpair*)
 fun MATCH (env, (rbinder,t,u), dpairs)
-        : (Envir.env * dpair list)Sequence.seq = 
+	: (Envir.env * dpair list)Sequence.seq = 
   let val (Var(v,T), targs) = strip_comb t;
       val Ts = binder_types env T;
       fun new_dset (u', (env',dpairs')) =
-          (*if v was updated to s, must unify s with u' *)
-          case Envir.lookup(env',v) of
-              None => (Envir.update ((v, types_abs(Ts, u')), env'),  dpairs')
-            | Some s => (env', ([], s, types_abs(Ts, u'))::dpairs')
+	  (*if v was updated to s, must unify s with u' *)
+	  case Envir.lookup(env',v) of
+	      None => (Envir.update ((v, types_abs(Ts, u')), env'),  dpairs')
+	    | Some s => (env', ([], s, types_abs(Ts, u'))::dpairs')
   in Sequence.maps new_dset
          (matchcopy (#1 v) (rbinder, targs, u, (env,dpairs)))
   end;
@@ -481,9 +481,9 @@
 let val (v,T) = get_eta_var(rbinder,0,t)
 in if occurs_terms (ref[], env, v, [u]) then raise ASSIGN
    else let val env = unify_types(body_type env T,
-                                  fastype env (rbinder,u),
-                                  env)
-        in Envir.vupdate ((v, rlist_abs(rbinder, u)), env) end
+				  fastype env (rbinder,u),
+				  env)
+	in Envir.vupdate ((v, rlist_abs(rbinder, u)), env) end
 end;
 
 
@@ -508,30 +508,30 @@
   let val (head,args) = strip_comb t 
   in  
       case head of
-          Bound i => (i-lev) mem banned  orelse
-                     exists (rigid_bound (lev, banned)) args
-        | Var _ => false        (*no rigid occurrences here!*)
-        | Abs (_,_,u) => 
-               rigid_bound(lev+1, banned) u  orelse
-               exists (rigid_bound (lev, banned)) args
-        | _ => exists (rigid_bound (lev, banned)) args
+	  Bound i => (i-lev) mem banned  orelse
+	      	     exists (rigid_bound (lev, banned)) args
+	| Var _ => false	(*no rigid occurrences here!*)
+	| Abs (_,_,u) => 
+	       rigid_bound(lev+1, banned) u  orelse
+	       exists (rigid_bound (lev, banned)) args
+	| _ => exists (rigid_bound (lev, banned)) args
   end;
 
 (*Squash down indices at level >=lev to delete the banned from a term.*)
 fun change_bnos banned =
   let fun change lev (Bound i) = 
-            if i<lev then Bound i
-            else  if (i-lev) mem banned  
-                  then raise CHANGE_FAIL (**flexible occurrence: give up**)
-            else  Bound (i - length (filter (fn j => j < i-lev) banned))
-        | change lev (Abs (a,T,t)) = Abs (a, T, change(lev+1) t)
-        | change lev (t$u) = change lev t $ change lev u
-        | change lev t = t
+	    if i<lev then Bound i
+	    else  if (i-lev) mem banned  
+		  then raise CHANGE_FAIL (**flexible occurrence: give up**)
+	    else  Bound (i - length (filter (fn j => j < i-lev) banned))
+	| change lev (Abs (a,T,t)) = Abs (a, T, change(lev+1) t)
+	| change lev (t$u) = change lev t $ change lev u
+	| change lev t = t
   in  change 0  end;
 
 (*Change indices, delete the argument if it contains a banned Bound*)
 fun change_arg banned ({j,t,T}, args) : flarg list =
-    if rigid_bound (0, banned) t  then  args    (*delete argument!*)
+    if rigid_bound (0, banned) t  then  args	(*delete argument!*)
     else  {j=j, t= change_bnos banned t, T=T} :: args;
 
 
@@ -550,19 +550,19 @@
   Update its head; squash indices in arguments. *)
 fun clean_term banned (env,t) =
     let val (Var(v,T), ts) = strip_comb t
-        val (Ts,U) = strip_type env T
-        and js = length ts - 1  downto 0
-        val args = sort arg_less
-                (foldr (change_arg banned) (flexargs (js,ts,Ts), []))
-        val ts' = map (#t) args
+	val (Ts,U) = strip_type env T
+	and js = length ts - 1  downto 0
+	val args = sort arg_less
+		(foldr (change_arg banned) (flexargs (js,ts,Ts), []))
+	val ts' = map (#t) args
     in
     if decreasing (length Ts) args then (env, (list_comb(Var(v,T), ts')))
     else let val (env',v') = Envir.genvar (#1v) (env, map (#T) args ---> U)
-             val body = list_comb(v', map (Bound o #j) args)
-             val env2 = Envir.vupdate (((v, types_abs(Ts, body)),   env'))
-             (*the vupdate affects ts' if they contain v*)
-         in  
-             (env2, Envir.norm_term env2 (list_comb(v',ts')))
+	     val body = list_comb(v', map (Bound o #j) args)
+	     val env2 = Envir.vupdate (((v, types_abs(Ts, body)),   env'))
+	     (*the vupdate affects ts' if they contain v*)
+	 in  
+	     (env2, Envir.norm_term env2 (list_comb(v',ts')))
          end
     end;
 
@@ -583,8 +583,8 @@
   let val loot = loose_bnos t  and  loou = loose_bnos u
       fun add_index (((a,T), j), (bnos, newbinder)) = 
             if  j mem loot  andalso  j mem loou 
-            then  (bnos, (a,T)::newbinder)      (*needed by both: keep*)
-            else  (j::bnos, newbinder);         (*remove*)
+            then  (bnos, (a,T)::newbinder)	(*needed by both: keep*)
+            else  (j::bnos, newbinder);		(*remove*)
       val indices = 0 upto (length rbinder - 1);
       val (banned,rbin') = foldr add_index (rbinder~~indices, ([],[]));
       val (env', t') = clean_term banned (env, t);
@@ -605,11 +605,11 @@
   let val t = Envir.norm_term env t0  and  u = Envir.norm_term env u0
   in  case  (head_of t, head_of u) of
       (Var(v,T), Var(w,U)) =>  (*Check for identical variables...*)
-        if v=w then             (*...occur check would falsely return true!*)
-            if T=U then (env, add_tpair (rbinder, (t,u), tpairs))
-            else raise TERM ("add_ffpair: Var name confusion", [t,u])
-        else if xless(v,w) then (*prefer to update the LARGER variable*)
-             clean_ffpair ((rbinder, u, t), (env,tpairs))
+	if v=w then		(*...occur check would falsely return true!*)
+	    if T=U then (env, add_tpair (rbinder, (t,u), tpairs))
+	    else raise TERM ("add_ffpair: Var name confusion", [t,u])
+	else if xless(v,w) then (*prefer to update the LARGER variable*)
+	     clean_ffpair ((rbinder, u, t), (env,tpairs))
         else clean_ffpair ((rbinder, t, u), (env,tpairs))
     | _ => raise TERM ("add_ffpair: Vars expected", [t,u])
   end;
@@ -634,26 +634,26 @@
 fun hounifiers (sg,env, tus : (term*term)list) 
   : (Envir.env * (term*term)list)Sequence.seq =
   let fun add_unify tdepth ((env,dpairs), reseq) =
-          Sequence.seqof (fn()=>
-          let val (env',flexflex,flexrigid) = 
-               (if tdepth> !trace_bound andalso !trace_simp
-                then print_dpairs "Enter SIMPL" (env,dpairs)  else ();
-                SIMPL (env,dpairs))
-          in case flexrigid of
-              [] => Some (foldr add_ffpair (flexflex, (env',[])), reseq)
-            | dp::frigid' => 
-                if tdepth > !search_bound then
-                    (prs"***Unification bound exceeded\n"; Sequence.pull reseq)
-                else
-                (if tdepth > !trace_bound then
-                    print_dpairs "Enter MATCH" (env',flexrigid@flexflex)
-                 else ();
-                 Sequence.pull (Sequence.its_right (add_unify (tdepth+1))
-                           (MATCH (env',dp, frigid'@flexflex), reseq)))
-          end
-          handle CANTUNIFY => 
-            (if tdepth > !trace_bound then writeln"Failure node" else ();
-             Sequence.pull reseq));
+	  Sequence.seqof (fn()=>
+	  let val (env',flexflex,flexrigid) = 
+	       (if tdepth> !trace_bound andalso !trace_simp
+		then print_dpairs "Enter SIMPL" (env,dpairs)  else ();
+		SIMPL (env,dpairs))
+	  in case flexrigid of
+	      [] => Some (foldr add_ffpair (flexflex, (env',[])), reseq)
+	    | dp::frigid' => 
+		if tdepth > !search_bound then
+		    (prs"***Unification bound exceeded\n"; Sequence.pull reseq)
+		else
+		(if tdepth > !trace_bound then
+		    print_dpairs "Enter MATCH" (env',flexrigid@flexflex)
+		 else ();
+		 Sequence.pull (Sequence.its_right (add_unify (tdepth+1))
+			   (MATCH (env',dp, frigid'@flexflex), reseq)))
+	  end
+	  handle CANTUNIFY => 
+	    (if tdepth > !trace_bound then writeln"Failure node" else ();
+	     Sequence.pull reseq));
      val dps = map (fn(t,u)=> ([],t,u)) tus
   in sgr := sg;
      add_unify 1 ((env,dps), Sequence.null) 
@@ -675,7 +675,7 @@
 (*Eliminate a flex-flex pair by the trivial substitution, see Huet (1975)
   Unifies ?f(t1...rm) with ?g(u1...un) by ?f -> %x1...xm.?a, ?g -> %x1...xn.?a
   Unfortunately, unifies ?f(t,u) with ?g(t,u) by ?f, ?g -> %(x,y)?a, 
-        though just ?g->?f is a more general unifier.
+	though just ?g->?f is a more general unifier.
   Unlike Huet (1975), does not smash together all variables of same type --
     requires more work yet gives a less general unifier (fewer variables).
   Handles ?f(t1...rm) with ?f(u1...um) to avoid multiple updates. *)