| author | wenzelm | 
| Wed, 31 Mar 2021 23:45:16 +0200 | |
| changeset 73525 | 419edc7f3726 | 
| parent 69593 | 3dda49e08b9d | 
| child 75274 | e89709b80b6e | 
| permissions | -rw-r--r-- | 
| 
43928
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Tools/lambda_lifting.ML  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
2  | 
Author: Sascha Boehme, TU Muenchen  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
4  | 
Lambda-lifting on terms, i.e., replacing (some) lambda-abstractions by  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
5  | 
fresh names accompanied with defining equations for these fresh names in  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
6  | 
terms of the lambda-abstractions' bodies.  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
9  | 
signature LAMBDA_LIFTING =  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
10  | 
sig  | 
| 
43929
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
11  | 
type context = (term * term) Termtab.table * Proof.context  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
12  | 
val init: Proof.context -> context  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
13  | 
val is_quantifier: term -> bool  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
14  | 
val lift_lambdas1: (term -> bool) -> string option -> term -> context ->  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
15  | 
term * context  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
16  | 
val finish: context -> term list * Proof.context  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
17  | 
val lift_lambdas: string option -> (term -> bool) -> term list ->  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
18  | 
Proof.context -> (term list * term list) * Proof.context  | 
| 
43928
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
19  | 
end  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
20  | 
|
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
21  | 
structure Lambda_Lifting: LAMBDA_LIFTING =  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
22  | 
struct  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
23  | 
|
| 
43929
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
24  | 
fun mk_def Ts T lhs rhs =  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
25  | 
let fun mk_all T t = HOLogic.all_const T $ Abs (Name.uu, T, t)  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
26  | 
in fold mk_all Ts (HOLogic.eq_const T $ lhs $ rhs) end  | 
| 
43928
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
27  | 
|
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
28  | 
fun mk_abs Ts = fold (fn T => fn t => Abs (Name.uu, T, t)) Ts  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
29  | 
|
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
30  | 
fun dest_abs Ts (Abs (_, T, t)) = dest_abs (T :: Ts) t  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
31  | 
| dest_abs Ts t = (Ts, t)  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
32  | 
|
| 
43929
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
33  | 
fun replace_lambda basename Us Ts t (cx as (defs, ctxt)) =  | 
| 
43928
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
34  | 
let  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
35  | 
val t1 = mk_abs Us t  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
36  | 
val bs = sort int_ord (Term.add_loose_bnos (t1, 0, []))  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
37  | 
fun rep i k = if member (op =) bs i then (Bound k, k+1) else (Bound i, k)  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
38  | 
val (rs, _) = fold_map rep (0 upto length Ts - 1) 0  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
39  | 
val t2 = Term.subst_bounds (rs, t1)  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
40  | 
val Ts' = map (nth Ts) bs  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
41  | 
val (_, t3) = dest_abs [] t2  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
42  | 
val t4 = mk_abs Ts' t2  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
43  | 
|
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
44  | 
val T = Term.fastype_of1 (Us @ Ts, t)  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
45  | 
fun app f = Term.list_comb (f, map Bound (rev bs))  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
46  | 
in  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
47  | 
(case Termtab.lookup defs t4 of  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
48  | 
SOME (f, _) => (app f, cx)  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
49  | 
| NONE =>  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
50  | 
let  | 
| 
43929
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
51  | 
val (n, ctxt') = yield_singleton Variable.variant_fixes basename ctxt  | 
| 
43928
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
52  | 
val (is, UTs) = split_list (map_index I (Us @ Ts'))  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
53  | 
val f = Free (n, rev UTs ---> T)  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
54  | 
val lhs = Term.list_comb (f, map Bound (rev is))  | 
| 
43929
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
55  | 
val def = mk_def UTs (Term.fastype_of1 (Us @ Ts, t)) lhs t3  | 
| 
43928
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
56  | 
in (app f, (Termtab.update (t4, (f, def)) defs, ctxt')) end)  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
57  | 
end  | 
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
58  | 
|
| 
43929
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
59  | 
type context = (term * term) Termtab.table * Proof.context  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
60  | 
|
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
61  | 
fun init ctxt = (Termtab.empty, ctxt)  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
62  | 
|
| 69593 | 63  | 
fun is_quantifier (Const (\<^const_name>\<open>All\<close>, _)) = true  | 
64  | 
| is_quantifier (Const (\<^const_name>\<open>Ex\<close>, _)) = true  | 
|
| 
43929
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
65  | 
| is_quantifier _ = false  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
66  | 
|
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
67  | 
fun lift_lambdas1 is_binder basename =  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
68  | 
let  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
69  | 
val basename' = the_default Name.uu basename  | 
| 
43928
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
70  | 
|
| 
43929
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
71  | 
fun traverse Ts (t $ (u as Abs (n, T, body))) =  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
72  | 
if is_binder t then  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
73  | 
traverse Ts t ##>> traverse (T :: Ts) body #>> (fn (t', body') =>  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
74  | 
t' $ Abs (n, T, body'))  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
75  | 
else traverse Ts t ##>> traverse Ts u #>> (op $)  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
76  | 
| traverse Ts (t as Abs _) =  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
77  | 
let val (Us, u) = dest_abs [] t  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
78  | 
in traverse (Us @ Ts) u #-> replace_lambda basename' Us Ts end  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
79  | 
| traverse Ts (t $ u) = traverse Ts t ##>> traverse Ts u #>> (op $)  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
80  | 
| traverse _ t = pair t  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
81  | 
in traverse [] end  | 
| 
43928
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
82  | 
|
| 
43929
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
83  | 
fun finish (defs, ctxt) = (Termtab.fold (cons o snd o snd) defs [], ctxt)  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
84  | 
|
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
85  | 
fun lift_lambdas basename is_binder ts ctxt =  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
86  | 
init ctxt  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
87  | 
|> fold_map (lift_lambdas1 is_binder basename) ts  | 
| 
 
61d432e51aff
generalized lambda-lifting such that it is less specifically tailored for SMT (it does not anymore dependent on any SMT-specific code)
 
boehmes 
parents: 
43928 
diff
changeset
 | 
88  | 
|-> (fn ts' => finish #>> pair ts')  | 
| 
43928
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
24d6e759753f
moved lambda-lifting on terms into a separate structure (for better re-use in tools other than SMT)
 
boehmes 
parents:  
diff
changeset
 | 
90  | 
end  |