| author | haftmann | 
| Wed, 18 Feb 2009 19:18:32 +0100 | |
| changeset 29970 | cbf46080ea3a | 
| parent 29269 | 5c25a2012975 | 
| child 30146 | a77fc0209723 | 
| permissions | -rw-r--r-- | 
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 1 | (* Title: Pure/envir.ML | 
| 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 10485 | 3 | |
| 15797 | 4 | Environments. The type of a term variable / sort of a type variable is | 
| 5 | part of its name. The lookup function must apply type substitutions, | |
| 6 | since they may change the identity of a variable. | |
| 0 | 7 | *) | 
| 8 | ||
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 9 | signature ENVIR = | 
| 0 | 10 | sig | 
| 15797 | 11 | type tenv | 
| 12 |   datatype env = Envir of {asol: tenv, iTs: Type.tyenv, maxidx: int}
 | |
| 13 | val type_env: env -> Type.tyenv | |
| 26638 | 14 | val insert_sorts: env -> sort list -> sort list | 
| 11513 | 15 | exception SAME | 
| 10485 | 16 | val genvars: string -> env * typ list -> env * term list | 
| 17 | val genvar: string -> env * typ -> env * term | |
| 15797 | 18 | val lookup: env * (indexname * typ) -> term option | 
| 16652 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 berghofe parents: 
15797diff
changeset | 19 | val lookup': tenv * (indexname * typ) -> term option | 
| 15797 | 20 | val update: ((indexname * typ) * term) * env -> env | 
| 10485 | 21 | val empty: int -> env | 
| 22 | val is_empty: env -> bool | |
| 19861 | 23 | val above: env -> int -> bool | 
| 15797 | 24 | val vupdate: ((indexname * typ) * term) * env -> env | 
| 25 | val alist_of: env -> (indexname * (typ * term)) list | |
| 10485 | 26 | val norm_term: env -> term -> term | 
| 11513 | 27 | val norm_term_same: env -> term -> term | 
| 15797 | 28 | val norm_type: Type.tyenv -> typ -> typ | 
| 29 | val norm_type_same: Type.tyenv -> typ -> typ | |
| 30 | val norm_types_same: Type.tyenv -> typ list -> typ list | |
| 10485 | 31 | val beta_norm: term -> term | 
| 12231 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 32 | val head_norm: env -> term -> term | 
| 18937 | 33 | val eta_contract: term -> term | 
| 34 | val beta_eta_contract: term -> term | |
| 12231 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 35 | val fastype: env -> typ list -> term -> typ | 
| 15797 | 36 | val typ_subst_TVars: Type.tyenv -> typ -> typ | 
| 37 | val subst_TVars: Type.tyenv -> term -> term | |
| 38 | val subst_Vars: tenv -> term -> term | |
| 39 | val subst_vars: Type.tyenv * tenv -> term -> term | |
| 19422 | 40 | val expand_atom: typ -> typ * term -> term | 
| 21695 | 41 | val expand_term: (term -> (typ * term) option) -> term -> term | 
| 21795 | 42 | val expand_term_frees: ((string * typ) * term) list -> term -> term | 
| 0 | 43 | end; | 
| 44 | ||
| 1500 | 45 | structure Envir : ENVIR = | 
| 0 | 46 | struct | 
| 47 | ||
| 48 | (*updating can destroy environment in 2 ways!! | |
| 49 | (1) variables out of range (2) circular assignments | |
| 50 | *) | |
| 15797 | 51 | type tenv = (typ * term) Vartab.table | 
| 52 | ||
| 0 | 53 | datatype env = Envir of | 
| 15797 | 54 |     {maxidx: int,      (*maximum index of vars*)
 | 
| 55 | asol: tenv, (*table of assignments to Vars*) | |
| 56 | iTs: Type.tyenv} (*table of assignments to TVars*) | |
| 0 | 57 | |
| 12496 | 58 | fun type_env (Envir {iTs, ...}) = iTs;
 | 
| 0 | 59 | |
| 26638 | 60 | (*NB: type unification may invent new sorts*) | 
| 61 | val insert_sorts = Vartab.fold (fn (_, (_, T)) => Sorts.insert_typ T) o type_env; | |
| 62 | ||
| 0 | 63 | (*Generate a list of distinct variables. | 
| 64 | Increments index to make them distinct from ALL present variables. *) | |
| 65 | fun genvars name (Envir{maxidx, asol, iTs}, Ts) : env * term list =
 | |
| 66 | let fun genvs (_, [] : typ list) : term list = [] | |
| 67 | | genvs (n, [T]) = [ Var((name, maxidx+1), T) ] | |
| 68 | | genvs (n, T::Ts) = | |
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 69 | Var((name ^ radixstring(26,"a",n), maxidx+1), T) | 
| 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 70 | :: genvs(n+1,Ts) | 
| 0 | 71 |   in  (Envir{maxidx=maxidx+1, asol=asol, iTs=iTs}, genvs (0,Ts))  end;
 | 
| 72 | ||
| 73 | (*Generate a variable.*) | |
| 74 | fun genvar name (env,T) : env * term = | |
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 75 | let val (env',[v]) = genvars name (env,[T]) | 
| 0 | 76 | in (env',v) end; | 
| 77 | ||
| 15797 | 78 | fun var_clash ixn T T' = raise TYPE ("Variable " ^
 | 
| 22678 | 79 | quote (Term.string_of_vname ixn) ^ " has two distinct types", | 
| 15797 | 80 | [T', T], []); | 
| 0 | 81 | |
| 15797 | 82 | fun gen_lookup f asol (xname, T) = | 
| 17412 | 83 | (case Vartab.lookup asol xname of | 
| 15797 | 84 | NONE => NONE | 
| 85 | | SOME (U, t) => if f (T, U) then SOME t | |
| 86 | else var_clash xname T U); | |
| 87 | ||
| 16652 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 berghofe parents: 
15797diff
changeset | 88 | (* When dealing with environments produced by matching instead *) | 
| 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 berghofe parents: 
15797diff
changeset | 89 | (* of unification, there is no need to chase assigned TVars. *) | 
| 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 berghofe parents: 
15797diff
changeset | 90 | (* In this case, we can simply ignore the type substitution *) | 
| 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 berghofe parents: 
15797diff
changeset | 91 | (* and use = instead of eq_type. *) | 
| 15797 | 92 | |
| 16652 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 berghofe parents: 
15797diff
changeset | 93 | fun lookup' (asol, p) = gen_lookup op = asol p; | 
| 15797 | 94 | |
| 16652 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 berghofe parents: 
15797diff
changeset | 95 | fun lookup2 (iTs, asol) p = | 
| 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 berghofe parents: 
15797diff
changeset | 96 | if Vartab.is_empty iTs then lookup' (asol, p) | 
| 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 berghofe parents: 
15797diff
changeset | 97 | else gen_lookup (Type.eq_type iTs) asol p; | 
| 15797 | 98 | |
| 99 | fun lookup (Envir {asol, iTs, ...}, p) = lookup2 (iTs, asol) p;
 | |
| 100 | ||
| 101 | fun update (((xname, T), t), Envir {maxidx, asol, iTs}) =
 | |
| 17412 | 102 |   Envir{maxidx=maxidx, asol=Vartab.update_new (xname, (T, t)) asol, iTs=iTs};
 | 
| 0 | 103 | |
| 5289 | 104 | (*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 | 105 | fun empty m = Envir{maxidx=m, asol=Vartab.empty, iTs=Vartab.empty};
 | 
| 0 | 106 | |
| 2142 
20f208ff085d
Deleted Olist constructor.  Replaced minidx by "above" function
 paulson parents: 
1500diff
changeset | 107 | (*Test for empty environment*) | 
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 108 | 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 | 109 | |
| 2142 
20f208ff085d
Deleted Olist constructor.  Replaced minidx by "above" function
 paulson parents: 
1500diff
changeset | 110 | (*Determine if the least index updated exceeds lim*) | 
| 19861 | 111 | fun above (Envir {asol, iTs, ...}) lim =
 | 
| 112 | (case Vartab.min_key asol of SOME (_, i) => i > lim | NONE => true) andalso | |
| 113 | (case Vartab.min_key iTs of SOME (_, i) => i > lim | NONE => true); | |
| 247 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 wenzelm parents: 
0diff
changeset | 114 | |
| 0 | 115 | (*Update, checking Var-Var assignments: try to suppress higher indexes*) | 
| 15797 | 116 | fun vupdate ((aU as (a, U), t), env as Envir {iTs, ...}) = case t of
 | 
| 117 | Var (nT as (name', T)) => | |
| 118 | if a = name' then env (*cycle!*) | |
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
26638diff
changeset | 119 | else if TermOrd.indexname_ord (a, name') = LESS then | 
| 15797 | 120 | (case lookup (env, nT) of (*if already assigned, chase*) | 
| 121 | NONE => update ((nT, Var (a, T)), env) | |
| 122 | | SOME u => vupdate ((aU, u), env)) | |
| 123 | else update ((aU, t), env) | |
| 124 | | _ => update ((aU, t), env); | |
| 0 | 125 | |
| 126 | ||
| 127 | (*Convert environment to alist*) | |
| 8407 
d522ad1809e9
Envir now uses Vartab instead of association lists.
 berghofe parents: 
5289diff
changeset | 128 | fun alist_of (Envir{asol,...}) = Vartab.dest asol;
 | 
| 0 | 129 | |
| 130 | ||
| 1500 | 131 | (*** Beta normal form for terms (not eta normal form). | 
| 132 | Chases variables in env; Does not exploit sharing of variable bindings | |
| 133 | Does not check types, so could loop. ***) | |
| 134 | ||
| 135 | (*raised when norm has no effect on a term, to do sharing instead of copying*) | |
| 136 | exception SAME; | |
| 0 | 137 | |
| 11513 | 138 | fun norm_term1 same (asol,t) : term = | 
| 15797 | 139 | let fun norm (Var wT) = | 
| 16652 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 berghofe parents: 
15797diff
changeset | 140 | (case lookup' (asol, wT) of | 
| 15531 | 141 | SOME u => (norm u handle SAME => u) | 
| 142 | | NONE => raise SAME) | |
| 10485 | 143 | | norm (Abs(a,T,body)) = Abs(a, T, norm body) | 
| 144 | | norm (Abs(_,_,body) $ t) = normh(subst_bound (t, body)) | |
| 145 | | norm (f $ t) = | |
| 146 | ((case norm f of | |
| 147 | Abs(_,_,body) => normh(subst_bound (t, body)) | |
| 148 | | nf => nf $ (norm t handle SAME => t)) | |
| 149 | handle SAME => f $ norm t) | |
| 150 | | norm _ = raise SAME | |
| 2191 | 151 | and normh t = norm t handle SAME => t | 
| 11513 | 152 | in (if same then norm else normh) t end | 
| 0 | 153 | |
| 11513 | 154 | fun normT iTs (Type (a, Ts)) = Type (a, normTs iTs Ts) | 
| 155 | | normT iTs (TFree _) = raise SAME | |
| 26328 | 156 | | normT iTs (TVar vS) = (case Type.lookup iTs vS of | 
| 15531 | 157 | SOME U => normTh iTs U | 
| 158 | | NONE => raise SAME) | |
| 11513 | 159 | and normTh iTs T = ((normT iTs T) handle SAME => T) | 
| 160 | and normTs iTs [] = raise SAME | |
| 161 | | normTs iTs (T :: Ts) = | |
| 162 | ((normT iTs T :: (normTs iTs Ts handle SAME => Ts)) | |
| 163 | handle SAME => T :: normTs iTs Ts); | |
| 164 | ||
| 165 | fun norm_term2 same (asol, iTs, t) : term = | |
| 166 | let fun norm (Const (a, T)) = Const(a, normT iTs T) | |
| 167 | | norm (Free (a, T)) = Free(a, normT iTs T) | |
| 168 | | norm (Var (w, T)) = | |
| 15797 | 169 | (case lookup2 (iTs, asol) (w, T) of | 
| 15531 | 170 | SOME u => normh u | 
| 171 | | NONE => Var(w, normT iTs T)) | |
| 11513 | 172 | | norm (Abs (a, T, body)) = | 
| 173 | (Abs (a, normT iTs T, normh body) handle SAME => Abs (a, T, norm body)) | |
| 174 | | norm (Abs(_, _, body) $ t) = normh (subst_bound (t, body)) | |
| 175 | | norm (f $ t) = | |
| 10485 | 176 | ((case norm f of | 
| 11513 | 177 | Abs(_, _, body) => normh (subst_bound (t, body)) | 
| 10485 | 178 | | nf => nf $ normh t) | 
| 179 | handle SAME => f $ norm t) | |
| 180 | | norm _ = raise SAME | |
| 1500 | 181 | and normh t = (norm t) handle SAME => t | 
| 11513 | 182 | in (if same then norm else normh) t end; | 
| 0 | 183 | |
| 11513 | 184 | fun gen_norm_term same (env as Envir{asol,iTs,...}) t : term =
 | 
| 185 | if Vartab.is_empty iTs then norm_term1 same (asol, t) | |
| 186 | else norm_term2 same (asol, iTs, t); | |
| 187 | ||
| 188 | val norm_term = gen_norm_term false; | |
| 189 | val norm_term_same = gen_norm_term true; | |
| 10485 | 190 | |
| 25471 
ca009b7ce693
Optimized beta_norm: only tries to normalize term when it contains
 berghofe parents: 
24670diff
changeset | 191 | fun beta_norm t = if Term.has_abs t then norm_term (empty 0) t else t; | 
| 719 
e3e1d1a6d408
Pure/envir/norm_term: replaced equality test for [] by null
 lcp parents: 
247diff
changeset | 192 | |
| 12496 | 193 | fun norm_type iTs = normTh iTs; | 
| 194 | fun norm_type_same iTs = | |
| 11513 | 195 | if Vartab.is_empty iTs then raise SAME else normT iTs; | 
| 196 | ||
| 12496 | 197 | fun norm_types_same iTs = | 
| 11513 | 198 | if Vartab.is_empty iTs then raise SAME else normTs iTs; | 
| 199 | ||
| 200 | ||
| 12231 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 201 | (*Put a term into head normal form for unification.*) | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 202 | |
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 203 | fun head_norm env t = | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 204 | let | 
| 15797 | 205 | fun hnorm (Var vT) = (case lookup (env, vT) of | 
| 15531 | 206 | SOME u => head_norm env u | 
| 207 | | NONE => raise SAME) | |
| 12231 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 208 | | hnorm (Abs (a, T, body)) = Abs (a, T, hnorm body) | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 209 | | hnorm (Abs (_, _, body) $ t) = | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 210 | head_norm env (subst_bound (t, body)) | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 211 | | hnorm (f $ t) = (case hnorm f of | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 212 | Abs (_, _, body) => head_norm env (subst_bound (t, body)) | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 213 | | nf => nf $ t) | 
| 20670 | 214 | | hnorm _ = raise SAME | 
| 12231 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 215 | in hnorm t handle SAME => t end; | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 216 | |
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 217 | |
| 18937 | 218 | (*Eta-contract a term (fully)*) | 
| 219 | ||
| 22174 | 220 | local | 
| 221 | ||
| 222 | fun decr lev (Bound i) = if i >= lev then Bound (i - 1) else raise SAME | |
| 223 | | decr lev (Abs (a, T, body)) = Abs (a, T, decr (lev + 1) body) | |
| 224 | | decr lev (t $ u) = (decr lev t $ decrh lev u handle SAME => t $ decr lev u) | |
| 225 | | decr _ _ = raise SAME | |
| 226 | and decrh lev t = (decr lev t handle SAME => t); | |
| 20670 | 227 | |
| 22174 | 228 | fun eta (Abs (a, T, body)) = | 
| 229 | ((case eta body of | |
| 230 | body' as (f $ Bound 0) => | |
| 231 | if loose_bvar1 (f, 0) then Abs (a, T, body') | |
| 232 | else decrh 0 f | |
| 233 | | body' => Abs (a, T, body')) handle SAME => | |
| 234 | (case body of | |
| 235 | f $ Bound 0 => | |
| 236 | if loose_bvar1 (f, 0) then raise SAME | |
| 237 | else decrh 0 f | |
| 238 | | _ => raise SAME)) | |
| 239 | | eta (t $ u) = (eta t $ etah u handle SAME => t $ eta u) | |
| 240 | | eta _ = raise SAME | |
| 241 | and etah t = (eta t handle SAME => t); | |
| 242 | ||
| 243 | in | |
| 244 | ||
| 245 | fun eta_contract t = | |
| 24670 | 246 | if Term.has_abs t then etah t else t; | 
| 18937 | 247 | |
| 248 | val beta_eta_contract = eta_contract o beta_norm; | |
| 249 | ||
| 22174 | 250 | end; | 
| 251 | ||
| 18937 | 252 | |
| 12231 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 253 | (*finds type of term without checking that combinations are consistent | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 254 | Ts holds types of bound variables*) | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 255 | fun fastype (Envir {iTs, ...}) =
 | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 256 | let val funerr = "fastype: expected function type"; | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 257 | fun fast Ts (f $ u) = | 
| 20670 | 258 | (case fast Ts f of | 
| 259 |            Type ("fun", [_, T]) => T
 | |
| 260 | | TVar ixnS => | |
| 26328 | 261 | (case Type.lookup iTs ixnS of | 
| 20670 | 262 |                    SOME (Type ("fun", [_, T])) => T
 | 
| 263 | | _ => raise TERM (funerr, [f $ u])) | |
| 264 | | _ => raise TERM (funerr, [f $ u])) | |
| 12231 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 265 | | fast Ts (Const (_, T)) = T | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 266 | | fast Ts (Free (_, T)) = T | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 267 | | fast Ts (Bound i) = | 
| 20670 | 268 | (List.nth (Ts, i) | 
| 269 |          handle Subscript => raise TERM ("fastype: Bound", [Bound i]))
 | |
| 270 | | fast Ts (Var (_, T)) = T | |
| 12231 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 271 | | fast Ts (Abs (_, T, u)) = T --> fast (T :: Ts) u | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 272 | in fast end; | 
| 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 berghofe parents: 
11513diff
changeset | 273 | |
| 15797 | 274 | |
| 275 | (*Substitute for type Vars in a type*) | |
| 276 | fun typ_subst_TVars iTs T = if Vartab.is_empty iTs then T else | |
| 277 | let fun subst(Type(a, Ts)) = Type(a, map subst Ts) | |
| 278 | | subst(T as TFree _) = T | |
| 279 | | subst(T as TVar ixnS) = | |
| 26328 | 280 | (case Type.lookup iTs ixnS of NONE => T | SOME(U) => U) | 
| 15797 | 281 | in subst T end; | 
| 282 | ||
| 283 | (*Substitute for type Vars in a term*) | |
| 20548 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 wenzelm parents: 
20098diff
changeset | 284 | val subst_TVars = map_types o typ_subst_TVars; | 
| 15797 | 285 | |
| 286 | (*Substitute for Vars in a term *) | |
| 287 | fun subst_Vars itms t = if Vartab.is_empty itms then t else | |
| 18937 | 288 | let fun subst (v as Var ixnT) = the_default v (lookup' (itms, ixnT)) | 
| 15797 | 289 | | subst (Abs (a, T, t)) = Abs (a, T, subst t) | 
| 290 | | subst (f $ t) = subst f $ subst t | |
| 291 | | subst t = t | |
| 292 | in subst t end; | |
| 293 | ||
| 294 | (*Substitute for type/term Vars in a term *) | |
| 16652 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 berghofe parents: 
15797diff
changeset | 295 | fun subst_vars (iTs, itms) = | 
| 15797 | 296 | if Vartab.is_empty iTs then subst_Vars itms else | 
| 297 | let fun subst (Const (a, T)) = Const(a, typ_subst_TVars iTs T) | |
| 298 | | subst (Free (a, T)) = Free (a, typ_subst_TVars iTs T) | |
| 16652 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 berghofe parents: 
15797diff
changeset | 299 | | subst (Var (ixn, T)) = (case lookup' (itms, (ixn, T)) of | 
| 15797 | 300 | NONE => Var (ixn, typ_subst_TVars iTs T) | 
| 301 | | SOME t => t) | |
| 302 | | subst (b as Bound _) = b | |
| 303 | | subst (Abs (a, T, t)) = Abs(a, typ_subst_TVars iTs T, subst t) | |
| 304 | | subst (f $ t) = subst f $ subst t | |
| 305 | in subst end; | |
| 306 | ||
| 18937 | 307 | |
| 21795 | 308 | (* expand defined atoms -- with local beta reduction *) | 
| 18937 | 309 | |
| 19422 | 310 | fun expand_atom T (U, u) = | 
| 311 | subst_TVars (Type.raw_match (U, T) Vartab.empty) u | |
| 18937 | 312 |   handle Type.TYPE_MATCH => raise TYPE ("expand_atom: ill-typed replacement", [T, U], [u]);
 | 
| 313 | ||
| 21795 | 314 | fun expand_term get = | 
| 21695 | 315 | let | 
| 316 | fun expand tm = | |
| 317 | let | |
| 318 | val (head, args) = Term.strip_comb tm; | |
| 319 | val args' = map expand args; | |
| 320 | fun comb head' = Term.list_comb (head', args'); | |
| 321 | in | |
| 322 | (case head of | |
| 323 | Abs (x, T, t) => comb (Abs (x, T, expand t)) | |
| 324 | | _ => | |
| 21795 | 325 | (case get head of | 
| 21695 | 326 | SOME def => Term.betapplys (expand_atom (Term.fastype_of head) def, args') | 
| 327 | | NONE => comb head) | |
| 328 | | _ => comb head) | |
| 329 | end; | |
| 330 | in expand end; | |
| 331 | ||
| 21795 | 332 | fun expand_term_frees defs = | 
| 333 | let | |
| 334 | val eqs = map (fn ((x, U), u) => (x, (U, u))) defs; | |
| 335 | val get = fn Free (x, _) => AList.lookup (op =) eqs x | _ => NONE; | |
| 336 | in expand_term get end; | |
| 337 | ||
| 0 | 338 | end; |