| author | fleury |
| Mon, 02 Jun 2014 15:10:18 +0200 | |
| changeset 57154 | f0eff6393a32 |
| parent 56981 | 3ef45ce002b5 |
| child 57156 | 3546a67226ea |
| permissions | -rw-r--r-- |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
1 |
(* Title: HOL/Tools/SMT2/smt2_solver.ML |
|
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 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
4 |
SMT solvers registry and SMT tactic. |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
5 |
*) |
|
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 |
signature SMT2_SOLVER = |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
8 |
sig |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
9 |
(*configuration*) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
10 |
datatype outcome = Unsat | Sat | Unknown |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
11 |
type solver_config = {
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
12 |
name: string, |
| 56090 | 13 |
class: Proof.context -> SMT2_Util.class, |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
14 |
avail: unit -> bool, |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
15 |
command: unit -> string list, |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
16 |
options: Proof.context -> string list, |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
17 |
default_max_relevant: int, |
| 56125 | 18 |
can_filter: bool, |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
19 |
outcome: string -> string list -> outcome * string list, |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
20 |
cex_parser: (Proof.context -> SMT2_Translate.replay_data -> string list -> |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
21 |
term list * term list) option, |
|
56083
b5d1d9c60341
have Sledgehammer generate Isar proofs from Z3 proofs
blanchet
parents:
56082
diff
changeset
|
22 |
replay: (Proof.context -> SMT2_Translate.replay_data -> string list -> |
| 56125 | 23 |
((int * (int * thm)) list * Z3_New_Proof.z3_step list) * thm) option} |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
24 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
25 |
(*registry*) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
26 |
val add_solver: solver_config -> theory -> theory |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
27 |
val default_max_relevant: Proof.context -> string -> int |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
28 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
29 |
(*filter*) |
|
56104
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents:
56099
diff
changeset
|
30 |
val smt2_filter: Proof.context -> thm -> ('a * (int option * thm)) list -> int -> Time.time ->
|
| 56128 | 31 |
{outcome: SMT2_Failure.failure option, rewrite_rules: thm list, conjecture_id: int,
|
|
56981
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents:
56132
diff
changeset
|
32 |
prem_ids: int list, helper_ids: (int * thm) list, fact_ids: (int * ('a * thm)) list,
|
| 56128 | 33 |
z3_proof: Z3_New_Proof.z3_step list} |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
34 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
35 |
(*tactic*) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
36 |
val smt2_tac: Proof.context -> thm list -> int -> tactic |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
37 |
val smt2_tac': Proof.context -> thm list -> int -> tactic |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
38 |
end |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
39 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
40 |
structure SMT2_Solver: SMT2_SOLVER = |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
41 |
struct |
|
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 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
44 |
(* interface to external solvers *) |
|
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 |
local |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
47 |
|
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
48 |
fun make_command command options problem_path proof_path = |
|
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
49 |
"(exec 2>&1;" :: map File.shell_quote (command () @ options) @ |
|
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
50 |
[File.shell_path problem_path, ")", ">", File.shell_path proof_path] |
|
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
51 |
|> space_implode " " |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
52 |
|
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
53 |
fun with_trace ctxt msg f x = |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
54 |
let val _ = SMT2_Config.trace_msg ctxt (fn () => msg) () |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
55 |
in f x end |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
56 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
57 |
fun run ctxt name mk_cmd input = |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
58 |
(case SMT2_Config.certificates_of ctxt of |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
59 |
NONE => |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
60 |
if not (SMT2_Config.is_available ctxt name) then |
| 56131 | 61 |
error ("The SMT solver " ^ quote name ^ " is not installed")
|
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
62 |
else if Config.get ctxt SMT2_Config.debug_files = "" then |
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
63 |
with_trace ctxt ("Invoking SMT solver " ^ quote name ^ " ...") (Cache_IO.run mk_cmd) input
|
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
64 |
else |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
65 |
let |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
66 |
val base_path = Path.explode (Config.get ctxt SMT2_Config.debug_files) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
67 |
val in_path = Path.ext "smt2_in" base_path |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
68 |
val out_path = Path.ext "smt2_out" base_path |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
69 |
in Cache_IO.raw_run mk_cmd input in_path out_path end |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
70 |
| SOME certs => |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
71 |
(case Cache_IO.lookup certs input of |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
72 |
(NONE, key) => |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
73 |
if Config.get ctxt SMT2_Config.read_only_certificates then |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
74 |
error ("Bad certificate cache: missing certificate")
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
75 |
else |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
76 |
Cache_IO.run_and_cache certs key mk_cmd input |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
77 |
| (SOME output, _) => |
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
78 |
with_trace ctxt ("Using cached certificate from " ^
|
|
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
79 |
File.shell_path (Cache_IO.cache_path_of certs) ^ " ...") I output)) |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
80 |
|
| 56087 | 81 |
(* Z3 returns 1 if "get-model" or "get-model" fails *) |
82 |
val normal_return_codes = [0, 1] |
|
83 |
||
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
84 |
fun run_solver ctxt name mk_cmd input = |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
85 |
let |
| 56087 | 86 |
fun pretty tag ls = Pretty.string_of (Pretty.big_list tag (map Pretty.str ls)) |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
87 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
88 |
val _ = SMT2_Config.trace_msg ctxt (pretty "Problem:" o split_lines) input |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
89 |
|
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
90 |
val {redirected_output = res, output = err, return_code} =
|
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
91 |
SMT2_Config.with_timeout ctxt (run ctxt name mk_cmd) input |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
92 |
val _ = SMT2_Config.trace_msg ctxt (pretty "Solver:") err |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
93 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
94 |
val output = fst (take_suffix (equal "") res) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
95 |
val _ = SMT2_Config.trace_msg ctxt (pretty "Result:") output |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
96 |
|
| 56087 | 97 |
val _ = member (op =) normal_return_codes return_code orelse |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
98 |
raise SMT2_Failure.SMT (SMT2_Failure.Abnormal_Termination return_code) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
99 |
in output end |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
100 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
101 |
fun trace_assms ctxt = |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
102 |
SMT2_Config.trace_msg ctxt (Pretty.string_of o |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
103 |
Pretty.big_list "Assertions:" o map (Display.pretty_thm ctxt o snd)) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
104 |
|
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
105 |
fun trace_replay_data ({context = ctxt, typs, terms, ...} : SMT2_Translate.replay_data) =
|
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
106 |
let |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
107 |
fun pretty_eq n p = Pretty.block [Pretty.str n, Pretty.str " = ", p] |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
108 |
fun p_typ (n, T) = pretty_eq n (Syntax.pretty_typ ctxt T) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
109 |
fun p_term (n, t) = pretty_eq n (Syntax.pretty_term ctxt t) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
110 |
in |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
111 |
SMT2_Config.trace_msg ctxt (fn () => |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
112 |
Pretty.string_of (Pretty.big_list "Names:" [ |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
113 |
Pretty.big_list "sorts:" (map p_typ (Symtab.dest typs)), |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
114 |
Pretty.big_list "functions:" (map p_term (Symtab.dest terms))])) () |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
115 |
end |
|
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 |
in |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
118 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
119 |
fun invoke name command ithms ctxt = |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
120 |
let |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
121 |
val options = SMT2_Config.solver_options_of ctxt |
| 56112 | 122 |
val comments = [space_implode " " options] |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
123 |
|
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
124 |
val (str, replay_data as {context = ctxt', ...}) =
|
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
125 |
ithms |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
126 |
|> tap (trace_assms ctxt) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
127 |
|> SMT2_Translate.translate ctxt comments |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
128 |
||> tap trace_replay_data |
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
129 |
in (run_solver ctxt' name (make_command command options) str, replay_data) end |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
130 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
131 |
end |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
132 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
133 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
134 |
(* configuration *) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
135 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
136 |
datatype outcome = Unsat | Sat | Unknown |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
137 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
138 |
type solver_config = {
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
139 |
name: string, |
| 56090 | 140 |
class: Proof.context -> SMT2_Util.class, |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
141 |
avail: unit -> bool, |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
142 |
command: unit -> string list, |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
143 |
options: Proof.context -> string list, |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
144 |
default_max_relevant: int, |
| 56125 | 145 |
can_filter: bool, |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
146 |
outcome: string -> string list -> outcome * string list, |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
147 |
cex_parser: (Proof.context -> SMT2_Translate.replay_data -> string list -> |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
148 |
term list * term list) option, |
|
56083
b5d1d9c60341
have Sledgehammer generate Isar proofs from Z3 proofs
blanchet
parents:
56082
diff
changeset
|
149 |
replay: (Proof.context -> SMT2_Translate.replay_data -> string list -> |
| 56125 | 150 |
((int * (int * thm)) list * Z3_New_Proof.z3_step list) * thm) option} |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
151 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
152 |
|
| 56106 | 153 |
(* check well-sortedness *) |
154 |
||
155 |
val has_topsort = Term.exists_type (Term.exists_subtype (fn |
|
156 |
TFree (_, []) => true |
|
157 |
| TVar (_, []) => true |
|
158 |
| _ => false)) |
|
159 |
||
160 |
(* top sorts cause problems with atomization *) |
|
161 |
fun check_topsort ctxt thm = |
|
162 |
if has_topsort (Thm.prop_of thm) then (SMT2_Normalize.drop_fact_warning ctxt thm; TrueI) else thm |
|
163 |
||
164 |
||
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
165 |
(* registry *) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
166 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
167 |
type solver_info = {
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
168 |
command: unit -> string list, |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
169 |
default_max_relevant: int, |
| 56125 | 170 |
can_filter: bool, |
171 |
finish: Proof.context -> SMT2_Translate.replay_data -> string list -> |
|
172 |
((int * (int * thm)) list * Z3_New_Proof.z3_step list) * thm} |
|
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
173 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
174 |
structure Solvers = Generic_Data |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
175 |
( |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
176 |
type T = solver_info Symtab.table |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
177 |
val empty = Symtab.empty |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
178 |
val extend = I |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
179 |
fun merge data = Symtab.merge (K true) data |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
180 |
) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
181 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
182 |
local |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
183 |
fun finish outcome cex_parser replay ocl outer_ctxt |
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
184 |
(replay_data as {context = ctxt, ...} : SMT2_Translate.replay_data) output =
|
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
185 |
(case outcome output of |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
186 |
(Unsat, ls) => |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
187 |
if not (Config.get ctxt SMT2_Config.oracle) andalso is_some replay |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
188 |
then the replay outer_ctxt replay_data ls |
|
56083
b5d1d9c60341
have Sledgehammer generate Isar proofs from Z3 proofs
blanchet
parents:
56082
diff
changeset
|
189 |
else (([], []), ocl ()) |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
190 |
| (result, ls) => |
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
191 |
let val (ts, us) = (case cex_parser of SOME f => f ctxt replay_data ls | _ => ([], [])) |
|
56083
b5d1d9c60341
have Sledgehammer generate Isar proofs from Z3 proofs
blanchet
parents:
56082
diff
changeset
|
192 |
in |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
193 |
raise SMT2_Failure.SMT (SMT2_Failure.Counterexample {
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
194 |
is_real_cex = (result = Sat), |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
195 |
free_constraints = ts, |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
196 |
const_defs = us}) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
197 |
end) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
198 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
199 |
val cfalse = Thm.cterm_of @{theory} (@{const Trueprop} $ @{const False})
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
200 |
in |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
201 |
|
| 56125 | 202 |
fun add_solver ({name, class, avail, command, options, default_max_relevant, can_filter,
|
203 |
outcome, cex_parser, replay} : solver_config) = |
|
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
204 |
let |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
205 |
fun solver ocl = {
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
206 |
command = command, |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
207 |
default_max_relevant = default_max_relevant, |
| 56125 | 208 |
can_filter = can_filter, |
209 |
finish = finish (outcome name) cex_parser replay ocl} |
|
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
210 |
|
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
211 |
val info = {name = name, class = class, avail = avail, options = options}
|
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
212 |
in |
| 56125 | 213 |
Thm.add_oracle (Binding.name name, K cfalse) #-> (fn (_, ocl) => |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
214 |
Context.theory_map (Solvers.map (Symtab.update_new (name, solver ocl)))) #> |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
215 |
Context.theory_map (SMT2_Config.add_solver info) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
216 |
end |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
217 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
218 |
end |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
219 |
|
| 56125 | 220 |
fun get_info ctxt name = the (Symtab.lookup (Solvers.get (Context.Proof ctxt)) name) |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
221 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
222 |
fun name_and_info_of ctxt = |
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
223 |
let val name = SMT2_Config.solver_of ctxt |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
224 |
in (name, get_info ctxt name) end |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
225 |
|
| 56106 | 226 |
fun apply_solver ctxt wthms0 = |
| 56082 | 227 |
let |
| 56106 | 228 |
val wthms = map (apsnd (check_topsort ctxt)) wthms0 |
| 56125 | 229 |
val (name, {command, finish, ...}) = name_and_info_of ctxt
|
230 |
val (output, replay_data) = invoke name command (SMT2_Normalize.normalize ctxt wthms) ctxt |
|
| 56128 | 231 |
in (finish ctxt replay_data output, replay_data) end |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
232 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
233 |
val default_max_relevant = #default_max_relevant oo get_info |
| 56125 | 234 |
val can_filter = #can_filter o snd o name_and_info_of |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
235 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
236 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
237 |
(* filter *) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
238 |
|
| 56106 | 239 |
val no_id = ~1 |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
240 |
|
|
56104
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents:
56099
diff
changeset
|
241 |
fun smt2_filter ctxt goal xwfacts i time_limit = |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
242 |
let |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
243 |
val ctxt = |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
244 |
ctxt |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
245 |
|> Config.put SMT2_Config.oracle false |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
246 |
|> Config.put SMT2_Config.filter_only_facts true |
| 56082 | 247 |
|> Config.put SMT2_Config.timeout (Time.toReal time_limit) |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
248 |
|
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
249 |
val ({context = ctxt, prems, concl, ...}, _) = Subgoal.focus ctxt i goal
|
| 56106 | 250 |
fun negate ct = Thm.dest_comb ct ||> Thm.apply @{cterm Not} |-> Thm.apply
|
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
251 |
val cprop = |
| 56082 | 252 |
(case try negate (Thm.rhs_of (SMT2_Normalize.atomize_conv ctxt concl)) of |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
253 |
SOME ct => ct |
| 56094 | 254 |
| NONE => raise SMT2_Failure.SMT (SMT2_Failure.Other_Failure "goal is not a HOL term")) |
| 56082 | 255 |
|
| 56106 | 256 |
val wconjecture = (NONE, Thm.assume cprop) |
257 |
val wprems = map (pair NONE) prems |
|
258 |
val wfacts = map snd xwfacts |
|
259 |
val wthms = wconjecture :: wprems @ wfacts |
|
260 |
val iwthms = map_index I wthms |
|
| 56082 | 261 |
|
| 56106 | 262 |
val conjecture_i = 0 |
|
56981
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents:
56132
diff
changeset
|
263 |
val prems_i = 1 |
|
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents:
56132
diff
changeset
|
264 |
val facts_i = prems_i + length wprems |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
265 |
in |
| 56106 | 266 |
wthms |
| 56082 | 267 |
|> apply_solver ctxt |
| 56128 | 268 |
|> (fn (((iidths0, z3_proof), _), {rewrite_rules, ...}) =>
|
|
56981
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents:
56132
diff
changeset
|
269 |
let |
|
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents:
56132
diff
changeset
|
270 |
val iidths = if can_filter ctxt then iidths0 else map (apsnd (apfst (K no_id))) iwthms |
|
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents:
56132
diff
changeset
|
271 |
fun id_of_index i = the_default no_id (Option.map fst (AList.lookup (op =) iidths i)) |
|
56104
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents:
56099
diff
changeset
|
272 |
in |
| 56128 | 273 |
{outcome = NONE,
|
274 |
rewrite_rules = rewrite_rules, |
|
|
56981
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents:
56132
diff
changeset
|
275 |
conjecture_id = id_of_index conjecture_i, |
|
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents:
56132
diff
changeset
|
276 |
prem_ids = map id_of_index (prems_i upto facts_i - 1), |
| 56106 | 277 |
helper_ids = map_filter (try (fn (~1, idth) => idth)) iidths, |
278 |
fact_ids = map_filter (fn (i, (id, _)) => |
|
279 |
try (apsnd (apsnd snd o nth xwfacts)) (id, i - facts_i)) iidths, |
|
|
56104
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents:
56099
diff
changeset
|
280 |
z3_proof = z3_proof} |
|
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents:
56099
diff
changeset
|
281 |
end) |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
282 |
end |
| 56128 | 283 |
handle SMT2_Failure.SMT fail => {outcome = SOME fail, rewrite_rules = [], conjecture_id = no_id,
|
|
56981
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents:
56132
diff
changeset
|
284 |
prem_ids = [], helper_ids = [], fact_ids = [], z3_proof = []} |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
285 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
286 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
287 |
(* SMT tactic *) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
288 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
289 |
local |
| 56106 | 290 |
fun trace_assumptions ctxt wfacts iidths = |
291 |
let val used = map_filter (try (snd o nth wfacts) o fst) iidths in |
|
|
56104
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents:
56099
diff
changeset
|
292 |
if Config.get ctxt SMT2_Config.trace_used_facts andalso length wfacts > 0 then |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
293 |
tracing (Pretty.string_of (Pretty.big_list "SMT used facts:" |
| 56106 | 294 |
(map (Display.pretty_thm ctxt) used))) |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
295 |
else () |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
296 |
end |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
297 |
|
| 56106 | 298 |
fun solve ctxt wfacts = |
299 |
wfacts |
|
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
300 |
|> apply_solver ctxt |
| 56128 | 301 |
|> fst |
| 56106 | 302 |
|>> apfst (trace_assumptions ctxt wfacts) |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
303 |
|> snd |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
304 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
305 |
fun str_of ctxt fail = |
|
56132
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
blanchet
parents:
56131
diff
changeset
|
306 |
"Solver " ^ SMT2_Config.solver_of ctxt ^ ": " ^ SMT2_Failure.string_of_failure ctxt fail |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
307 |
|
| 56106 | 308 |
fun safe_solve ctxt wfacts = SOME (solve ctxt wfacts) |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
309 |
handle |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
310 |
SMT2_Failure.SMT (fail as SMT2_Failure.Counterexample _) => |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
311 |
(SMT2_Config.verbose_msg ctxt (str_of ctxt) fail; NONE) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
312 |
| SMT2_Failure.SMT (fail as SMT2_Failure.Time_Out) => |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
313 |
error ("SMT: Solver " ^ quote (SMT2_Config.solver_of ctxt) ^ ": " ^
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
314 |
SMT2_Failure.string_of_failure ctxt fail ^ " (setting the " ^ |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
315 |
"configuration option " ^ quote (Config.name_of SMT2_Config.timeout) ^ " might help)") |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
316 |
| SMT2_Failure.SMT fail => error (str_of ctxt fail) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
317 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
318 |
fun resolve (SOME thm) = rtac thm 1 |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
319 |
| resolve NONE = no_tac |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
320 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
321 |
fun tac prove ctxt rules = |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
322 |
CONVERSION (SMT2_Normalize.atomize_conv ctxt) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
323 |
THEN' rtac @{thm ccontr}
|
| 56106 | 324 |
THEN' SUBPROOF (fn {context = ctxt, prems, ...} =>
|
325 |
resolve (prove ctxt (map (pair NONE) (rules @ prems)))) ctxt |
|
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
326 |
in |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
327 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
328 |
val smt2_tac = tac safe_solve |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
329 |
val smt2_tac' = tac (SOME oo solve) |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
330 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
331 |
end |
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
332 |
|
|
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff
changeset
|
333 |
end |