| author | blanchet | 
| Fri, 06 Aug 2010 11:37:33 +0200 | |
| changeset 38216 | 17d9808ed626 | 
| parent 35408 | b48ab741683b | 
| child 42083 | e1209fc7ecdc | 
| permissions | -rw-r--r-- | 
| 
247
 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 
wenzelm 
parents: 
0 
diff
changeset
 | 
1  | 
(* Title: Pure/envir.ML  | 
| 
 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 
wenzelm 
parents: 
0 
diff
changeset
 | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 10485 | 3  | 
|
| 32031 | 4  | 
Free-form 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,  | 
|
| 15797 | 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: 
0 
diff
changeset
 | 
9  | 
signature ENVIR =  | 
| 0 | 10  | 
sig  | 
| 32018 | 11  | 
type tenv = (typ * term) Vartab.table  | 
| 32031 | 12  | 
  datatype env = Envir of {maxidx: int, tenv: tenv, tyenv: Type.tyenv}
 | 
13  | 
val maxidx_of: env -> int  | 
|
14  | 
val term_env: env -> tenv  | 
|
| 15797 | 15  | 
val type_env: env -> Type.tyenv  | 
| 32031 | 16  | 
val is_empty: env -> bool  | 
17  | 
val empty: int -> env  | 
|
18  | 
val merge: env * env -> env  | 
|
| 26638 | 19  | 
val insert_sorts: env -> sort list -> sort list  | 
| 10485 | 20  | 
val genvars: string -> env * typ list -> env * term list  | 
21  | 
val genvar: string -> env * typ -> env * term  | 
|
| 15797 | 22  | 
val lookup: env * (indexname * typ) -> term option  | 
| 
16652
 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 
berghofe 
parents: 
15797 
diff
changeset
 | 
23  | 
val lookup': tenv * (indexname * typ) -> term option  | 
| 15797 | 24  | 
val update: ((indexname * typ) * term) * env -> env  | 
| 19861 | 25  | 
val above: env -> int -> bool  | 
| 15797 | 26  | 
val vupdate: ((indexname * typ) * term) * env -> env  | 
| 32018 | 27  | 
val norm_type_same: Type.tyenv -> typ Same.operation  | 
28  | 
val norm_types_same: Type.tyenv -> typ list Same.operation  | 
|
| 15797 | 29  | 
val norm_type: Type.tyenv -> typ -> typ  | 
| 32018 | 30  | 
val norm_term_same: env -> term Same.operation  | 
31  | 
val norm_term: env -> term -> term  | 
|
| 10485 | 32  | 
val beta_norm: term -> term  | 
| 
12231
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
33  | 
val head_norm: env -> term -> term  | 
| 18937 | 34  | 
val eta_contract: term -> term  | 
35  | 
val beta_eta_contract: term -> term  | 
|
| 
12231
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
36  | 
val fastype: env -> typ list -> term -> typ  | 
| 32034 | 37  | 
val subst_type_same: Type.tyenv -> typ Same.operation  | 
38  | 
val subst_term_types_same: Type.tyenv -> term Same.operation  | 
|
39  | 
val subst_term_same: Type.tyenv * tenv -> term Same.operation  | 
|
40  | 
val subst_type: Type.tyenv -> typ -> typ  | 
|
41  | 
val subst_term_types: Type.tyenv -> term -> term  | 
|
42  | 
val subst_term: Type.tyenv * tenv -> term -> term  | 
|
| 19422 | 43  | 
val expand_atom: typ -> typ * term -> term  | 
| 21695 | 44  | 
val expand_term: (term -> (typ * term) option) -> term -> term  | 
| 21795 | 45  | 
val expand_term_frees: ((string * typ) * term) list -> term -> term  | 
| 0 | 46  | 
end;  | 
47  | 
||
| 32031 | 48  | 
structure Envir: ENVIR =  | 
| 0 | 49  | 
struct  | 
50  | 
||
| 32031 | 51  | 
(** datatype env **)  | 
52  | 
||
53  | 
(*Updating can destroy environment in 2 ways!  | 
|
54  | 
(1) variables out of range  | 
|
55  | 
(2) circular assignments  | 
|
| 0 | 56  | 
*)  | 
| 32031 | 57  | 
|
| 32018 | 58  | 
type tenv = (typ * term) Vartab.table;  | 
| 15797 | 59  | 
|
| 0 | 60  | 
datatype env = Envir of  | 
| 32031 | 61  | 
 {maxidx: int,          (*upper bound of maximum index of vars*)
 | 
62  | 
tenv: tenv, (*assignments to Vars*)  | 
|
63  | 
tyenv: Type.tyenv}; (*assignments to TVars*)  | 
|
64  | 
||
| 32796 | 65  | 
fun make_env (maxidx, tenv, tyenv) =  | 
66  | 
  Envir {maxidx = maxidx, tenv = tenv, tyenv = tyenv};
 | 
|
| 32031 | 67  | 
|
68  | 
fun maxidx_of (Envir {maxidx, ...}) = maxidx;
 | 
|
69  | 
fun term_env (Envir {tenv, ...}) = tenv;
 | 
|
70  | 
fun type_env (Envir {tyenv, ...}) = tyenv;
 | 
|
71  | 
||
72  | 
fun is_empty env =  | 
|
73  | 
Vartab.is_empty (term_env env) andalso  | 
|
74  | 
Vartab.is_empty (type_env env);  | 
|
| 0 | 75  | 
|
| 32031 | 76  | 
|
77  | 
(* build env *)  | 
|
78  | 
||
79  | 
fun empty maxidx = make_env (maxidx, Vartab.empty, Vartab.empty);  | 
|
| 0 | 80  | 
|
| 32031 | 81  | 
fun merge  | 
82  | 
   (Envir {maxidx = maxidx1, tenv = tenv1, tyenv = tyenv1},
 | 
|
83  | 
    Envir {maxidx = maxidx2, tenv = tenv2, tyenv = tyenv2}) =
 | 
|
84  | 
make_env (Int.max (maxidx1, maxidx2),  | 
|
85  | 
Vartab.merge (op =) (tenv1, tenv2),  | 
|
86  | 
Vartab.merge (op =) (tyenv1, tyenv2));  | 
|
87  | 
||
88  | 
||
89  | 
(*NB: type unification may invent new sorts*) (* FIXME tenv!? *)  | 
|
| 26638 | 90  | 
val insert_sorts = Vartab.fold (fn (_, (_, T)) => Sorts.insert_typ T) o type_env;  | 
91  | 
||
| 0 | 92  | 
(*Generate a list of distinct variables.  | 
93  | 
Increments index to make them distinct from ALL present variables. *)  | 
|
| 32031 | 94  | 
fun genvars name (Envir {maxidx, tenv, tyenv}, Ts) : env * term list =
 | 
| 32018 | 95  | 
let  | 
96  | 
fun genvs (_, [] : typ list) : term list = []  | 
|
97  | 
| genvs (n, [T]) = [Var ((name, maxidx + 1), T)]  | 
|
98  | 
| genvs (n, T :: Ts) =  | 
|
99  | 
Var ((name ^ radixstring (26, "a" , n), maxidx + 1), T)  | 
|
100  | 
:: genvs (n + 1, Ts);  | 
|
| 32031 | 101  | 
  in (Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}, genvs (0, Ts)) end;
 | 
| 0 | 102  | 
|
103  | 
(*Generate a variable.*)  | 
|
| 32018 | 104  | 
fun genvar name (env, T) : env * term =  | 
105  | 
let val (env', [v]) = genvars name (env, [T])  | 
|
106  | 
in (env', v) end;  | 
|
| 0 | 107  | 
|
| 32031 | 108  | 
fun var_clash xi T T' =  | 
109  | 
  raise TYPE ("Variable " ^ quote (Term.string_of_vname xi) ^ " has two distinct types",
 | 
|
| 32018 | 110  | 
[T', T], []);  | 
| 0 | 111  | 
|
| 32031 | 112  | 
fun lookup_check check tenv (xi, T) =  | 
113  | 
(case Vartab.lookup tenv xi of  | 
|
| 32018 | 114  | 
NONE => NONE  | 
| 32031 | 115  | 
| SOME (U, t) => if check (T, U) then SOME t else var_clash xi T U);  | 
| 15797 | 116  | 
|
| 
16652
 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 
berghofe 
parents: 
15797 
diff
changeset
 | 
117  | 
(* When dealing with environments produced by matching instead *)  | 
| 
 
4ecf94235ec7
Fixed bug: lookup' must use = instead of eq_type to compare types of
 
berghofe 
parents: 
15797 
diff
changeset
 | 
118  | 
(* 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: 
15797 
diff
changeset
 | 
119  | 
(* 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: 
15797 
diff
changeset
 | 
120  | 
(* and use = instead of eq_type. *)  | 
| 15797 | 121  | 
|
| 32031 | 122  | 
fun lookup' (tenv, p) = lookup_check (op =) tenv p;  | 
| 15797 | 123  | 
|
| 32031 | 124  | 
fun lookup2 (tyenv, tenv) =  | 
125  | 
lookup_check (Type.eq_type tyenv) tenv;  | 
|
| 15797 | 126  | 
|
| 32031 | 127  | 
fun lookup (Envir {tenv, tyenv, ...}, p) = lookup2 (tyenv, tenv) p;
 | 
| 0 | 128  | 
|
| 32031 | 129  | 
fun update (((xi, T), t), Envir {maxidx, tenv, tyenv}) =
 | 
130  | 
  Envir {maxidx = maxidx, tenv = Vartab.update_new (xi, (T, t)) tenv, tyenv = tyenv};
 | 
|
| 
247
 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 
wenzelm 
parents: 
0 
diff
changeset
 | 
131  | 
|
| 
2142
 
20f208ff085d
Deleted Olist constructor.  Replaced minidx by "above" function
 
paulson 
parents: 
1500 
diff
changeset
 | 
132  | 
(*Determine if the least index updated exceeds lim*)  | 
| 32031 | 133  | 
fun above (Envir {tenv, tyenv, ...}) lim =
 | 
134  | 
(case Vartab.min_key tenv of SOME (_, i) => i > lim | NONE => true) andalso  | 
|
135  | 
(case Vartab.min_key tyenv of SOME (_, i) => i > lim | NONE => true);  | 
|
| 
247
 
bc10568855ee
added is_empty: env -> bool, minidx: env -> int option;
 
wenzelm 
parents: 
0 
diff
changeset
 | 
136  | 
|
| 0 | 137  | 
(*Update, checking Var-Var assignments: try to suppress higher indexes*)  | 
| 32031 | 138  | 
fun vupdate ((aU as (a, U), t), env as Envir {tyenv, ...}) =
 | 
139  | 
(case t of  | 
|
140  | 
Var (nT as (name', T)) =>  | 
|
141  | 
if a = name' then env (*cycle!*)  | 
|
| 35408 | 142  | 
else if Term_Ord.indexname_ord (a, name') = LESS then  | 
| 32031 | 143  | 
(case lookup (env, nT) of (*if already assigned, chase*)  | 
144  | 
NONE => update ((nT, Var (a, T)), env)  | 
|
145  | 
| SOME u => vupdate ((aU, u), env))  | 
|
146  | 
else update ((aU, t), env)  | 
|
147  | 
| _ => update ((aU, t), env));  | 
|
| 0 | 148  | 
|
149  | 
||
150  | 
||
| 32031 | 151  | 
(** beta normalization wrt. environment **)  | 
| 0 | 152  | 
|
| 32031 | 153  | 
(*Chases variables in env. Does not exploit sharing of variable bindings  | 
154  | 
Does not check types, so could loop.*)  | 
|
| 1500 | 155  | 
|
| 32018 | 156  | 
local  | 
157  | 
||
| 32031 | 158  | 
fun norm_type0 tyenv : typ Same.operation =  | 
| 32018 | 159  | 
let  | 
160  | 
fun norm (Type (a, Ts)) = Type (a, Same.map norm Ts)  | 
|
161  | 
| norm (TFree _) = raise Same.SAME  | 
|
162  | 
| norm (TVar v) =  | 
|
| 32031 | 163  | 
(case Type.lookup tyenv v of  | 
| 32018 | 164  | 
SOME U => Same.commit norm U  | 
165  | 
| NONE => raise Same.SAME);  | 
|
166  | 
in norm end;  | 
|
| 0 | 167  | 
|
| 32031 | 168  | 
fun norm_term1 tenv : term Same.operation =  | 
| 32018 | 169  | 
let  | 
170  | 
fun norm (Var v) =  | 
|
| 32031 | 171  | 
(case lookup' (tenv, v) of  | 
| 32018 | 172  | 
SOME u => Same.commit norm u  | 
173  | 
| NONE => raise Same.SAME)  | 
|
| 32031 | 174  | 
| norm (Abs (a, T, body)) = Abs (a, T, norm body)  | 
| 32018 | 175  | 
| norm (Abs (_, _, body) $ t) = Same.commit norm (subst_bound (t, body))  | 
176  | 
| norm (f $ t) =  | 
|
177  | 
((case norm f of  | 
|
178  | 
Abs (_, _, body) => Same.commit norm (subst_bound (t, body))  | 
|
179  | 
| nf => nf $ Same.commit norm t)  | 
|
180  | 
handle Same.SAME => f $ norm t)  | 
|
181  | 
| norm _ = raise Same.SAME;  | 
|
182  | 
in norm end;  | 
|
| 0 | 183  | 
|
| 32031 | 184  | 
fun norm_term2 tenv tyenv : term Same.operation =  | 
| 32018 | 185  | 
let  | 
| 32031 | 186  | 
val normT = norm_type0 tyenv;  | 
| 32018 | 187  | 
fun norm (Const (a, T)) = Const (a, normT T)  | 
188  | 
| norm (Free (a, T)) = Free (a, normT T)  | 
|
189  | 
| norm (Var (xi, T)) =  | 
|
| 32031 | 190  | 
(case lookup2 (tyenv, tenv) (xi, T) of  | 
| 32018 | 191  | 
SOME u => Same.commit norm u  | 
192  | 
| NONE => Var (xi, normT T))  | 
|
193  | 
| norm (Abs (a, T, body)) =  | 
|
194  | 
(Abs (a, normT T, Same.commit norm body)  | 
|
195  | 
handle Same.SAME => Abs (a, T, norm body))  | 
|
196  | 
| norm (Abs (_, _, body) $ t) = Same.commit norm (subst_bound (t, body))  | 
|
197  | 
| norm (f $ t) =  | 
|
198  | 
((case norm f of  | 
|
199  | 
Abs (_, _, body) => Same.commit norm (subst_bound (t, body))  | 
|
200  | 
| nf => nf $ Same.commit norm t)  | 
|
201  | 
handle Same.SAME => f $ norm t)  | 
|
202  | 
| norm _ = raise Same.SAME;  | 
|
203  | 
in norm end;  | 
|
| 11513 | 204  | 
|
| 32018 | 205  | 
in  | 
206  | 
||
| 32031 | 207  | 
fun norm_type_same tyenv T =  | 
208  | 
if Vartab.is_empty tyenv then raise Same.SAME  | 
|
209  | 
else norm_type0 tyenv T;  | 
|
| 0 | 210  | 
|
| 32031 | 211  | 
fun norm_types_same tyenv Ts =  | 
212  | 
if Vartab.is_empty tyenv then raise Same.SAME  | 
|
213  | 
else Same.map (norm_type0 tyenv) Ts;  | 
|
| 32018 | 214  | 
|
| 32031 | 215  | 
fun norm_type tyenv T = norm_type_same tyenv T handle Same.SAME => T;  | 
| 11513 | 216  | 
|
| 32031 | 217  | 
fun norm_term_same (Envir {tenv, tyenv, ...}) =
 | 
218  | 
if Vartab.is_empty tyenv then norm_term1 tenv  | 
|
219  | 
else norm_term2 tenv tyenv;  | 
|
| 10485 | 220  | 
|
| 32018 | 221  | 
fun norm_term envir t = norm_term_same envir t handle Same.SAME => t;  | 
| 
25471
 
ca009b7ce693
Optimized beta_norm: only tries to normalize term when it contains
 
berghofe 
parents: 
24670 
diff
changeset
 | 
222  | 
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: 
247 
diff
changeset
 | 
223  | 
|
| 32018 | 224  | 
end;  | 
| 11513 | 225  | 
|
226  | 
||
| 
12231
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
227  | 
(*Put a term into head normal form for unification.*)  | 
| 
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
228  | 
|
| 32018 | 229  | 
fun head_norm env =  | 
| 
12231
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
230  | 
let  | 
| 32018 | 231  | 
fun norm (Var v) =  | 
232  | 
(case lookup (env, v) of  | 
|
| 15531 | 233  | 
SOME u => head_norm env u  | 
| 32018 | 234  | 
| NONE => raise Same.SAME)  | 
235  | 
| norm (Abs (a, T, body)) = Abs (a, T, norm body)  | 
|
236  | 
| norm (Abs (_, _, body) $ t) = Same.commit norm (subst_bound (t, body))  | 
|
237  | 
| norm (f $ t) =  | 
|
238  | 
(case norm f of  | 
|
239  | 
Abs (_, _, body) => Same.commit norm (subst_bound (t, body))  | 
|
240  | 
| nf => nf $ t)  | 
|
241  | 
| norm _ = raise Same.SAME;  | 
|
242  | 
in Same.commit norm end;  | 
|
| 
12231
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
243  | 
|
| 
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
244  | 
|
| 18937 | 245  | 
(*Eta-contract a term (fully)*)  | 
246  | 
||
| 22174 | 247  | 
local  | 
248  | 
||
| 32018 | 249  | 
fun decr lev (Bound i) = if i >= lev then Bound (i - 1) else raise Same.SAME  | 
| 22174 | 250  | 
| decr lev (Abs (a, T, body)) = Abs (a, T, decr (lev + 1) body)  | 
| 32018 | 251  | 
| decr lev (t $ u) = (decr lev t $ decrh lev u handle Same.SAME => t $ decr lev u)  | 
252  | 
| decr _ _ = raise Same.SAME  | 
|
253  | 
and decrh lev t = (decr lev t handle Same.SAME => t);  | 
|
| 20670 | 254  | 
|
| 22174 | 255  | 
fun eta (Abs (a, T, body)) =  | 
256  | 
((case eta body of  | 
|
257  | 
body' as (f $ Bound 0) =>  | 
|
258  | 
if loose_bvar1 (f, 0) then Abs (a, T, body')  | 
|
259  | 
else decrh 0 f  | 
|
| 32018 | 260  | 
| body' => Abs (a, T, body')) handle Same.SAME =>  | 
| 22174 | 261  | 
(case body of  | 
262  | 
f $ Bound 0 =>  | 
|
| 32018 | 263  | 
if loose_bvar1 (f, 0) then raise Same.SAME  | 
| 22174 | 264  | 
else decrh 0 f  | 
| 32018 | 265  | 
| _ => raise Same.SAME))  | 
266  | 
| eta (t $ u) = (eta t $ Same.commit eta u handle Same.SAME => t $ eta u)  | 
|
267  | 
| eta _ = raise Same.SAME;  | 
|
| 22174 | 268  | 
|
269  | 
in  | 
|
270  | 
||
271  | 
fun eta_contract t =  | 
|
| 32018 | 272  | 
if Term.has_abs t then Same.commit eta t else t;  | 
| 18937 | 273  | 
|
274  | 
val beta_eta_contract = eta_contract o beta_norm;  | 
|
275  | 
||
| 22174 | 276  | 
end;  | 
277  | 
||
| 18937 | 278  | 
|
| 
12231
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
279  | 
(*finds type of term without checking that combinations are consistent  | 
| 
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
280  | 
Ts holds types of bound variables*)  | 
| 32031 | 281  | 
fun fastype (Envir {tyenv, ...}) =
 | 
282  | 
let  | 
|
283  | 
val funerr = "fastype: expected function type";  | 
|
| 
12231
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
284  | 
fun fast Ts (f $ u) =  | 
| 
32648
 
143e0b0a6b33
Correct chasing of type variable instantiations during type unification.
 
paulson 
parents: 
32034 
diff
changeset
 | 
285  | 
(case Type.devar tyenv (fast Ts f) of  | 
| 32031 | 286  | 
            Type ("fun", [_, T]) => T
 | 
| 
32648
 
143e0b0a6b33
Correct chasing of type variable instantiations during type unification.
 
paulson 
parents: 
32034 
diff
changeset
 | 
287  | 
| TVar v => raise TERM (funerr, [f $ u])  | 
| 32031 | 288  | 
| _ => raise TERM (funerr, [f $ u]))  | 
| 
12231
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
289  | 
| fast Ts (Const (_, T)) = T  | 
| 
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
290  | 
| fast Ts (Free (_, T)) = T  | 
| 
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
291  | 
| fast Ts (Bound i) =  | 
| 32031 | 292  | 
          (nth Ts i handle Subscript => raise TERM ("fastype: Bound", [Bound i]))
 | 
| 20670 | 293  | 
| fast Ts (Var (_, T)) = T  | 
| 32031 | 294  | 
| fast Ts (Abs (_, T, u)) = T --> fast (T :: Ts) u;  | 
295  | 
in fast end;  | 
|
| 
12231
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
11513 
diff
changeset
 | 
296  | 
|
| 15797 | 297  | 
|
| 32034 | 298  | 
|
299  | 
(** plain substitution -- without variable chasing **)  | 
|
300  | 
||
301  | 
local  | 
|
| 15797 | 302  | 
|
| 32034 | 303  | 
fun subst_type0 tyenv = Term_Subst.map_atypsT_same  | 
304  | 
(fn TVar v =>  | 
|
305  | 
(case Type.lookup tyenv v of  | 
|
306  | 
SOME U => U  | 
|
307  | 
| NONE => raise Same.SAME)  | 
|
308  | 
| _ => raise Same.SAME);  | 
|
309  | 
||
310  | 
fun subst_term1 tenv = Term_Subst.map_aterms_same  | 
|
311  | 
(fn Var v =>  | 
|
312  | 
(case lookup' (tenv, v) of  | 
|
313  | 
SOME u => u  | 
|
314  | 
| NONE => raise Same.SAME)  | 
|
315  | 
| _ => raise Same.SAME);  | 
|
| 15797 | 316  | 
|
| 32034 | 317  | 
fun subst_term2 tenv tyenv : term Same.operation =  | 
318  | 
let  | 
|
319  | 
val substT = subst_type0 tyenv;  | 
|
320  | 
fun subst (Const (a, T)) = Const (a, substT T)  | 
|
321  | 
| subst (Free (a, T)) = Free (a, substT T)  | 
|
322  | 
| subst (Var (xi, T)) =  | 
|
323  | 
(case lookup' (tenv, (xi, T)) of  | 
|
324  | 
SOME u => u  | 
|
325  | 
| NONE => Var (xi, substT T))  | 
|
326  | 
| subst (Bound _) = raise Same.SAME  | 
|
327  | 
| subst (Abs (a, T, t)) =  | 
|
328  | 
(Abs (a, substT T, Same.commit subst t)  | 
|
329  | 
handle Same.SAME => Abs (a, T, subst t))  | 
|
330  | 
| subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u);  | 
|
331  | 
in subst end;  | 
|
332  | 
||
333  | 
in  | 
|
| 15797 | 334  | 
|
| 32034 | 335  | 
fun subst_type_same tyenv T =  | 
336  | 
if Vartab.is_empty tyenv then raise Same.SAME  | 
|
337  | 
else subst_type0 tyenv T;  | 
|
338  | 
||
339  | 
fun subst_term_types_same tyenv t =  | 
|
340  | 
if Vartab.is_empty tyenv then raise Same.SAME  | 
|
341  | 
else Term_Subst.map_types_same (subst_type0 tyenv) t;  | 
|
342  | 
||
343  | 
fun subst_term_same (tyenv, tenv) =  | 
|
344  | 
if Vartab.is_empty tenv then subst_term_types_same tyenv  | 
|
345  | 
else if Vartab.is_empty tyenv then subst_term1 tenv  | 
|
346  | 
else subst_term2 tenv tyenv;  | 
|
347  | 
||
348  | 
fun subst_type tyenv T = subst_type_same tyenv T handle Same.SAME => T;  | 
|
349  | 
fun subst_term_types tyenv t = subst_term_types_same tyenv t handle Same.SAME => t;  | 
|
350  | 
fun subst_term envs t = subst_term_same envs t handle Same.SAME => t;  | 
|
351  | 
||
352  | 
end;  | 
|
| 15797 | 353  | 
|
| 18937 | 354  | 
|
| 32034 | 355  | 
|
356  | 
(** expand defined atoms -- with local beta reduction **)  | 
|
| 18937 | 357  | 
|
| 19422 | 358  | 
fun expand_atom T (U, u) =  | 
| 32034 | 359  | 
subst_term_types (Type.raw_match (U, T) Vartab.empty) u  | 
| 32031 | 360  | 
    handle Type.TYPE_MATCH => raise TYPE ("expand_atom: ill-typed replacement", [T, U], [u]);
 | 
| 18937 | 361  | 
|
| 21795 | 362  | 
fun expand_term get =  | 
| 21695 | 363  | 
let  | 
364  | 
fun expand tm =  | 
|
365  | 
let  | 
|
366  | 
val (head, args) = Term.strip_comb tm;  | 
|
367  | 
val args' = map expand args;  | 
|
368  | 
fun comb head' = Term.list_comb (head', args');  | 
|
369  | 
in  | 
|
370  | 
(case head of  | 
|
371  | 
Abs (x, T, t) => comb (Abs (x, T, expand t))  | 
|
372  | 
| _ =>  | 
|
| 32031 | 373  | 
(case get head of  | 
374  | 
SOME def => Term.betapplys (expand_atom (Term.fastype_of head) def, args')  | 
|
375  | 
| NONE => comb head))  | 
|
| 21695 | 376  | 
end;  | 
377  | 
in expand end;  | 
|
378  | 
||
| 21795 | 379  | 
fun expand_term_frees defs =  | 
380  | 
let  | 
|
381  | 
val eqs = map (fn ((x, U), u) => (x, (U, u))) defs;  | 
|
382  | 
val get = fn Free (x, _) => AList.lookup (op =) eqs x | _ => NONE;  | 
|
383  | 
in expand_term get end;  | 
|
384  | 
||
| 0 | 385  | 
end;  |