src/HOL/Tools/SMT/smt_solver.ML
author blanchet
Thu, 26 Jul 2012 11:08:16 +0200
changeset 49549 2307efbfc554
parent 49547 c0f44941e674
child 49917 44a6967240b7
permissions -rw-r--r--
Z3 prints so many warnings that the very informative abnormal termination exception hardly ever gets raised -- better be more aggressive here
boehmes@36890
     1
(*  Title:      HOL/Tools/SMT/smt_solver.ML
boehmes@36890
     2
    Author:     Sascha Boehme, TU Muenchen
boehmes@36890
     3
boehmes@36890
     4
SMT solvers registry and SMT tactic.
boehmes@36890
     5
*)
boehmes@36890
     6
boehmes@36890
     7
signature SMT_SOLVER =
boehmes@36890
     8
sig
boehmes@40651
     9
  (*configuration*)
boehmes@40403
    10
  datatype outcome = Unsat | Sat | Unknown
boehmes@36890
    11
  type solver_config = {
boehmes@40403
    12
    name: string,
boehmes@49058
    13
    class: Proof.context -> SMT_Utils.class,
boehmes@41680
    14
    avail: unit -> bool,
boehmes@41680
    15
    command: unit -> string list,
boehmes@40403
    16
    options: Proof.context -> string list,
boehmes@41680
    17
    default_max_relevant: int,
boehmes@41680
    18
    supports_filter: bool,
boehmes@40403
    19
    outcome: string -> string list -> outcome * string list,
boehmes@40403
    20
    cex_parser: (Proof.context -> SMT_Translate.recon -> string list ->
boehmes@41076
    21
      term list * term list) option,
boehmes@40403
    22
    reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
boehmes@41680
    23
      int list * thm) option }
boehmes@36890
    24
boehmes@40651
    25
  (*registry*)
boehmes@40403
    26
  val add_solver: solver_config -> theory -> theory
blanchet@41188
    27
  val solver_name_of: Proof.context -> string
blanchet@41188
    28
  val available_solvers_of: Proof.context -> string list
boehmes@41680
    29
  val apply_solver: Proof.context -> (int * (int option * thm)) list ->
boehmes@41680
    30
    int list * thm
blanchet@41227
    31
  val default_max_relevant: Proof.context -> string -> int
boehmes@36890
    32
boehmes@40403
    33
  (*filter*)
boehmes@41680
    34
  type 'a smt_filter_data =
boehmes@41680
    35
    ('a * thm) list * ((int * thm) list * Proof.context)
boehmes@41680
    36
  val smt_filter_preprocess: Proof.state ->
boehmes@41680
    37
    ('a * (int option * thm)) list -> int -> 'a smt_filter_data
boehmes@41680
    38
  val smt_filter_apply: Time.time -> 'a smt_filter_data ->
blanchet@41480
    39
    {outcome: SMT_Failure.failure option, used_facts: ('a * thm) list}
boehmes@40402
    40
boehmes@36890
    41
  (*tactic*)
wenzelm@47340
    42
  val smt_tac: Proof.context -> thm list -> int -> tactic
wenzelm@47340
    43
  val smt_tac': Proof.context -> thm list -> int -> tactic
boehmes@36890
    44
end
boehmes@36890
    45
boehmes@36890
    46
structure SMT_Solver: SMT_SOLVER =
boehmes@36890
    47
struct
boehmes@36890
    48
boehmes@40403
    49
boehmes@36890
    50
(* interface to external solvers *)
boehmes@36890
    51
boehmes@36890
    52
local
boehmes@36890
    53
boehmes@41680
    54
fun make_cmd command options problem_path proof_path = space_implode " " (
blanchet@49547
    55
  "(exec 2>&1;" :: map File.shell_quote (command () @ options) @
blanchet@49547
    56
  [File.shell_path problem_path, ")", ">", File.shell_path proof_path])
boehmes@36890
    57
boehmes@41680
    58
fun trace_and ctxt msg f x =
boehmes@41680
    59
  let val _ = SMT_Config.trace_msg ctxt (fn () => msg) ()
boehmes@41680
    60
  in f x end
boehmes@41680
    61
boehmes@41680
    62
fun run ctxt name mk_cmd input =
boehmes@41576
    63
  (case SMT_Config.certificates_of ctxt of
boehmes@40816
    64
    NONE =>
boehmes@47608
    65
      if not (SMT_Config.is_available ctxt name) then
boehmes@47608
    66
        error ("The SMT solver " ^ quote name ^ " is not installed.")
boehmes@47608
    67
      else if Config.get ctxt SMT_Config.debug_files = "" then
boehmes@41680
    68
        trace_and ctxt ("Invoking SMT solver " ^ quote name ^ " ...")
boehmes@41680
    69
          (Cache_IO.run mk_cmd) input
boehmes@40816
    70
      else
boehmes@40816
    71
        let
boehmes@41576
    72
          val base_path = Path.explode (Config.get ctxt SMT_Config.debug_files)
boehmes@40816
    73
          val in_path = Path.ext "smt_in" base_path
boehmes@40816
    74
          val out_path = Path.ext "smt_out" base_path
boehmes@41680
    75
        in Cache_IO.raw_run mk_cmd input in_path out_path end
boehmes@36890
    76
  | SOME certs =>
boehmes@36890
    77
      (case Cache_IO.lookup certs input of
boehmes@36890
    78
        (NONE, key) =>
blanchet@48023
    79
          if Config.get ctxt SMT_Config.read_only_certificates then
blanchet@48023
    80
            error ("Bad certificate cache: missing certificate")
boehmes@40781
    81
          else
boehmes@41680
    82
            Cache_IO.run_and_cache certs key mk_cmd input
boehmes@36890
    83
      | (SOME output, _) =>
boehmes@41680
    84
          trace_and ctxt ("Using cached certificate from " ^
boehmes@41680
    85
            File.shell_path (Cache_IO.cache_path_of certs) ^ " ...")
boehmes@41680
    86
            I output))
boehmes@36890
    87
boehmes@41680
    88
fun run_solver ctxt name mk_cmd input =
boehmes@36890
    89
  let
boehmes@36890
    90
    fun pretty tag ls = Pretty.string_of (Pretty.big_list tag
boehmes@36890
    91
      (map Pretty.str ls))
boehmes@36890
    92
boehmes@41576
    93
    val _ = SMT_Config.trace_msg ctxt (pretty "Problem:" o split_lines) input
boehmes@36890
    94
boehmes@40792
    95
    val {redirected_output=res, output=err, return_code} =
boehmes@41680
    96
      SMT_Config.with_timeout ctxt (run ctxt name mk_cmd) input
boehmes@41576
    97
    val _ = SMT_Config.trace_msg ctxt (pretty "Solver:") err
boehmes@36890
    98
haftmann@40057
    99
    val ls = rev (snd (chop_while (equal "") (rev res)))
boehmes@41576
   100
    val _ = SMT_Config.trace_msg ctxt (pretty "Result:") ls
boehmes@40792
   101
blanchet@49549
   102
    val _ = return_code <> 0 andalso
boehmes@40802
   103
      raise SMT_Failure.SMT (SMT_Failure.Abnormal_Termination return_code)
boehmes@36890
   104
  in ls end
boehmes@36890
   105
boehmes@41576
   106
fun trace_assms ctxt =
boehmes@41576
   107
  SMT_Config.trace_msg ctxt (Pretty.string_of o
boehmes@41576
   108
    Pretty.big_list "Assertions:" o map (Display.pretty_thm ctxt o snd))
boehmes@40439
   109
boehmes@41375
   110
fun trace_recon_data ({context=ctxt, typs, terms, ...} : SMT_Translate.recon) =
boehmes@36890
   111
  let
boehmes@36890
   112
    fun pretty_eq n p = Pretty.block [Pretty.str n, Pretty.str " = ", p]
boehmes@40651
   113
    fun p_typ (n, T) = pretty_eq n (Syntax.pretty_typ ctxt T)
boehmes@40651
   114
    fun p_term (n, t) = pretty_eq n (Syntax.pretty_term ctxt t)
boehmes@36890
   115
  in
boehmes@41576
   116
    SMT_Config.trace_msg ctxt (fn () =>
boehmes@40651
   117
      Pretty.string_of (Pretty.big_list "Names:" [
boehmes@40651
   118
        Pretty.big_list "sorts:" (map p_typ (Symtab.dest typs)),
boehmes@40651
   119
        Pretty.big_list "functions:" (map p_term (Symtab.dest terms))])) ()
boehmes@36890
   120
  end
boehmes@36890
   121
boehmes@41680
   122
in
boehmes@41680
   123
boehmes@41680
   124
fun invoke name command ithms ctxt =
boehmes@40403
   125
  let
boehmes@41680
   126
    val options = SMT_Config.solver_options_of ctxt
boehmes@40403
   127
    val comments = ("solver: " ^ name) ::
boehmes@41576
   128
      ("timeout: " ^ string_of_real (Config.get ctxt SMT_Config.timeout)) ::
boehmes@41576
   129
      ("random seed: " ^
boehmes@41576
   130
        string_of_int (Config.get ctxt SMT_Config.random_seed)) ::
boehmes@41680
   131
      "arguments:" :: options
boehmes@36890
   132
boehmes@41375
   133
    val (str, recon as {context=ctxt', ...}) =
boehmes@41375
   134
      ithms
boehmes@41375
   135
      |> tap (trace_assms ctxt)
boehmes@41375
   136
      |> SMT_Translate.translate ctxt comments
boehmes@41375
   137
      ||> tap trace_recon_data
boehmes@41846
   138
  in (run_solver ctxt' name (make_cmd command options) str, recon) end
boehmes@41374
   139
boehmes@41680
   140
end
boehmes@40405
   141
blanchet@41289
   142
boehmes@41680
   143
(* configuration *)
boehmes@41680
   144
boehmes@41680
   145
datatype outcome = Unsat | Sat | Unknown
boehmes@41680
   146
boehmes@41680
   147
type solver_config = {
boehmes@41680
   148
  name: string,
boehmes@49058
   149
  class: Proof.context -> SMT_Utils.class,
boehmes@41680
   150
  avail: unit -> bool,
boehmes@41680
   151
  command: unit -> string list,
boehmes@41680
   152
  options: Proof.context -> string list,
boehmes@41680
   153
  default_max_relevant: int,
boehmes@41680
   154
  supports_filter: bool,
boehmes@41680
   155
  outcome: string -> string list -> outcome * string list,
boehmes@41680
   156
  cex_parser: (Proof.context -> SMT_Translate.recon -> string list ->
boehmes@41680
   157
    term list * term list) option,
boehmes@41680
   158
  reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
boehmes@41680
   159
    int list * thm) option }
boehmes@41680
   160
blanchet@41289
   161
blanchet@41289
   162
(* registry *)
blanchet@41289
   163
blanchet@41289
   164
type solver_info = {
boehmes@41680
   165
  command: unit -> string list,
boehmes@41680
   166
  default_max_relevant: int,
boehmes@41680
   167
  supports_filter: bool,
boehmes@41375
   168
  reconstruct: Proof.context -> string list * SMT_Translate.recon ->
boehmes@41680
   169
    int list * thm }
boehmes@36890
   170
boehmes@36891
   171
structure Solvers = Generic_Data
boehmes@36890
   172
(
boehmes@40403
   173
  type T = solver_info Symtab.table
boehmes@36890
   174
  val empty = Symtab.empty
boehmes@36890
   175
  val extend = I
boehmes@36890
   176
  fun merge data = Symtab.merge (K true) data
boehmes@36890
   177
)
boehmes@36890
   178
boehmes@40403
   179
local
boehmes@41375
   180
  fun finish outcome cex_parser reconstruct ocl outer_ctxt
boehmes@41375
   181
      (output, (recon as {context=ctxt, ...} : SMT_Translate.recon)) =
boehmes@40403
   182
    (case outcome output of
boehmes@40403
   183
      (Unsat, ls) =>
boehmes@41576
   184
        if not (Config.get ctxt SMT_Config.oracle) andalso is_some reconstruct
boehmes@41375
   185
        then the reconstruct outer_ctxt recon ls
boehmes@41375
   186
        else ([], ocl ())
boehmes@40403
   187
    | (result, ls) =>
boehmes@41076
   188
        let
boehmes@41076
   189
          val (ts, us) =
boehmes@41076
   190
            (case cex_parser of SOME f => f ctxt recon ls | _ => ([], []))
boehmes@41076
   191
         in
boehmes@41076
   192
          raise SMT_Failure.SMT (SMT_Failure.Counterexample {
boehmes@41076
   193
            is_real_cex = (result = Sat),
boehmes@41076
   194
            free_constraints = ts,
boehmes@41076
   195
            const_defs = us})
boehmes@40651
   196
        end)
boehmes@40817
   197
boehmes@40817
   198
  val cfalse = Thm.cterm_of @{theory} (@{const Trueprop} $ @{const False})
boehmes@40403
   199
in
boehmes@40403
   200
boehmes@40403
   201
fun add_solver cfg =
boehmes@40403
   202
  let
boehmes@41680
   203
    val {name, class, avail, command, options, default_max_relevant,
boehmes@41680
   204
      supports_filter, outcome, cex_parser, reconstruct} = cfg
boehmes@40403
   205
boehmes@40817
   206
    fun core_oracle () = cfalse
boehmes@40403
   207
boehmes@41680
   208
    fun solver ocl = {
boehmes@41680
   209
      command = command,
boehmes@41680
   210
      default_max_relevant = default_max_relevant,
boehmes@41680
   211
      supports_filter = supports_filter,
boehmes@41680
   212
      reconstruct = finish (outcome name) cex_parser reconstruct ocl }
boehmes@41680
   213
boehmes@41680
   214
    val info = {name=name, class=class, avail=avail, options=options}
boehmes@40403
   215
  in
boehmes@40403
   216
    Thm.add_oracle (Binding.name name, core_oracle) #-> (fn (_, ocl) =>
boehmes@40403
   217
    Context.theory_map (Solvers.map (Symtab.update_new (name, solver ocl)))) #>
boehmes@41680
   218
    Context.theory_map (SMT_Config.add_solver info)
boehmes@40403
   219
  end
boehmes@40403
   220
boehmes@40403
   221
end
boehmes@40403
   222
blanchet@41227
   223
fun get_info ctxt name =
blanchet@41227
   224
  the (Symtab.lookup (Solvers.get (Context.Proof ctxt)) name)
blanchet@41227
   225
boehmes@41680
   226
val solver_name_of = SMT_Config.solver_of
boehmes@41680
   227
boehmes@41680
   228
val available_solvers_of = SMT_Config.available_solvers_of
boehmes@41680
   229
boehmes@41680
   230
fun name_and_info_of ctxt =
boehmes@41680
   231
  let val name = solver_name_of ctxt
blanchet@41227
   232
  in (name, get_info ctxt name) end
boehmes@36890
   233
boehmes@41680
   234
fun gen_preprocess ctxt iwthms = SMT_Normalize.normalize iwthms ctxt
boehmes@36890
   235
boehmes@41680
   236
fun gen_apply (ithms, ctxt) =
boehmes@41680
   237
  let val (name, {command, reconstruct, ...}) = name_and_info_of ctxt
boehmes@41680
   238
  in
boehmes@47608
   239
    (ithms, ctxt)
boehmes@47608
   240
    |-> invoke name command
boehmes@47608
   241
    |> reconstruct ctxt
boehmes@47608
   242
    |>> distinct (op =)
boehmes@41680
   243
  end
blanchet@41188
   244
boehmes@41680
   245
fun apply_solver ctxt = gen_apply o gen_preprocess ctxt
blanchet@41227
   246
blanchet@41227
   247
val default_max_relevant = #default_max_relevant oo get_info
boehmes@36890
   248
boehmes@41680
   249
val supports_filter = #supports_filter o snd o name_and_info_of 
boehmes@36890
   250
boehmes@41680
   251
boehmes@41680
   252
(* check well-sortedness *)
boehmes@40402
   253
boehmes@40402
   254
val has_topsort = Term.exists_type (Term.exists_subtype (fn
boehmes@40402
   255
    TFree (_, []) => true
boehmes@40402
   256
  | TVar (_, []) => true
boehmes@40402
   257
  | _ => false))
boehmes@40402
   258
blanchet@41480
   259
(* without this test, we would run into problems when atomizing the rules: *)
blanchet@41480
   260
fun check_topsort iwthms =
boehmes@41374
   261
  if exists (has_topsort o Thm.prop_of o snd o snd) iwthms then
boehmes@40651
   262
    raise SMT_Failure.SMT (SMT_Failure.Other_Failure ("proof state " ^
boehmes@40651
   263
      "contains the universal sort {}"))
boehmes@41680
   264
  else ()
boehmes@41680
   265
boehmes@41680
   266
boehmes@41680
   267
(* filter *)
boehmes@40403
   268
boehmes@40817
   269
val cnot = Thm.cterm_of @{theory} @{const Not}
boehmes@40817
   270
blanchet@41480
   271
fun mk_result outcome xrules = { outcome = outcome, used_facts = xrules }
boehmes@41374
   272
boehmes@41680
   273
type 'a smt_filter_data = ('a * thm) list * ((int * thm) list * Proof.context)
blanchet@41480
   274
boehmes@41680
   275
fun smt_filter_preprocess st xwthms i =
boehmes@40403
   276
  let
boehmes@40405
   277
    val ctxt =
boehmes@40440
   278
      Proof.context_of st
boehmes@41576
   279
      |> Config.put SMT_Config.oracle false
boehmes@41576
   280
      |> Config.put SMT_Config.drop_bad_facts true
boehmes@41576
   281
      |> Config.put SMT_Config.filter_only_facts true
boehmes@41374
   282
boehmes@41374
   283
    val {facts, goal, ...} = Proof.goal st
boehmes@40603
   284
    val ({context=ctxt', prems, concl, ...}, _) = Subgoal.focus ctxt i goal
wenzelm@47368
   285
    fun negate ct = Thm.dest_comb ct ||> Thm.apply cnot |-> Thm.apply
boehmes@43180
   286
    val cprop =
boehmes@43180
   287
      (case try negate (Thm.rhs_of (SMT_Normalize.atomize_conv ctxt' concl)) of
boehmes@43180
   288
        SOME ct => ct
boehmes@43180
   289
      | NONE => raise SMT_Failure.SMT (SMT_Failure.Other_Failure (
boehmes@43180
   290
          "goal is not a HOL term")))
boehmes@40403
   291
  in
boehmes@41680
   292
    map snd xwthms
boehmes@41680
   293
    |> map_index I
boehmes@41680
   294
    |> append (map (pair ~1 o pair NONE) (Thm.assume cprop :: prems @ facts))
boehmes@41680
   295
    |> tap check_topsort
boehmes@41680
   296
    |> gen_preprocess ctxt'
boehmes@41680
   297
    |> pair (map (apsnd snd) xwthms)
blanchet@41480
   298
  end
blanchet@41480
   299
boehmes@41680
   300
fun smt_filter_apply time_limit (xthms, (ithms, ctxt)) =
blanchet@41482
   301
  let
boehmes@41680
   302
    val ctxt' =
boehmes@41680
   303
      ctxt
boehmes@41680
   304
      |> Config.put SMT_Config.timeout (Time.toReal time_limit)
boehmes@41680
   305
boehmes@41680
   306
    fun filter_thms false = K xthms
boehmes@41680
   307
      | filter_thms true = map_filter (try (nth xthms)) o fst
blanchet@41482
   308
  in
boehmes@41680
   309
    (ithms, ctxt')
boehmes@41680
   310
    |> gen_apply
boehmes@41680
   311
    |> filter_thms (supports_filter ctxt')
boehmes@41374
   312
    |> mk_result NONE
boehmes@40403
   313
  end
boehmes@41374
   314
  handle SMT_Failure.SMT fail => mk_result (SOME fail) []
boehmes@40402
   315
boehmes@40402
   316
boehmes@36890
   317
(* SMT tactic *)
boehmes@36890
   318
boehmes@41680
   319
local
boehmes@41680
   320
  fun trace_assumptions ctxt iwthms idxs =
boehmes@40406
   321
    let
boehmes@41680
   322
      val wthms =
boehmes@41680
   323
        idxs
boehmes@41680
   324
        |> filter (fn i => i >= 0)
boehmes@41680
   325
        |> map_filter (AList.lookup (op =) iwthms)
boehmes@40406
   326
    in
boehmes@41680
   327
      if Config.get ctxt SMT_Config.trace_used_facts andalso length wthms > 0
boehmes@41680
   328
      then
boehmes@41680
   329
        tracing (Pretty.string_of (Pretty.big_list "SMT used facts:"
boehmes@41680
   330
          (map (Display.pretty_thm ctxt o snd) wthms)))
boehmes@41680
   331
      else ()
boehmes@41680
   332
    end
boehmes@36890
   333
boehmes@41680
   334
  fun solve ctxt iwthms =
boehmes@41680
   335
    iwthms
boehmes@41680
   336
    |> tap check_topsort
boehmes@41680
   337
    |> apply_solver ctxt
boehmes@41680
   338
    |>> trace_assumptions ctxt iwthms
boehmes@41680
   339
    |> snd
boehmes@41680
   340
boehmes@41680
   341
  fun str_of ctxt fail =
boehmes@41680
   342
    SMT_Failure.string_of_failure ctxt fail
boehmes@41680
   343
    |> prefix ("Solver " ^ SMT_Config.solver_of ctxt ^ ": ")
boehmes@41680
   344
boehmes@41680
   345
  fun safe_solve ctxt iwthms = SOME (solve ctxt iwthms)
boehmes@41680
   346
    handle
boehmes@41680
   347
      SMT_Failure.SMT (fail as SMT_Failure.Counterexample _) =>
boehmes@41680
   348
        (SMT_Config.verbose_msg ctxt (str_of ctxt) fail; NONE)
boehmes@42632
   349
    | SMT_Failure.SMT (fail as SMT_Failure.Time_Out) =>
boehmes@42632
   350
        error ("SMT: Solver " ^ quote (SMT_Config.solver_of ctxt) ^ ": " ^
boehmes@42632
   351
          SMT_Failure.string_of_failure ctxt fail ^ " (setting the " ^
wenzelm@43487
   352
          "configuration option " ^ quote (Config.name_of SMT_Config.timeout) ^ " might help)")
boehmes@42632
   353
    | SMT_Failure.SMT fail => error (str_of ctxt fail)
boehmes@41680
   354
boehmes@41680
   355
  fun tag_rules thms = map_index (apsnd (pair NONE)) thms
boehmes@41680
   356
  fun tag_prems thms = map (pair ~1 o pair NONE) thms
boehmes@41680
   357
boehmes@41680
   358
  fun resolve (SOME thm) = Tactic.rtac thm 1
wenzelm@47340
   359
    | resolve NONE = no_tac
boehmes@41680
   360
boehmes@41680
   361
  fun tac prove ctxt rules =
boehmes@41680
   362
    CONVERSION (SMT_Normalize.atomize_conv ctxt)
boehmes@41680
   363
    THEN' Tactic.rtac @{thm ccontr}
boehmes@41680
   364
    THEN' SUBPROOF (fn {context, prems, ...} =>
boehmes@41680
   365
      resolve (prove context (tag_rules rules @ tag_prems prems))) ctxt
boehmes@41680
   366
in
boehmes@41680
   367
boehmes@41680
   368
val smt_tac = tac safe_solve
boehmes@41680
   369
val smt_tac' = tac (SOME oo solve)
boehmes@41680
   370
boehmes@41680
   371
end
boehmes@41680
   372
boehmes@36890
   373
end