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