| author | wenzelm | 
| Fri, 08 Oct 1993 12:33:17 +0100 | |
| changeset 40 | 3f9f8395519e | 
| parent 0 | a5a9c433f639 | 
| child 61 | f8c1922b78e3 | 
| permissions | -rw-r--r-- | 
| 0 | 1  | 
(* Title: term.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
4  | 
Copyright Cambridge University 1992  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
||
8  | 
(*Simply typed lambda-calculus: types, terms, and basic operations*)  | 
|
9  | 
||
10  | 
||
11  | 
(*Indexnames can be quickly renamed by adding an offset to the integer part,  | 
|
12  | 
for resolution.*)  | 
|
13  | 
type indexname = string*int;  | 
|
14  | 
||
15  | 
(* Types are classified by classes. *)  | 
|
16  | 
type class = string;  | 
|
17  | 
type sort = class list;  | 
|
18  | 
||
19  | 
(* The sorts attached to TFrees and TVars specify the sort of that variable *)  | 
|
20  | 
datatype typ = Type of string * typ list  | 
|
21  | 
| TFree of string * sort  | 
|
22  | 
| TVar of indexname * sort;  | 
|
23  | 
||
24  | 
infixr 5 -->;  | 
|
25  | 
fun S --> T = Type("fun",[S,T]);
 | 
|
26  | 
||
27  | 
(*handy for multiple args: [T1,...,Tn]--->T gives T1-->(T2--> ... -->T)*)  | 
|
28  | 
infixr --->;  | 
|
29  | 
val op ---> = foldr (op -->);  | 
|
30  | 
||
31  | 
||
32  | 
(*terms. Bound variables are indicated by depth number.  | 
|
33  | 
Free variables, (scheme) variables and constants have names.  | 
|
34  | 
An term is "closed" if there every bound variable of level "lev"  | 
|
35  | 
is enclosed by at least "lev" abstractions.  | 
|
36  | 
||
37  | 
It is possible to create meaningless terms containing loose bound vars  | 
|
38  | 
or type mismatches. But such terms are not allowed in rules. *)  | 
|
39  | 
||
40  | 
||
41  | 
||
42  | 
infix 9 $; (*application binds tightly!*)  | 
|
43  | 
datatype term =  | 
|
44  | 
Const of string * typ  | 
|
45  | 
| Free of string * typ  | 
|
46  | 
| Var of indexname * typ  | 
|
47  | 
| Bound of int  | 
|
48  | 
| Abs of string*typ*term  | 
|
49  | 
| op $ of term*term;  | 
|
50  | 
||
51  | 
||
52  | 
(*For errors involving type mismatches*)  | 
|
53  | 
exception TYPE of string * typ list * term list;  | 
|
54  | 
||
| 
40
 
3f9f8395519e
added raise_type: string -> typ list -> term list -> 'a;
 
wenzelm 
parents: 
0 
diff
changeset
 | 
55  | 
fun raise_type msg tys ts = raise TYPE (msg, tys, ts);  | 
| 
 
3f9f8395519e
added raise_type: string -> typ list -> term list -> 'a;
 
wenzelm 
parents: 
0 
diff
changeset
 | 
56  | 
|
| 0 | 57  | 
(*For system errors involving terms*)  | 
58  | 
exception TERM of string * term list;  | 
|
59  | 
||
| 
40
 
3f9f8395519e
added raise_type: string -> typ list -> term list -> 'a;
 
wenzelm 
parents: 
0 
diff
changeset
 | 
60  | 
fun raise_term msg ts = raise TERM (msg, ts);  | 
| 
 
3f9f8395519e
added raise_type: string -> typ list -> term list -> 'a;
 
wenzelm 
parents: 
0 
diff
changeset
 | 
61  | 
|
| 0 | 62  | 
|
63  | 
(*Note variable naming conventions!  | 
|
64  | 
a,b,c: string  | 
|
65  | 
f,g,h: functions (including terms of function type)  | 
|
66  | 
i,j,m,n: int  | 
|
67  | 
t,u: term  | 
|
68  | 
v,w: indexnames  | 
|
69  | 
x,y: any  | 
|
70  | 
A,B,C: term (denoting formulae)  | 
|
71  | 
T,U: typ  | 
|
72  | 
*)  | 
|
73  | 
||
74  | 
||
75  | 
(** Discriminators **)  | 
|
76  | 
||
77  | 
fun is_Const (Const _) = true  | 
|
78  | 
| is_Const _ = false;  | 
|
79  | 
||
80  | 
fun is_Free (Free _) = true  | 
|
81  | 
| is_Free _ = false;  | 
|
82  | 
||
83  | 
fun is_Var (Var _) = true  | 
|
84  | 
| is_Var _ = false;  | 
|
85  | 
||
86  | 
fun is_TVar (TVar _) = true  | 
|
87  | 
| is_TVar _ = false;  | 
|
88  | 
||
89  | 
(** Destructors **)  | 
|
90  | 
||
91  | 
fun dest_Const (Const x) = x  | 
|
92  | 
  | dest_Const t = raise TERM("dest_Const", [t]);
 | 
|
93  | 
||
94  | 
fun dest_Free (Free x) = x  | 
|
95  | 
  | dest_Free t = raise TERM("dest_Free", [t]);
 | 
|
96  | 
||
97  | 
fun dest_Var (Var x) = x  | 
|
98  | 
  | dest_Var t = raise TERM("dest_Var", [t]);
 | 
|
99  | 
||
100  | 
||
101  | 
(* maps [T1,...,Tn]--->T to the list [T1,T2,...,Tn]*)  | 
|
102  | 
fun binder_types (Type("fun",[S,T])) = S :: binder_types T
 | 
|
103  | 
| binder_types _ = [];  | 
|
104  | 
||
105  | 
(* maps [T1,...,Tn]--->T to T*)  | 
|
106  | 
fun body_type (Type("fun",[S,T])) = body_type T
 | 
|
107  | 
| body_type T = T;  | 
|
108  | 
||
109  | 
(* maps [T1,...,Tn]--->T to ([T1,T2,...,Tn], T) *)  | 
|
110  | 
fun strip_type T : typ list * typ =  | 
|
111  | 
(binder_types T, body_type T);  | 
|
112  | 
||
113  | 
||
114  | 
(*Compute the type of the term, checking that combinations are well-typed  | 
|
115  | 
Ts = [T0,T1,...] holds types of bound variables 0, 1, ...*)  | 
|
116  | 
fun type_of1 (Ts, Const (_,T)) = T  | 
|
117  | 
| type_of1 (Ts, Free (_,T)) = T  | 
|
118  | 
| type_of1 (Ts, Bound i) = (nth_elem (i,Ts)  | 
|
119  | 
  	handle LIST _ => raise TYPE("type_of: bound variable", [], [Bound i]))
 | 
|
120  | 
| type_of1 (Ts, Var (_,T)) = T  | 
|
121  | 
| type_of1 (Ts, Abs (_,T,body)) = T --> type_of1(T::Ts, body)  | 
|
122  | 
| type_of1 (Ts, f$u) =  | 
|
123  | 
let val U = type_of1(Ts,u)  | 
|
124  | 
and T = type_of1(Ts,f)  | 
|
125  | 
in case T of  | 
|
126  | 
	    Type("fun",[T1,T2]) =>
 | 
|
127  | 
if T1=U then T2 else raise TYPE  | 
|
128  | 
	         ("type_of: type mismatch in application", [T1,U], [f$u])
 | 
|
129  | 
	  | _ => raise TYPE ("type_of: Rator must have function type",
 | 
|
130  | 
[T,U], [f$u])  | 
|
131  | 
end;  | 
|
132  | 
||
133  | 
||
134  | 
fun type_of t : typ = type_of1 ([],t);  | 
|
135  | 
||
136  | 
(*Determines the type of a term, with minimal checking*)  | 
|
137  | 
fun fastype_of(Ts, f$u) = (case fastype_of(Ts,f) of  | 
|
138  | 
	Type("fun",[_,T]) => T
 | 
|
139  | 
	| _ => raise TERM("fastype_of: expected function type", [f$u]))
 | 
|
140  | 
| fastype_of(_, Const (_,T)) = T  | 
|
141  | 
| fastype_of(_, Free (_,T)) = T  | 
|
142  | 
| fastype_of(Ts, Bound i) = (nth_elem(i,Ts)  | 
|
143  | 
  	 handle LIST _ => raise TERM("fastype_of: Bound", [Bound i]))
 | 
|
144  | 
| fastype_of(_, Var (_,T)) = T  | 
|
145  | 
| fastype_of(Ts, Abs (_,T,u)) = T --> fastype_of(T::Ts, u);  | 
|
146  | 
||
147  | 
||
148  | 
(* maps (x1,...,xn)t to t *)  | 
|
149  | 
fun strip_abs_body (Abs(_,_,t)) = strip_abs_body t  | 
|
150  | 
| strip_abs_body u = u;  | 
|
151  | 
||
152  | 
||
153  | 
(* maps (x1,...,xn)t to [x1, ..., xn] *)  | 
|
154  | 
fun strip_abs_vars (Abs(a,T,t)) = (a,T) :: strip_abs_vars t  | 
|
155  | 
| strip_abs_vars u = [] : (string*typ) list;  | 
|
156  | 
||
157  | 
||
158  | 
fun strip_qnt_body qnt =  | 
|
159  | 
let fun strip(tm as Const(c,_)$Abs(_,_,t)) = if c=qnt then strip t else tm  | 
|
160  | 
| strip t = t  | 
|
161  | 
in strip end;  | 
|
162  | 
||
163  | 
fun strip_qnt_vars qnt =  | 
|
164  | 
let fun strip(Const(c,_)$Abs(a,T,t)) = if c=qnt then (a,T)::strip t else []  | 
|
165  | 
| strip t = [] : (string*typ) list  | 
|
166  | 
in strip end;  | 
|
167  | 
||
168  | 
||
169  | 
(* maps (f, [t1,...,tn]) to f(t1,...,tn) *)  | 
|
170  | 
val list_comb : term * term list -> term = foldl (op $);  | 
|
171  | 
||
172  | 
||
173  | 
(* maps f(t1,...,tn) to (f, [t1,...,tn]) ; naturally tail-recursive*)  | 
|
174  | 
fun strip_comb u : term * term list =  | 
|
175  | 
let fun stripc (f$t, ts) = stripc (f, t::ts)  | 
|
176  | 
| stripc x = x  | 
|
177  | 
in stripc(u,[]) end;  | 
|
178  | 
||
179  | 
||
180  | 
(* maps f(t1,...,tn) to f , which is never a combination *)  | 
|
181  | 
fun head_of (f$t) = head_of f  | 
|
182  | 
| head_of u = u;  | 
|
183  | 
||
184  | 
||
185  | 
(*Number of atoms and abstractions in a term*)  | 
|
186  | 
fun size_of_term (Abs (_,_,body)) = 1 + size_of_term body  | 
|
187  | 
| size_of_term (f$t) = size_of_term f + size_of_term t  | 
|
188  | 
| size_of_term _ = 1;  | 
|
189  | 
||
190  | 
||
191  | 
(* apply a function to all types in a term *)  | 
|
192  | 
fun map_term_types f =  | 
|
193  | 
let fun map(Const(a,T)) = Const(a, f T)  | 
|
194  | 
| map(Free(a,T)) = Free(a, f T)  | 
|
195  | 
| map(Var(v,T)) = Var(v, f T)  | 
|
196  | 
| map(t as Bound _) = t  | 
|
197  | 
| map(Abs(a,T,t)) = Abs(a, f T, map t)  | 
|
198  | 
| map(f$t) = map f $ map t;  | 
|
199  | 
in map end;  | 
|
200  | 
||
201  | 
(* iterate a function over all types in a term *)  | 
|
202  | 
fun it_term_types f =  | 
|
203  | 
let fun iter(Const(_,T), a) = f(T,a)  | 
|
204  | 
| iter(Free(_,T), a) = f(T,a)  | 
|
205  | 
| iter(Var(_,T), a) = f(T,a)  | 
|
206  | 
| iter(Abs(_,T,t), a) = iter(t,f(T,a))  | 
|
207  | 
| iter(f$u, a) = iter(f, iter(u, a))  | 
|
208  | 
| iter(Bound _, a) = a  | 
|
209  | 
in iter end  | 
|
210  | 
||
211  | 
||
212  | 
(** Connectives of higher order logic **)  | 
|
213  | 
||
214  | 
val propT : typ = Type("prop",[]);
 | 
|
215  | 
||
216  | 
val implies = Const("==>", propT-->propT-->propT);
 | 
|
217  | 
||
218  | 
fun all T = Const("all", (T-->propT)-->propT);
 | 
|
219  | 
||
220  | 
fun equals T = Const("==", T-->T-->propT);
 | 
|
221  | 
||
222  | 
fun flexpair T = Const("=?=", T-->T-->propT);
 | 
|
223  | 
||
224  | 
(* maps !!x1...xn. t to t *)  | 
|
225  | 
fun strip_all_body (Const("all",_)$Abs(_,_,t))  =  strip_all_body t  
 | 
|
226  | 
| strip_all_body t = t;  | 
|
227  | 
||
228  | 
(* maps !!x1...xn. t to [x1, ..., xn] *)  | 
|
229  | 
fun strip_all_vars (Const("all",_)$Abs(a,T,t))  =
 | 
|
230  | 
(a,T) :: strip_all_vars t  | 
|
231  | 
| strip_all_vars t = [] : (string*typ) list;  | 
|
232  | 
||
233  | 
(*increments a term's non-local bound variables  | 
|
234  | 
required when moving a term within abstractions  | 
|
235  | 
inc is increment for bound variables  | 
|
236  | 
lev is level at which a bound variable is considered 'loose'*)  | 
|
237  | 
fun incr_bv (inc, lev, u as Bound i) = if i>=lev then Bound(i+inc) else u  | 
|
238  | 
| incr_bv (inc, lev, Abs(a,T,body)) =  | 
|
239  | 
Abs(a, T, incr_bv(inc,lev+1,body))  | 
|
240  | 
| incr_bv (inc, lev, f$t) =  | 
|
241  | 
incr_bv(inc,lev,f) $ incr_bv(inc,lev,t)  | 
|
242  | 
| incr_bv (inc, lev, u) = u;  | 
|
243  | 
||
244  | 
fun incr_boundvars 0 t = t  | 
|
245  | 
| incr_boundvars inc t = incr_bv(inc,0,t);  | 
|
246  | 
||
247  | 
||
248  | 
(*Accumulate all 'loose' bound vars referring to level 'lev' or beyond.  | 
|
249  | 
(Bound 0) is loose at level 0 *)  | 
|
250  | 
fun add_loose_bnos (Bound i, lev, js) =  | 
|
251  | 
if i<lev then js else (i-lev) :: js  | 
|
252  | 
| add_loose_bnos (Abs (_,_,t), lev, js) = add_loose_bnos (t, lev+1, js)  | 
|
253  | 
| add_loose_bnos (f$t, lev, js) =  | 
|
254  | 
add_loose_bnos (f, lev, add_loose_bnos (t, lev, js))  | 
|
255  | 
| add_loose_bnos (_, _, js) = js;  | 
|
256  | 
||
257  | 
fun loose_bnos t = add_loose_bnos (t, 0, []);  | 
|
258  | 
||
259  | 
(* loose_bvar(t,k) iff t contains a 'loose' bound variable referring to  | 
|
260  | 
level k or beyond. *)  | 
|
261  | 
fun loose_bvar(Bound i,k) = i >= k  | 
|
262  | 
| loose_bvar(f$t, k) = loose_bvar(f,k) orelse loose_bvar(t,k)  | 
|
263  | 
| loose_bvar(Abs(_,_,t),k) = loose_bvar(t,k+1)  | 
|
264  | 
| loose_bvar _ = false;  | 
|
265  | 
||
266  | 
||
267  | 
(*Substitute arguments for loose bound variables.  | 
|
268  | 
Beta-reduction of arg(n-1)...arg0 into t replacing (Bound i) with (argi).  | 
|
269  | 
Note that for ((x,y)c)(a,b), the bound vars in c are x=1 and y=0  | 
|
270  | 
and the appropriate call is subst_bounds([b,a], c) .  | 
|
271  | 
Loose bound variables >=n are reduced by "n" to  | 
|
272  | 
compensate for the disappearance of lambdas.  | 
|
273  | 
*)  | 
|
274  | 
fun subst_bounds (args: term list, t) : term =  | 
|
275  | 
let val n = length args;  | 
|
276  | 
fun subst (t as Bound i, lev) =  | 
|
277  | 
if i<lev then t (*var is locally bound*)  | 
|
278  | 
else (case (drop (i-lev,args)) of  | 
|
279  | 
[] => Bound(i-n) (*loose: change it*)  | 
|
280  | 
| arg::_ => incr_boundvars lev arg)  | 
|
281  | 
| subst (Abs(a,T,body), lev) = Abs(a, T, subst(body,lev+1))  | 
|
282  | 
| subst (f$t, lev) = subst(f,lev) $ subst(t,lev)  | 
|
283  | 
| subst (t,lev) = t  | 
|
284  | 
in case args of [] => t | _ => subst (t,0) end;  | 
|
285  | 
||
286  | 
(*beta-reduce if possible, else form application*)  | 
|
287  | 
fun betapply (Abs(_,_,t), u) = subst_bounds([u],t)  | 
|
288  | 
| betapply (f,u) = f$u;  | 
|
289  | 
||
290  | 
(*Tests whether 2 terms are alpha-convertible and have same type.  | 
|
291  | 
Note that constants and Vars may have more than one type.*)  | 
|
292  | 
infix aconv;  | 
|
293  | 
fun (Const(a,T)) aconv (Const(b,U)) = a=b andalso T=U  | 
|
294  | 
| (Free(a,T)) aconv (Free(b,U)) = a=b andalso T=U  | 
|
295  | 
| (Var(v,T)) aconv (Var(w,U)) = v=w andalso T=U  | 
|
296  | 
| (Bound i) aconv (Bound j) = i=j  | 
|
297  | 
| (Abs(_,T,t)) aconv (Abs(_,U,u)) = t aconv u andalso T=U  | 
|
298  | 
| (f$t) aconv (g$u) = (f aconv g) andalso (t aconv u)  | 
|
299  | 
| _ aconv _ = false;  | 
|
300  | 
||
301  | 
(*are two term lists alpha-convertible in corresponding elements?*)  | 
|
302  | 
fun aconvs ([],[]) = true  | 
|
303  | 
| aconvs (t::ts, u::us) = t aconv u andalso aconvs(ts,us)  | 
|
304  | 
| aconvs _ = false;  | 
|
305  | 
||
306  | 
(*A fast unification filter: true unless the two terms cannot be unified.  | 
|
307  | 
Terms must be NORMAL. Treats all Vars as distinct. *)  | 
|
308  | 
fun could_unify (t,u) =  | 
|
309  | 
let fun matchrands (f$t, g$u) = could_unify(t,u) andalso matchrands(f,g)  | 
|
310  | 
| matchrands _ = true  | 
|
311  | 
in case (head_of t , head_of u) of  | 
|
312  | 
(_, Var _) => true  | 
|
313  | 
| (Var _, _) => true  | 
|
314  | 
| (Const(a,_), Const(b,_)) => a=b andalso matchrands(t,u)  | 
|
315  | 
| (Free(a,_), Free(b,_)) => a=b andalso matchrands(t,u)  | 
|
316  | 
| (Bound i, Bound j) => i=j andalso matchrands(t,u)  | 
|
317  | 
| (Abs _, _) => true (*because of possible eta equality*)  | 
|
318  | 
| (_, Abs _) => true  | 
|
319  | 
| _ => false  | 
|
320  | 
end;  | 
|
321  | 
||
322  | 
(*Substitute new for free occurrences of old in a term*)  | 
|
323  | 
fun subst_free [] = (fn t=>t)  | 
|
324  | 
| subst_free pairs =  | 
|
325  | 
let fun substf u =  | 
|
326  | 
case gen_assoc (op aconv) (pairs, u) of  | 
|
327  | 
Some u' => u'  | 
|
328  | 
| None => (case u of Abs(a,T,t) => Abs(a, T, substf t)  | 
|
329  | 
| t$u' => substf t $ substf u'  | 
|
330  | 
| _ => u)  | 
|
331  | 
in substf end;  | 
|
332  | 
||
333  | 
(*a total, irreflexive ordering on index names*)  | 
|
334  | 
fun xless ((a,i), (b,j): indexname) = i<j orelse (i=j andalso a<b);  | 
|
335  | 
||
336  | 
||
337  | 
(*Abstraction of the term "body" over its occurrences of v,  | 
|
338  | 
which must contain no loose bound variables.  | 
|
339  | 
The resulting term is ready to become the body of an Abs.*)  | 
|
340  | 
fun abstract_over (v,body) =  | 
|
341  | 
let fun abst (lev,u) = if (v aconv u) then (Bound lev) else  | 
|
342  | 
(case u of  | 
|
343  | 
Abs(a,T,t) => Abs(a, T, abst(lev+1, t))  | 
|
344  | 
| f$rand => abst(lev,f) $ abst(lev,rand)  | 
|
345  | 
| _ => u)  | 
|
346  | 
in abst(0,body) end;  | 
|
347  | 
||
348  | 
||
349  | 
(*Form an abstraction over a free variable.*)  | 
|
350  | 
fun absfree (a,T,body) = Abs(a, T, abstract_over (Free(a,T), body));  | 
|
351  | 
||
352  | 
(*Abstraction over a list of free variables*)  | 
|
353  | 
fun list_abs_free ([ ] , t) = t  | 
|
354  | 
| list_abs_free ((a,T)::vars, t) =  | 
|
355  | 
absfree(a, T, list_abs_free(vars,t));  | 
|
356  | 
||
357  | 
(*Quantification over a list of free variables*)  | 
|
358  | 
fun list_all_free ([], t: term) = t  | 
|
359  | 
| list_all_free ((a,T)::vars, t) =  | 
|
360  | 
(all T) $ (absfree(a, T, list_all_free(vars,t)));  | 
|
361  | 
||
362  | 
(*Quantification over a list of variables (already bound in body) *)  | 
|
363  | 
fun list_all ([], t) = t  | 
|
364  | 
| list_all ((a,T)::vars, t) =  | 
|
365  | 
(all T) $ (Abs(a, T, list_all(vars,t)));  | 
|
366  | 
||
367  | 
(*Replace the ATOMIC term ti by ui; instl = [(t1,u1), ..., (tn,un)].  | 
|
368  | 
A simultaneous substitution: [ (a,b), (b,a) ] swaps a and b. *)  | 
|
369  | 
fun subst_atomic [] t = t : term  | 
|
370  | 
| subst_atomic (instl: (term*term) list) t =  | 
|
371  | 
let fun subst (Abs(a,T,body)) = Abs(a, T, subst body)  | 
|
372  | 
| subst (f$t') = subst f $ subst t'  | 
|
373  | 
| subst t = (case assoc(instl,t) of  | 
|
374  | 
Some u => u | None => t)  | 
|
375  | 
in subst t end;  | 
|
376  | 
||
377  | 
fun typ_subst_TVars iTs T = if null iTs then T else  | 
|
378  | 
let fun subst(Type(a,Ts)) = Type(a, map subst Ts)  | 
|
379  | 
| subst(T as TFree _) = T  | 
|
380  | 
| subst(T as TVar(ixn,_)) =  | 
|
381  | 
(case assoc(iTs,ixn) of None => T | Some(U) => U)  | 
|
382  | 
in subst T end;  | 
|
383  | 
||
384  | 
val subst_TVars = map_term_types o typ_subst_TVars;  | 
|
385  | 
||
386  | 
fun subst_Vars itms t = if null itms then t else  | 
|
387  | 
let fun subst(v as Var(ixn,_)) =  | 
|
388  | 
(case assoc(itms,ixn) of None => v | Some t => t)  | 
|
389  | 
| subst(Abs(a,T,t)) = Abs(a,T,subst t)  | 
|
390  | 
| subst(f$t) = subst f $ subst t  | 
|
391  | 
| subst(t) = t  | 
|
392  | 
in subst t end;  | 
|
393  | 
||
394  | 
fun subst_vars(iTs,itms) = if null iTs then subst_Vars itms else  | 
|
395  | 
let fun subst(Const(a,T)) = Const(a,typ_subst_TVars iTs T)  | 
|
396  | 
| subst(Free(a,T)) = Free(a,typ_subst_TVars iTs T)  | 
|
397  | 
| subst(v as Var(ixn,T)) = (case assoc(itms,ixn) of  | 
|
398  | 
None => Var(ixn,typ_subst_TVars iTs T)  | 
|
399  | 
| Some t => t)  | 
|
400  | 
| subst(b as Bound _) = b  | 
|
401  | 
| subst(Abs(a,T,t)) = Abs(a,typ_subst_TVars iTs T,subst t)  | 
|
402  | 
| subst(f$t) = subst f $ subst t  | 
|
403  | 
in subst end;  | 
|
404  | 
||
405  | 
||
406  | 
(*Computing the maximum index of a typ*)  | 
|
407  | 
fun maxidx_of_typ(Type(_,Ts)) =  | 
|
408  | 
if Ts=[] then ~1 else max(map maxidx_of_typ Ts)  | 
|
409  | 
| maxidx_of_typ(TFree _) = ~1  | 
|
410  | 
| maxidx_of_typ(TVar((_,i),_)) = i;  | 
|
411  | 
||
412  | 
||
413  | 
(*Computing the maximum index of a term*)  | 
|
414  | 
fun maxidx_of_term (Const(_,T)) = maxidx_of_typ T  | 
|
415  | 
| maxidx_of_term (Bound _) = ~1  | 
|
416  | 
| maxidx_of_term (Free(_,T)) = maxidx_of_typ T  | 
|
417  | 
| maxidx_of_term (Var ((_,i), T)) = max[i, maxidx_of_typ T]  | 
|
418  | 
| maxidx_of_term (Abs (_,T,body)) = max[maxidx_of_term body, maxidx_of_typ T]  | 
|
419  | 
| maxidx_of_term (f$t) = max [maxidx_of_term f, maxidx_of_term t];  | 
|
420  | 
||
421  | 
||
422  | 
(* Increment the index of all Poly's in T by k *)  | 
|
423  | 
fun incr_tvar k (Type(a,Ts)) = Type(a, map (incr_tvar k) Ts)  | 
|
424  | 
| incr_tvar k (T as TFree _) = T  | 
|
425  | 
| incr_tvar k (TVar((a,i),rs)) = TVar((a,i+k),rs);  | 
|
426  | 
||
427  | 
||
428  | 
(**** Syntax-related declarations ****)  | 
|
429  | 
||
430  | 
||
431  | 
(*Dummy type for parsing. Will be replaced during type inference. *)  | 
|
432  | 
val dummyT = Type("dummy",[]);
 | 
|
433  | 
||
434  | 
(*scan a numeral of the given radix, normally 10*)  | 
|
435  | 
fun scan_radixint (radix: int, cs) : int * string list =  | 
|
436  | 
let val zero = ord"0"  | 
|
437  | 
val limit = zero+radix  | 
|
438  | 
fun scan (num,[]) = (num,[])  | 
|
439  | 
| scan (num, c::cs) =  | 
|
440  | 
if zero <= ord c andalso ord c < limit  | 
|
441  | 
then scan(radix*num + ord c - zero, cs)  | 
|
442  | 
else (num, c::cs)  | 
|
443  | 
in scan(0,cs) end;  | 
|
444  | 
||
445  | 
fun scan_int cs = scan_radixint(10,cs);  | 
|
446  | 
||
447  | 
||
448  | 
(*** Printing ***)  | 
|
449  | 
||
450  | 
||
451  | 
(*Makes a variant of the name c distinct from the names in bs.  | 
|
452  | 
First attaches the suffix "a" and then increments this. *)  | 
|
453  | 
fun variant bs c : string =  | 
|
454  | 
let fun vary2 c = if (c mem bs) then vary2 (bump_string c) else c  | 
|
455  | 
fun vary1 c = if (c mem bs) then vary2 (c ^ "a") else c  | 
|
456  | 
in vary1 (if c="" then "u" else c) end;  | 
|
457  | 
||
458  | 
(*Create variants of the list of names, with priority to the first ones*)  | 
|
459  | 
fun variantlist ([], used) = []  | 
|
460  | 
| variantlist(b::bs, used) =  | 
|
461  | 
let val b' = variant used b  | 
|
462  | 
in b' :: variantlist (bs, b'::used) end;  | 
|
463  | 
||
464  | 
(** TFrees and TVars **)  | 
|
465  | 
||
466  | 
(*maps (bs,v) to v'::bs this reverses the identifiers bs*)  | 
|
467  | 
fun add_new_id (bs, c) : string list = variant bs c :: bs;  | 
|
468  | 
||
469  | 
(*Accumulates the names in the term, suppressing duplicates.  | 
|
470  | 
Includes Frees and Consts. For choosing unambiguous bound var names.*)  | 
|
471  | 
fun add_term_names (Const(a,_), bs) = a ins bs  | 
|
472  | 
| add_term_names (Free(a,_), bs) = a ins bs  | 
|
473  | 
| add_term_names (f$u, bs) = add_term_names (f, add_term_names(u, bs))  | 
|
474  | 
| add_term_names (Abs(_,_,t), bs) = add_term_names(t,bs)  | 
|
475  | 
| add_term_names (_, bs) = bs;  | 
|
476  | 
||
477  | 
(*Accumulates the TVars in a type, suppressing duplicates. *)  | 
|
478  | 
fun add_typ_tvars(Type(_,Ts),vs) = foldr add_typ_tvars (Ts,vs)  | 
|
479  | 
| add_typ_tvars(TFree(_),vs) = vs  | 
|
480  | 
| add_typ_tvars(TVar(v),vs) = v ins vs;  | 
|
481  | 
||
482  | 
(*Accumulates the TFrees in a type, suppressing duplicates. *)  | 
|
483  | 
fun add_typ_tfree_names(Type(_,Ts),fs) = foldr add_typ_tfree_names (Ts,fs)  | 
|
484  | 
| add_typ_tfree_names(TFree(f,_),fs) = f ins fs  | 
|
485  | 
| add_typ_tfree_names(TVar(_),fs) = fs;  | 
|
486  | 
||
487  | 
fun add_typ_tfrees(Type(_,Ts),fs) = foldr add_typ_tfrees (Ts,fs)  | 
|
488  | 
| add_typ_tfrees(TFree(f),fs) = f ins fs  | 
|
489  | 
| add_typ_tfrees(TVar(_),fs) = fs;  | 
|
490  | 
||
491  | 
(*Accumulates the TVars in a term, suppressing duplicates. *)  | 
|
492  | 
val add_term_tvars = it_term_types add_typ_tvars;  | 
|
493  | 
val add_term_tvar_ixns = (map #1) o (it_term_types add_typ_tvars);  | 
|
494  | 
||
495  | 
(*Accumulates the TFrees in a term, suppressing duplicates. *)  | 
|
496  | 
val add_term_tfrees = it_term_types add_typ_tfrees;  | 
|
497  | 
val add_term_tfree_names = it_term_types add_typ_tfree_names;  | 
|
498  | 
||
499  | 
(*Non-list versions*)  | 
|
500  | 
fun typ_tfrees T = add_typ_tfrees(T,[]);  | 
|
501  | 
fun typ_tvars T = add_typ_tvars(T,[]);  | 
|
502  | 
fun term_tfrees t = add_term_tfrees(t,[]);  | 
|
503  | 
fun term_tvars t = add_term_tvars(t,[]);  | 
|
504  | 
||
505  | 
(** Frees and Vars **)  | 
|
506  | 
||
507  | 
(*a partial ordering (not reflexive) for atomic terms*)  | 
|
508  | 
fun atless (Const (a,_), Const (b,_)) = a<b  | 
|
509  | 
| atless (Free (a,_), Free (b,_)) = a<b  | 
|
510  | 
| atless (Var(v,_), Var(w,_)) = xless(v,w)  | 
|
511  | 
| atless (Bound i, Bound j) = i<j  | 
|
512  | 
| atless _ = false;  | 
|
513  | 
||
514  | 
(*insert atomic term into partially sorted list, suppressing duplicates (?)*)  | 
|
515  | 
fun insert_aterm (t,us) =  | 
|
516  | 
let fun inserta [] = [t]  | 
|
517  | 
| inserta (us as u::us') =  | 
|
518  | 
if atless(t,u) then t::us  | 
|
519  | 
else if t=u then us (*duplicate*)  | 
|
520  | 
else u :: inserta(us')  | 
|
521  | 
in inserta us end;  | 
|
522  | 
||
523  | 
(*Accumulates the Vars in the term, suppressing duplicates*)  | 
|
524  | 
fun add_term_vars (t, vars: term list) = case t of  | 
|
525  | 
Var _ => insert_aterm(t,vars)  | 
|
526  | 
| Abs (_,_,body) => add_term_vars(body,vars)  | 
|
527  | 
| f$t => add_term_vars (f, add_term_vars(t, vars))  | 
|
528  | 
| _ => vars;  | 
|
529  | 
||
530  | 
fun term_vars t = add_term_vars(t,[]);  | 
|
531  | 
||
532  | 
(*Accumulates the Frees in the term, suppressing duplicates*)  | 
|
533  | 
fun add_term_frees (t, frees: term list) = case t of  | 
|
534  | 
Free _ => insert_aterm(t,frees)  | 
|
535  | 
| Abs (_,_,body) => add_term_frees(body,frees)  | 
|
536  | 
| f$t => add_term_frees (f, add_term_frees(t, frees))  | 
|
537  | 
| _ => frees;  | 
|
538  | 
||
539  | 
fun term_frees t = add_term_frees(t,[]);  | 
|
540  | 
||
541  | 
(*Given an abstraction over P, replaces the bound variable by a Free variable  | 
|
542  | 
having a unique name. *)  | 
|
543  | 
fun variant_abs (a,T,P) =  | 
|
544  | 
let val b = variant (add_term_names(P,[])) a  | 
|
545  | 
in (b, subst_bounds ([Free(b,T)], P)) end;  | 
|
546  | 
||
547  | 
(* renames and reverses the strings in vars away from names *)  | 
|
548  | 
fun rename_aTs names vars : (string*typ)list =  | 
|
549  | 
let fun rename_aT (vars,(a,T)) =  | 
|
550  | 
(variant (map #1 vars @ names) a, T) :: vars  | 
|
551  | 
in foldl rename_aT ([],vars) end;  | 
|
552  | 
||
553  | 
fun rename_wrt_term t = rename_aTs (add_term_names(t,[]));  |