author | wenzelm |
Mon, 24 Oct 2022 20:37:32 +0200 | |
changeset 76371 | 1ac2416e8432 |
parent 75806 | 2b106aae897c |
child 78177 | ea7a3cc64df5 |
permissions | -rw-r--r-- |
75806
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75365
diff
changeset
|
1 |
(* Title: HOL/Tools/SMT/cvc_proof_parse.ML |
59015 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
3 |
||
75806
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75365
diff
changeset
|
4 |
CVC4 and cvc5 proof (actually, unsat core) parsing. |
59015 | 5 |
*) |
6 |
||
75806
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75365
diff
changeset
|
7 |
signature CVC_PROOF_PARSE = |
59015 | 8 |
sig |
9 |
val parse_proof: SMT_Translate.replay_data -> |
|
10 |
((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list -> |
|
11 |
SMT_Solver.parsed_proof |
|
12 |
end; |
|
13 |
||
75806
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75365
diff
changeset
|
14 |
structure CVC_Proof_Parse: CVC_PROOF_PARSE = |
59015 | 15 |
struct |
16 |
||
17 |
fun parse_proof ({ll_defs, assms, ...} : SMT_Translate.replay_data) xfacts prems _ output = |
|
60201 | 18 |
if exists (String.isPrefix "(error \"This build of CVC4 doesn't have proof support") output then |
19 |
{outcome = NONE, fact_ids = NONE, atp_proof = K []} |
|
20 |
else |
|
21 |
let |
|
22 |
val num_ll_defs = length ll_defs |
|
59015 | 23 |
|
60201 | 24 |
val id_of_index = Integer.add num_ll_defs |
25 |
val index_of_id = Integer.add (~ num_ll_defs) |
|
59015 | 26 |
|
75274
e89709b80b6e
used more descriptive assert names in SMT-Lib output
desharna
parents:
60201
diff
changeset
|
27 |
val used_assert_ids = |
e89709b80b6e
used more descriptive assert names in SMT-Lib output
desharna
parents:
60201
diff
changeset
|
28 |
map_filter (try (snd o SMTLIB_Interface.role_and_index_of_assert_name)) output |
60201 | 29 |
val used_assm_js = |
30 |
map_filter (fn id => let val i = index_of_id id in if i >= 0 then SOME i else NONE end) |
|
31 |
used_assert_ids |
|
59015 | 32 |
|
60201 | 33 |
val conjecture_i = 0 |
34 |
val prems_i = conjecture_i + 1 |
|
35 |
val num_prems = length prems |
|
36 |
val facts_i = prems_i + num_prems |
|
59015 | 37 |
|
60201 | 38 |
val fact_ids' = |
39 |
map_filter (fn j => |
|
75365 | 40 |
let val ((i, _), _) = nth assms j in |
60201 | 41 |
try (apsnd (nth xfacts)) (id_of_index j, i - facts_i) |
42 |
end) used_assm_js |
|
43 |
in |
|
44 |
{outcome = NONE, fact_ids = SOME fact_ids', atp_proof = K []} |
|
45 |
end |
|
59015 | 46 |
|
47 |
end; |