1(*  Title:      HOL/Tools/SMT/smt_solver.ML
2    Author:     Sascha Boehme, TU Muenchen
3
4SMT solvers registry and SMT tactic.
5*)
6
7signature SMT_SOLVER =
8sig
9  (*configuration*)
10  datatype outcome = Unsat | Sat | Unknown | Time_Out
11
12  type parsed_proof =
13    {outcome: SMT_Failure.failure option,
14     fact_ids: (int * ((string * ATP_Problem_Generate.stature) * thm)) list option,
15     atp_proof: unit -> (term, string) ATP_Proof.atp_step list}
16
17  type solver_config =
18    {name: string,
19     class: Proof.context -> SMT_Util.class,
20     avail: unit -> bool,
21     command: unit -> string list,
22     options: Proof.context -> string list,
23     smt_options: (string * string) list,
24     default_max_relevant: int,
25     outcome: string -> string list -> outcome * string list,
26     parse_proof: (Proof.context -> SMT_Translate.replay_data ->
27       ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->
28       parsed_proof) option,
29     replay: (Proof.context -> SMT_Translate.replay_data -> string list -> thm) option}
30
31  (*registry*)
32  val add_solver: solver_config -> theory -> theory
33  val default_max_relevant: Proof.context -> string -> int
34
35  (*filter*)
36  val smt_filter: Proof.context -> thm -> ((string * ATP_Problem_Generate.stature) * thm) list ->
37    int -> Time.time -> parsed_proof
38
39  (*tactic*)
40  val smt_tac: Proof.context -> thm list -> int -> tactic
41  val smt_tac': Proof.context -> thm list -> int -> tactic
42end;
43
44structure SMT_Solver: SMT_SOLVER =
45struct
46
47(* interface to external solvers *)
48
49local
50
51fun make_command command options problem_path proof_path =
52  Bash.strings (command () @ options) ^ " " ^
53    Bash.string (File.platform_path problem_path) ^
54    " > " ^ File.bash_path proof_path ^ " 2>&1"
55
56fun with_trace ctxt msg f x =
57  let val _ = SMT_Config.trace_msg ctxt (fn () => msg) ()
58  in f x end
59
60fun run ctxt name mk_cmd input =
61  (case SMT_Config.certificates_of ctxt of
62    NONE =>
63      if not (SMT_Config.is_available ctxt name) then
64        error ("The SMT solver " ^ quote name ^ " is not installed")
65      else if Config.get ctxt SMT_Config.debug_files = "" then
66        with_trace ctxt ("Invoking SMT solver " ^ quote name ^ " ...") (Cache_IO.run mk_cmd) input
67      else
68        let
69          val base_path = Path.explode (Config.get ctxt SMT_Config.debug_files)
70          val in_path = Path.ext "smt_in" base_path
71          val out_path = Path.ext "smt_out" base_path
72        in Cache_IO.raw_run mk_cmd input in_path out_path end
73  | SOME certs =>
74      (case Cache_IO.lookup certs input of
75        (NONE, key) =>
76          if Config.get ctxt SMT_Config.read_only_certificates then
77            error ("Bad certificate cache: missing certificate")
78          else
79            Cache_IO.run_and_cache certs key mk_cmd input
80      | (SOME output, _) =>
81          with_trace ctxt ("Using cached certificate from " ^
82            Path.print (Cache_IO.cache_path_of certs) ^ " ...") I output))
83
84(* Z3 returns 1 if "get-proof" or "get-model" fails. veriT returns 255. *)
85val normal_return_codes = [0, 1, 255]
86
87fun run_solver ctxt name mk_cmd input =
88  let
89    fun pretty tag lines = Pretty.string_of (Pretty.big_list tag (map Pretty.str lines))
90
91    val _ = SMT_Config.trace_msg ctxt (pretty "Problem:" o split_lines) input
92
93    val ({elapsed, ...}, {redirected_output = res, output = err, return_code}) =
94      Timing.timing (SMT_Config.with_timeout ctxt (run ctxt name mk_cmd)) input
95    val _ = SMT_Config.trace_msg ctxt (pretty "Solver:") err
96
97    val output = drop_suffix (equal "") res
98    val _ = SMT_Config.trace_msg ctxt (pretty "Result:") output
99    val _ = SMT_Config.trace_msg ctxt (pretty "Time (ms):") [\<^make_string> (Time.toMilliseconds elapsed)]
100    val _ = SMT_Config.statistics_msg ctxt (pretty "Time (ms):") [\<^make_string> (Time.toMilliseconds elapsed)]
101
102    val _ = member (op =) normal_return_codes return_code orelse
103      raise SMT_Failure.SMT (SMT_Failure.Abnormal_Termination return_code)
104  in output end
105
106fun trace_assms ctxt =
107  SMT_Config.trace_msg ctxt (Pretty.string_of o
108    Pretty.big_list "Assertions:" o map (Thm.pretty_thm ctxt o snd))
109
110fun trace_replay_data ({context = ctxt, typs, terms, ...} : SMT_Translate.replay_data) =
111  let
112    fun pretty_eq n p = Pretty.block [Pretty.str n, Pretty.str " = ", p]
113    fun p_typ (n, T) = pretty_eq n (Syntax.pretty_typ ctxt T)
114    fun p_term (n, t) = pretty_eq n (Syntax.pretty_term ctxt t)
115  in
116    SMT_Config.trace_msg ctxt (fn () =>
117      Pretty.string_of (Pretty.big_list "Names:" [
118        Pretty.big_list "sorts:" (map p_typ (Symtab.dest typs)),
119        Pretty.big_list "functions:" (map p_term (Symtab.dest terms))])) ()
120  end
121
122in
123
124fun invoke name command smt_options ithms ctxt =
125  let
126    val options = SMT_Config.solver_options_of ctxt
127    val comments = [space_implode " " options]
128
129    val (str, replay_data as {context = ctxt', ...}) =
130      ithms
131      |> tap (trace_assms ctxt)
132      |> SMT_Translate.translate ctxt smt_options comments
133      ||> tap trace_replay_data
134  in (run_solver ctxt' name (make_command command options) str, replay_data) end
135
136end
137
138
139(* configuration *)
140
141datatype outcome = Unsat | Sat | Unknown | Time_Out
142
143type parsed_proof =
144  {outcome: SMT_Failure.failure option,
145   fact_ids: (int * ((string * ATP_Problem_Generate.stature) * thm)) list option,
146   atp_proof: unit -> (term, string) ATP_Proof.atp_step list}
147
148type solver_config =
149  {name: string,
150   class: Proof.context -> SMT_Util.class,
151   avail: unit -> bool,
152   command: unit -> string list,
153   options: Proof.context -> string list,
154   smt_options: (string * string) list,
155   default_max_relevant: int,
156   outcome: string -> string list -> outcome * string list,
157   parse_proof: (Proof.context -> SMT_Translate.replay_data ->
158     ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->
159     parsed_proof) option,
160   replay: (Proof.context -> SMT_Translate.replay_data -> string list -> thm) option}
161
162
163(* check well-sortedness *)
164
165val has_topsort = Term.exists_type (Term.exists_subtype (fn
166    TFree (_, []) => true
167  | TVar (_, []) => true
168  | _ => false))
169
170(* top sorts cause problems with atomization *)
171fun check_topsort ctxt thm =
172  if has_topsort (Thm.prop_of thm) then (SMT_Normalize.drop_fact_warning ctxt thm; TrueI) else thm
173
174
175(* registry *)
176
177type solver_info = {
178  command: unit -> string list,
179  smt_options: (string * string) list,
180  default_max_relevant: int,
181  parse_proof: Proof.context -> SMT_Translate.replay_data ->
182    ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->
183    parsed_proof,
184  replay: Proof.context -> SMT_Translate.replay_data -> string list -> thm}
185
186structure Solvers = Generic_Data
187(
188  type T = solver_info Symtab.table
189  val empty = Symtab.empty
190  val extend = I
191  fun merge data = Symtab.merge (K true) data
192)
193
194local
195  fun parse_proof outcome parse_proof0 outer_ctxt replay_data xfacts prems concl output =
196    (case outcome output of
197      (Unsat, lines) =>
198        (case parse_proof0 of
199          SOME pp => pp outer_ctxt replay_data xfacts prems concl lines
200        | NONE => {outcome = NONE, fact_ids = NONE, atp_proof = K []})
201    | (Time_Out, _) => raise SMT_Failure.SMT (SMT_Failure.Time_Out)
202    | (result, _) => raise SMT_Failure.SMT (SMT_Failure.Counterexample (result = Sat)))
203
204  fun replay outcome replay0 oracle outer_ctxt
205      (replay_data as {context = ctxt, ...} : SMT_Translate.replay_data) output =
206    (case outcome output of
207      (Unsat, lines) =>
208        if Config.get ctxt SMT_Config.oracle then
209          oracle ()
210        else
211          (case replay0 of
212            SOME replay => replay outer_ctxt replay_data lines
213          | NONE => error "No proof reconstruction for solver -- \
214            \declare [[smt_oracle]] to allow oracle")
215    | (Time_Out, _) => raise SMT_Failure.SMT (SMT_Failure.Time_Out)
216    | (result, _) => raise SMT_Failure.SMT (SMT_Failure.Counterexample (result = Sat)))
217
218  val cfalse = Thm.cterm_of \<^context> \<^prop>\<open>False\<close>
219in
220
221fun add_solver ({name, class, avail, command, options, smt_options, default_max_relevant, outcome,
222    parse_proof = parse_proof0, replay = replay0} : solver_config) =
223  let
224    fun solver oracle = {
225      command = command,
226      smt_options = smt_options,
227      default_max_relevant = default_max_relevant,
228      parse_proof = parse_proof (outcome name) parse_proof0,
229      replay = replay (outcome name) replay0 oracle}
230
231    val info = {name = name, class = class, avail = avail, options = options}
232  in
233    Thm.add_oracle (Binding.name name, K cfalse) #-> (fn (_, oracle) =>
234    Context.theory_map (Solvers.map (Symtab.update_new (name, solver oracle)))) #>
235    Context.theory_map (SMT_Config.add_solver info)
236  end
237
238end
239
240fun get_info ctxt name = the (Symtab.lookup (Solvers.get (Context.Proof ctxt)) name)
241
242fun name_and_info_of ctxt =
243  let val name = SMT_Config.solver_of ctxt
244  in (name, get_info ctxt name) end
245
246val default_max_relevant = #default_max_relevant oo get_info
247
248fun apply_solver_and_replay ctxt thms0 =
249  let
250    val thms = map (check_topsort ctxt) thms0
251    val (name, {command, smt_options, replay, ...}) = name_and_info_of ctxt
252    val (output, replay_data) =
253      invoke name command smt_options (SMT_Normalize.normalize ctxt thms) ctxt
254  in replay ctxt replay_data output end
255
256
257(* filter *)
258
259fun smt_filter ctxt0 goal xfacts i time_limit =
260  let
261    val ctxt = ctxt0 |> Config.put SMT_Config.timeout (Time.toReal time_limit)
262
263    val ({context = ctxt, prems, concl, ...}, _) = Subgoal.focus ctxt i NONE goal
264    fun negate ct = Thm.dest_comb ct ||> Thm.apply \<^cterm>\<open>Not\<close> |-> Thm.apply
265    val cprop =
266      (case try negate (Thm.rhs_of (SMT_Normalize.atomize_conv ctxt concl)) of
267        SOME ct => ct
268      | NONE => raise SMT_Failure.SMT (SMT_Failure.Other_Failure "cannot atomize goal"))
269
270    val conjecture = Thm.assume cprop
271    val facts = map snd xfacts
272    val thms = conjecture :: prems @ facts
273    val thms' = map (check_topsort ctxt) thms
274
275    val (name, {command, smt_options, parse_proof, ...}) = name_and_info_of ctxt
276    val (output, replay_data) =
277      invoke name command smt_options (SMT_Normalize.normalize ctxt thms') ctxt
278  in
279    parse_proof ctxt replay_data xfacts (map Thm.prop_of prems) (Thm.term_of concl) output
280  end
281  handle SMT_Failure.SMT fail => {outcome = SOME fail, fact_ids = NONE, atp_proof = K []}
282
283
284(* SMT tactic *)
285
286local
287  fun str_of ctxt fail =
288    "Solver " ^ SMT_Config.solver_of ctxt ^ ": " ^ SMT_Failure.string_of_failure fail
289
290  fun safe_solve ctxt facts = SOME (apply_solver_and_replay ctxt facts)
291    handle
292      SMT_Failure.SMT (fail as SMT_Failure.Counterexample _) =>
293        (SMT_Config.verbose_msg ctxt (str_of ctxt) fail; NONE)
294    | SMT_Failure.SMT (fail as SMT_Failure.Time_Out) =>
295        error ("SMT: Solver " ^ quote (SMT_Config.solver_of ctxt) ^ ": " ^
296          SMT_Failure.string_of_failure fail ^ " (setting the " ^
297          "configuration option " ^ quote (Config.name_of SMT_Config.timeout) ^ " might help)")
298    | SMT_Failure.SMT fail => error (str_of ctxt fail)
299
300  fun resolve ctxt (SOME thm) = resolve_tac ctxt [thm] 1
301    | resolve _ NONE = no_tac
302
303  fun tac prove ctxt rules =
304    CONVERSION (SMT_Normalize.atomize_conv ctxt)
305    THEN' resolve_tac ctxt @{thms ccontr}
306    THEN' SUBPROOF (fn {context = ctxt', prems, ...} =>
307      resolve ctxt' (prove ctxt' (rules @ prems))) ctxt
308in
309
310val smt_tac = tac safe_solve
311val smt_tac' = tac (SOME oo apply_solver_and_replay)
312
313end
314
315end;
316