author | mueller |
Fri, 30 May 1997 16:37:20 +0200 | |
changeset 3380 | 2986e3b1f86a |
parent 3353 | 9112a2efb9a3 |
child 3391 | 5e45dd3b64e9 |
permissions | -rw-r--r-- |
3302 | 1 |
(* Title: TFL/usyntax |
2 |
ID: $Id$ |
|
3 |
Author: Konrad Slind, Cambridge University Computer Laboratory |
|
4 |
Copyright 1997 University of Cambridge |
|
5 |
||
6 |
Emulation of HOL's abstract syntax functions |
|
7 |
*) |
|
8 |
||
2112 | 9 |
structure USyntax : USyntax_sig = |
10 |
struct |
|
11 |
||
12 |
structure Utils = Utils; |
|
13 |
open Utils; |
|
14 |
||
15 |
infix 4 ##; |
|
16 |
||
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
17 |
fun USYN_ERR{func,mesg} = Utils.ERR{module = "USyntax", func = func, mesg = mesg}; |
2112 | 18 |
|
19 |
||
20 |
(*--------------------------------------------------------------------------- |
|
21 |
* |
|
22 |
* Types |
|
23 |
* |
|
24 |
*---------------------------------------------------------------------------*) |
|
25 |
val mk_prim_vartype = TVar; |
|
26 |
fun mk_vartype s = mk_prim_vartype((s,0),["term"]); |
|
27 |
||
28 |
(* But internally, it's useful *) |
|
29 |
fun dest_vtype (TVar x) = x |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
30 |
| dest_vtype _ = raise USYN_ERR{func = "dest_vtype", |
2112 | 31 |
mesg = "not a flexible type variable"}; |
32 |
||
33 |
val is_vartype = Utils.can dest_vtype; |
|
34 |
||
35 |
val type_vars = map mk_prim_vartype o typ_tvars |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
36 |
fun type_varsl L = Utils.mk_set (curry op=) |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
37 |
(Utils.rev_itlist (curry op @ o type_vars) L []); |
2112 | 38 |
|
39 |
val alpha = mk_vartype "'a" |
|
40 |
val beta = mk_vartype "'b" |
|
41 |
||
42 |
||
43 |
||
44 |
(* What nonsense *) |
|
45 |
nonfix -->; |
|
46 |
val --> = -->; |
|
47 |
infixr 3 -->; |
|
48 |
||
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
49 |
fun strip_type ty = (binder_types ty, body_type ty); |
2112 | 50 |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
51 |
fun strip_prod_type (Type("*", [ty1,ty2])) = |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
52 |
strip_prod_type ty1 @ strip_prod_type ty2 |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
53 |
| strip_prod_type ty = [ty]; |
2112 | 54 |
|
55 |
||
56 |
||
57 |
(*--------------------------------------------------------------------------- |
|
58 |
* |
|
59 |
* Terms |
|
60 |
* |
|
61 |
*---------------------------------------------------------------------------*) |
|
62 |
nonfix aconv; |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
63 |
val aconv = curry (op aconv); |
2112 | 64 |
|
65 |
fun free_vars tm = add_term_frees(tm,[]); |
|
66 |
||
67 |
||
68 |
(* Free variables, in order of occurrence, from left to right in the |
|
69 |
* syntax tree. *) |
|
70 |
fun free_vars_lr tm = |
|
71 |
let fun memb x = let fun m[] = false | m(y::rst) = (x=y)orelse m rst in m end |
|
72 |
fun add (t, frees) = case t of |
|
73 |
Free _ => if (memb t frees) then frees else t::frees |
|
74 |
| Abs (_,_,body) => add(body,frees) |
|
75 |
| f$t => add(t, add(f, frees)) |
|
76 |
| _ => frees |
|
77 |
in rev(add(tm,[])) |
|
78 |
end; |
|
79 |
||
80 |
||
81 |
||
82 |
fun free_varsl L = Utils.mk_set aconv |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
83 |
(Utils.rev_itlist (curry op @ o free_vars) L []); |
2112 | 84 |
|
85 |
val type_vars_in_term = map mk_prim_vartype o term_tvars; |
|
86 |
||
87 |
(* Can't really be very exact in Isabelle *) |
|
88 |
fun all_vars tm = |
|
89 |
let fun memb x = let fun m[] = false | m(y::rst) = (x=y)orelse m rst in m end |
|
90 |
fun add (t, A) = case t of |
|
91 |
Free _ => if (memb t A) then A else t::A |
|
92 |
| Abs (s,ty,body) => add(body, add(Free(s,ty),A)) |
|
93 |
| f$t => add(t, add(f, A)) |
|
94 |
| _ => A |
|
95 |
in rev(add(tm,[])) |
|
96 |
end; |
|
97 |
fun all_varsl L = Utils.mk_set aconv |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
98 |
(Utils.rev_itlist (curry op @ o all_vars) L []); |
2112 | 99 |
|
100 |
||
101 |
(* Prelogic *) |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
102 |
fun dest_tybinding (v,ty) = (#1(dest_vtype v),ty) |
2112 | 103 |
fun inst theta = subst_vars (map dest_tybinding theta,[]) |
104 |
||
105 |
fun beta_conv((t1 as Abs _ ) $ t2) = betapply(t1,t2) |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
106 |
| beta_conv _ = raise USYN_ERR{func = "beta_conv", mesg = "Not a beta-redex"}; |
2112 | 107 |
|
108 |
||
109 |
(* Construction routines *) |
|
110 |
fun mk_var{Name,Ty} = Free(Name,Ty); |
|
111 |
val mk_prim_var = Var; |
|
112 |
||
113 |
val string_variant = variant; |
|
114 |
||
115 |
local fun var_name(Var((Name,_),_)) = Name |
|
116 |
| var_name(Free(s,_)) = s |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
117 |
| var_name _ = raise USYN_ERR{func = "variant", |
2112 | 118 |
mesg = "list elem. is not a variable"} |
119 |
in |
|
120 |
fun variant [] v = v |
|
121 |
| variant vlist (Var((Name,i),ty)) = |
|
122 |
Var((string_variant (map var_name vlist) Name,i),ty) |
|
123 |
| variant vlist (Free(Name,ty)) = |
|
124 |
Free(string_variant (map var_name vlist) Name,ty) |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
125 |
| variant _ _ = raise USYN_ERR{func = "variant", |
2112 | 126 |
mesg = "2nd arg. should be a variable"} |
127 |
end; |
|
128 |
||
129 |
fun mk_const{Name,Ty} = Const(Name,Ty) |
|
130 |
fun mk_comb{Rator,Rand} = Rator $ Rand; |
|
131 |
||
132 |
fun mk_abs{Bvar as Var((s,_),ty),Body} = Abs(s,ty,abstract_over(Bvar,Body)) |
|
133 |
| mk_abs{Bvar as Free(s,ty),Body} = Abs(s,ty,abstract_over(Bvar,Body)) |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
134 |
| mk_abs _ = raise USYN_ERR{func = "mk_abs", mesg = "Bvar is not a variable"}; |
2112 | 135 |
|
136 |
||
137 |
fun mk_imp{ant,conseq} = |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
138 |
let val c = mk_const{Name = "op -->", Ty = HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT} |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
139 |
in list_comb(c,[ant,conseq]) |
2112 | 140 |
end; |
141 |
||
142 |
fun mk_select (r as {Bvar,Body}) = |
|
143 |
let val ty = type_of Bvar |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
144 |
val c = mk_const{Name = "Eps", Ty = (ty --> HOLogic.boolT) --> ty} |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
145 |
in list_comb(c,[mk_abs r]) |
2112 | 146 |
end; |
147 |
||
148 |
fun mk_forall (r as {Bvar,Body}) = |
|
149 |
let val ty = type_of Bvar |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
150 |
val c = mk_const{Name = "All", Ty = (ty --> HOLogic.boolT) --> HOLogic.boolT} |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
151 |
in list_comb(c,[mk_abs r]) |
2112 | 152 |
end; |
153 |
||
154 |
fun mk_exists (r as {Bvar,Body}) = |
|
155 |
let val ty = type_of Bvar |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
156 |
val c = mk_const{Name = "Ex", Ty = (ty --> HOLogic.boolT) --> HOLogic.boolT} |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
157 |
in list_comb(c,[mk_abs r]) |
2112 | 158 |
end; |
159 |
||
160 |
||
161 |
fun mk_conj{conj1,conj2} = |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
162 |
let val c = mk_const{Name = "op &", Ty = HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT} |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
163 |
in list_comb(c,[conj1,conj2]) |
2112 | 164 |
end; |
165 |
||
166 |
fun mk_disj{disj1,disj2} = |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
167 |
let val c = mk_const{Name = "op |", Ty = HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT} |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
168 |
in list_comb(c,[disj1,disj2]) |
2112 | 169 |
end; |
170 |
||
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
171 |
fun prod_ty ty1 ty2 = Type("*", [ty1,ty2]); |
2112 | 172 |
|
173 |
local |
|
174 |
fun mk_uncurry(xt,yt,zt) = |
|
175 |
mk_const{Name = "split", Ty = (xt --> yt --> zt) --> prod_ty xt yt --> zt} |
|
176 |
fun dest_pair(Const("Pair",_) $ M $ N) = {fst=M, snd=N} |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
177 |
| dest_pair _ = raise USYN_ERR{func = "dest_pair", mesg = "not a pair"} |
2112 | 178 |
fun is_var(Var(_)) = true | is_var (Free _) = true | is_var _ = false |
179 |
in |
|
180 |
fun mk_pabs{varstruct,body} = |
|
181 |
let fun mpa(varstruct,body) = |
|
182 |
if (is_var varstruct) |
|
183 |
then mk_abs{Bvar = varstruct, Body = body} |
|
184 |
else let val {fst,snd} = dest_pair varstruct |
|
185 |
in mk_comb{Rator= mk_uncurry(type_of fst,type_of snd,type_of body), |
|
186 |
Rand = mpa(fst,mpa(snd,body))} |
|
187 |
end |
|
188 |
in mpa(varstruct,body) |
|
189 |
end |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
190 |
handle _ => raise USYN_ERR{func = "mk_pabs", mesg = ""}; |
2112 | 191 |
end; |
192 |
||
193 |
(* Destruction routines *) |
|
194 |
||
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
195 |
datatype lambda = VAR of {Name : string, Ty : typ} |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
196 |
| CONST of {Name : string, Ty : typ} |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
197 |
| COMB of {Rator: term, Rand : term} |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
198 |
| LAMB of {Bvar : term, Body : term}; |
2112 | 199 |
|
200 |
||
201 |
fun dest_term(Var((s,i),ty)) = VAR{Name = s, Ty = ty} |
|
202 |
| dest_term(Free(s,ty)) = VAR{Name = s, Ty = ty} |
|
203 |
| dest_term(Const(s,ty)) = CONST{Name = s, Ty = ty} |
|
204 |
| dest_term(M$N) = COMB{Rator=M,Rand=N} |
|
205 |
| dest_term(Abs(s,ty,M)) = let val v = mk_var{Name = s, Ty = ty} |
|
206 |
in LAMB{Bvar = v, Body = betapply (M,v)} |
|
207 |
end |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
208 |
| dest_term(Bound _) = raise USYN_ERR{func = "dest_term",mesg = "Bound"}; |
2112 | 209 |
|
210 |
fun dest_const(Const(s,ty)) = {Name = s, Ty = ty} |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
211 |
| dest_const _ = raise USYN_ERR{func = "dest_const", mesg = "not a constant"}; |
2112 | 212 |
|
213 |
fun dest_comb(t1 $ t2) = {Rator = t1, Rand = t2} |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
214 |
| dest_comb _ = raise USYN_ERR{func = "dest_comb", mesg = "not a comb"}; |
2112 | 215 |
|
216 |
fun dest_abs(a as Abs(s,ty,M)) = |
|
217 |
let val v = mk_var{Name = s, Ty = ty} |
|
218 |
in {Bvar = v, Body = betapply (a,v)} |
|
219 |
end |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
220 |
| dest_abs _ = raise USYN_ERR{func = "dest_abs", mesg = "not an abstraction"}; |
2112 | 221 |
|
222 |
fun dest_eq(Const("op =",_) $ M $ N) = {lhs=M, rhs=N} |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
223 |
| dest_eq _ = raise USYN_ERR{func = "dest_eq", mesg = "not an equality"}; |
2112 | 224 |
|
225 |
fun dest_imp(Const("op -->",_) $ M $ N) = {ant=M, conseq=N} |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
226 |
| dest_imp _ = raise USYN_ERR{func = "dest_imp", mesg = "not an implication"}; |
2112 | 227 |
|
228 |
fun dest_select(Const("Eps",_) $ (a as Abs _)) = dest_abs a |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
229 |
| dest_select _ = raise USYN_ERR{func = "dest_select", mesg = "not a select"}; |
2112 | 230 |
|
231 |
fun dest_forall(Const("All",_) $ (a as Abs _)) = dest_abs a |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
232 |
| dest_forall _ = raise USYN_ERR{func = "dest_forall", mesg = "not a forall"}; |
2112 | 233 |
|
234 |
fun dest_exists(Const("Ex",_) $ (a as Abs _)) = dest_abs a |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
235 |
| dest_exists _ = raise USYN_ERR{func = "dest_exists", mesg="not an existential"}; |
2112 | 236 |
|
237 |
fun dest_neg(Const("not",_) $ M) = M |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
238 |
| dest_neg _ = raise USYN_ERR{func = "dest_neg", mesg = "not a negation"}; |
2112 | 239 |
|
240 |
fun dest_conj(Const("op &",_) $ M $ N) = {conj1=M, conj2=N} |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
241 |
| dest_conj _ = raise USYN_ERR{func = "dest_conj", mesg = "not a conjunction"}; |
2112 | 242 |
|
243 |
fun dest_disj(Const("op |",_) $ M $ N) = {disj1=M, disj2=N} |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
244 |
| dest_disj _ = raise USYN_ERR{func = "dest_disj", mesg = "not a disjunction"}; |
2112 | 245 |
|
246 |
fun mk_pair{fst,snd} = |
|
247 |
let val ty1 = type_of fst |
|
248 |
val ty2 = type_of snd |
|
249 |
val c = mk_const{Name = "Pair", Ty = ty1 --> ty2 --> prod_ty ty1 ty2} |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
250 |
in list_comb(c,[fst,snd]) |
2112 | 251 |
end; |
252 |
||
253 |
fun dest_pair(Const("Pair",_) $ M $ N) = {fst=M, snd=N} |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
254 |
| dest_pair _ = raise USYN_ERR{func = "dest_pair", mesg = "not a pair"}; |
2112 | 255 |
|
256 |
||
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
257 |
local fun ucheck t = (if #Name(dest_const t) = "split" then t |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
258 |
else raise Match) |
2112 | 259 |
in |
260 |
fun dest_pabs tm = |
|
261 |
let val {Bvar,Body} = dest_abs tm |
|
262 |
in {varstruct = Bvar, body = Body} |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
263 |
end |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
264 |
handle |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
265 |
_ => let val {Rator,Rand} = dest_comb tm |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
266 |
val _ = ucheck Rator |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
267 |
val {varstruct = lv,body} = dest_pabs Rand |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
268 |
val {varstruct = rv,body} = dest_pabs body |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
269 |
in {varstruct = mk_pair{fst = lv, snd = rv}, body = body} |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
270 |
end |
2112 | 271 |
end; |
272 |
||
273 |
||
274 |
(* Garbage - ought to be dropped *) |
|
275 |
val lhs = #lhs o dest_eq |
|
276 |
val rhs = #rhs o dest_eq |
|
277 |
val rator = #Rator o dest_comb |
|
278 |
val rand = #Rand o dest_comb |
|
279 |
val bvar = #Bvar o dest_abs |
|
280 |
val body = #Body o dest_abs |
|
281 |
||
282 |
||
283 |
(* Query routines *) |
|
284 |
val is_comb = can dest_comb |
|
285 |
val is_abs = can dest_abs |
|
286 |
val is_eq = can dest_eq |
|
287 |
val is_imp = can dest_imp |
|
288 |
val is_forall = can dest_forall |
|
289 |
val is_exists = can dest_exists |
|
290 |
val is_neg = can dest_neg |
|
291 |
val is_conj = can dest_conj |
|
292 |
val is_disj = can dest_disj |
|
293 |
val is_pair = can dest_pair |
|
294 |
val is_pabs = can dest_pabs |
|
295 |
||
296 |
||
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
297 |
(* Construction of a cterm from a list of Terms *) |
2112 | 298 |
|
299 |
fun list_mk_abs(L,tm) = itlist (fn v => fn M => mk_abs{Bvar=v, Body=M}) L tm; |
|
300 |
||
301 |
(* These others are almost never used *) |
|
302 |
fun list_mk_imp(A,c) = itlist(fn a => fn tm => mk_imp{ant=a,conseq=tm}) A c; |
|
303 |
fun list_mk_exists(V,t) = itlist(fn v => fn b => mk_exists{Bvar=v, Body=b})V t; |
|
304 |
fun list_mk_forall(V,t) = itlist(fn v => fn b => mk_forall{Bvar=v, Body=b})V t; |
|
305 |
val list_mk_conj = end_itlist(fn c1 => fn tm => mk_conj{conj1=c1, conj2=tm}) |
|
306 |
val list_mk_disj = end_itlist(fn d1 => fn tm => mk_disj{disj1=d1, disj2=tm}) |
|
307 |
||
308 |
||
309 |
(* Need to reverse? *) |
|
310 |
fun gen_all tm = list_mk_forall(free_vars tm, tm); |
|
311 |
||
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
312 |
(* Destructing a cterm to a list of Terms *) |
2112 | 313 |
fun strip_comb tm = |
314 |
let fun dest(M$N, A) = dest(M, N::A) |
|
315 |
| dest x = x |
|
316 |
in dest(tm,[]) |
|
317 |
end; |
|
318 |
||
319 |
fun strip_abs(tm as Abs _) = |
|
320 |
let val {Bvar,Body} = dest_abs tm |
|
321 |
val (bvs, core) = strip_abs Body |
|
322 |
in (Bvar::bvs, core) |
|
323 |
end |
|
324 |
| strip_abs M = ([],M); |
|
325 |
||
326 |
||
327 |
fun strip_imp fm = |
|
328 |
if (is_imp fm) |
|
329 |
then let val {ant,conseq} = dest_imp fm |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
330 |
val (was,wb) = strip_imp conseq |
2112 | 331 |
in ((ant::was), wb) |
332 |
end |
|
333 |
else ([],fm); |
|
334 |
||
335 |
fun strip_forall fm = |
|
336 |
if (is_forall fm) |
|
337 |
then let val {Bvar,Body} = dest_forall fm |
|
338 |
val (bvs,core) = strip_forall Body |
|
339 |
in ((Bvar::bvs), core) |
|
340 |
end |
|
341 |
else ([],fm); |
|
342 |
||
343 |
||
344 |
fun strip_exists fm = |
|
345 |
if (is_exists fm) |
|
346 |
then let val {Bvar, Body} = dest_exists fm |
|
347 |
val (bvs,core) = strip_exists Body |
|
348 |
in (Bvar::bvs, core) |
|
349 |
end |
|
350 |
else ([],fm); |
|
351 |
||
352 |
fun strip_conj w = |
|
353 |
if (is_conj w) |
|
354 |
then let val {conj1,conj2} = dest_conj w |
|
355 |
in (strip_conj conj1@strip_conj conj2) |
|
356 |
end |
|
357 |
else [w]; |
|
358 |
||
359 |
fun strip_disj w = |
|
360 |
if (is_disj w) |
|
361 |
then let val {disj1,disj2} = dest_disj w |
|
362 |
in (strip_disj disj1@strip_disj disj2) |
|
363 |
end |
|
364 |
else [w]; |
|
365 |
||
366 |
fun strip_pair tm = |
|
367 |
if (is_pair tm) |
|
368 |
then let val {fst,snd} = dest_pair tm |
|
369 |
fun dtuple t = |
|
370 |
if (is_pair t) |
|
371 |
then let val{fst,snd} = dest_pair t |
|
372 |
in (fst :: dtuple snd) |
|
373 |
end |
|
374 |
else [t] |
|
375 |
in fst::dtuple snd |
|
376 |
end |
|
377 |
else [tm]; |
|
378 |
||
379 |
||
380 |
fun mk_preterm tm = #t(rep_cterm tm); |
|
381 |
||
382 |
(* Miscellaneous *) |
|
383 |
||
384 |
fun mk_vstruct ty V = |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
385 |
let fun follow_prod_type (Type("*",[ty1,ty2])) vs = |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
386 |
let val (ltm,vs1) = follow_prod_type ty1 vs |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
387 |
val (rtm,vs2) = follow_prod_type ty2 vs1 |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
388 |
in (mk_pair{fst=ltm, snd=rtm}, vs2) end |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
389 |
| follow_prod_type _ (v::vs) = (v,vs) |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
390 |
in #1 (follow_prod_type ty V) end; |
2112 | 391 |
|
392 |
||
393 |
(* Search a term for a sub-term satisfying the predicate p. *) |
|
394 |
fun find_term p = |
|
395 |
let fun find tm = |
|
396 |
if (p tm) |
|
397 |
then tm |
|
398 |
else if (is_abs tm) |
|
399 |
then find (#Body(dest_abs tm)) |
|
400 |
else let val {Rator,Rand} = dest_comb tm |
|
401 |
in find Rator handle _ => find Rand |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
402 |
end handle _ => raise USYN_ERR{func = "find_term",mesg = ""} |
2112 | 403 |
in find |
404 |
end; |
|
405 |
||
406 |
(******************************************************************* |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
407 |
* find_terms: (term -> HOLogic.boolT) -> term -> term list |
2112 | 408 |
* |
409 |
* Find all subterms in a term that satisfy a given predicate p. |
|
410 |
* |
|
411 |
*******************************************************************) |
|
412 |
fun find_terms p = |
|
413 |
let fun accum tl tm = |
|
414 |
let val tl' = if (p tm) then (tm::tl) else tl |
|
415 |
in if (is_abs tm) |
|
416 |
then accum tl' (#Body(dest_abs tm)) |
|
417 |
else let val {Rator,Rand} = dest_comb tm |
|
418 |
in accum (accum tl' Rator) Rand |
|
419 |
end handle _ => tl' |
|
420 |
end |
|
421 |
in accum [] |
|
422 |
end; |
|
423 |
||
424 |
||
425 |
fun dest_relation tm = |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
426 |
if (type_of tm = HOLogic.boolT) |
2112 | 427 |
then let val (Const("op :",_) $ (Const("Pair",_)$y$x) $ R) = tm |
428 |
in (R,y,x) |
|
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
429 |
end handle _ => raise USYN_ERR{func="dest_relation", |
2112 | 430 |
mesg="unexpected term structure"} |
3353
9112a2efb9a3
Removal of module Mask and datatype binding with its constructor |->
paulson
parents:
3330
diff
changeset
|
431 |
else raise USYN_ERR{func="dest_relation",mesg="not a boolean term"}; |
2112 | 432 |
|
433 |
fun is_WFR tm = (#Name(dest_const(rator tm)) = "wf") handle _ => false; |
|
434 |
||
435 |
fun ARB ty = mk_select{Bvar=mk_var{Name="v",Ty=ty}, |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
2112
diff
changeset
|
436 |
Body=mk_const{Name="True",Ty=HOLogic.boolT}}; |
2112 | 437 |
|
438 |
end; (* Syntax *) |