| author | paulson | 
| Fri, 10 Mar 2000 17:51:59 +0100 | |
| changeset 8413 | 09db77a084aa | 
| parent 8407 | d522ad1809e9 | 
| child 10485 | f1576723371f | 
| permissions | -rw-r--r-- | 
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 1 | (* Title: Pure/envir.ML | 
| 0 | 2 | ID: $Id$ | 
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1988 University of Cambridge | 
| 5 | *) | |
| 6 | ||
| 2142 
20f208ff085d
Deleted Olist constructor.  Replaced minidx by "above" function
 paulson parents: 
1500diff
changeset | 7 | (*Environments. They don't take account that typ is part of variable name. | 
| 
20f208ff085d
Deleted Olist constructor.  Replaced minidx by "above" function
 paulson parents: 
1500diff
changeset | 8 | Therefore we check elsewhere that two variables with same | 
| 
20f208ff085d
Deleted Olist constructor.  Replaced minidx by "above" function
 paulson parents: 
1500diff
changeset | 9 | names and different types cannot occur. | 
| 0 | 10 | *) | 
| 11 | ||
| 12 | ||
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 13 | signature ENVIR = | 
| 0 | 14 | sig | 
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 15 |   datatype env = Envir of {asol: term Vartab.table, iTs: typ Vartab.table, maxidx: int}
 | 
| 1460 | 16 | val alist_of : env -> (indexname * term) list | 
| 17 | val empty : int->env | |
| 18 | val is_empty : env -> bool | |
| 2142 
20f208ff085d
Deleted Olist constructor.  Replaced minidx by "above" function
 paulson parents: 
1500diff
changeset | 19 | val above : int * env -> bool | 
| 1460 | 20 | val genvar : string -> env * typ -> env * term | 
| 21 | val genvars : string -> env * typ list -> env * term list | |
| 22 | val lookup : env * indexname -> term option | |
| 23 | val norm_term : env -> term -> term | |
| 24 | val update : (indexname * term) * env -> env | |
| 25 | val vupdate : (indexname * term) * env -> env | |
| 0 | 26 | end; | 
| 27 | ||
| 1500 | 28 | structure Envir : ENVIR = | 
| 0 | 29 | struct | 
| 30 | ||
| 31 | (*updating can destroy environment in 2 ways!! | |
| 32 | (1) variables out of range (2) circular assignments | |
| 33 | *) | |
| 34 | datatype env = Envir of | |
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 35 |     {maxidx: int,               (*maximum index of vars*)
 | 
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 36 | asol: term Vartab.table, (*table of assignments to Vars*) | 
| 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 37 | iTs: typ Vartab.table} (*table of assignments to TVars*) | 
| 0 | 38 | |
| 39 | ||
| 40 | (*Generate a list of distinct variables. | |
| 41 | Increments index to make them distinct from ALL present variables. *) | |
| 42 | fun genvars name (Envir{maxidx, asol, iTs}, Ts) : env * term list =
 | |
| 43 | let fun genvs (_, [] : typ list) : term list = [] | |
| 44 | | genvs (n, [T]) = [ Var((name, maxidx+1), T) ] | |
| 45 | | genvs (n, T::Ts) = | |
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 46 | Var((name ^ radixstring(26,"a",n), maxidx+1), T) | 
| 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 47 | :: genvs(n+1,Ts) | 
| 0 | 48 |   in  (Envir{maxidx=maxidx+1, asol=asol, iTs=iTs}, genvs (0,Ts))  end;
 | 
| 49 | ||
| 50 | (*Generate a variable.*) | |
| 51 | fun genvar name (env,T) : env * term = | |
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 52 | let val (env',[v]) = genvars name (env,[T]) | 
| 0 | 53 | in (env',v) end; | 
| 54 | ||
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 55 | fun lookup (Envir{asol,...}, xname) : term option = Vartab.lookup (asol, xname);
 | 
| 0 | 56 | |
| 57 | fun update ((xname, t), Envir{maxidx, asol, iTs}) =
 | |
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 58 |   Envir{maxidx=maxidx, asol=Vartab.update_new ((xname,t), asol), iTs=iTs};
 | 
| 0 | 59 | |
| 5289 | 60 | (*The empty environment. New variables will start with the given index+1.*) | 
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 61 | fun empty m = Envir{maxidx=m, asol=Vartab.empty, iTs=Vartab.empty};
 | 
| 0 | 62 | |
| 2142 
20f208ff085d
Deleted Olist constructor.  Replaced minidx by "above" function
 paulson parents: 
1500diff
changeset | 63 | (*Test for empty environment*) | 
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 64 | fun is_empty (Envir {asol, iTs, ...}) = Vartab.is_empty asol andalso Vartab.is_empty iTs;
 | 
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 65 | |
| 2142 
20f208ff085d
Deleted Olist constructor.  Replaced minidx by "above" function
 paulson parents: 
1500diff
changeset | 66 | (*Determine if the least index updated exceeds lim*) | 
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 67 | fun above (lim, Envir {asol, iTs, ...}) =
 | 
| 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 68 | (case (Vartab.min_key asol, Vartab.min_key iTs) of | 
| 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 69 | (None, None) => true | 
| 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 70 | | (Some (_, i), None) => i > lim | 
| 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 71 | | (None, Some (_, i')) => i' > lim | 
| 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 72 | | (Some (_, i), Some (_, i')) => i > lim andalso i' > lim); | 
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 73 | |
| 0 | 74 | (*Update, checking Var-Var assignments: try to suppress higher indexes*) | 
| 75 | fun vupdate((a,t), env) = case t of | |
| 76 | Var(name',T) => | |
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 77 | if a=name' then env (*cycle!*) | 
| 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 78 | else if xless(a, name') then | 
| 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 79 | (case lookup(env,name') of (*if already assigned, chase*) | 
| 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 80 | None => update((name', Var(a,T)), env) | 
| 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 81 | | Some u => vupdate((a,u), env)) | 
| 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 82 | else update((a,t), env) | 
| 0 | 83 | | _ => update((a,t), env); | 
| 84 | ||
| 85 | ||
| 86 | (*Convert environment to alist*) | |
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 87 | fun alist_of (Envir{asol,...}) = Vartab.dest asol;
 | 
| 0 | 88 | |
| 89 | ||
| 1500 | 90 | (*** Beta normal form for terms (not eta normal form). | 
| 91 | Chases variables in env; Does not exploit sharing of variable bindings | |
| 92 | Does not check types, so could loop. ***) | |
| 93 | ||
| 94 | (*raised when norm has no effect on a term, to do sharing instead of copying*) | |
| 95 | exception SAME; | |
| 0 | 96 | |
| 1500 | 97 | fun norm_term1 (asol,t) : term = | 
| 98 | let fun norm (Var (w,T)) = | |
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 99 | (case Vartab.lookup (asol, w) of | 
| 2191 | 100 | Some u => (norm u handle SAME => u) | 
| 1500 | 101 | | None => raise SAME) | 
| 102 | | norm (Abs(a,T,body)) = Abs(a, T, norm body) | |
| 2191 | 103 | | norm (Abs(_,_,body) $ t) = normh(subst_bound (t, body)) | 
| 1500 | 104 | | norm (f $ t) = | 
| 105 | ((case norm f of | |
| 2191 | 106 | Abs(_,_,body) => normh(subst_bound (t, body)) | 
| 107 | | nf => nf $ (norm t handle SAME => t)) | |
| 1500 | 108 | handle SAME => f $ norm t) | 
| 109 | | norm _ = raise SAME | |
| 2191 | 110 | and normh t = norm t handle SAME => t | 
| 1500 | 111 | in normh t end | 
| 0 | 112 | |
| 1500 | 113 | and norm_term2(asol,iTs,t) : term = | 
| 114 | let fun normT(Type(a,Ts)) = Type(a, normTs Ts) | |
| 115 | | normT(TFree _) = raise SAME | |
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 116 | | normT(TVar(v,_)) = (case Vartab.lookup (iTs, v) of | 
| 1500 | 117 | Some(U) => normTh U | 
| 118 | | None => raise SAME) | |
| 119 | and normTh T = ((normT T) handle SAME => T) | |
| 120 | and normTs([]) = raise SAME | |
| 2181 | 121 | | normTs(T::Ts) = ((normT T :: (normTs Ts handle SAME => Ts)) | 
| 1500 | 122 | handle SAME => T :: normTs Ts) | 
| 123 | and norm(Const(a,T)) = Const(a, normT T) | |
| 124 | | norm(Free(a,T)) = Free(a, normT T) | |
| 125 | | norm(Var (w,T)) = | |
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 126 | (case Vartab.lookup (asol, w) of | 
| 1500 | 127 | Some u => normh u | 
| 128 | | None => Var(w,normT T)) | |
| 129 | | norm(Abs(a,T,body)) = | |
| 130 | (Abs(a,normT T,normh body) handle SAME => Abs(a, T, norm body)) | |
| 2191 | 131 | | norm(Abs(_,_,body) $ t) = normh(subst_bound (t, body)) | 
| 1500 | 132 | | norm(f $ t) = | 
| 133 | ((case norm f of | |
| 2191 | 134 | Abs(_,_,body) => normh(subst_bound (t, body)) | 
| 1500 | 135 | | nf => nf $ normh t) | 
| 136 | handle SAME => f $ norm t) | |
| 137 | | norm _ = raise SAME | |
| 138 | and normh t = (norm t) handle SAME => t | |
| 139 | in normh t end; | |
| 0 | 140 | |
| 141 | (*curried version might be slower in recursive calls*) | |
| 142 | fun norm_term (env as Envir{asol,iTs,...}) t : term =
 | |
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 143 | if Vartab.is_empty iTs then norm_term1(asol, t) else norm_term2(asol,iTs, t) | 
| 719 
e3e1d1a6d408
Pure/envir/norm_term: replaced equality test for [] by null
 lcp parents: 
247diff
changeset | 144 | |
| 0 | 145 | end; | 
| 146 |