author | wenzelm |
Wed, 25 Jun 2025 16:35:25 +0200 | |
changeset 82768 | 8f866fd6fae1 |
parent 75365 | 83940294cc67 |
permissions | -rw-r--r-- |
58061 | 1 |
(* Title: HOL/Tools/SMT/z3_replay.ML |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
2 |
Author: Sascha Boehme, TU Muenchen |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
3 |
Author: Jasmin Blanchette, TU Muenchen |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
4 |
|
57219
34018603e0d0
factor out SMT-LIB 2 type/term parsing from Z3-specific code
blanchet
parents:
57165
diff
changeset
|
5 |
Z3 proof parsing and replay. |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
6 |
*) |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
7 |
|
58061 | 8 |
signature Z3_REPLAY = |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
9 |
sig |
58061 | 10 |
val parse_proof: Proof.context -> SMT_Translate.replay_data -> |
57164 | 11 |
((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list -> |
58061 | 12 |
SMT_Solver.parsed_proof |
13 |
val replay: Proof.context -> SMT_Translate.replay_data -> string list -> thm |
|
57229 | 14 |
end; |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
15 |
|
58061 | 16 |
structure Z3_Replay: Z3_REPLAY = |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
17 |
struct |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
18 |
|
69204
d5ab1636660b
split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
67091
diff
changeset
|
19 |
local |
d5ab1636660b
split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
67091
diff
changeset
|
20 |
fun extract (Z3_Proof.Z3_Step {id, rule, concl, fixes, ...}) = (id, rule, concl, fixes) |
d5ab1636660b
split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
67091
diff
changeset
|
21 |
fun cond rule = Z3_Proof.is_assumption rule andalso rule <> Z3_Proof.Hypothesis |
d5ab1636660b
split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
67091
diff
changeset
|
22 |
in |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
23 |
|
69204
d5ab1636660b
split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
67091
diff
changeset
|
24 |
val add_asserted = SMT_Replay.add_asserted Inttab.update Inttab.empty extract cond |
d5ab1636660b
split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
67091
diff
changeset
|
25 |
|
d5ab1636660b
split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
67091
diff
changeset
|
26 |
end |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
27 |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
28 |
fun add_paramTs names t = |
69204
d5ab1636660b
split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
67091
diff
changeset
|
29 |
fold2 (fn n => fn (_, T) => AList.update (op =) (n, T)) names (SMT_Replay.params_of t) |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
30 |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
31 |
fun new_fixes ctxt nTs = |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
32 |
let |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
33 |
val (ns, ctxt') = Variable.variant_fixes (replicate (length nTs) "") ctxt |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59582
diff
changeset
|
34 |
fun mk (n, T) n' = (n, Thm.cterm_of ctxt' (Free (n', T))) |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
35 |
in (ctxt', Symtab.make (map2 mk nTs ns)) end |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
36 |
|
69593 | 37 |
fun forall_elim_term ct (Const (\<^const_name>\<open>Pure.all\<close>, _) $ (a as Abs _)) = |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
38 |
Term.betapply (a, Thm.term_of ct) |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
39 |
| forall_elim_term _ qt = raise TERM ("forall_elim'", [qt]) |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
40 |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
41 |
fun apply_fixes elim env = fold (elim o the o Symtab.lookup env) |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
42 |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
43 |
val apply_fixes_prem = uncurry o apply_fixes Thm.forall_elim |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
44 |
val apply_fixes_concl = apply_fixes forall_elim_term |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
45 |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
46 |
fun export_fixes env names = Drule.forall_intr_list (map (the o Symtab.lookup env) names) |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
47 |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
48 |
fun under_fixes f ctxt (prems, nthms) names concl = |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
49 |
let |
69204
d5ab1636660b
split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
67091
diff
changeset
|
50 |
val thms1 = map (SMT_Replay.varify ctxt) prems |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
51 |
val (ctxt', env) = |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
52 |
add_paramTs names concl [] |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
53 |
|> fold (uncurry add_paramTs o apsnd Thm.prop_of) nthms |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
54 |
|> new_fixes ctxt |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
55 |
val thms2 = map (apply_fixes_prem env) nthms |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
56 |
val t = apply_fixes_concl env names concl |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
57 |
in export_fixes env names (f ctxt' (thms1 @ thms2) t) end |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
58 |
|
58061 | 59 |
fun replay_thm ctxt assumed nthms (Z3_Proof.Z3_Step {id, rule, concl, fixes, is_fix_step, ...}) = |
60 |
if Z3_Proof.is_assumption rule then |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
61 |
(case Inttab.lookup assumed id of |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
62 |
SOME (_, thm) => thm |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59582
diff
changeset
|
63 |
| NONE => Thm.assume (Thm.cterm_of ctxt concl)) |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
64 |
else |
58061 | 65 |
under_fixes (Z3_Replay_Methods.method_for rule) ctxt |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
66 |
(if is_fix_step then (map snd nthms, []) else ([], nthms)) fixes concl |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
67 |
|
59215
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
boehmes
parents:
59213
diff
changeset
|
68 |
fun replay_step ctxt assumed (step as Z3_Proof.Z3_Step {id, rule, prems, fixes, ...}) state = |
59213
ef5e68575bc4
limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
boehmes
parents:
58482
diff
changeset
|
69 |
let |
59215
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
boehmes
parents:
59213
diff
changeset
|
70 |
val (proofs, stats) = state |
59213
ef5e68575bc4
limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
boehmes
parents:
58482
diff
changeset
|
71 |
val nthms = map (the o Inttab.lookup proofs) prems |
59215
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
boehmes
parents:
59213
diff
changeset
|
72 |
val replay = Timing.timing (replay_thm ctxt assumed nthms) |
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
boehmes
parents:
59213
diff
changeset
|
73 |
val ({elapsed, ...}, thm) = |
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
boehmes
parents:
59213
diff
changeset
|
74 |
SMT_Config.with_time_limit ctxt SMT_Config.reconstruction_step_timeout replay step |
62519 | 75 |
handle Timeout.TIMEOUT _ => raise SMT_Failure.SMT SMT_Failure.Time_Out |
59215
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
boehmes
parents:
59213
diff
changeset
|
76 |
val stats' = Symtab.cons_list (Z3_Proof.string_of_rule rule, Time.toMilliseconds elapsed) stats |
35c13f4995b1
optionally display statistics for Z3 proof reconstruction
boehmes
parents:
59213
diff
changeset
|
77 |
in (Inttab.update (id, (fixes, thm)) proofs, stats') end |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
78 |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
79 |
(* |- (EX x. P x) = P c |- ~ (ALL x. P x) = ~ P c *) |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
80 |
local |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
81 |
val sk_rules = @{lemma |
67091 | 82 |
"c = (SOME x. P x) \<Longrightarrow> (\<exists>x. P x) = P c" |
83 |
"c = (SOME x. \<not> P x) \<Longrightarrow> (\<not> (\<forall>x. P x)) = (\<not> P c)" |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
84 |
by (metis someI_ex)+} |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
85 |
in |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
86 |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59381
diff
changeset
|
87 |
fun discharge_sk_tac ctxt i st = |
60752 | 88 |
(resolve_tac ctxt @{thms trans} i |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59381
diff
changeset
|
89 |
THEN resolve_tac ctxt sk_rules i |
60752 | 90 |
THEN (resolve_tac ctxt @{thms refl} ORELSE' discharge_sk_tac ctxt) (i+1) |
91 |
THEN resolve_tac ctxt @{thms refl} i) st |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
92 |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
93 |
end |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
94 |
|
67091 | 95 |
val true_thm = @{lemma "\<not>False" by simp} |
59381
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
boehmes
parents:
59374
diff
changeset
|
96 |
fun make_discharge_rules rules = rules @ [@{thm allI}, @{thm refl}, @{thm reflexive}, true_thm] |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
97 |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
98 |
val intro_def_rules = @{lemma |
67091 | 99 |
"(\<not> P \<or> P) \<and> (P \<or> \<not> P)" |
100 |
"(P \<or> \<not> P) \<and> (\<not> P \<or> P)" |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
101 |
by fast+} |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
102 |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59381
diff
changeset
|
103 |
fun discharge_assms_tac ctxt rules = |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59381
diff
changeset
|
104 |
REPEAT |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59381
diff
changeset
|
105 |
(HEADGOAL (resolve_tac ctxt (intro_def_rules @ rules) ORELSE' |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59381
diff
changeset
|
106 |
SOLVED' (discharge_sk_tac ctxt))) |
57230 | 107 |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
108 |
fun discharge_assms ctxt rules thm = |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
109 |
(if Thm.nprems_of thm = 0 then |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
110 |
thm |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
111 |
else |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59381
diff
changeset
|
112 |
(case Seq.pull (discharge_assms_tac ctxt rules thm) of |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
113 |
SOME (thm', _) => thm' |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
114 |
| NONE => raise THM ("failed to discharge premise", 1, [thm]))) |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
115 |
|> Goal.norm_result ctxt |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
116 |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
117 |
fun discharge rules outer_ctxt inner_ctxt = |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
118 |
singleton (Proof_Context.export inner_ctxt outer_ctxt) |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
119 |
#> discharge_assms outer_ctxt (make_discharge_rules rules) |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
120 |
|
57157 | 121 |
fun parse_proof outer_ctxt |
58061 | 122 |
({context = ctxt, typs, terms, ll_defs, rewrite_rules, assms} : SMT_Translate.replay_data) |
57541 | 123 |
xfacts prems concl output = |
57157 | 124 |
let |
58061 | 125 |
val (steps, ctxt2) = Z3_Proof.parse typs terms output ctxt |
75365 | 126 |
val ((iidths, _), _) = add_asserted outer_ctxt rewrite_rules (map (apfst fst) assms) steps ctxt2 |
57164 | 127 |
|
128 |
fun id_of_index i = the_default ~1 (Option.map fst (AList.lookup (op =) iidths i)) |
|
129 |
||
130 |
val conjecture_i = 0 |
|
131 |
val prems_i = 1 |
|
132 |
val facts_i = prems_i + length prems |
|
133 |
||
134 |
val conjecture_id = id_of_index conjecture_i |
|
135 |
val prem_ids = map id_of_index (prems_i upto facts_i - 1) |
|
58482
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58061
diff
changeset
|
136 |
val fact_ids' = |
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58061
diff
changeset
|
137 |
map_filter (fn (i, (id, _)) => try (apsnd (nth xfacts)) (id, i - facts_i)) iidths |
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58061
diff
changeset
|
138 |
val helper_ids' = map_filter (try (fn (~1, idth) => idth)) iidths |
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58061
diff
changeset
|
139 |
|
57164 | 140 |
val fact_helper_ts = |
59582 | 141 |
map (fn (_, th) => (ATP_Util.short_thm_name ctxt th, Thm.prop_of th)) helper_ids' @ |
142 |
map (fn (_, ((s, _), th)) => (s, Thm.prop_of th)) fact_ids' |
|
58482
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58061
diff
changeset
|
143 |
val fact_helper_ids' = |
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58061
diff
changeset
|
144 |
map (apsnd (ATP_Util.short_thm_name ctxt)) helper_ids' @ map (apsnd (fst o fst)) fact_ids' |
57157 | 145 |
in |
60201 | 146 |
{outcome = NONE, fact_ids = SOME fact_ids', |
58061 | 147 |
atp_proof = fn () => Z3_Isar.atp_proof_of_z3_proof ctxt ll_defs rewrite_rules prems concl |
58482
7836013951e6
simplified and repaired veriT index handling code
blanchet
parents:
58061
diff
changeset
|
148 |
fact_helper_ts prem_ids conjecture_id fact_helper_ids' steps} |
57157 | 149 |
end |
150 |
||
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
151 |
fun replay outer_ctxt |
58061 | 152 |
({context = ctxt, typs, terms, rewrite_rules, assms, ...} : SMT_Translate.replay_data) output = |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
153 |
let |
58061 | 154 |
val (steps, ctxt2) = Z3_Proof.parse typs terms output ctxt |
69204
d5ab1636660b
split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
67091
diff
changeset
|
155 |
val ((_, rules), (ctxt3, assumed)) = |
75365 | 156 |
add_asserted outer_ctxt rewrite_rules (map (apfst fst) assms) steps ctxt2 |
57157 | 157 |
val ctxt4 = |
158 |
ctxt3 |
|
69204
d5ab1636660b
split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
67091
diff
changeset
|
159 |
|> put_simpset (SMT_Replay.make_simpset ctxt3 []) |
58061 | 160 |
|> Config.put SAT.solver (Config.get ctxt3 SMT_Config.sat_solver) |
59374
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
boehmes
parents:
59215
diff
changeset
|
161 |
val len = length steps |
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
boehmes
parents:
59215
diff
changeset
|
162 |
val start = Timing.start () |
69204
d5ab1636660b
split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
67091
diff
changeset
|
163 |
val print_runtime_statistics = SMT_Replay.intermediate_statistics ctxt4 start len |
59374
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
boehmes
parents:
59215
diff
changeset
|
164 |
fun blockwise f (i, x) y = |
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
boehmes
parents:
59215
diff
changeset
|
165 |
(if i > 0 andalso i mod 100 = 0 then print_runtime_statistics i else (); f x y) |
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
boehmes
parents:
59215
diff
changeset
|
166 |
val (proofs, stats) = |
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
boehmes
parents:
59215
diff
changeset
|
167 |
fold_index (blockwise (replay_step ctxt4 assumed)) steps (assumed, Symtab.empty) |
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
boehmes
parents:
59215
diff
changeset
|
168 |
val _ = print_runtime_statistics len |
2f5447b764f9
more detailed runtime statistics for Z3 proof reconstruction
boehmes
parents:
59215
diff
changeset
|
169 |
val total = Time.toMilliseconds (#elapsed (Timing.result start)) |
58061 | 170 |
val (_, Z3_Proof.Z3_Step {id, ...}) = split_last steps |
69204
d5ab1636660b
split SMT reconstruction into library
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
67091
diff
changeset
|
171 |
val _ = SMT_Config.statistics_msg ctxt4 (Pretty.string_of o SMT_Replay.pretty_statistics "Z3" total) stats |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
172 |
in |
57157 | 173 |
Inttab.lookup proofs id |> the |> snd |> discharge rules outer_ctxt ctxt4 |
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
174 |
end |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
175 |
|
57229 | 176 |
end; |