author | boehmes |
Fri, 08 Apr 2011 19:04:08 +0200 | |
changeset 42321 | ce83c1654b86 |
parent 42319 | 9a8ba59aed06 |
child 43154 | 72e4753a6677 |
permissions | -rw-r--r-- |
36898 | 1 |
(* Title: HOL/Tools/SMT/smt_translate.ML |
2 |
Author: Sascha Boehme, TU Muenchen |
|
3 |
||
4 |
Translate theorems into an SMT intermediate format and serialize them. |
|
5 |
*) |
|
6 |
||
7 |
signature SMT_TRANSLATE = |
|
8 |
sig |
|
41123 | 9 |
(*intermediate term structure*) |
36898 | 10 |
datatype squant = SForall | SExists |
11 |
datatype 'a spattern = SPat of 'a list | SNoPat of 'a list |
|
12 |
datatype sterm = |
|
13 |
SVar of int | |
|
14 |
SApp of string * sterm list | |
|
15 |
SLet of string * sterm * sterm | |
|
40664 | 16 |
SQua of squant * string list * sterm spattern list * int option * sterm |
36898 | 17 |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
18 |
(*translation configuration*) |
36898 | 19 |
type prefixes = {sort_prefix: string, func_prefix: string} |
20 |
type sign = { |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
21 |
header: string list, |
36898 | 22 |
sorts: string list, |
39298
5aefb5bc8a93
added preliminary support for SMT datatypes (for now restricted to tuples and lists); only the Z3 interface (in oracle mode) makes use of it, there is especially no Z3 proof reconstruction support for datatypes yet
boehmes
parents:
37124
diff
changeset
|
23 |
dtyps: (string * (string * (string * string) list) list) list list, |
36898 | 24 |
funcs: (string * (string list * string)) list } |
25 |
type config = { |
|
26 |
prefixes: prefixes, |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
27 |
header: term list -> string list, |
41059
d2b1fc1b8e19
centralized handling of built-in types and constants;
boehmes
parents:
41057
diff
changeset
|
28 |
is_fol: bool, |
d2b1fc1b8e19
centralized handling of built-in types and constants;
boehmes
parents:
41057
diff
changeset
|
29 |
has_datatypes: bool, |
36898 | 30 |
serialize: string list -> sign -> sterm list -> string } |
31 |
type recon = { |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
32 |
context: Proof.context, |
36898 | 33 |
typs: typ Symtab.table, |
34 |
terms: term Symtab.table, |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
35 |
rewrite_rules: thm list, |
40161
539d07b00e5f
keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents:
39535
diff
changeset
|
36 |
assms: (int * thm) list } |
36898 | 37 |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
38 |
(*translation*) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
39 |
val add_config: SMT_Utils.class * (Proof.context -> config) -> |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
40 |
Context.generic -> Context.generic |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
41 |
val translate: Proof.context -> string list -> (int * thm) list -> |
36898 | 42 |
string * recon |
43 |
end |
|
44 |
||
45 |
structure SMT_Translate: SMT_TRANSLATE = |
|
46 |
struct |
|
47 |
||
40663 | 48 |
|
36898 | 49 |
(* intermediate term structure *) |
50 |
||
51 |
datatype squant = SForall | SExists |
|
52 |
||
53 |
datatype 'a spattern = SPat of 'a list | SNoPat of 'a list |
|
54 |
||
55 |
datatype sterm = |
|
56 |
SVar of int | |
|
57 |
SApp of string * sterm list | |
|
58 |
SLet of string * sterm * sterm | |
|
40664 | 59 |
SQua of squant * string list * sterm spattern list * int option * sterm |
36898 | 60 |
|
61 |
||
62 |
||
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
63 |
(* translation configuration *) |
36898 | 64 |
|
65 |
type prefixes = {sort_prefix: string, func_prefix: string} |
|
66 |
||
67 |
type sign = { |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
68 |
header: string list, |
36898 | 69 |
sorts: string list, |
39298
5aefb5bc8a93
added preliminary support for SMT datatypes (for now restricted to tuples and lists); only the Z3 interface (in oracle mode) makes use of it, there is especially no Z3 proof reconstruction support for datatypes yet
boehmes
parents:
37124
diff
changeset
|
70 |
dtyps: (string * (string * (string * string) list) list) list list, |
36898 | 71 |
funcs: (string * (string list * string)) list } |
72 |
||
73 |
type config = { |
|
74 |
prefixes: prefixes, |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
75 |
header: term list -> string list, |
41059
d2b1fc1b8e19
centralized handling of built-in types and constants;
boehmes
parents:
41057
diff
changeset
|
76 |
is_fol: bool, |
d2b1fc1b8e19
centralized handling of built-in types and constants;
boehmes
parents:
41057
diff
changeset
|
77 |
has_datatypes: bool, |
36898 | 78 |
serialize: string list -> sign -> sterm list -> string } |
79 |
||
80 |
type recon = { |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
81 |
context: Proof.context, |
36898 | 82 |
typs: typ Symtab.table, |
83 |
terms: term Symtab.table, |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
84 |
rewrite_rules: thm list, |
40161
539d07b00e5f
keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents:
39535
diff
changeset
|
85 |
assms: (int * thm) list } |
36898 | 86 |
|
87 |
||
88 |
||
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
89 |
(* translation context *) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
90 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
91 |
fun make_tr_context {sort_prefix, func_prefix} = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
92 |
(sort_prefix, 1, Typtab.empty, func_prefix, 1, Termtab.empty) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
93 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
94 |
fun string_of_index pre i = pre ^ string_of_int i |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
95 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
96 |
fun add_typ T proper (cx as (sp, Tidx, typs, fp, idx, terms)) = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
97 |
(case Typtab.lookup typs T of |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
98 |
SOME (n, _) => (n, cx) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
99 |
| NONE => |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
100 |
let |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
101 |
val n = string_of_index sp Tidx |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
102 |
val typs' = Typtab.update (T, (n, proper)) typs |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
103 |
in (n, (sp, Tidx+1, typs', fp, idx, terms)) end) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
104 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
105 |
fun add_fun t sort (cx as (sp, Tidx, typs, fp, idx, terms)) = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
106 |
(case Termtab.lookup terms t of |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
107 |
SOME (n, _) => (n, cx) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
108 |
| NONE => |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
109 |
let |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
110 |
val n = string_of_index fp idx |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
111 |
val terms' = Termtab.update (t, (n, sort)) terms |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
112 |
in (n, (sp, Tidx, typs, fp, idx+1, terms')) end) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
113 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
114 |
fun sign_of header dtyps (_, _, typs, _, _, terms) = { |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
115 |
header = header, |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
116 |
sorts = Typtab.fold (fn (_, (n, true)) => cons n | _ => I) typs [], |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
117 |
dtyps = dtyps, |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
118 |
funcs = Termtab.fold (fn (_, (n, SOME ss)) => cons (n,ss) | _ => I) terms []} |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
119 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
120 |
fun recon_of ctxt rules thms ithms (_, _, typs, _, _, terms) = |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
121 |
let |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
122 |
fun add_typ (T, (n, _)) = Symtab.update (n, T) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
123 |
val typs' = Typtab.fold add_typ typs Symtab.empty |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
124 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
125 |
fun add_fun (t, (n, _)) = Symtab.update (n, t) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
126 |
val terms' = Termtab.fold add_fun terms Symtab.empty |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
127 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
128 |
val assms = map (pair ~1) thms @ ithms |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
129 |
in |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
130 |
{context=ctxt, typs=typs', terms=terms', rewrite_rules=rules, assms=assms} |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
131 |
end |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
132 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
133 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
134 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
135 |
(* preprocessing *) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
136 |
|
41426
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
137 |
(** datatype declarations **) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
138 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
139 |
fun collect_datatypes_and_records (tr_context, ctxt) ts = |
41426
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
140 |
let |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
141 |
val (declss, ctxt') = |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
142 |
fold (Term.fold_types SMT_Datatypes.add_decls) ts ([], ctxt) |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
143 |
|
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
144 |
fun is_decl_typ T = exists (exists (equal T o fst)) declss |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
145 |
|
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
146 |
fun add_typ' T proper = |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
147 |
(case SMT_Builtin.dest_builtin_typ ctxt' T of |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
148 |
SOME n => pair n |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
149 |
| NONE => add_typ T proper) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
150 |
|
41426
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
151 |
fun tr_select sel = |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
152 |
let val T = Term.range_type (Term.fastype_of sel) |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
153 |
in add_fun sel NONE ##>> add_typ' T (not (is_decl_typ T)) end |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
154 |
fun tr_constr (constr, selects) = |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
155 |
add_fun constr NONE ##>> fold_map tr_select selects |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
156 |
fun tr_typ (T, cases) = add_typ' T false ##>> fold_map tr_constr cases |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
157 |
val (declss', tr_context') = fold_map (fold_map tr_typ) declss tr_context |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
158 |
|
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
159 |
fun add (constr, selects) = |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
160 |
Termtab.update (constr, length selects) #> |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
161 |
fold (Termtab.update o rpair 1) selects |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
162 |
val funcs = fold (fold (fold add o snd)) declss Termtab.empty |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
163 |
|
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
164 |
in ((funcs, declss', tr_context', ctxt'), ts) end |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
165 |
(* FIXME: also return necessary datatype and record theorems *) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
166 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
167 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
168 |
(** eta-expand quantifiers, let expressions and built-ins *) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
169 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
170 |
local |
42319
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
171 |
fun eta f T t = Abs (Name.uu, T, f (Term.incr_boundvars 1 t $ Bound 0)) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
172 |
|
42319
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
173 |
fun exp f T = eta f (Term.domain_type (Term.domain_type T)) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
174 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
175 |
fun exp2 T q = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
176 |
let val U = Term.domain_type T |
42319
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
177 |
in Abs (Name.uu, U, q $ eta I (Term.domain_type U) (Bound 0)) end |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
178 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
179 |
fun exp2' T l = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
180 |
let val (U1, U2) = Term.dest_funT T ||> Term.domain_type |
42319
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
181 |
in Abs (Name.uu, U1, eta I U2 (l $ Bound 0)) end |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
182 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
183 |
fun expf k i T t = |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
184 |
let val Ts = drop i (fst (SMT_Utils.dest_funT k T)) |
41195
f59491d56327
fixed eta-expansion: introduce a couple of abstractions at once
boehmes
parents:
41173
diff
changeset
|
185 |
in |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
186 |
Term.incr_boundvars (length Ts) t |
42321
ce83c1654b86
fixed eta-expansion: use correct order to apply new bound variables
boehmes
parents:
42319
diff
changeset
|
187 |
|> fold_rev (fn i => fn u => u $ Bound i) (0 upto length Ts - 1) |
41195
f59491d56327
fixed eta-expansion: introduce a couple of abstractions at once
boehmes
parents:
41173
diff
changeset
|
188 |
|> fold_rev (fn T => fn u => Abs (Name.uu, T, u)) Ts |
f59491d56327
fixed eta-expansion: introduce a couple of abstractions at once
boehmes
parents:
41173
diff
changeset
|
189 |
end |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
190 |
in |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
191 |
|
42319
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
192 |
fun eta_expand ctxt is_fol funcs = |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
193 |
let |
41426
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
194 |
fun exp_func t T ts = |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
195 |
(case Termtab.lookup funcs t of |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
196 |
SOME k => |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
197 |
Term.list_comb (t, ts) |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
198 |
|> k <> length ts ? expf k (length ts) T |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
199 |
| NONE => Term.list_comb (t, ts)) |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
200 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
201 |
fun expand ((q as Const (@{const_name All}, _)) $ Abs a) = q $ abs_expand a |
42319
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
202 |
| expand ((q as Const (@{const_name All}, T)) $ t) = q $ exp expand T t |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
203 |
| expand (q as Const (@{const_name All}, T)) = exp2 T q |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
204 |
| expand ((q as Const (@{const_name Ex}, _)) $ Abs a) = q $ abs_expand a |
42319
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
205 |
| expand ((q as Const (@{const_name Ex}, T)) $ t) = q $ exp expand T t |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
206 |
| expand (q as Const (@{const_name Ex}, T)) = exp2 T q |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
207 |
| expand ((l as Const (@{const_name Let}, _)) $ t $ Abs a) = |
42319
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
208 |
if is_fol then expand (Term.betapply (Abs a, t)) |
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
209 |
else l $ expand t $ abs_expand a |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
210 |
| expand ((l as Const (@{const_name Let}, T)) $ t $ u) = |
42319
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
211 |
if is_fol then expand (u $ t) |
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
212 |
else l $ expand t $ exp expand (Term.range_type T) u |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
213 |
| expand ((l as Const (@{const_name Let}, T)) $ t) = |
42319
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
214 |
if is_fol then |
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
215 |
let val U = Term.domain_type (Term.range_type T) |
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
216 |
in Abs (Name.uu, U, Bound 0 $ Term.incr_boundvars 1 t) end |
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
217 |
else exp2 T (l $ expand t) |
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
218 |
| expand (l as Const (@{const_name Let}, T)) = |
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
219 |
if is_fol then |
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
220 |
let val U = Term.domain_type (Term.range_type T) |
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
221 |
in |
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
222 |
Abs (Name.uu, Term.domain_type T, Abs (Name.uu, U, |
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
223 |
Bound 0 $ Bound 1)) |
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
224 |
end |
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
225 |
else exp2' T l |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
226 |
| expand t = |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
227 |
(case Term.strip_comb t of |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
228 |
(u as Const (c as (_, T)), ts) => |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
229 |
(case SMT_Builtin.dest_builtin ctxt c ts of |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
230 |
SOME (_, k, us, mk) => |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
231 |
if k = length us then mk (map expand us) |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
232 |
else expf k (length ts) T (mk (map expand us)) |
41426
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
233 |
| NONE => exp_func u T (map expand ts)) |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
234 |
| (u as Free (_, T), ts) => exp_func u T (map expand ts) |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
235 |
| (Abs a, ts) => Term.list_comb (abs_expand a, map expand ts) |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
236 |
| (u, ts) => Term.list_comb (u, map expand ts)) |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
237 |
|
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
238 |
and abs_expand (n, T, t) = Abs (n, T, expand t) |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
239 |
|
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
240 |
in map expand end |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
241 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
242 |
end |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
243 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
244 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
245 |
(** lambda-lifting **) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
246 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
247 |
local |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
248 |
fun mk_def Ts T lhs rhs = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
249 |
let |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
250 |
val eq = HOLogic.eq_const T $ lhs $ rhs |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
251 |
val trigger = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
252 |
[[Const (@{const_name SMT.pat}, T --> @{typ SMT.pattern}) $ lhs]] |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
253 |
|> map (HOLogic.mk_list @{typ SMT.pattern}) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
254 |
|> HOLogic.mk_list @{typ "SMT.pattern list"} |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
255 |
fun mk_all T t = HOLogic.all_const T $ Abs (Name.uu, T, t) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
256 |
in fold mk_all Ts (@{const SMT.trigger} $ trigger $ eq) end |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
257 |
|
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
258 |
fun mk_abs Ts = fold (fn T => fn t => Abs (Name.uu, T, t)) Ts |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
259 |
|
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
260 |
fun dest_abs Ts (Abs (_, T, t)) = dest_abs (T :: Ts) t |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
261 |
| dest_abs Ts t = (Ts, t) |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
262 |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
263 |
fun replace_lambda Us Ts t (cx as (defs, ctxt)) = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
264 |
let |
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
265 |
val t1 = mk_abs Us t |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
266 |
val bs = sort int_ord (Term.add_loose_bnos (t1, 0, [])) |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
267 |
fun rep i k = if member (op =) bs i then (Bound k, k+1) else (Bound i, k) |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
268 |
val (rs, _) = fold_map rep (0 upto length Ts - 1) 0 |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
269 |
val t2 = Term.subst_bounds (rs, t1) |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
270 |
val Ts' = map (nth Ts) bs |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
271 |
val (_, t3) = dest_abs [] t2 |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
272 |
val t4 = mk_abs Ts' t2 |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
273 |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
274 |
val T = Term.fastype_of1 (Us @ Ts, t) |
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
275 |
fun app f = Term.list_comb (f, map Bound (rev bs)) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
276 |
in |
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
277 |
(case Termtab.lookup defs t4 of |
41197
edab1efe0a70
fix lambda-lifting: take level of bound variables into account and also apply bound variables from outer scope
boehmes
parents:
41196
diff
changeset
|
278 |
SOME (f, _) => (app f, cx) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
279 |
| NONE => |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
280 |
let |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
281 |
val (n, ctxt') = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
282 |
yield_singleton Variable.variant_fixes Name.uu ctxt |
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
283 |
val (is, UTs) = split_list (map_index I (Us @ Ts')) |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
284 |
val f = Free (n, rev UTs ---> T) |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
285 |
val lhs = Term.list_comb (f, map Bound (rev is)) |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
286 |
val def = mk_def UTs (Term.fastype_of1 (Us @ Ts, t)) lhs t3 |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
287 |
in (app f, (Termtab.update (t4, (f, def)) defs, ctxt')) end) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
288 |
end |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
289 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
290 |
fun traverse Ts t = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
291 |
(case t of |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
292 |
(q as Const (@{const_name All}, _)) $ Abs a => |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
293 |
abs_traverse Ts a #>> (fn a' => q $ Abs a') |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
294 |
| (q as Const (@{const_name Ex}, _)) $ Abs a => |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
295 |
abs_traverse Ts a #>> (fn a' => q $ Abs a') |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
296 |
| (l as Const (@{const_name Let}, _)) $ u $ Abs a => |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
297 |
traverse Ts u ##>> abs_traverse Ts a #>> |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
298 |
(fn (u', a') => l $ u' $ Abs a') |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
299 |
| Abs _ => |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
300 |
let val (Us, u) = dest_abs [] t |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
301 |
in traverse (Us @ Ts) u #-> replace_lambda Us Ts end |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
302 |
| u1 $ u2 => traverse Ts u1 ##>> traverse Ts u2 #>> (op $) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
303 |
| _ => pair t) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
304 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
305 |
and abs_traverse Ts (n, T, t) = traverse (T::Ts) t #>> (fn t' => (n, T, t')) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
306 |
in |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
307 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
308 |
fun lift_lambdas ctxt ts = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
309 |
(Termtab.empty, ctxt) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
310 |
|> fold_map (traverse []) ts |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
311 |
|> (fn (us, (defs, ctxt')) => |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
312 |
(ctxt', Termtab.fold (cons o snd o snd) defs us)) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
313 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
314 |
end |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
315 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
316 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
317 |
(** introduce explicit applications **) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
318 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
319 |
local |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
320 |
(* |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
321 |
Make application explicit for functions with varying number of arguments. |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
322 |
*) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
323 |
|
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
324 |
fun add t i = Termtab.map_default (t, i) (Integer.min i) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
325 |
|
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
326 |
fun min_arities t = |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
327 |
(case Term.strip_comb t of |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
328 |
(u as Const _, ts) => add u (length ts) #> fold min_arities ts |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
329 |
| (u as Free _, ts) => add u (length ts) #> fold min_arities ts |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
330 |
| (Abs (_, _, u), ts) => min_arities u #> fold min_arities ts |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
331 |
| (_, ts) => fold min_arities ts) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
332 |
|
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
333 |
fun app u (t, T) = |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
334 |
(Const (@{const_name SMT.fun_app}, T --> T) $ t $ u, Term.range_type T) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
335 |
|
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
336 |
fun apply i t T ts = |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
337 |
let |
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
338 |
val (ts1, ts2) = chop i ts |
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
339 |
val (_, U) = SMT_Utils.dest_funT i T |
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
340 |
in fst (fold app ts2 (Term.list_comb (t, ts1), U)) end |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
341 |
in |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
342 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
343 |
fun intro_explicit_application ts = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
344 |
let |
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
345 |
val arities = fold min_arities ts Termtab.empty |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
346 |
fun apply' t = apply (the (Termtab.lookup arities t)) t |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
347 |
|
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
348 |
fun traverse Ts t = |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
349 |
(case Term.strip_comb t of |
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
350 |
(u as Const (_, T), ts) => apply' u T (map (traverse Ts) ts) |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
351 |
| (u as Free (_, T), ts) => apply' u T (map (traverse Ts) ts) |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
352 |
| (u as Bound i, ts) => apply 0 u (nth Ts i) (map (traverse Ts) ts) |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
353 |
| (Abs (n, T, u), ts) => traverses Ts (Abs (n, T, traverse (T::Ts) u)) ts |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
354 |
| (u, ts) => traverses Ts u ts) |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
355 |
and traverses Ts t ts = Term.list_comb (t, map (traverse Ts) ts) |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
356 |
in map (traverse []) ts end |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
357 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
358 |
val fun_app_eq = mk_meta_eq @{thm SMT.fun_app_def} |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
359 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
360 |
end |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
361 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
362 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
363 |
(** map HOL formulas to FOL formulas (i.e., separate formulas froms terms) **) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
364 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
365 |
local |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
366 |
val term_bool = @{lemma "SMT.term_true ~= SMT.term_false" |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
367 |
by (simp add: SMT.term_true_def SMT.term_false_def)} |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
368 |
|
41785
77dcc197df9a
wrap occurrences of quantifiers in first-order terms (i.e., outside first-order formulas) in If-expressions
boehmes
parents:
41426
diff
changeset
|
369 |
val is_quant = member (op =) [@{const_name All}, @{const_name Ex}] |
77dcc197df9a
wrap occurrences of quantifiers in first-order terms (i.e., outside first-order formulas) in If-expressions
boehmes
parents:
41426
diff
changeset
|
370 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
371 |
val fol_rules = [ |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
372 |
Let_def, |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
373 |
mk_meta_eq @{thm SMT.term_true_def}, |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
374 |
mk_meta_eq @{thm SMT.term_false_def}, |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
375 |
@{lemma "P = True == P" by (rule eq_reflection) simp}, |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
376 |
@{lemma "if P then True else False == P" by (rule eq_reflection) simp}] |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
377 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
378 |
fun as_term t = @{const HOL.eq (bool)} $ t $ @{const SMT.term_true} |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
379 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
380 |
fun wrap_in_if t = |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
381 |
@{const If (bool)} $ t $ @{const SMT.term_true} $ @{const SMT.term_false} |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
382 |
|
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
383 |
fun is_builtin_conn_or_pred ctxt c ts = |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
384 |
is_some (SMT_Builtin.dest_builtin_conn ctxt c ts) orelse |
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
385 |
is_some (SMT_Builtin.dest_builtin_pred ctxt c ts) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
386 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
387 |
fun builtin b ctxt c ts = |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
388 |
(case (Const c, ts) of |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
389 |
(@{const HOL.eq (bool)}, [t, u]) => |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
390 |
if t = @{const SMT.term_true} orelse u = @{const SMT.term_true} then |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
391 |
SMT_Builtin.dest_builtin_eq ctxt t u |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
392 |
else b ctxt c ts |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
393 |
| _ => b ctxt c ts) |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
394 |
in |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
395 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
396 |
fun folify ctxt = |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
397 |
let |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
398 |
fun in_list T f t = HOLogic.mk_list T (map f (HOLogic.dest_list t)) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
399 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
400 |
fun in_term t = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
401 |
(case Term.strip_comb t of |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
402 |
(@{const True}, []) => @{const SMT.term_true} |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
403 |
| (@{const False}, []) => @{const SMT.term_false} |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
404 |
| (u as Const (@{const_name If}, _), [t1, t2, t3]) => |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
405 |
u $ in_form t1 $ in_term t2 $ in_term t3 |
41785
77dcc197df9a
wrap occurrences of quantifiers in first-order terms (i.e., outside first-order formulas) in If-expressions
boehmes
parents:
41426
diff
changeset
|
406 |
| (Const (c as (n, _)), ts) => |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
407 |
if is_builtin_conn_or_pred ctxt c ts then wrap_in_if (in_form t) |
41785
77dcc197df9a
wrap occurrences of quantifiers in first-order terms (i.e., outside first-order formulas) in If-expressions
boehmes
parents:
41426
diff
changeset
|
408 |
else if is_quant n then wrap_in_if (in_form t) |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
409 |
else Term.list_comb (Const c, map in_term ts) |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
410 |
| (Free c, ts) => Term.list_comb (Free c, map in_term ts) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
411 |
| _ => t) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
412 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
413 |
and in_weight ((c as @{const SMT.weight}) $ w $ t) = c $ w $ in_form t |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
414 |
| in_weight t = in_form t |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
415 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
416 |
and in_pat ((p as Const (@{const_name SMT.pat}, _)) $ t) = p $ in_term t |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
417 |
| in_pat ((p as Const (@{const_name SMT.nopat}, _)) $ t) = p $ in_term t |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
418 |
| in_pat t = raise TERM ("bad pattern", [t]) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
419 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
420 |
and in_pats ps = |
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
421 |
in_list @{typ "SMT.pattern list"} (in_list @{typ SMT.pattern} in_pat) ps |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
422 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
423 |
and in_trigger ((c as @{const SMT.trigger}) $ p $ t) = |
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
424 |
c $ in_pats p $ in_weight t |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
425 |
| in_trigger t = in_weight t |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
426 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
427 |
and in_form t = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
428 |
(case Term.strip_comb t of |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
429 |
(q as Const (qn, _), [Abs (n, T, u)]) => |
41785
77dcc197df9a
wrap occurrences of quantifiers in first-order terms (i.e., outside first-order formulas) in If-expressions
boehmes
parents:
41426
diff
changeset
|
430 |
if is_quant qn then q $ Abs (n, T, in_trigger u) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
431 |
else as_term (in_term t) |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
432 |
| (Const c, ts) => |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
433 |
(case SMT_Builtin.dest_builtin_conn ctxt c ts of |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
434 |
SOME (_, _, us, mk) => mk (map in_form us) |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
435 |
| NONE => |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
436 |
(case SMT_Builtin.dest_builtin_pred ctxt c ts of |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
437 |
SOME (_, _, us, mk) => mk (map in_term us) |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
438 |
| NONE => as_term (in_term t))) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
439 |
| _ => as_term (in_term t)) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
440 |
in |
42319
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
441 |
map in_form #> |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
442 |
cons (SMT_Utils.prop_of term_bool) #> |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
443 |
pair (fol_rules, [term_bool], builtin) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
444 |
end |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
445 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
446 |
end |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
447 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
448 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
449 |
(* translation into intermediate format *) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
450 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
451 |
(** utility functions **) |
36898 | 452 |
|
453 |
val quantifier = (fn |
|
454 |
@{const_name All} => SOME SForall |
|
455 |
| @{const_name Ex} => SOME SExists |
|
456 |
| _ => NONE) |
|
457 |
||
458 |
fun group_quant qname Ts (t as Const (q, _) $ Abs (_, T, u)) = |
|
459 |
if q = qname then group_quant qname (T :: Ts) u else (Ts, t) |
|
460 |
| group_quant _ Ts t = (Ts, t) |
|
461 |
||
40664 | 462 |
fun dest_weight (@{const SMT.weight} $ w $ t) = |
41173
7c6178d45cc8
fixed checking and translation of weights (previously, weights occurring in terms were rejected, and weight numbers were unintended translated into Vars)
boehmes
parents:
41172
diff
changeset
|
463 |
(SOME (snd (HOLogic.dest_number w)), t) |
40664 | 464 |
| dest_weight t = (NONE, t) |
465 |
||
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
466 |
fun dest_pat (Const (@{const_name SMT.pat}, _) $ t) = (t, true) |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
467 |
| dest_pat (Const (@{const_name SMT.nopat}, _) $ t) = (t, false) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
468 |
| dest_pat t = raise TERM ("bad pattern", [t]) |
37124 | 469 |
|
470 |
fun dest_pats [] = I |
|
471 |
| dest_pats ts = |
|
472 |
(case map dest_pat ts |> split_list ||> distinct (op =) of |
|
473 |
(ps, [true]) => cons (SPat ps) |
|
474 |
| (ps, [false]) => cons (SNoPat ps) |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
475 |
| _ => raise TERM ("bad multi-pattern", ts)) |
36898 | 476 |
|
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
477 |
fun dest_trigger (@{const SMT.trigger} $ tl $ t) = |
37124 | 478 |
(rev (fold (dest_pats o HOLogic.dest_list) (HOLogic.dest_list tl) []), t) |
36898 | 479 |
| dest_trigger t = ([], t) |
480 |
||
481 |
fun dest_quant qn T t = quantifier qn |> Option.map (fn q => |
|
482 |
let |
|
483 |
val (Ts, u) = group_quant qn [T] t |
|
40664 | 484 |
val (ps, p) = dest_trigger u |
485 |
val (w, b) = dest_weight p |
|
486 |
in (q, rev Ts, ps, w, b) end) |
|
36898 | 487 |
|
488 |
fun fold_map_pat f (SPat ts) = fold_map f ts #>> SPat |
|
489 |
| fold_map_pat f (SNoPat ts) = fold_map f ts #>> SNoPat |
|
490 |
||
491 |
||
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
492 |
(** translation from Isabelle terms into SMT intermediate terms **) |
36898 | 493 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
494 |
fun intermediate header dtyps builtin ctxt ts trx = |
41059
d2b1fc1b8e19
centralized handling of built-in types and constants;
boehmes
parents:
41057
diff
changeset
|
495 |
let |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
496 |
fun transT (T as TFree _) = add_typ T true |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
497 |
| transT (T as TVar _) = (fn _ => raise TYPE ("bad SMT type", [T], [])) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
498 |
| transT (T as Type _) = |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
499 |
(case SMT_Builtin.dest_builtin_typ ctxt T of |
39298
5aefb5bc8a93
added preliminary support for SMT datatypes (for now restricted to tuples and lists); only the Z3 interface (in oracle mode) makes use of it, there is especially no Z3 proof reconstruction support for datatypes yet
boehmes
parents:
37124
diff
changeset
|
500 |
SOME n => pair n |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
501 |
| NONE => add_typ T true) |
36898 | 502 |
|
503 |
fun app n ts = SApp (n, ts) |
|
504 |
||
505 |
fun trans t = |
|
506 |
(case Term.strip_comb t of |
|
507 |
(Const (qn, _), [Abs (_, T, t1)]) => |
|
508 |
(case dest_quant qn T t1 of |
|
40664 | 509 |
SOME (q, Ts, ps, w, b) => |
36898 | 510 |
fold_map transT Ts ##>> fold_map (fold_map_pat trans) ps ##>> |
40664 | 511 |
trans b #>> (fn ((Ts', ps'), b') => SQua (q, Ts', ps', w, b')) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
512 |
| NONE => raise TERM ("unsupported quantifier", [t])) |
36898 | 513 |
| (Const (@{const_name Let}, _), [t1, Abs (_, T, t2)]) => |
514 |
transT T ##>> trans t1 ##>> trans t2 #>> |
|
515 |
(fn ((U, u1), u2) => SLet (U, u1, u2)) |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
516 |
| (u as Const (c as (_, T)), ts) => |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
517 |
(case builtin ctxt c ts of |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
518 |
SOME (n, _, us, _) => fold_map trans us #>> app n |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
519 |
| NONE => transs u T ts) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
520 |
| (u as Free (_, T), ts) => transs u T ts |
36898 | 521 |
| (Bound i, []) => pair (SVar i) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
522 |
| _ => raise TERM ("bad SMT term", [t])) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
523 |
|
36898 | 524 |
and transs t T ts = |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
525 |
let val (Us, U) = SMT_Utils.dest_funT (length ts) T |
36898 | 526 |
in |
527 |
fold_map transT Us ##>> transT U #-> (fn Up => |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
528 |
add_fun t (SOME Up) ##>> fold_map trans ts #>> SApp) |
36898 | 529 |
end |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
530 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
531 |
val (us, trx') = fold_map trans ts trx |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
532 |
in ((sign_of (header ts) dtyps trx', us), trx') end |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
533 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
534 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
535 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
536 |
(* translation *) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
537 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
538 |
structure Configs = Generic_Data |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
539 |
( |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
540 |
type T = (Proof.context -> config) SMT_Utils.dict |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
541 |
val empty = [] |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
542 |
val extend = I |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
543 |
fun merge data = SMT_Utils.dict_merge fst data |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
544 |
) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
545 |
|
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
546 |
fun add_config (cs, cfg) = Configs.map (SMT_Utils.dict_update (cs, cfg)) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
547 |
|
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
548 |
fun get_config ctxt = |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
549 |
let val cs = SMT_Config.solver_class_of ctxt |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
550 |
in |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
551 |
(case SMT_Utils.dict_get (Configs.get (Context.Proof ctxt)) cs of |
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
552 |
SOME cfg => cfg ctxt |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
553 |
| NONE => error ("SMT: no translation configuration found " ^ |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
554 |
"for solver class " ^ quote (SMT_Utils.string_of_class cs))) |
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
555 |
end |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
556 |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
557 |
fun translate ctxt comments ithms = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
558 |
let |
41232
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
559 |
val {prefixes, is_fol, header, has_datatypes, serialize} = get_config ctxt |
4ea9f2a8c093
fixed lambda-lifting: shift indices of bound variables correctly (after locking the required bound variables) and apply bound variables to the new function symbol in the right order;
boehmes
parents:
41198
diff
changeset
|
560 |
|
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
561 |
val with_datatypes = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
562 |
has_datatypes andalso Config.get ctxt SMT_Config.datatypes |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
563 |
|
41426
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
564 |
fun no_dtyps (tr_context, ctxt) ts = |
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
565 |
((Termtab.empty, [], tr_context, ctxt), ts) |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
566 |
|
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
567 |
val ts1 = map (Envir.beta_eta_contract o SMT_Utils.prop_of o snd) ithms |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
568 |
|
41426
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
569 |
val ((funcs, dtyps, tr_context, ctxt1), ts2) = |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
570 |
((make_tr_context prefixes, ctxt), ts1) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
571 |
|-> (if with_datatypes then collect_datatypes_and_records else no_dtyps) |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
572 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
573 |
val (ctxt2, ts3) = |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
574 |
ts2 |
42319
9a8ba59aed06
unfold and eta-contract let expressions before lambda-lifting to avoid bad terms
boehmes
parents:
41785
diff
changeset
|
575 |
|> eta_expand ctxt1 is_fol funcs |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
576 |
|> lift_lambdas ctxt1 |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
577 |
||> intro_explicit_application |
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
578 |
|
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
579 |
val ((rewrite_rules, extra_thms, builtin), ts4) = |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
580 |
(if is_fol then folify ctxt2 else pair ([], [], I)) ts3 |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
581 |
|
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
582 |
val rewrite_rules' = fun_app_eq :: rewrite_rules |
36898 | 583 |
in |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
584 |
(ts4, tr_context) |
41328
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents:
41281
diff
changeset
|
585 |
|-> intermediate header dtyps (builtin SMT_Builtin.dest_builtin) ctxt2 |
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41123
diff
changeset
|
586 |
|>> uncurry (serialize comments) |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41250
diff
changeset
|
587 |
||> recon_of ctxt2 rewrite_rules' extra_thms ithms |
36898 | 588 |
end |
589 |
||
590 |
end |