author | dixon |
Thu, 05 May 2005 11:58:59 +0200 | |
changeset 15928 | 66b165ee016c |
parent 15915 | b0e8b37642a4 |
child 16179 | fa7e70be26b0 |
permissions | -rw-r--r-- |
15481 | 1 |
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) |
2 |
(* Title: libs/term_lib.ML |
|
3 |
Author: Lucas Dixon, University of Edinburgh |
|
4 |
lucasd@dai.ed.ac.uk |
|
5 |
Created: 17 Aug 2002 |
|
6 |
*) |
|
7 |
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) |
|
8 |
(* DESCRIPTION: |
|
9 |
||
10 |
Additional code to work with terms. |
|
11 |
||
12 |
*) |
|
13 |
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) |
|
14 |
signature TERM_LIB = |
|
15 |
sig |
|
16 |
val fo_term_height : Term.term -> int |
|
17 |
val ho_term_height : Term.term -> int |
|
18 |
||
19 |
val current_sign : unit -> Sign.sg |
|
20 |
||
21 |
structure fvartabS : TABLE |
|
22 |
||
23 |
(* Matching/unification with exceptions handled *) |
|
24 |
val clean_match : Type.tsig -> Term.term * Term.term |
|
15915
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15814
diff
changeset
|
25 |
-> ((Term.indexname * (Term.sort * Term.typ)) list |
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15814
diff
changeset
|
26 |
* (Term.indexname * (Term.typ * Term.term)) list) option |
15481 | 27 |
val clean_unify : Sign.sg -> int -> Term.term * Term.term |
15915
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15814
diff
changeset
|
28 |
-> ((Term.indexname * (Term.sort * Term.typ)) list |
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15814
diff
changeset
|
29 |
* (Term.indexname * (Term.typ * Term.term)) list) Seq.seq |
15481 | 30 |
|
31 |
(* managing variables in terms, can doing conversions *) |
|
32 |
val bounds_to_frees : Term.term -> Term.term |
|
33 |
val bounds_to_frees_with_vars : |
|
34 |
(string * Term.typ) list -> Term.term -> Term.term |
|
35 |
||
36 |
val change_bounds_to_frees : Term.term -> Term.term |
|
37 |
val change_frees_to_vars : Term.term -> Term.term |
|
38 |
val change_vars_to_frees : Term.term -> Term.term |
|
39 |
val loose_bnds_to_frees : |
|
40 |
(string * Term.typ) list -> Term.term -> Term.term |
|
41 |
||
42 |
(* make all variables named uniquely *) |
|
43 |
val unique_namify : Term.term -> Term.term |
|
44 |
||
45 |
(* breasking up a term and putting it into a normal form |
|
46 |
independent of internal term context *) |
|
47 |
val cleaned_term_conc : Term.term -> Term.term |
|
48 |
val cleaned_term_parts : Term.term -> Term.term list * Term.term |
|
49 |
val cterm_of_term : Term.term -> Thm.cterm |
|
50 |
||
51 |
(* terms of theorems and subgoals *) |
|
52 |
val term_of_thm : Thm.thm -> Term.term |
|
53 |
val get_prems_of_sg_term : Term.term -> Term.term list |
|
54 |
val triv_thm_from_string : string -> Thm.thm |
|
55 |
(* val make_term_into_simp_thm : |
|
56 |
(string * Term.typ) list -> Sign.sg -> Term.term -> Thm.thm |
|
57 |
val thms_of_prems_in_goal : int -> Thm.thm -> Thm.thm list |
|
58 |
*) |
|
59 |
||
60 |
||
61 |
(* some common term manipulations *) |
|
62 |
val try_dest_Goal : Term.term -> Term.term |
|
63 |
val try_dest_Trueprop : Term.term -> Term.term |
|
64 |
||
65 |
val bot_left_leaf_of : Term.term -> Term.term |
|
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
66 |
val bot_left_leaf_noabs_of : Term.term -> Term.term |
15481 | 67 |
|
68 |
(* term containing another term - an embedding check *) |
|
69 |
val term_contains : Term.term -> Term.term -> bool |
|
70 |
||
71 |
(* name-convertable checking - like ae-convertable, but allows for |
|
72 |
var's and free's to be mixed - and doesn't used buggy code. :-) *) |
|
73 |
val get_name_eq_member : Term.term -> Term.term list -> Term.term option |
|
74 |
val name_eq_member : Term.term -> Term.term list -> bool |
|
75 |
val term_name_eq : |
|
76 |
Term.term -> |
|
77 |
Term.term -> |
|
78 |
(Term.term * Term.term) list -> |
|
15531 | 79 |
(Term.term * Term.term) list option |
15481 | 80 |
|
81 |
(* is the typ a function or is it atomic *) |
|
82 |
val is_fun_typ : Term.typ -> bool |
|
83 |
val is_atomic_typ : Term.typ -> bool |
|
84 |
||
85 |
(* variable analysis *) |
|
86 |
val is_some_kind_of_var : Term.term -> bool |
|
87 |
val same_var_check : |
|
15531 | 88 |
''a -> ''b -> (''a * ''b) list -> (''a * ''b) list option |
15481 | 89 |
val has_new_vars : Term.term * Term.term -> bool |
90 |
val has_new_typ_vars : Term.term * Term.term -> bool |
|
91 |
(* checks to see if the lhs -> rhs is a invalid rewrite rule *) |
|
92 |
val is_not_valid_rwrule : Type.tsig -> (Term.term * Term.term) -> bool |
|
93 |
||
94 |
(* get the frees in a term that are of atomic type, ie non-functions *) |
|
95 |
val atomic_frees_of_term : Term.term -> (string * Term.typ) list |
|
96 |
||
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
97 |
(* get used names in a theorem to avoid upon instantiation. *) |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
98 |
val usednames_of_term : Term.term -> string list |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
99 |
val usednames_of_thm : Thm.thm -> string list |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
100 |
|
15481 | 101 |
(* Isar term skolemisationm and unsolemisation *) |
102 |
(* I think this is written also in IsarRTechn and also in somewhere else *) |
|
103 |
(* val skolemise_term : (string,Term.typ) list -> Term.term -> Term.term |
|
104 |
val unskolemise_term : (string,Term.typ) list -> Term.term -> Term.term *) |
|
105 |
val unskolemise_all_term : |
|
106 |
Term.term -> |
|
107 |
(((string * Term.typ) * string) list * Term.term) |
|
108 |
||
109 |
(* given a string describing term then a string for its type, returns |
|
110 |
read term *) |
|
111 |
val readwty : string -> string -> Term.term |
|
112 |
||
113 |
(* pretty stuff *) |
|
114 |
val pp_term : Term.term -> unit |
|
115 |
val pretty_print_sort : string list -> string |
|
116 |
val pretty_print_term : Term.term -> string |
|
117 |
val pretty_print_type : Term.typ -> string |
|
118 |
val pretty_print_typelist : |
|
119 |
Term.typ list -> (Term.typ -> string) -> string |
|
120 |
||
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
121 |
(* for debugging: quickly get a string of a term w.r.t the_context() *) |
15481 | 122 |
val string_of_term : Term.term -> string |
123 |
||
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
124 |
(* Pretty printing in a way that allows them to be parsed by ML. |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
125 |
these are perhaps redundent, check the standard basis |
15481 | 126 |
lib for generic versions for any datatype? *) |
127 |
val writesort : string list -> unit |
|
128 |
val writeterm : Term.term -> unit |
|
129 |
val writetype : Term.typ -> unit |
|
130 |
end |
|
131 |
||
132 |
||
133 |
structure TermLib : TERM_LIB = |
|
134 |
struct |
|
135 |
||
136 |
(* Two kinds of depth measure for HOAS terms, a first order (flat) and a |
|
137 |
higher order one. |
|
138 |
Note: not stable of eta-contraction: embedding eta-expands term, |
|
139 |
thus we assume eta-expanded *) |
|
140 |
fun fo_term_height (a $ b) = |
|
141 |
IsaPLib.max (1 + fo_term_height b, (fo_term_height a)) |
|
142 |
| fo_term_height (Abs(_,_,t)) = fo_term_height t |
|
143 |
| fo_term_height _ = 0; |
|
144 |
||
145 |
fun ho_term_height (a $ b) = |
|
146 |
1 + (IsaPLib.max (ho_term_height b, ho_term_height a)) |
|
147 |
| ho_term_height (Abs(_,_,t)) = ho_term_height t |
|
148 |
| ho_term_height _ = 0; |
|
149 |
||
15928
66b165ee016c
lucas - made clean unify smash unifiers so that when we get flex-flex constraints subst does not barf. Also added fix_vars_upto_idx to IsaND.
dixon
parents:
15915
diff
changeset
|
150 |
|
15481 | 151 |
(* Higher order matching with exception handled *) |
152 |
(* returns optional instantiation *) |
|
153 |
fun clean_match tsig (a as (pat, t)) = |
|
15798
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents:
15570
diff
changeset
|
154 |
let val (tyenv, tenv) = Pattern.match tsig a |
15915
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15814
diff
changeset
|
155 |
in SOME (Vartab.dest tyenv, Vartab.dest tenv) |
15798
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents:
15570
diff
changeset
|
156 |
end handle Pattern.MATCH => NONE; |
15481 | 157 |
(* Higher order unification with exception handled, return the instantiations *) |
158 |
(* given Signature, max var index, pat, tgt *) |
|
159 |
(* returns Seq of instantiations *) |
|
160 |
fun clean_unify sgn ix (a as (pat, tgt)) = |
|
161 |
let |
|
162 |
(* type info will be re-derived, maybe this can be cached |
|
163 |
for efficiency? *) |
|
164 |
val pat_ty = Term.type_of pat; |
|
165 |
val tgt_ty = Term.type_of tgt; |
|
166 |
(* is it OK to ignore the type instantiation info? |
|
167 |
or should I be using it? *) |
|
168 |
val typs_unify = |
|
15531 | 169 |
(SOME (Type.unify (Sign.tsig_of sgn) (Term.Vartab.empty, ix) |
15481 | 170 |
(pat_ty, tgt_ty))) |
15531 | 171 |
handle Type.TUNIFY => NONE; |
15481 | 172 |
in |
173 |
case typs_unify of |
|
15531 | 174 |
SOME (typinsttab, ix2) => |
15481 | 175 |
let |
15570 | 176 |
(* is it right to throw away the flexes? |
15481 | 177 |
or should I be using them somehow? *) |
15928
66b165ee016c
lucas - made clean unify smash unifiers so that when we get flex-flex constraints subst does not barf. Also added fix_vars_upto_idx to IsaND.
dixon
parents:
15915
diff
changeset
|
178 |
fun mk_insts env = |
15915
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15814
diff
changeset
|
179 |
(Vartab.dest (Envir.type_env env), |
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15814
diff
changeset
|
180 |
Envir.alist_of env); |
15481 | 181 |
val initenv = Envir.Envir {asol = Vartab.empty, |
182 |
iTs = typinsttab, maxidx = ix2}; |
|
15928
66b165ee016c
lucas - made clean unify smash unifiers so that when we get flex-flex constraints subst does not barf. Also added fix_vars_upto_idx to IsaND.
dixon
parents:
15915
diff
changeset
|
183 |
val useq = (Unify.smash_unifiers (sgn,initenv,[a])) |
66b165ee016c
lucas - made clean unify smash unifiers so that when we get flex-flex constraints subst does not barf. Also added fix_vars_upto_idx to IsaND.
dixon
parents:
15915
diff
changeset
|
184 |
handle UnequalLengths => Seq.empty |
66b165ee016c
lucas - made clean unify smash unifiers so that when we get flex-flex constraints subst does not barf. Also added fix_vars_upto_idx to IsaND.
dixon
parents:
15915
diff
changeset
|
185 |
| Term.TERM _ => Seq.empty; |
15481 | 186 |
fun clean_unify' useq () = |
187 |
(case (Seq.pull useq) of |
|
15531 | 188 |
NONE => NONE |
189 |
| SOME (h,t) => SOME (mk_insts h, Seq.make (clean_unify' t))) |
|
15570 | 190 |
handle UnequalLengths => NONE |
15531 | 191 |
| Term.TERM _ => NONE; |
15481 | 192 |
in |
193 |
(Seq.make (clean_unify' useq)) |
|
194 |
end |
|
15531 | 195 |
| NONE => Seq.empty |
15481 | 196 |
end; |
197 |
||
198 |
fun asm_mk t = (assume (cterm_of (Theory.sign_of (the_context())) t)); |
|
199 |
fun asm_read s = |
|
200 |
(Thm.assume (read_cterm (Theory.sign_of (Context.the_context())) (s,propT))); |
|
201 |
||
202 |
||
203 |
(* more pretty printing code for Isabelle terms etc *) |
|
204 |
||
205 |
||
206 |
(* pretty_print_typelist l f = print a typelist. |
|
207 |
l = list of types to print : typ list |
|
208 |
f = function used to print a single type : typ -> string |
|
209 |
*) |
|
210 |
fun pretty_print_typelist [] f = "" |
|
211 |
| pretty_print_typelist [(h: typ)] (f : typ -> string) = (f h) |
|
212 |
| pretty_print_typelist ((h: typ) :: t) (f : typ -> string) = |
|
213 |
(f h) ^ ", " ^ (pretty_print_typelist t f); |
|
214 |
||
215 |
||
216 |
||
217 |
(* pretty_print_sort s = print a sort |
|
218 |
s = sort to print : string list |
|
219 |
*) |
|
220 |
fun pretty_print_sort [] = "" |
|
221 |
| pretty_print_sort ([h]) = "\"" ^ h ^ "\"" |
|
222 |
| pretty_print_sort (h :: t) = "\"" ^ h ^ "\"," ^ (pretty_print_sort t); |
|
223 |
||
224 |
||
225 |
(* pretty_print_type t = print a type |
|
226 |
t = type to print : type |
|
227 |
*) |
|
228 |
fun pretty_print_type (Type (n, l)) = |
|
229 |
"Type(\"" ^ n ^ "\", [" ^ (pretty_print_typelist l pretty_print_type) ^ "])" |
|
230 |
| pretty_print_type (TFree (n, s)) = |
|
231 |
"TFree(\"" ^ n ^ "\", [" ^ (pretty_print_sort s) ^ "])" |
|
232 |
| pretty_print_type (TVar ((n, i), s)) = |
|
233 |
"TVar( (\"" ^ n ^ "\", " ^ (string_of_int i) ^ "), [" ^ (pretty_print_sort s) ^ "])"; |
|
234 |
||
235 |
||
236 |
(* pretty_print_term t = print a term prints types and sorts too. |
|
237 |
t = term to print : term |
|
238 |
*) |
|
239 |
fun pretty_print_term (Const (s, t)) = |
|
240 |
"Const(\"" ^ s ^ "\", " ^ (pretty_print_type t) ^ ")" |
|
241 |
| pretty_print_term (Free (s, t)) = |
|
242 |
"Free(\"" ^ s ^ "\", " ^ (pretty_print_type t) ^ ")" |
|
243 |
| pretty_print_term (Var ((n, i), t)) = |
|
244 |
"Var( (\"" ^ n ^ "\"," ^ (string_of_int i) ^ "), " ^ (pretty_print_type t) ^ ")" |
|
245 |
| pretty_print_term (Bound i) = |
|
246 |
"Bound(" ^ (string_of_int i) ^ ")" |
|
247 |
| pretty_print_term (Abs (s, t, r)) = |
|
248 |
"Abs(\"" ^ s ^ "\"," ^ (pretty_print_type t) ^ ", \n " ^ (pretty_print_term r) ^ ")" |
|
249 |
| pretty_print_term (op $ (t1, t2)) = |
|
250 |
"(" ^ (pretty_print_term t1) ^ ") $\n (" ^ (pretty_print_term t2) ^ ")"; |
|
251 |
||
252 |
(* Write the term out nicly instead of just creating a string for it *) |
|
253 |
fun writeterm t = writeln (pretty_print_term t); |
|
254 |
fun writetype t = writeln (pretty_print_type t); |
|
255 |
fun writesort s = writeln (pretty_print_sort s); |
|
256 |
||
257 |
||
258 |
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) |
|
259 |
(* Turn on show types *) |
|
260 |
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) |
|
261 |
||
262 |
(* if (!show_types) then true else set show_types; *) |
|
263 |
||
264 |
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) |
|
265 |
(* functions *) |
|
266 |
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) |
|
267 |
||
15915
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15814
diff
changeset
|
268 |
fun current_sign () = Theory.sign_of (the_context()); |
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15814
diff
changeset
|
269 |
fun cterm_of_term (t : term) = Thm.cterm_of (current_sign()) t; |
15481 | 270 |
fun term_of_thm t = (Thm.prop_of t); |
271 |
||
272 |
(* little function to make a trueprop of anything by putting a P |
|
273 |
function in front of it... |
|
274 |
fun HOL_mk_term_prop t = |
|
275 |
((Const("Trueprop", Type("fun", |
|
276 |
[Type("bool", []), Type("prop", [])]))) $ |
|
277 |
((Free("P", Type("fun", [type_of t, |
|
278 |
Type("bool", [])]))) $ t)); |
|
279 |
||
280 |
*) |
|
281 |
||
282 |
fun string_of_term t = |
|
283 |
(Sign.string_of_term (current_sign()) t); |
|
284 |
||
285 |
(* |
|
286 |
(allt as (Const("Trueprop", ty) $ m)) = |
|
287 |
(string_of_cterm (cterm_of_term allt)) |
|
288 |
| string_of_term (t : term) = |
|
289 |
string_of_cterm (cterm_of_term (HOL_mk_term_prop t)); |
|
290 |
*) |
|
291 |
||
292 |
(* creates a simple cterm of the term if it's not already a prop by |
|
293 |
prefixing it with a polymorphic function P, then create the cterm |
|
294 |
and print that. Handy tool for printing terms that are not |
|
295 |
already propositions, or not cterms. |
|
296 |
*) |
|
297 |
fun pp_term t = writeln (string_of_term t); |
|
298 |
||
299 |
(* create a trivial HOL thm from anything... *) |
|
300 |
fun triv_thm_from_string s = |
|
301 |
Thm.trivial (cterm_of (current_sign()) (read s)); |
|
302 |
||
303 |
(* Checks if vars could be the same - alpha convertable |
|
304 |
w.r.t. previous vars, a and b are assumed to be vars, |
|
305 |
free vars, but not bound vars, |
|
306 |
Note frees and vars must all have unique names. *) |
|
307 |
fun same_var_check a b L = |
|
308 |
let |
|
15531 | 309 |
fun bterm_from t [] = NONE |
15481 | 310 |
| bterm_from t ((a,b)::m) = |
15531 | 311 |
if t = a then SOME b else bterm_from t m |
15481 | 312 |
|
313 |
val bl_opt = bterm_from a L |
|
314 |
in |
|
315 |
case bterm_from a L of |
|
15531 | 316 |
NONE => SOME ((a,b)::L) |
317 |
| SOME b2 => if b2 = b then SOME L else NONE |
|
15481 | 318 |
end; |
319 |
||
320 |
(* FIXME: make more efficient, only require a single new var! *) |
|
321 |
(* check if the new term has any meta variables not in the old term *) |
|
322 |
fun has_new_vars (old, new) = |
|
323 |
(case IsaPLib.lrem (op =) (Term.term_vars old) (Term.term_vars new) of |
|
324 |
[] => false |
|
325 |
| (_::_) => true); |
|
326 |
(* check if the new term has any meta variables not in the old term *) |
|
327 |
fun has_new_typ_vars (old, new) = |
|
328 |
(case IsaPLib.lrem (op =) (Term.term_tvars old) (Term.term_tvars new) of |
|
329 |
[] => false |
|
330 |
| (_::_) => true); |
|
331 |
||
332 |
(* working with Isar terms and their skolemisation(s) *) |
|
333 |
(* val skolemise_term : (string,Term.typ) list -> Term.term -> Term.term |
|
334 |
val unskolemise_term : (string,Term.typ) list -> Term.term -> Term.term |
|
335 |
*) |
|
336 |
||
337 |
(* cunning version *) |
|
338 |
||
15570 | 339 |
(* This version avoids name conflicts that might be introduced by |
15481 | 340 |
unskolemisation, and returns a list of (string * Term.typ) * string, |
341 |
where the outer string is the original name, and the inner string is |
|
342 |
the new name, and the type is the type of the free variable that was |
|
343 |
renamed. *) |
|
344 |
local |
|
345 |
fun myadd (n,ty) (L as (renames, names)) = |
|
346 |
let val n' = Syntax.dest_skolem n in |
|
347 |
case (Library.find_first (fn (_,n2) => (n2 = n)) renames) of |
|
15531 | 348 |
NONE => |
15481 | 349 |
let val renamedn = Term.variant names n' in |
350 |
(renamedn, (((renamedn, ty), n) :: renames, |
|
351 |
renamedn :: names)) end |
|
15531 | 352 |
| (SOME ((renamedn, _), _)) => (renamedn, L) |
15481 | 353 |
end |
15570 | 354 |
handle Fail _ => (n, L); |
15481 | 355 |
|
356 |
fun unsk (L,f) (t1 $ t2) = |
|
357 |
let val (L', t1') = unsk (L, I) t1 |
|
358 |
in unsk (L', (fn x => f (t1' $ x))) t2 end |
|
359 |
| unsk (L,f) (Abs(n,ty,t)) = |
|
360 |
unsk (L, (fn x => Abs(n,ty,x))) t |
|
361 |
| unsk (L, f) (Free (n,ty)) = |
|
362 |
let val (renamed_n, L') = myadd (n ,ty) L |
|
363 |
in (L', f (Free (renamed_n, ty))) end |
|
364 |
| unsk (L, f) l = (L, f l); |
|
365 |
in |
|
366 |
fun unskolemise_all_term t = |
|
367 |
let |
|
368 |
val names = Term.add_term_names (t,[]) |
|
369 |
val ((renames,names'),t') = unsk (([], names),I) t |
|
370 |
in (renames,t') end; |
|
371 |
end |
|
372 |
||
373 |
(* fun unskolemise_all_term t = |
|
374 |
let |
|
375 |
fun fmap fv = |
|
376 |
let (n,ty) = (Term.dest_Free fv) in |
|
15531 | 377 |
SOME (fv, Free (Syntax.dest_skolem n, ty)) handle LIST _ => NONE |
15481 | 378 |
end |
15570 | 379 |
val substfrees = List.mapPartial fmap (Term.term_frees t) |
15481 | 380 |
in |
381 |
Term.subst_free substfrees t |
|
382 |
end; *) |
|
383 |
||
384 |
(* true if the type t is a function *) |
|
385 |
fun is_fun_typ (Type(s, l)) = |
|
386 |
if s = "fun" then true else false |
|
387 |
| is_fun_typ _ = false; |
|
388 |
||
389 |
val is_atomic_typ = not o is_fun_typ; |
|
390 |
||
391 |
(* get the frees in a term that are of atomic type, ie non-functions *) |
|
392 |
val atomic_frees_of_term = |
|
15570 | 393 |
List.filter (is_atomic_typ o snd) |
15481 | 394 |
o map Term.dest_Free |
395 |
o Term.term_frees; |
|
396 |
||
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
397 |
fun usednames_of_term t = Term.add_term_names (t,[]); |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
398 |
fun usednames_of_thm th = |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
399 |
let val rep = Thm.rep_thm th |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
400 |
val hyps = #hyps rep |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
401 |
val (tpairl,tpairr) = Library.split_list (#tpairs rep) |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
402 |
val prop = #prop rep |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
403 |
in |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
404 |
List.foldr Term.add_term_names [] (prop :: (tpairl @ (tpairr @ hyps))) |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
405 |
end; |
15481 | 406 |
|
407 |
(* read in a string and a top-level type and this will give back a term *) |
|
408 |
fun readwty tstr tystr = |
|
409 |
let |
|
410 |
val sgn = Theory.sign_of (the_context()) |
|
411 |
in |
|
15531 | 412 |
Sign.simple_read_term sgn (Sign.read_typ (sgn, K NONE) tystr) tstr |
15481 | 413 |
end; |
414 |
||
415 |
||
416 |
(* first term is equal to the second in some name convertable |
|
417 |
way... Note: This differs from the aeconv in the Term.ML file of |
|
418 |
Isabelle in that it allows a var to be alpha-equiv to a free var. |
|
419 |
||
420 |
Also, the isabelle term.ML version of aeconv uses a |
|
421 |
function that it claims doesn't work! *) |
|
422 |
||
423 |
fun term_name_eq (Abs(_,ty1,t1)) (Abs(_,ty2,t2)) l = |
|
424 |
if ty1 = ty2 then term_name_eq t1 t2 l |
|
15531 | 425 |
else NONE |
15481 | 426 |
|
427 |
| term_name_eq (ah $ at) (bh $ bt) l = |
|
428 |
let |
|
429 |
val lopt = (term_name_eq ah bh l) |
|
430 |
in |
|
15570 | 431 |
if isSome lopt then |
432 |
term_name_eq at bt (valOf lopt) |
|
15481 | 433 |
else |
15531 | 434 |
NONE |
15481 | 435 |
end |
436 |
||
437 |
| term_name_eq (Const(a,T)) (Const(b,U)) l = |
|
15531 | 438 |
if a=b andalso T=U then SOME l |
439 |
else NONE |
|
15481 | 440 |
|
441 |
| term_name_eq (a as Free(s1,ty1)) (b as Free(s2,ty2)) l = |
|
442 |
same_var_check a b l |
|
443 |
||
444 |
| term_name_eq (a as Free(s1,ty1)) (b as Var(n2,ty2)) l = |
|
445 |
same_var_check a b l |
|
446 |
||
447 |
| term_name_eq (a as Var(n1,ty1)) (b as Free(s2,ty2)) l = |
|
448 |
same_var_check a b l |
|
449 |
||
450 |
| term_name_eq (a as Var(n1,ty1)) (b as Var(n2,ty2)) l = |
|
451 |
same_var_check a b l |
|
452 |
||
453 |
| term_name_eq (Bound i) (Bound j) l = |
|
15531 | 454 |
if i = j then SOME l else NONE |
15481 | 455 |
|
15531 | 456 |
| term_name_eq a b l = ((*writeln ("unchecked case:\n" ^ "a:\n" ^ (pretty_print_term a) ^ "\nb:\n" ^ (pretty_print_term b));*) NONE); |
15481 | 457 |
|
458 |
(* checks to see if the term in a name-equivalent member of the list of terms. *) |
|
15531 | 459 |
fun get_name_eq_member a [] = NONE |
15481 | 460 |
| get_name_eq_member a (h :: t) = |
15570 | 461 |
if isSome (term_name_eq a h []) then SOME h |
15481 | 462 |
else get_name_eq_member a t; |
463 |
||
464 |
fun name_eq_member a [] = false |
|
465 |
| name_eq_member a (h :: t) = |
|
15570 | 466 |
if isSome (term_name_eq a h []) then true |
15481 | 467 |
else name_eq_member a t; |
468 |
||
469 |
(* true if term is a variable *) |
|
470 |
fun is_some_kind_of_var (Free(s, ty)) = true |
|
471 |
| is_some_kind_of_var (Var(i, ty)) = true |
|
472 |
| is_some_kind_of_var (Bound(i)) = true |
|
473 |
| is_some_kind_of_var _ = false; |
|
474 |
||
475 |
(* checks to see if the lhs -> rhs is a invalid rewrite rule *) |
|
476 |
(* FIXME: we should really check that there is a subterm on the lhs |
|
477 |
which embeds into the rhs, this would be much closer to the normal |
|
478 |
notion of valid wave rule - ie there exists at least one case where it |
|
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
479 |
is a valid wave rule... *) |
15481 | 480 |
fun is_not_valid_rwrule tysig (lhs, rhs) = |
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
481 |
Term.is_Var (Term.head_of lhs) (* if lhs is essentially just a var *) |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
482 |
orelse has_new_vars (lhs,rhs) |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
483 |
orelse has_new_typ_vars (lhs,rhs) |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
484 |
orelse Pattern.matches_subterm tysig (lhs, rhs); |
15481 | 485 |
|
486 |
||
487 |
(* first term contains the second in some name convertable way... *) |
|
488 |
(* note: this is equivalent to an alpha-convertable emedding *) |
|
489 |
(* takes as args: term containing a, term contained b, |
|
490 |
(bound vars of a, bound vars of b), |
|
491 |
current alpha-conversion-pairs, |
|
492 |
returns: bool:can convert, new alpha-conversion table *) |
|
493 |
(* in bellow: we *don't* use: a loose notion that only requires |
|
494 |
variables to match variables, and doesn't worry about the actual |
|
495 |
pattern in the variables. *) |
|
496 |
fun term_contains1 (Bs, FVs) (Abs(s,ty,t)) (Abs(s2,ty2,t2)) = |
|
497 |
if ty = ty2 then |
|
15531 | 498 |
term_contains1 ((SOME(s,s2,ty)::Bs), FVs) t t2 |
15481 | 499 |
else [] |
500 |
||
501 |
| term_contains1 T t1 (Abs(s2,ty2,t2)) = [] |
|
502 |
||
503 |
| term_contains1 (Bs, FVs) (Abs(s,ty,t)) t2 = |
|
15531 | 504 |
term_contains1 (NONE::Bs, FVs) t t2 |
15481 | 505 |
|
506 |
| term_contains1 T (ah $ at) (bh $ bt) = |
|
507 |
(term_contains1 T ah (bh $ bt)) @ |
|
508 |
(term_contains1 T at (bh $ bt)) @ |
|
15570 | 509 |
(List.concat (map (fn inT => (term_contains1 inT at bt)) |
15481 | 510 |
(term_contains1 T ah bh))) |
511 |
||
512 |
| term_contains1 T a (bh $ bt) = [] |
|
513 |
||
514 |
| term_contains1 T (ah $ at) b = |
|
515 |
(term_contains1 T ah b) @ (term_contains1 T at b) |
|
516 |
||
517 |
| term_contains1 T a b = |
|
518 |
(* simple list table lookup to check if a named variable has been |
|
519 |
mapped to a variable, if not adds the mapping and return some new |
|
520 |
list mapping, if it is then it checks that the pair are mapped to |
|
521 |
each other, if so returns the current mapping list, else none. *) |
|
522 |
let |
|
15531 | 523 |
fun bterm_from t [] = NONE |
15481 | 524 |
| bterm_from t ((a,b)::m) = |
15531 | 525 |
if t = a then SOME b else bterm_from t m |
15481 | 526 |
|
527 |
(* check to see if, w.r.t. the variable mapping, two terms are leaf |
|
528 |
terms and are mapped to each other. Note constants are only mapped |
|
529 |
to the same constant. *) |
|
530 |
fun same_leaf_check (T as (Bs,FVs)) (Bound i) (Bound j) = |
|
531 |
let |
|
532 |
fun aux_chk (i1,i2) [] = false |
|
15531 | 533 |
| aux_chk (0,0) ((SOME _) :: bnds) = true |
534 |
| aux_chk (i1,0) (NONE :: bnds) = false |
|
535 |
| aux_chk (i1,i2) ((SOME _) :: bnds) = |
|
15481 | 536 |
aux_chk (i1 - 1,i2 - 1) bnds |
15531 | 537 |
| aux_chk (i1,i2) (NONE :: bnds) = |
15481 | 538 |
aux_chk (i1,i2 - 1) bnds |
539 |
in |
|
540 |
if (aux_chk (i,j) Bs) then [T] |
|
541 |
else [] |
|
542 |
end |
|
543 |
| same_leaf_check (T as (Bs,(Fs,Vs))) |
|
544 |
(a as (Free (an,aty))) |
|
545 |
(b as (Free (bn,bty))) = |
|
546 |
(case bterm_from an Fs of |
|
15531 | 547 |
SOME b2n => if bn = b2n then [T] |
15481 | 548 |
else [] (* conflict of var name *) |
15531 | 549 |
| NONE => [(Bs,((an,bn)::Fs,Vs))]) |
15481 | 550 |
| same_leaf_check (T as (Bs,(Fs,Vs))) |
551 |
(a as (Var (an,aty))) |
|
552 |
(b as (Var (bn,bty))) = |
|
553 |
(case bterm_from an Vs of |
|
15531 | 554 |
SOME b2n => if bn = b2n then [T] |
15481 | 555 |
else [] (* conflict of var name *) |
15531 | 556 |
| NONE => [(Bs,(Fs,(an,bn)::Vs))]) |
15481 | 557 |
| same_leaf_check T (a as (Const _)) (b as (Const _)) = |
558 |
if a = b then [T] else [] |
|
559 |
| same_leaf_check T _ _ = [] |
|
560 |
||
561 |
in |
|
562 |
same_leaf_check T a b |
|
563 |
end; |
|
564 |
||
565 |
(* wrapper for term_contains1: checks if the term "a" contains in |
|
566 |
some embedded way, (w.r.t. name -convertable) the term "b" *) |
|
567 |
fun term_contains a b = |
|
568 |
case term_contains1 ([],([],[])) a b of |
|
569 |
(_ :: _) => true |
|
570 |
| [] => false; |
|
571 |
||
572 |
||
573 |
(* pp_term a; pp_term b; writeln "EQTerm matches are: "; map (fn (a,b) => writeln ("(" ^ (string_of_term a) ^ ", " ^ (string_of_term b) ^ ")")) L;*) |
|
574 |
||
575 |
||
576 |
(* change all bound variables to see ones with appropriate name and |
|
577 |
type *) |
|
578 |
||
579 |
(* naming convention is OK as we use 'variant' from term.ML *) |
|
580 |
||
581 |
(* Note "bounds_to_frees" defined below, its better and quicker, but |
|
582 |
keeps the quantifiers handing about, and changes all bounds, not |
|
583 |
just universally quantified ones. *) |
|
584 |
fun change_bounds_to_frees t = |
|
585 |
let |
|
586 |
val vars = strip_all_vars t |
|
587 |
val frees_names = map (fn Free(s,n) => s | _ => "") (term_frees t) |
|
588 |
val body = strip_all_body t |
|
589 |
||
590 |
fun bnds_to_frees [] _ acc = acc |
|
591 |
| bnds_to_frees ((name,ty)::more) names acc = |
|
592 |
let |
|
593 |
val new_name = variant names name |
|
594 |
in |
|
595 |
bnds_to_frees more (new_name::names) (Free(new_name,ty)::acc) |
|
596 |
end |
|
597 |
in |
|
598 |
(subst_bounds ((bnds_to_frees vars frees_names []), body)) |
|
599 |
end; |
|
600 |
||
601 |
||
602 |
||
603 |
structure fvartabS = TableFun(type key = string val ord = string_ord); |
|
604 |
||
605 |
(* a runtime-quick function which makes sure that every variable has a |
|
606 |
unique name *) |
|
607 |
fun unique_namify_aux (ntab,(Abs(s,ty,t))) = |
|
608 |
(case (fvartabS.lookup (ntab,s)) of |
|
15531 | 609 |
NONE => |
15481 | 610 |
let |
611 |
val (ntab2,t2) = unique_namify_aux ((fvartabS.update ((s,s),ntab)), t) |
|
612 |
in |
|
613 |
(ntab2,Abs(s,ty,t2)) |
|
614 |
end |
|
15531 | 615 |
| SOME s2 => |
15481 | 616 |
let |
617 |
val s_new = (Term.variant (fvartabS.keys ntab) s) |
|
618 |
val (ntab2,t2) = |
|
619 |
unique_namify_aux ((fvartabS.update ((s_new,s_new),ntab)), t) |
|
620 |
in |
|
621 |
(ntab2,Abs(s_new,ty,t2)) |
|
622 |
end) |
|
623 |
| unique_namify_aux (ntab,(a $ b)) = |
|
624 |
let |
|
625 |
val (ntab1,t1) = unique_namify_aux (ntab,a) |
|
626 |
val (ntab2,t2) = unique_namify_aux (ntab1,b) |
|
627 |
in |
|
628 |
(ntab2, t1$t2) |
|
629 |
end |
|
630 |
| unique_namify_aux (nt as (ntab,Const x)) = nt |
|
631 |
| unique_namify_aux (nt as (ntab,f as Free (s,ty))) = |
|
632 |
(case (fvartabS.lookup (ntab,s)) of |
|
15531 | 633 |
NONE => ((fvartabS.update ((s,s),ntab)), f) |
634 |
| SOME _ => nt) |
|
15481 | 635 |
| unique_namify_aux (nt as (ntab,v as Var ((s,i),ty))) = |
636 |
(case (fvartabS.lookup (ntab,s)) of |
|
15531 | 637 |
NONE => ((fvartabS.update ((s,s),ntab)), v) |
638 |
| SOME _ => nt) |
|
15481 | 639 |
| unique_namify_aux (nt as (ntab, Bound i)) = nt; |
640 |
||
641 |
fun unique_namify t = |
|
642 |
#2 (unique_namify_aux (fvartabS.empty, t)); |
|
643 |
||
644 |
(* a runtime-quick function which makes sure that every variable has a |
|
645 |
unique name and also changes bound variables to free variables, used |
|
646 |
for embedding checks, Note that this is a pretty naughty term |
|
647 |
manipulation, which doesn't have necessary relation to the original |
|
648 |
sematics of the term. This is really a trick for our embedding code. *) |
|
649 |
||
650 |
fun bounds_to_frees_aux T (ntab,(Abs(s,ty,t))) = |
|
651 |
(case (fvartabS.lookup (ntab,s)) of |
|
15531 | 652 |
NONE => |
15481 | 653 |
let |
654 |
val (ntab2,t2) = bounds_to_frees_aux ((s,ty)::T) |
|
655 |
((fvartabS.update ((s,s),ntab)), t) |
|
656 |
in |
|
657 |
(ntab2,Abs(s,ty,t2)) |
|
658 |
end |
|
15531 | 659 |
| SOME s2 => |
15481 | 660 |
let |
661 |
val s_new = (Term.variant (fvartabS.keys ntab) s) |
|
662 |
val (ntab2,t2) = |
|
663 |
bounds_to_frees_aux ((s_new,ty)::T) |
|
664 |
(fvartabS.update (((s_new,s_new),ntab)), t) |
|
665 |
in |
|
666 |
(ntab2,Abs(s_new,ty,t2)) |
|
667 |
end) |
|
668 |
| bounds_to_frees_aux T (ntab,(a $ b)) = |
|
669 |
let |
|
670 |
val (ntab1,t1) = bounds_to_frees_aux T (ntab,a) |
|
671 |
val (ntab2,t2) = bounds_to_frees_aux T (ntab1,b) |
|
672 |
in |
|
673 |
(ntab2, t1$t2) |
|
674 |
end |
|
675 |
| bounds_to_frees_aux T (nt as (ntab,Const x)) = nt |
|
676 |
| bounds_to_frees_aux T (nt as (ntab,f as Free (s,ty))) = |
|
677 |
(case (fvartabS.lookup (ntab,s)) of |
|
15531 | 678 |
NONE => ((fvartabS.update ((s,s),ntab)), f) |
679 |
| SOME _ => nt) |
|
15481 | 680 |
| bounds_to_frees_aux T (nt as (ntab,v as Var ((s,i),ty))) = |
681 |
(case (fvartabS.lookup (ntab,s)) of |
|
15531 | 682 |
NONE => ((fvartabS.update ((s,s),ntab)), v) |
683 |
| SOME _ => nt) |
|
15481 | 684 |
| bounds_to_frees_aux T (nt as (ntab, Bound i)) = |
685 |
let |
|
15570 | 686 |
val (s,ty) = List.nth (T,i) |
15481 | 687 |
in |
688 |
(ntab, Free (s,ty)) |
|
689 |
end; |
|
690 |
||
691 |
||
692 |
fun bounds_to_frees t = |
|
693 |
#2 (bounds_to_frees_aux [] (fvartabS.empty,t)); |
|
694 |
||
695 |
fun bounds_to_frees_with_vars vars t = |
|
696 |
#2 (bounds_to_frees_aux vars (fvartabS.empty,t)); |
|
697 |
||
698 |
||
699 |
||
700 |
(* loose bounds to frees *) |
|
701 |
fun loose_bnds_to_frees_aux (bnds,vars) (Abs(s,ty,t)) = |
|
702 |
Abs(s,ty,loose_bnds_to_frees_aux (bnds + 1,vars) t) |
|
703 |
| loose_bnds_to_frees_aux (bnds,vars) (a $ b) = |
|
704 |
(loose_bnds_to_frees_aux (bnds,vars) a) |
|
705 |
$ (loose_bnds_to_frees_aux (bnds,vars) b) |
|
706 |
| loose_bnds_to_frees_aux (bnds,vars) (t as (Bound i)) = |
|
15570 | 707 |
if (bnds <= i) then Free (List.nth (vars,i - bnds)) |
15481 | 708 |
else t |
709 |
| loose_bnds_to_frees_aux (bnds,vars) t = t; |
|
710 |
||
711 |
||
712 |
fun loose_bnds_to_frees vars t = |
|
713 |
loose_bnds_to_frees_aux (0,vars) t; |
|
714 |
||
715 |
(* puts prems of a theorem into a useful form, (rulify) |
|
716 |
Note that any loose bound variables are treated as free vars |
|
717 |
*) |
|
718 |
(* fun make_term_into_simp_thm vars sgn t = |
|
719 |
let |
|
720 |
val triv = |
|
721 |
Thm.trivial (Thm.cterm_of |
|
722 |
sgn (loose_bnds_to_frees vars t)) |
|
723 |
in |
|
724 |
SplitterData.mk_eq (Drule.standard (ObjectLogic.rulify_no_asm triv)) |
|
725 |
end; |
|
726 |
||
727 |
fun thms_of_prems_in_goal i tm= |
|
728 |
let |
|
15570 | 729 |
val goal = (List.nth (Thm.prems_of tm,i-1)) |
15481 | 730 |
val vars = rev (strip_all_vars goal) |
731 |
val prems = Logic.strip_assums_hyp (strip_all_body goal) |
|
732 |
val simp_prem_thms = |
|
733 |
map (make_term_into_simp_thm vars (Thm.sign_of_thm tm)) prems |
|
734 |
in |
|
735 |
simp_prem_thms |
|
736 |
end; |
|
737 |
*) |
|
738 |
||
739 |
(* replace all universally quantifief variables (at HOL object level) |
|
740 |
with free vars |
|
741 |
fun HOL_Alls_to_frees TL (Const("All", _) $ Abs(v, ty, t)) = |
|
742 |
HOL_Alls_to_frees ((Free (v, ty)) :: TL) t |
|
743 |
| HOL_Alls_to_frees TL t = |
|
744 |
(TL,(subst_bounds (TL, t))); |
|
745 |
*) |
|
746 |
||
747 |
(* this is just a hack for mixing with interactive mode, and using topthm() *) |
|
748 |
||
749 |
fun try_dest_Goal (Const("Goal", _) $ T) = T |
|
750 |
| try_dest_Goal T = T; |
|
751 |
||
752 |
fun try_dest_Trueprop (Const("Trueprop", _) $ T) = T |
|
753 |
| try_dest_Trueprop T = T; |
|
754 |
||
755 |
fun bot_left_leaf_of (l $ r) = bot_left_leaf_of l |
|
756 |
| bot_left_leaf_of (Abs(s,ty,t)) = bot_left_leaf_of t |
|
757 |
| bot_left_leaf_of x = x; |
|
758 |
||
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
759 |
fun bot_left_leaf_noabs_of (l $ r) = bot_left_leaf_noabs_of l |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15798
diff
changeset
|
760 |
| bot_left_leaf_noabs_of x = x; |
15481 | 761 |
|
762 |
(* cleaned up normal form version of the subgoal terms conclusion *) |
|
763 |
fun cleaned_term_conc t = |
|
764 |
let |
|
765 |
val concl = Logic.strip_imp_concl (change_bounds_to_frees t) |
|
766 |
in |
|
767 |
(try_dest_Trueprop (try_dest_Goal concl)) |
|
768 |
end; |
|
769 |
||
770 |
(* fun get_prems_of_sg_term t = |
|
771 |
map opt_dest_Trueprop (Logic.strip_imp_prems t); *) |
|
772 |
||
773 |
fun get_prems_of_sg_term t = |
|
774 |
map try_dest_Trueprop (Logic.strip_assums_hyp t); |
|
775 |
||
776 |
||
777 |
(* drop premices, clean bound var stuff, and make a trueprop... *) |
|
778 |
fun cleaned_term_parts t = |
|
779 |
let |
|
780 |
val t2 = (change_bounds_to_frees t) |
|
781 |
val concl = Logic.strip_imp_concl t2 |
|
782 |
val prems = map try_dest_Trueprop (Logic.strip_imp_prems t2) |
|
783 |
in |
|
784 |
(prems, (try_dest_Trueprop (try_dest_Goal concl))) |
|
785 |
end; |
|
786 |
||
787 |
(* old version |
|
788 |
fun cleaned_term t = |
|
789 |
let |
|
790 |
val concl = (HOLogic.dest_Trueprop (Logic.strip_imp_concl |
|
791 |
(change_bounds_to_frees t) )) |
|
792 |
in |
|
793 |
concl |
|
794 |
end; |
|
795 |
*) |
|
796 |
||
797 |
(* change free variables to vars and visa versa *) |
|
798 |
(* *** check naming is OK, can we just use the vasr old name? *) |
|
799 |
(* *** Check: Logic.varify and Logic.unvarify *) |
|
800 |
fun change_vars_to_frees (a$b) = |
|
801 |
(change_vars_to_frees a) $ (change_vars_to_frees b) |
|
802 |
| change_vars_to_frees (Abs(s,ty,t)) = |
|
803 |
(Abs(s,Type.varifyT ty,change_vars_to_frees t)) |
|
804 |
| change_vars_to_frees (Var((s,i),ty)) = (Free(s,Type.unvarifyT ty)) |
|
805 |
| change_vars_to_frees l = l; |
|
806 |
||
807 |
fun change_frees_to_vars (a$b) = |
|
808 |
(change_frees_to_vars a) $ (change_frees_to_vars b) |
|
809 |
| change_frees_to_vars (Abs(s,ty,t)) = |
|
810 |
(Abs(s,Type.varifyT ty,change_frees_to_vars t)) |
|
811 |
| change_frees_to_vars (Free(s,ty)) = (Var((s,0),Type.varifyT ty)) |
|
812 |
| change_frees_to_vars l = l; |
|
813 |
||
814 |
||
815 |
end; |
|
816 |
||
817 |
||
818 |
||
819 |
(* ignores lambda abstractions, ie (\x y = y) |
|
820 |
same as embedding check code, ie. (f a b) = can "a" be embedded in "b" |
|
821 |
*) |
|
822 |
(* fun term_is_same_or_simpler_than (Abs(s,ty,t)) b = |
|
823 |
term_is_same_or_simpler_than t b |
|
824 |
||
825 |
| term_is_same_or_simpler_than a (Abs(s,ty,t)) = |
|
826 |
term_is_same_or_simpler_than a t |
|
827 |
||
828 |
| term_is_same_or_simpler_than (ah $ at) (bh $ bt) = |
|
829 |
(term_is_same_or_simpler_than (ah $ at) bh) orelse |
|
830 |
(term_is_same_or_simpler_than (ah $ at) bt) orelse |
|
831 |
((term_is_same_or_simpler_than ah bh) andalso |
|
832 |
(term_is_same_or_simpler_than at bt)) |
|
833 |
||
834 |
| term_is_same_or_simpler_than (ah $ at) b = false |
|
835 |
||
836 |
| term_is_same_or_simpler_than a (bh $ bt) = |
|
837 |
(term_is_same_or_simpler_than a bh) orelse |
|
838 |
(term_is_same_or_simpler_than a bt) |
|
839 |
||
840 |
| term_is_same_or_simpler_than a b = |
|
841 |
if a = b then true else false; |
|
842 |
*) |
|
843 |
||
844 |
(* fun term_is_simpler_than t1 t2 = |
|
845 |
(are_different_terms t1 t2) andalso |
|
846 |
(term_is_same_or_simpler_than t1 t2); |
|
847 |
*) |