author | wenzelm |
Thu, 02 Sep 2010 16:31:50 +0200 | |
changeset 39046 | 5b38730f3e12 |
parent 37124 | fe22fc54b876 |
child 39298 | 5aefb5bc8a93 |
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 |
|
9 |
(* intermediate term structure *) |
|
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 | |
|
16 |
SQua of squant * string list * sterm spattern list * sterm |
|
17 |
||
18 |
(* configuration options *) |
|
19 |
type prefixes = {sort_prefix: string, func_prefix: string} |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
20 |
type header = Proof.context -> term list -> string list |
36898 | 21 |
type strict = { |
22 |
is_builtin_conn: string * typ -> bool, |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
23 |
is_builtin_pred: Proof.context -> string * typ -> bool, |
36898 | 24 |
is_builtin_distinct: bool} |
25 |
type builtins = { |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
26 |
builtin_typ: Proof.context -> typ -> string option, |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
27 |
builtin_num: Proof.context -> typ -> int -> string option, |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
28 |
builtin_fun: Proof.context -> string * typ -> term list -> |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
29 |
(string * term list) option } |
36898 | 30 |
type sign = { |
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
31 |
header: string list, |
36898 | 32 |
sorts: string list, |
33 |
funcs: (string * (string list * string)) list } |
|
34 |
type config = { |
|
35 |
prefixes: prefixes, |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
36 |
header: header, |
36898 | 37 |
strict: strict option, |
38 |
builtins: builtins, |
|
39 |
serialize: string list -> sign -> sterm list -> string } |
|
40 |
type recon = { |
|
41 |
typs: typ Symtab.table, |
|
42 |
terms: term Symtab.table, |
|
43 |
unfolds: thm list, |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
44 |
assms: thm list } |
36898 | 45 |
|
46 |
val translate: config -> Proof.context -> string list -> thm list -> |
|
47 |
string * recon |
|
48 |
end |
|
49 |
||
50 |
structure SMT_Translate: SMT_TRANSLATE = |
|
51 |
struct |
|
52 |
||
53 |
(* intermediate term structure *) |
|
54 |
||
55 |
datatype squant = SForall | SExists |
|
56 |
||
57 |
datatype 'a spattern = SPat of 'a list | SNoPat of 'a list |
|
58 |
||
59 |
datatype sterm = |
|
60 |
SVar of int | |
|
61 |
SApp of string * sterm list | |
|
62 |
SLet of string * sterm * sterm | |
|
63 |
SQua of squant * string list * sterm spattern list * sterm |
|
64 |
||
65 |
||
66 |
||
67 |
(* configuration options *) |
|
68 |
||
69 |
type prefixes = {sort_prefix: string, func_prefix: string} |
|
70 |
||
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
71 |
type header = Proof.context -> term list -> string list |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
72 |
|
36898 | 73 |
type strict = { |
74 |
is_builtin_conn: string * typ -> bool, |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
75 |
is_builtin_pred: Proof.context -> string * typ -> bool, |
36898 | 76 |
is_builtin_distinct: bool} |
77 |
||
78 |
type builtins = { |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
79 |
builtin_typ: Proof.context -> typ -> string option, |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
80 |
builtin_num: Proof.context -> typ -> int -> string option, |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
81 |
builtin_fun: Proof.context -> string * typ -> term list -> |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
82 |
(string * term list) option } |
36898 | 83 |
|
84 |
type sign = { |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
85 |
header: string list, |
36898 | 86 |
sorts: string list, |
87 |
funcs: (string * (string list * string)) list } |
|
88 |
||
89 |
type config = { |
|
90 |
prefixes: prefixes, |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
91 |
header: header, |
36898 | 92 |
strict: strict option, |
93 |
builtins: builtins, |
|
94 |
serialize: string list -> sign -> sterm list -> string } |
|
95 |
||
96 |
type recon = { |
|
97 |
typs: typ Symtab.table, |
|
98 |
terms: term Symtab.table, |
|
99 |
unfolds: thm list, |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
100 |
assms: thm list } |
36898 | 101 |
|
102 |
||
103 |
||
104 |
(* utility functions *) |
|
105 |
||
106 |
val dest_funT = |
|
107 |
let |
|
108 |
fun dest Ts 0 T = (rev Ts, T) |
|
109 |
| dest Ts i (Type ("fun", [T, U])) = dest (T::Ts) (i-1) U |
|
110 |
| dest _ _ T = raise TYPE ("dest_funT", [T], []) |
|
111 |
in dest [] end |
|
112 |
||
113 |
val quantifier = (fn |
|
114 |
@{const_name All} => SOME SForall |
|
115 |
| @{const_name Ex} => SOME SExists |
|
116 |
| _ => NONE) |
|
117 |
||
118 |
fun group_quant qname Ts (t as Const (q, _) $ Abs (_, T, u)) = |
|
119 |
if q = qname then group_quant qname (T :: Ts) u else (Ts, t) |
|
120 |
| group_quant _ Ts t = (Ts, t) |
|
121 |
||
37124 | 122 |
fun dest_pat (Const (@{const_name pat}, _) $ t) = (t, true) |
123 |
| dest_pat (Const (@{const_name nopat}, _) $ t) = (t, false) |
|
124 |
| dest_pat t = raise TERM ("dest_pat", [t]) |
|
125 |
||
126 |
fun dest_pats [] = I |
|
127 |
| dest_pats ts = |
|
128 |
(case map dest_pat ts |> split_list ||> distinct (op =) of |
|
129 |
(ps, [true]) => cons (SPat ps) |
|
130 |
| (ps, [false]) => cons (SNoPat ps) |
|
131 |
| _ => raise TERM ("dest_pats", ts)) |
|
36898 | 132 |
|
133 |
fun dest_trigger (@{term trigger} $ tl $ t) = |
|
37124 | 134 |
(rev (fold (dest_pats o HOLogic.dest_list) (HOLogic.dest_list tl) []), t) |
36898 | 135 |
| dest_trigger t = ([], t) |
136 |
||
137 |
fun dest_quant qn T t = quantifier qn |> Option.map (fn q => |
|
138 |
let |
|
139 |
val (Ts, u) = group_quant qn [T] t |
|
140 |
val (ps, b) = dest_trigger u |
|
141 |
in (q, rev Ts, ps, b) end) |
|
142 |
||
143 |
fun fold_map_pat f (SPat ts) = fold_map f ts #>> SPat |
|
144 |
| fold_map_pat f (SNoPat ts) = fold_map f ts #>> SNoPat |
|
145 |
||
146 |
fun prop_of thm = HOLogic.dest_Trueprop (Thm.prop_of thm) |
|
147 |
||
148 |
||
149 |
||
150 |
(* enforce a strict separation between formulas and terms *) |
|
151 |
||
37124 | 152 |
val term_eq_rewr = @{lemma "term_eq x y == x = y" by (simp add: term_eq_def)} |
36898 | 153 |
|
37124 | 154 |
val term_bool = @{lemma "~(term_eq True False)" by (simp add: term_eq_def)} |
36898 | 155 |
val term_bool' = Simplifier.rewrite_rule [term_eq_rewr] term_bool |
156 |
||
157 |
||
158 |
val needs_rewrite = Thm.prop_of #> Term.exists_subterm (fn |
|
159 |
Const (@{const_name Let}, _) => true |
|
160 |
| @{term "op = :: bool => _"} $ _ $ @{term True} => true |
|
161 |
| Const (@{const_name If}, _) $ _ $ @{term True} $ @{term False} => true |
|
162 |
| _ => false) |
|
163 |
||
164 |
val rewrite_rules = [ |
|
165 |
Let_def, |
|
166 |
@{lemma "P = True == P" by (rule eq_reflection) simp}, |
|
167 |
@{lemma "if P then True else False == P" by (rule eq_reflection) simp}] |
|
168 |
||
169 |
fun rewrite ctxt = Simplifier.full_rewrite |
|
170 |
(Simplifier.context ctxt empty_ss addsimps rewrite_rules) |
|
171 |
||
172 |
fun normalize ctxt thm = |
|
173 |
if needs_rewrite thm then Conv.fconv_rule (rewrite ctxt) thm else thm |
|
174 |
||
175 |
val unfold_rules = term_eq_rewr :: rewrite_rules |
|
176 |
||
177 |
||
178 |
val revert_types = |
|
179 |
let |
|
180 |
fun revert @{typ prop} = @{typ bool} |
|
181 |
| revert (Type (n, Ts)) = Type (n, map revert Ts) |
|
182 |
| revert T = T |
|
183 |
in Term.map_types revert end |
|
184 |
||
185 |
||
186 |
fun strictify {is_builtin_conn, is_builtin_pred, is_builtin_distinct} ctxt = |
|
187 |
let |
|
188 |
fun is_builtin_conn' (@{const_name True}, _) = false |
|
189 |
| is_builtin_conn' (@{const_name False}, _) = false |
|
190 |
| is_builtin_conn' c = is_builtin_conn c |
|
191 |
||
192 |
val propT = @{typ prop} and boolT = @{typ bool} |
|
193 |
val as_propT = (fn @{typ bool} => propT | T => T) |
|
194 |
fun mapTs f g = Term.strip_type #> (fn (Ts, T) => map f Ts ---> g T) |
|
195 |
fun conn (n, T) = (n, mapTs as_propT as_propT T) |
|
196 |
fun pred (n, T) = (n, mapTs I as_propT T) |
|
197 |
||
198 |
val term_eq = @{term "op = :: bool => _"} |> Term.dest_Const |> pred |
|
199 |
fun as_term t = Const term_eq $ t $ @{term True} |
|
200 |
||
201 |
val if_term = Const (@{const_name If}, [propT, boolT, boolT] ---> boolT) |
|
202 |
fun wrap_in_if t = if_term $ t $ @{term True} $ @{term False} |
|
203 |
||
204 |
fun in_list T f t = HOLogic.mk_list T (map f (HOLogic.dest_list t)) |
|
205 |
||
206 |
fun in_term t = |
|
207 |
(case Term.strip_comb t of |
|
208 |
(c as Const (@{const_name If}, _), [t1, t2, t3]) => |
|
209 |
c $ in_form t1 $ in_term t2 $ in_term t3 |
|
210 |
| (h as Const c, ts) => |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
211 |
if is_builtin_conn' (conn c) orelse is_builtin_pred ctxt (pred c) |
36898 | 212 |
then wrap_in_if (in_form t) |
213 |
else Term.list_comb (h, map in_term ts) |
|
214 |
| (h as Free _, ts) => Term.list_comb (h, map in_term ts) |
|
215 |
| _ => t) |
|
216 |
||
217 |
and in_pat ((c as Const (@{const_name pat}, _)) $ t) = c $ in_term t |
|
218 |
| in_pat ((c as Const (@{const_name nopat}, _)) $ t) = c $ in_term t |
|
219 |
| in_pat t = raise TERM ("in_pat", [t]) |
|
220 |
||
37124 | 221 |
and in_pats ps = |
222 |
in_list @{typ "pattern list"} (in_list @{typ pattern} in_pat) ps |
|
36898 | 223 |
|
224 |
and in_trig ((c as @{term trigger}) $ p $ t) = c $ in_pats p $ in_form t |
|
225 |
| in_trig t = in_form t |
|
226 |
||
227 |
and in_form t = |
|
228 |
(case Term.strip_comb t of |
|
229 |
(q as Const (qn, _), [Abs (n, T, t')]) => |
|
230 |
if is_some (quantifier qn) then q $ Abs (n, T, in_trig t') |
|
231 |
else as_term (in_term t) |
|
232 |
| (Const (c as (@{const_name distinct}, T)), [t']) => |
|
233 |
if is_builtin_distinct then Const (pred c) $ in_list T in_term t' |
|
234 |
else as_term (in_term t) |
|
235 |
| (Const c, ts) => |
|
236 |
if is_builtin_conn (conn c) |
|
237 |
then Term.list_comb (Const (conn c), map in_form ts) |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
238 |
else if is_builtin_pred ctxt (pred c) |
36898 | 239 |
then Term.list_comb (Const (pred c), map in_term ts) |
240 |
else as_term (in_term t) |
|
241 |
| _ => as_term (in_term t)) |
|
242 |
in |
|
243 |
map (normalize ctxt) #> (fn thms => ((unfold_rules, term_bool' :: thms), |
|
244 |
map (in_form o prop_of) (term_bool :: thms))) |
|
245 |
end |
|
246 |
||
247 |
||
248 |
||
249 |
(* translation from Isabelle terms into SMT intermediate terms *) |
|
250 |
||
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
251 |
val empty_context = (1, Typtab.empty, 1, Termtab.empty) |
36898 | 252 |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
253 |
fun make_sign header (_, typs, _, terms) = { |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
254 |
header = header, |
36898 | 255 |
sorts = Typtab.fold (cons o snd) typs [], |
256 |
funcs = Termtab.fold (cons o snd) terms [] } |
|
257 |
||
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
258 |
fun make_recon (unfolds, assms) (_, typs, _, terms) = { |
36898 | 259 |
typs = Symtab.make (map swap (Typtab.dest typs)), |
260 |
terms = Symtab.make (map (fn (t, (n, _)) => (n, t)) (Termtab.dest terms)), |
|
261 |
unfolds = unfolds, |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
262 |
assms = assms } |
36898 | 263 |
|
264 |
fun string_of_index pre i = pre ^ string_of_int i |
|
265 |
||
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
266 |
fun fresh_typ sort_prefix T (cx as (Tidx, typs, idx, terms)) = |
36898 | 267 |
(case Typtab.lookup typs T of |
268 |
SOME s => (s, cx) |
|
269 |
| NONE => |
|
270 |
let |
|
271 |
val s = string_of_index sort_prefix Tidx |
|
272 |
val typs' = Typtab.update (T, s) typs |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
273 |
in (s, (Tidx+1, typs', idx, terms)) end) |
36898 | 274 |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
275 |
fun fresh_fun func_prefix t ss (cx as (Tidx, typs, idx, terms)) = |
36898 | 276 |
(case Termtab.lookup terms t of |
277 |
SOME (f, _) => (f, cx) |
|
278 |
| NONE => |
|
279 |
let |
|
280 |
val f = string_of_index func_prefix idx |
|
281 |
val terms' = Termtab.update (revert_types t, (f, ss)) terms |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
282 |
in (f, (Tidx, typs, idx+1, terms')) end) |
36898 | 283 |
|
284 |
fun relaxed thms = (([], thms), map prop_of thms) |
|
285 |
||
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
286 |
fun with_context header f (ths, ts) = |
36898 | 287 |
let val (us, context) = fold_map f ts empty_context |
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
288 |
in ((make_sign (header ts) context, us), make_recon ths context) end |
36898 | 289 |
|
290 |
||
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
291 |
fun translate {prefixes, strict, header, builtins, serialize} ctxt comments = |
36898 | 292 |
let |
293 |
val {sort_prefix, func_prefix} = prefixes |
|
294 |
val {builtin_typ, builtin_num, builtin_fun} = builtins |
|
295 |
||
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
296 |
fun transT T = |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
297 |
(case builtin_typ ctxt T of |
36898 | 298 |
SOME n => pair n |
299 |
| NONE => fresh_typ sort_prefix T) |
|
300 |
||
301 |
fun app n ts = SApp (n, ts) |
|
302 |
||
303 |
fun trans t = |
|
304 |
(case Term.strip_comb t of |
|
305 |
(Const (qn, _), [Abs (_, T, t1)]) => |
|
306 |
(case dest_quant qn T t1 of |
|
307 |
SOME (q, Ts, ps, b) => |
|
308 |
fold_map transT Ts ##>> fold_map (fold_map_pat trans) ps ##>> |
|
309 |
trans b #>> (fn ((Ts', ps'), b') => SQua (q, Ts', ps', b')) |
|
310 |
| NONE => raise TERM ("intermediate", [t])) |
|
311 |
| (Const (@{const_name Let}, _), [t1, Abs (_, T, t2)]) => |
|
312 |
transT T ##>> trans t1 ##>> trans t2 #>> |
|
313 |
(fn ((U, u1), u2) => SLet (U, u1, u2)) |
|
314 |
| (h as Const (c as (@{const_name distinct}, T)), [t1]) => |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
315 |
(case builtin_fun ctxt c (HOLogic.dest_list t1) of |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
316 |
SOME (n, ts) => fold_map trans ts #>> app n |
36898 | 317 |
| NONE => transs h T [t1]) |
318 |
| (h as Const (c as (_, T)), ts) => |
|
319 |
(case try HOLogic.dest_number t of |
|
320 |
SOME (T, i) => |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
321 |
(case builtin_num ctxt T i of |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
322 |
SOME n => pair (SApp (n, [])) |
36898 | 323 |
| NONE => transs t T []) |
324 |
| NONE => |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
325 |
(case builtin_fun ctxt c ts of |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
326 |
SOME (n, ts') => fold_map trans ts' #>> app n |
36898 | 327 |
| NONE => transs h T ts)) |
328 |
| (h as Free (_, T), ts) => transs h T ts |
|
329 |
| (Bound i, []) => pair (SVar i) |
|
330 |
| _ => raise TERM ("intermediate", [t])) |
|
331 |
||
332 |
and transs t T ts = |
|
333 |
let val (Us, U) = dest_funT (length ts) T |
|
334 |
in |
|
335 |
fold_map transT Us ##>> transT U #-> (fn Up => |
|
336 |
fresh_fun func_prefix t Up ##>> fold_map trans ts #>> SApp) |
|
337 |
end |
|
338 |
in |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
339 |
(case strict of SOME strct => strictify strct ctxt | NONE => relaxed) #> |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
340 |
with_context (header ctxt) trans #>> uncurry (serialize comments) |
36898 | 341 |
end |
342 |
||
343 |
end |