author | paulson |
Fri, 18 Apr 1997 11:48:16 +0200 | |
changeset 2983 | f914a1663b2a |
parent 2508 | ce48daa388a7 |
child 3408 | 98a2d517cabe |
permissions | -rw-r--r-- |
1460 | 1 |
(* Title: Pure/logic.ML |
0 | 2 |
ID: $Id$ |
1460 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright Cambridge University 1992 |
5 |
||
6 |
Supporting code for defining the abstract type "thm" |
|
7 |
*) |
|
8 |
||
9 |
infix occs; |
|
10 |
||
11 |
signature LOGIC = |
|
12 |
sig |
|
1460 | 13 |
val assum_pairs : term -> (term*term)list |
14 |
val auto_rename : bool ref |
|
15 |
val close_form : term -> term |
|
16 |
val count_prems : term * int -> int |
|
17 |
val dest_equals : term -> term * term |
|
18 |
val dest_flexpair : term -> term * term |
|
19 |
val dest_implies : term -> term * term |
|
20 |
val dest_inclass : term -> typ * class |
|
21 |
val dest_type : term -> typ |
|
22 |
val flatten_params : int -> term -> term |
|
23 |
val freeze_vars : term -> term |
|
24 |
val incr_indexes : typ list * int -> term -> term |
|
25 |
val lift_fns : term * int -> (term -> term) * (term -> term) |
|
26 |
val list_flexpairs : (term*term)list * term -> term |
|
27 |
val list_implies : term list * term -> term |
|
0 | 28 |
val list_rename_params: string list * term -> term |
637 | 29 |
val is_equals : term -> bool |
1460 | 30 |
val mk_equals : term * term -> term |
31 |
val mk_flexpair : term * term -> term |
|
32 |
val mk_implies : term * term -> term |
|
33 |
val mk_inclass : typ * class -> term |
|
34 |
val mk_type : typ -> term |
|
35 |
val occs : term * term -> bool |
|
36 |
val rule_of : (term*term)list * term list * term -> term |
|
37 |
val set_rename_prefix : string -> unit |
|
38 |
val skip_flexpairs : term -> term |
|
0 | 39 |
val strip_assums_concl: term -> term |
1460 | 40 |
val strip_assums_hyp : term -> term list |
41 |
val strip_flexpairs : term -> (term*term)list * term |
|
42 |
val strip_horn : term -> (term*term)list * term list * term |
|
43 |
val strip_imp_concl : term -> term |
|
44 |
val strip_imp_prems : term -> term list |
|
45 |
val strip_params : term -> (string * typ) list |
|
46 |
val strip_prems : int * term list * term -> term list * term |
|
47 |
val thaw_vars : term -> term |
|
2508 | 48 |
val unvarify : term -> term |
49 |
val varify : term -> term |
|
50 |
val termord : term * term -> order |
|
51 |
val lextermord : term list * term list -> order |
|
52 |
val termless : term * term -> bool |
|
0 | 53 |
end; |
54 |
||
1500 | 55 |
structure Logic : LOGIC = |
0 | 56 |
struct |
398
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
57 |
|
0 | 58 |
(*** Abstract syntax operations on the meta-connectives ***) |
59 |
||
60 |
(** equality **) |
|
61 |
||
1835 | 62 |
(*Make an equality. DOES NOT CHECK TYPE OF u*) |
64
0bbe5d86cb38
logic/mk_equals,mk_flexpair: now calls fastype_of instead of type_of.
lcp
parents:
0
diff
changeset
|
63 |
fun mk_equals(t,u) = equals(fastype_of t) $ t $ u; |
0 | 64 |
|
65 |
fun dest_equals (Const("==",_) $ t $ u) = (t,u) |
|
66 |
| dest_equals t = raise TERM("dest_equals", [t]); |
|
67 |
||
637 | 68 |
fun is_equals (Const ("==", _) $ _ $ _) = true |
69 |
| is_equals _ = false; |
|
70 |
||
71 |
||
0 | 72 |
(** implies **) |
73 |
||
74 |
fun mk_implies(A,B) = implies $ A $ B; |
|
75 |
||
76 |
fun dest_implies (Const("==>",_) $ A $ B) = (A,B) |
|
77 |
| dest_implies A = raise TERM("dest_implies", [A]); |
|
78 |
||
79 |
(** nested implications **) |
|
80 |
||
81 |
(* [A1,...,An], B goes to A1==>...An==>B *) |
|
82 |
fun list_implies ([], B) = B : term |
|
83 |
| list_implies (A::AS, B) = implies $ A $ list_implies(AS,B); |
|
84 |
||
85 |
(* A1==>...An==>B goes to [A1,...,An], where B is not an implication *) |
|
86 |
fun strip_imp_prems (Const("==>", _) $ A $ B) = A :: strip_imp_prems B |
|
87 |
| strip_imp_prems _ = []; |
|
88 |
||
89 |
(* A1==>...An==>B goes to B, where B is not an implication *) |
|
90 |
fun strip_imp_concl (Const("==>", _) $ A $ B) = strip_imp_concl B |
|
91 |
| strip_imp_concl A = A : term; |
|
92 |
||
93 |
(*Strip and return premises: (i, [], A1==>...Ai==>B) |
|
1460 | 94 |
goes to ([Ai, A(i-1),...,A1] , B) (REVERSED) |
0 | 95 |
if i<0 or else i too big then raises TERM*) |
96 |
fun strip_prems (0, As, B) = (As, B) |
|
97 |
| strip_prems (i, As, Const("==>", _) $ A $ B) = |
|
1460 | 98 |
strip_prems (i-1, A::As, B) |
0 | 99 |
| strip_prems (_, As, A) = raise TERM("strip_prems", A::As); |
100 |
||
101 |
(*Count premises -- quicker than (length ostrip_prems) *) |
|
102 |
fun count_prems (Const("==>", _) $ A $ B, n) = count_prems (B,n+1) |
|
103 |
| count_prems (_,n) = n; |
|
104 |
||
105 |
(** flex-flex constraints **) |
|
106 |
||
64
0bbe5d86cb38
logic/mk_equals,mk_flexpair: now calls fastype_of instead of type_of.
lcp
parents:
0
diff
changeset
|
107 |
(*Make a constraint.*) |
0bbe5d86cb38
logic/mk_equals,mk_flexpair: now calls fastype_of instead of type_of.
lcp
parents:
0
diff
changeset
|
108 |
fun mk_flexpair(t,u) = flexpair(fastype_of t) $ t $ u; |
0 | 109 |
|
110 |
fun dest_flexpair (Const("=?=",_) $ t $ u) = (t,u) |
|
111 |
| dest_flexpair t = raise TERM("dest_flexpair", [t]); |
|
112 |
||
113 |
(*make flexflex antecedents: ( [(a1,b1),...,(an,bn)] , C ) |
|
114 |
goes to (a1=?=b1) ==>...(an=?=bn)==>C *) |
|
115 |
fun list_flexpairs ([], A) = A |
|
116 |
| list_flexpairs ((t,u)::pairs, A) = |
|
1460 | 117 |
implies $ (mk_flexpair(t,u)) $ list_flexpairs(pairs,A); |
0 | 118 |
|
119 |
(*Make the object-rule tpairs==>As==>B *) |
|
120 |
fun rule_of (tpairs, As, B) = list_flexpairs(tpairs, list_implies(As, B)); |
|
121 |
||
122 |
(*Remove and return flexflex pairs: |
|
1460 | 123 |
(a1=?=b1)==>...(an=?=bn)==>C to ( [(a1,b1),...,(an,bn)] , C ) |
0 | 124 |
[Tail recursive in order to return a pair of results] *) |
125 |
fun strip_flex_aux (pairs, Const("==>", _) $ (Const("=?=",_)$t$u) $ C) = |
|
126 |
strip_flex_aux ((t,u)::pairs, C) |
|
127 |
| strip_flex_aux (pairs,C) = (rev pairs, C); |
|
128 |
||
129 |
fun strip_flexpairs A = strip_flex_aux([], A); |
|
130 |
||
131 |
(*Discard flexflex pairs*) |
|
132 |
fun skip_flexpairs (Const("==>", _) $ (Const("=?=",_)$_$_) $ C) = |
|
1460 | 133 |
skip_flexpairs C |
0 | 134 |
| skip_flexpairs C = C; |
135 |
||
136 |
(*strip a proof state (Horn clause): |
|
137 |
(a1==b1)==>...(am==bm)==>B1==>...Bn==>C |
|
138 |
goes to ( [(a1,b1),...,(am,bm)] , [B1,...,Bn] , C) *) |
|
139 |
fun strip_horn A = |
|
140 |
let val (tpairs,horn) = strip_flexpairs A |
|
141 |
in (tpairs, strip_imp_prems horn, strip_imp_concl horn) end; |
|
142 |
||
398
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
143 |
(** types as terms **) |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
144 |
|
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
145 |
fun mk_type ty = Const ("TYPE", itselfT ty); |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
146 |
|
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
147 |
fun dest_type (Const ("TYPE", Type ("itself", [ty]))) = ty |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
148 |
| dest_type t = raise TERM ("dest_type", [t]); |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
149 |
|
447 | 150 |
(** class constraints **) |
398
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
151 |
|
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
152 |
fun mk_inclass (ty, c) = |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
153 |
Const (Sign.const_of_class c, itselfT ty --> propT) $ mk_type ty; |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
154 |
|
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
155 |
fun dest_inclass (t as Const (c_class, _) $ ty) = |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
156 |
((dest_type ty, Sign.class_of_const c_class) |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
157 |
handle TERM _ => raise TERM ("dest_inclass", [t])) |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
158 |
| dest_inclass t = raise TERM ("dest_inclass", [t]); |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
159 |
|
0 | 160 |
|
161 |
(*** Low-level term operations ***) |
|
162 |
||
163 |
(*Does t occur in u? Or is alpha-convertible to u? |
|
164 |
The term t must contain no loose bound variables*) |
|
165 |
fun t occs u = (t aconv u) orelse |
|
166 |
(case u of |
|
167 |
Abs(_,_,body) => t occs body |
|
1460 | 168 |
| f$t' => t occs f orelse t occs t' |
169 |
| _ => false); |
|
0 | 170 |
|
171 |
(*Close up a formula over all free variables by quantification*) |
|
172 |
fun close_form A = |
|
173 |
list_all_free (map dest_Free (sort atless (term_frees A)), |
|
1460 | 174 |
A); |
0 | 175 |
|
176 |
||
177 |
(*Freeze all (T)Vars by turning them into (T)Frees*) |
|
178 |
fun freeze_vars(Var(ixn,T)) = Free(Syntax.string_of_vname ixn, |
|
179 |
Type.freeze_vars T) |
|
180 |
| freeze_vars(Const(a,T)) = Const(a,Type.freeze_vars T) |
|
181 |
| freeze_vars(Free(a,T)) = Free(a,Type.freeze_vars T) |
|
182 |
| freeze_vars(s$t) = freeze_vars s $ freeze_vars t |
|
183 |
| freeze_vars(Abs(a,T,t)) = Abs(a,Type.freeze_vars T,freeze_vars t) |
|
184 |
| freeze_vars(b) = b; |
|
185 |
||
186 |
(*Reverse the effect of freeze_vars*) |
|
187 |
fun thaw_vars(Const(a,T)) = Const(a,Type.thaw_vars T) |
|
188 |
| thaw_vars(Free(a,T)) = |
|
189 |
let val T' = Type.thaw_vars T |
|
190 |
in case explode a of |
|
1460 | 191 |
"?"::vn => let val (ixn,_) = Syntax.scan_varname vn |
0 | 192 |
in Var(ixn,T') end |
1460 | 193 |
| _ => Free(a,T') |
0 | 194 |
end |
195 |
| thaw_vars(Abs(a,T,t)) = Abs(a,Type.thaw_vars T, thaw_vars t) |
|
196 |
| thaw_vars(s$t) = thaw_vars s $ thaw_vars t |
|
197 |
| thaw_vars(b) = b; |
|
198 |
||
199 |
||
200 |
(*** Specialized operations for resolution... ***) |
|
201 |
||
202 |
(*For all variables in the term, increment indexnames and lift over the Us |
|
203 |
result is ?Gidx(B.(lev+n-1),...,B.lev) where lev is abstraction level *) |
|
204 |
fun incr_indexes (Us: typ list, inc:int) t = |
|
205 |
let fun incr (Var ((a,i), T), lev) = |
|
1460 | 206 |
Unify.combound (Var((a, i+inc), Us---> incr_tvar inc T), |
207 |
lev, length Us) |
|
208 |
| incr (Abs (a,T,body), lev) = |
|
209 |
Abs (a, incr_tvar inc T, incr(body,lev+1)) |
|
210 |
| incr (Const(a,T),_) = Const(a, incr_tvar inc T) |
|
211 |
| incr (Free(a,T),_) = Free(a, incr_tvar inc T) |
|
212 |
| incr (f$t, lev) = incr(f,lev) $ incr(t,lev) |
|
213 |
| incr (t,lev) = t |
|
0 | 214 |
in incr(t,0) end; |
215 |
||
216 |
(*Make lifting functions from subgoal and increment. |
|
217 |
lift_abs operates on tpairs (unification constraints) |
|
218 |
lift_all operates on propositions *) |
|
219 |
fun lift_fns (B,inc) = |
|
220 |
let fun lift_abs (Us, Const("==>", _) $ _ $ B) u = lift_abs (Us,B) u |
|
1460 | 221 |
| lift_abs (Us, Const("all",_)$Abs(a,T,t)) u = |
222 |
Abs(a, T, lift_abs (T::Us, t) u) |
|
223 |
| lift_abs (Us, _) u = incr_indexes(rev Us, inc) u |
|
0 | 224 |
fun lift_all (Us, Const("==>", _) $ A $ B) u = |
1460 | 225 |
implies $ A $ lift_all (Us,B) u |
226 |
| lift_all (Us, Const("all",_)$Abs(a,T,t)) u = |
|
227 |
all T $ Abs(a, T, lift_all (T::Us,t) u) |
|
228 |
| lift_all (Us, _) u = incr_indexes(rev Us, inc) u; |
|
0 | 229 |
in (lift_abs([],B), lift_all([],B)) end; |
230 |
||
231 |
(*Strips assumptions in goal, yielding list of hypotheses. *) |
|
232 |
fun strip_assums_hyp (Const("==>", _) $ H $ B) = H :: strip_assums_hyp B |
|
233 |
| strip_assums_hyp (Const("all",_)$Abs(a,T,t)) = strip_assums_hyp t |
|
234 |
| strip_assums_hyp B = []; |
|
235 |
||
236 |
(*Strips assumptions in goal, yielding conclusion. *) |
|
237 |
fun strip_assums_concl (Const("==>", _) $ H $ B) = strip_assums_concl B |
|
238 |
| strip_assums_concl (Const("all",_)$Abs(a,T,t)) = strip_assums_concl t |
|
239 |
| strip_assums_concl B = B; |
|
240 |
||
241 |
(*Make a list of all the parameters in a subgoal, even if nested*) |
|
242 |
fun strip_params (Const("==>", _) $ H $ B) = strip_params B |
|
243 |
| strip_params (Const("all",_)$Abs(a,T,t)) = (a,T) :: strip_params t |
|
244 |
| strip_params B = []; |
|
245 |
||
246 |
(*Removes the parameters from a subgoal and renumber bvars in hypotheses, |
|
247 |
where j is the total number of parameters (precomputed) |
|
248 |
If n>0 then deletes assumption n. *) |
|
249 |
fun remove_params j n A = |
|
250 |
if j=0 andalso n<=0 then A (*nothing left to do...*) |
|
251 |
else case A of |
|
252 |
Const("==>", _) $ H $ B => |
|
1460 | 253 |
if n=1 then (remove_params j (n-1) B) |
254 |
else implies $ (incr_boundvars j H) $ (remove_params j (n-1) B) |
|
0 | 255 |
| Const("all",_)$Abs(a,T,t) => remove_params (j-1) n t |
256 |
| _ => if n>0 then raise TERM("remove_params", [A]) |
|
257 |
else A; |
|
258 |
||
259 |
(** Auto-renaming of parameters in subgoals **) |
|
260 |
||
261 |
val auto_rename = ref false |
|
262 |
and rename_prefix = ref "ka"; |
|
263 |
||
264 |
(*rename_prefix is not exported; it is set by this function.*) |
|
265 |
fun set_rename_prefix a = |
|
266 |
if a<>"" andalso forall is_letter (explode a) |
|
267 |
then (rename_prefix := a; auto_rename := true) |
|
268 |
else error"rename prefix must be nonempty and consist of letters"; |
|
269 |
||
270 |
(*Makes parameters in a goal have distinctive names (not guaranteed unique!) |
|
271 |
A name clash could cause the printer to rename bound vars; |
|
272 |
then res_inst_tac would not work properly.*) |
|
273 |
fun rename_vars (a, []) = [] |
|
274 |
| rename_vars (a, (_,T)::vars) = |
|
275 |
(a,T) :: rename_vars (bump_string a, vars); |
|
276 |
||
277 |
(*Move all parameters to the front of the subgoal, renaming them apart; |
|
278 |
if n>0 then deletes assumption n. *) |
|
279 |
fun flatten_params n A = |
|
280 |
let val params = strip_params A; |
|
1460 | 281 |
val vars = if !auto_rename |
282 |
then rename_vars (!rename_prefix, params) |
|
2266 | 283 |
else ListPair.zip (variantlist(map #1 params,[]), |
284 |
map #2 params) |
|
0 | 285 |
in list_all (vars, remove_params (length vars) n A) |
286 |
end; |
|
287 |
||
288 |
(*Makes parameters in a goal have the names supplied by the list cs.*) |
|
289 |
fun list_rename_params (cs, Const("==>", _) $ A $ B) = |
|
290 |
implies $ A $ list_rename_params (cs, B) |
|
291 |
| list_rename_params (c::cs, Const("all",_)$Abs(_,T,t)) = |
|
292 |
all T $ Abs(c, T, list_rename_params (cs, t)) |
|
293 |
| list_rename_params (cs, B) = B; |
|
294 |
||
295 |
(*Strips assumptions in goal yielding ( [Hn,...,H1], [xm,...,x1], B ) |
|
296 |
where H1,...,Hn are the hypotheses and x1...xm are the parameters. *) |
|
297 |
fun strip_assums_aux (Hs, params, Const("==>", _) $ H $ B) = |
|
1460 | 298 |
strip_assums_aux (H::Hs, params, B) |
0 | 299 |
| strip_assums_aux (Hs, params, Const("all",_)$Abs(a,T,t)) = |
1460 | 300 |
strip_assums_aux (Hs, (a,T)::params, t) |
0 | 301 |
| strip_assums_aux (Hs, params, B) = (Hs, params, B); |
302 |
||
303 |
fun strip_assums A = strip_assums_aux ([],[],A); |
|
304 |
||
305 |
||
306 |
(*Produces disagreement pairs, one for each assumption proof, in order. |
|
307 |
A is the first premise of the lifted rule, and thus has the form |
|
308 |
H1 ==> ... Hk ==> B and the pairs are (H1,B),...,(Hk,B) *) |
|
309 |
fun assum_pairs A = |
|
310 |
let val (Hs, params, B) = strip_assums A |
|
311 |
val D = Unify.rlist_abs(params, B) |
|
312 |
fun pairrev ([],pairs) = pairs |
|
313 |
| pairrev (H::Hs,pairs) = |
|
1460 | 314 |
pairrev(Hs, (Unify.rlist_abs(params,H), D) :: pairs) |
0 | 315 |
in pairrev (Hs,[]) (*WAS: map pair (rev Hs) *) |
316 |
end; |
|
317 |
||
318 |
||
319 |
(*Converts Frees to Vars and TFrees to TVars so that axioms can be written |
|
320 |
without (?) everywhere*) |
|
321 |
fun varify (Const(a,T)) = Const(a, Type.varifyT T) |
|
322 |
| varify (Free(a,T)) = Var((a,0), Type.varifyT T) |
|
323 |
| varify (Var(ixn,T)) = Var(ixn, Type.varifyT T) |
|
324 |
| varify (Abs (a,T,body)) = Abs (a, Type.varifyT T, varify body) |
|
325 |
| varify (f$t) = varify f $ varify t |
|
326 |
| varify t = t; |
|
327 |
||
546 | 328 |
(*Inverse of varify. Converts axioms back to their original form.*) |
585 | 329 |
fun unvarify (Const(a,T)) = Const(a, Type.unvarifyT T) |
330 |
| unvarify (Var((a,0), T)) = Free(a, Type.unvarifyT T) |
|
331 |
| unvarify (Var(ixn,T)) = Var(ixn, Type.unvarifyT T) (*non-0 index!*) |
|
332 |
| unvarify (Abs (a,T,body)) = Abs (a, Type.unvarifyT T, unvarify body) |
|
546 | 333 |
| unvarify (f$t) = unvarify f $ unvarify t |
334 |
| unvarify t = t; |
|
335 |
||
2508 | 336 |
|
337 |
(*** term order ***) |
|
338 |
||
339 |
(* NB: non-linearity of the ordering is not a soundness problem *) |
|
340 |
||
341 |
(* FIXME: "***ABSTRACTION***" is a hack and makes the ordering non-linear *) |
|
342 |
fun string_of_hd(Const(a,_)) = a |
|
343 |
| string_of_hd(Free(a,_)) = a |
|
344 |
| string_of_hd(Var(v,_)) = Syntax.string_of_vname v |
|
345 |
| string_of_hd(Bound i) = string_of_int i |
|
346 |
| string_of_hd(Abs _) = "***ABSTRACTION***"; |
|
347 |
||
348 |
(* a strict (not reflexive) linear well-founded AC-compatible ordering |
|
349 |
* for terms: |
|
350 |
* s < t <=> 1. size(s) < size(t) or |
|
351 |
2. size(s) = size(t) and s=f(...) and t = g(...) and f<g or |
|
352 |
3. size(s) = size(t) and s=f(s1..sn) and t=f(t1..tn) and |
|
353 |
(s1..sn) < (t1..tn) (lexicographically) |
|
354 |
*) |
|
355 |
||
356 |
(* FIXME: should really take types into account as well. |
|
357 |
* Otherwise non-linear *) |
|
358 |
fun termord(Abs(_,_,t),Abs(_,_,u)) = termord(t,u) |
|
359 |
| termord(t,u) = |
|
360 |
(case intord(size_of_term t,size_of_term u) of |
|
361 |
EQUAL => let val (f,ts) = strip_comb t and (g,us) = strip_comb u |
|
362 |
in case stringord(string_of_hd f, string_of_hd g) of |
|
363 |
EQUAL => lextermord(ts,us) |
|
364 |
| ord => ord |
|
365 |
end |
|
366 |
| ord => ord) |
|
367 |
and lextermord(t::ts,u::us) = |
|
368 |
(case termord(t,u) of |
|
369 |
EQUAL => lextermord(ts,us) |
|
370 |
| ord => ord) |
|
371 |
| lextermord([],[]) = EQUAL |
|
372 |
| lextermord _ = error("lextermord"); |
|
373 |
||
374 |
fun termless tu = (termord tu = LESS); |
|
375 |
||
0 | 376 |
end; |