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