src/HOL/Tools/SMT/smt_solver.ML
author boehmes
Thu, 04 Nov 2010 18:01:55 +0100
changeset 40603 82ebdd19c4a4
parent 40578 5edeb5d269fa
child 40651 7550b2cba1cb
permissions -rw-r--r--
simulate more closely the behaviour of the tactic
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@40403
     9
  datatype failure =
boehmes@40403
    10
    Counterexample of bool * term list |
boehmes@40403
    11
    Time_Out |
boehmes@40438
    12
    Out_Of_Memory |
boehmes@40403
    13
    Other_Failure of string
boehmes@40438
    14
  val string_of_failure: Proof.context -> string -> failure -> string
boehmes@40403
    15
  exception SMT of failure
boehmes@36890
    16
boehmes@36890
    17
  type interface = {
boehmes@36890
    18
    extra_norm: SMT_Normalize.extra_norm,
boehmes@36890
    19
    translate: SMT_Translate.config }
boehmes@40403
    20
  datatype outcome = Unsat | Sat | Unknown
boehmes@36890
    21
  type solver_config = {
boehmes@40403
    22
    name: string,
boehmes@40403
    23
    env_var: string,
boehmes@40403
    24
    is_remote: bool,
boehmes@40403
    25
    options: Proof.context -> string list,
boehmes@36890
    26
    interface: interface,
boehmes@40403
    27
    outcome: string -> string list -> outcome * string list,
boehmes@40403
    28
    cex_parser: (Proof.context -> SMT_Translate.recon -> string list ->
boehmes@40403
    29
      term list) option,
boehmes@40403
    30
    reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
boehmes@40403
    31
      (int list * thm) * Proof.context) option }
boehmes@36890
    32
boehmes@36890
    33
  (*options*)
boehmes@40403
    34
  val oracle: bool Config.T
boehmes@40405
    35
  val filter_only: bool Config.T
boehmes@40403
    36
  val datatypes: bool Config.T
boehmes@40524
    37
  val keep_assms: bool Config.T
boehmes@40578
    38
  val timeout: real Config.T
boehmes@36890
    39
  val with_timeout: Proof.context -> ('a -> 'b) -> 'a -> 'b
boehmes@40522
    40
  val traceN: string
boehmes@36890
    41
  val trace: bool Config.T
boehmes@36890
    42
  val trace_msg: Proof.context -> ('a -> string) -> 'a -> unit
boehmes@40403
    43
  val trace_used_facts: bool Config.T
boehmes@36890
    44
boehmes@36890
    45
  (*certificates*)
boehmes@36890
    46
  val fixed_certificates: bool Config.T
boehmes@36890
    47
  val select_certificates: string -> Context.generic -> Context.generic
boehmes@36890
    48
boehmes@36890
    49
  (*solvers*)
boehmes@40437
    50
  type solver = bool option -> Proof.context -> (int * thm) list ->
boehmes@40437
    51
    int list * thm
boehmes@40403
    52
  val add_solver: solver_config -> theory -> theory
boehmes@40403
    53
  val set_solver_options: string -> string -> Context.generic ->
boehmes@40403
    54
    Context.generic
boehmes@36891
    55
  val all_solver_names_of: Context.generic -> string list
boehmes@36890
    56
  val solver_name_of: Context.generic -> string
boehmes@36890
    57
  val select_solver: string -> Context.generic -> Context.generic
boehmes@36890
    58
  val solver_of: Context.generic -> solver
boehmes@40407
    59
  val is_locally_installed: Proof.context -> bool
boehmes@36890
    60
boehmes@40403
    61
  (*filter*)
boehmes@40405
    62
  val smt_filter: bool -> Time.time -> Proof.state -> ('a * thm) list -> int ->
boehmes@40438
    63
    {outcome: failure option, used_facts: 'a list,
boehmes@40438
    64
    run_time_in_msecs: int option}
boehmes@40402
    65
boehmes@36890
    66
  (*tactic*)
boehmes@36890
    67
  val smt_tac': bool -> Proof.context -> thm list -> int -> Tactical.tactic
boehmes@36890
    68
  val smt_tac: Proof.context -> thm list -> int -> Tactical.tactic
boehmes@36890
    69
boehmes@36890
    70
  (*setup*)
boehmes@36890
    71
  val setup: theory -> theory
boehmes@36890
    72
  val print_setup: Context.generic -> unit
boehmes@36890
    73
end
boehmes@36890
    74
boehmes@36890
    75
structure SMT_Solver: SMT_SOLVER =
boehmes@36890
    76
struct
boehmes@36890
    77
boehmes@40403
    78
datatype failure =
boehmes@40403
    79
  Counterexample of bool * term list |
boehmes@40403
    80
  Time_Out |
boehmes@40438
    81
  Out_Of_Memory |
boehmes@40403
    82
  Other_Failure of string
boehmes@36890
    83
boehmes@40438
    84
fun string_of_failure ctxt _ (Counterexample (real, ex)) =
boehmes@40403
    85
      let
boehmes@40406
    86
        val msg = (if real then "C" else "Potential c") ^ "ounterexample found"
boehmes@40403
    87
      in
boehmes@40403
    88
        if null ex then msg ^ "."
boehmes@40403
    89
        else Pretty.string_of (Pretty.big_list (msg ^ ":")
boehmes@40403
    90
          (map (Syntax.pretty_term ctxt) ex))
boehmes@40403
    91
      end
boehmes@40438
    92
  | string_of_failure _ name Time_Out = name ^ " timed out."
boehmes@40438
    93
  | string_of_failure _ name Out_Of_Memory = name ^ " ran out of memory."
boehmes@40438
    94
  | string_of_failure _ _ (Other_Failure msg) = msg
boehmes@40403
    95
boehmes@40403
    96
exception SMT of failure
boehmes@36890
    97
boehmes@36890
    98
type interface = {
boehmes@36890
    99
  extra_norm: SMT_Normalize.extra_norm,
boehmes@36890
   100
  translate: SMT_Translate.config }
boehmes@36890
   101
boehmes@40403
   102
datatype outcome = Unsat | Sat | Unknown
boehmes@40403
   103
boehmes@36890
   104
type solver_config = {
boehmes@40403
   105
  name: string,
boehmes@40403
   106
  env_var: string,
boehmes@40403
   107
  is_remote: bool,
boehmes@40403
   108
  options: Proof.context -> string list,
boehmes@36890
   109
  interface: interface,
boehmes@40403
   110
  outcome: string -> string list -> outcome * string list,
boehmes@40403
   111
  cex_parser: (Proof.context -> SMT_Translate.recon -> string list ->
boehmes@40403
   112
    term list) option,
boehmes@40403
   113
  reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
boehmes@40403
   114
    (int list * thm) * Proof.context) option }
boehmes@36890
   115
boehmes@36890
   116
boehmes@36890
   117
boehmes@36890
   118
(* SMT options *)
boehmes@36890
   119
boehmes@40403
   120
val (oracle, setup_oracle) = Attrib.config_bool "smt_oracle" (K true)
boehmes@40403
   121
boehmes@40405
   122
val (filter_only, setup_filter_only) =
boehmes@40405
   123
  Attrib.config_bool "smt_filter_only" (K false)
boehmes@40405
   124
boehmes@40403
   125
val (datatypes, setup_datatypes) = Attrib.config_bool "smt_datatypes" (K false)
boehmes@40403
   126
boehmes@40524
   127
val (keep_assms, setup_keep_assms) =
boehmes@40524
   128
  Attrib.config_bool "smt_keep_assms" (K true)
boehmes@40524
   129
boehmes@40578
   130
val (timeout, setup_timeout) = Attrib.config_real "smt_timeout" (K 30.0)
boehmes@36890
   131
boehmes@36890
   132
fun with_timeout ctxt f x =
boehmes@40578
   133
  TimeLimit.timeLimit (seconds (Config.get ctxt timeout)) f x
boehmes@40403
   134
  handle TimeLimit.TimeOut => raise SMT Time_Out
boehmes@36890
   135
boehmes@40522
   136
val traceN = "smt_trace"
boehmes@40522
   137
val (trace, setup_trace) = Attrib.config_bool traceN (K false)
boehmes@36890
   138
boehmes@36890
   139
fun trace_msg ctxt f x =
boehmes@36890
   140
  if Config.get ctxt trace then tracing (f x) else ()
boehmes@36890
   141
boehmes@40403
   142
val (trace_used_facts, setup_trace_used_facts) =
boehmes@40403
   143
  Attrib.config_bool "smt_trace_used_facts" (K false)
boehmes@36890
   144
boehmes@36890
   145
boehmes@36890
   146
(* SMT certificates *)
boehmes@36890
   147
boehmes@36890
   148
val (fixed_certificates, setup_fixed_certificates) =
boehmes@36890
   149
  Attrib.config_bool "smt_fixed" (K false)
boehmes@36890
   150
boehmes@36890
   151
structure Certificates = Generic_Data
boehmes@36890
   152
(
boehmes@36890
   153
  type T = Cache_IO.cache option
boehmes@36890
   154
  val empty = NONE
boehmes@36890
   155
  val extend = I
boehmes@36890
   156
  fun merge (s, _) = s
boehmes@36890
   157
)
boehmes@36890
   158
boehmes@36890
   159
val get_certificates_path =
boehmes@36890
   160
  Option.map (Cache_IO.cache_path_of) o Certificates.get
boehmes@36890
   161
boehmes@36890
   162
fun select_certificates name = Certificates.put (
boehmes@36890
   163
  if name = "" then NONE
boehmes@36890
   164
  else SOME (Cache_IO.make (Path.explode name)))
boehmes@36890
   165
boehmes@36890
   166
boehmes@36890
   167
boehmes@36890
   168
(* interface to external solvers *)
boehmes@36890
   169
boehmes@40407
   170
fun get_local_solver env_var =
boehmes@40407
   171
  let val local_solver = getenv env_var
boehmes@40407
   172
  in if local_solver <> "" then SOME local_solver else NONE end
boehmes@40407
   173
boehmes@36890
   174
local
boehmes@36890
   175
boehmes@40437
   176
fun choose (rm, env_var, is_remote, name) =
boehmes@36890
   177
  let
boehmes@40437
   178
    val force_local = (case rm of SOME false => true | _ => false)
boehmes@40437
   179
    val force_remote = (case rm of SOME true => true | _ => false)
boehmes@40407
   180
    val lsolver = get_local_solver env_var
boehmes@36890
   181
    val remote_url = getenv "REMOTE_SMT_URL"
boehmes@40437
   182
    val trace = if is_some rm then K () else tracing
boehmes@36890
   183
  in
boehmes@40407
   184
    if not force_remote andalso is_some lsolver
boehmes@36890
   185
    then 
boehmes@40437
   186
     (trace ("Invoking local SMT solver " ^ quote (the lsolver) ^ " ...");
boehmes@40407
   187
      [the lsolver])
boehmes@40437
   188
    else if not force_local andalso is_remote
boehmes@36890
   189
    then
boehmes@40437
   190
     (trace ("Invoking remote SMT solver " ^ quote name ^ " at " ^
boehmes@36890
   191
        quote remote_url ^ " ...");
boehmes@40407
   192
      [getenv "REMOTE_SMT", name])
boehmes@40407
   193
    else if force_remote
boehmes@40437
   194
    then error ("The SMT solver " ^ quote name ^ " is not remotely available.")
boehmes@40437
   195
    else error ("The SMT solver " ^ quote name ^ " has not been found " ^
boehmes@40437
   196
      "on this computer. Please set the Isabelle environment variable " ^
boehmes@40437
   197
      quote env_var ^ ".")
boehmes@36890
   198
  end
boehmes@36890
   199
boehmes@36890
   200
fun make_cmd solver args problem_path proof_path = space_implode " " (
boehmes@36890
   201
  map File.shell_quote (solver @ args) @
boehmes@36890
   202
  [File.shell_path problem_path, "2>&1", ">", File.shell_path proof_path])
boehmes@36890
   203
boehmes@36890
   204
fun run ctxt cmd args input =
boehmes@36890
   205
  (case Certificates.get (Context.Proof ctxt) of
boehmes@36890
   206
    NONE => Cache_IO.run (make_cmd (choose cmd) args) input
boehmes@36890
   207
  | SOME certs =>
boehmes@36890
   208
      (case Cache_IO.lookup certs input of
boehmes@36890
   209
        (NONE, key) =>
boehmes@36890
   210
          if Config.get ctxt fixed_certificates
boehmes@36890
   211
          then error ("Bad certificates cache: missing certificate")
boehmes@36890
   212
          else Cache_IO.run_and_cache certs key (make_cmd (choose cmd) args)
boehmes@36890
   213
            input
boehmes@36890
   214
      | (SOME output, _) =>
boehmes@36890
   215
         (tracing ("Using cached certificate from " ^
boehmes@36890
   216
            File.shell_path (Cache_IO.cache_path_of certs) ^ " ...");
boehmes@36890
   217
          output)))
boehmes@36890
   218
boehmes@36890
   219
in
boehmes@36890
   220
boehmes@36890
   221
fun run_solver ctxt cmd args input =
boehmes@36890
   222
  let
boehmes@36890
   223
    fun pretty tag ls = Pretty.string_of (Pretty.big_list tag
boehmes@36890
   224
      (map Pretty.str ls))
boehmes@36890
   225
boehmes@36890
   226
    val _ = trace_msg ctxt (pretty "SMT problem:" o split_lines) input
boehmes@36890
   227
boehmes@36890
   228
    val (res, err) = with_timeout ctxt (run ctxt cmd args) input
boehmes@36890
   229
    val _ = trace_msg ctxt (pretty "SMT solver:") err
boehmes@36890
   230
haftmann@40057
   231
    val ls = rev (snd (chop_while (equal "") (rev res)))
boehmes@36890
   232
    val _ = trace_msg ctxt (pretty "SMT result:") ls
boehmes@36890
   233
  in ls end
boehmes@36890
   234
boehmes@36890
   235
end
boehmes@36890
   236
boehmes@40439
   237
fun trace_assms ctxt = trace_msg ctxt (Pretty.string_of o
boehmes@40439
   238
  Pretty.big_list "SMT assertions:" o map (Display.pretty_thm ctxt o snd))
boehmes@40439
   239
blanchet@36933
   240
fun trace_recon_data ctxt ({typs, terms, ...} : SMT_Translate.recon) =
boehmes@36890
   241
  let
boehmes@36890
   242
    fun pretty_eq n p = Pretty.block [Pretty.str n, Pretty.str " = ", p]
boehmes@36890
   243
    fun pretty_typ (n, T) = pretty_eq n (Syntax.pretty_typ ctxt T)
boehmes@36890
   244
    fun pretty_term (n, t) = pretty_eq n (Syntax.pretty_term ctxt t)
boehmes@36890
   245
  in
boehmes@36890
   246
    trace_msg ctxt (fn () => Pretty.string_of (Pretty.big_list "SMT names:" [
boehmes@36890
   247
      Pretty.big_list "sorts:" (map pretty_typ (Symtab.dest typs)),
boehmes@36890
   248
      Pretty.big_list "functions:" (map pretty_term (Symtab.dest terms))])) ()
boehmes@36890
   249
  end
boehmes@36890
   250
boehmes@40403
   251
fun invoke translate_config name cmd more_opts options irules ctxt =
boehmes@40403
   252
  let
boehmes@40403
   253
    val args = more_opts @ options ctxt
boehmes@40403
   254
    val comments = ("solver: " ^ name) ::
boehmes@40578
   255
      ("timeout: " ^ Time.toString (seconds (Config.get ctxt timeout))) ::
boehmes@40403
   256
      "arguments:" :: args
boehmes@40403
   257
  in
boehmes@40403
   258
    irules
boehmes@40439
   259
    |> tap (trace_assms ctxt)
boehmes@40403
   260
    |> SMT_Translate.translate translate_config ctxt comments
boehmes@40403
   261
    ||> tap (trace_recon_data ctxt)
boehmes@40403
   262
    |>> run_solver ctxt cmd args
boehmes@40403
   263
    |> rpair ctxt
boehmes@40403
   264
  end
boehmes@36890
   265
boehmes@36890
   266
fun discharge_definitions thm =
boehmes@36890
   267
  if Thm.nprems_of thm = 0 then thm
boehmes@36890
   268
  else discharge_definitions (@{thm reflexive} RS thm)
boehmes@36890
   269
boehmes@40403
   270
fun set_has_datatypes with_datatypes translate =
boehmes@36890
   271
  let
boehmes@40403
   272
    val {prefixes, header, strict, builtins, serialize} = translate
boehmes@40403
   273
    val {builtin_typ, builtin_num, builtin_fun, has_datatypes} = builtins
boehmes@40403
   274
    val with_datatypes' = has_datatypes andalso with_datatypes
boehmes@40403
   275
    val builtins' = {builtin_typ=builtin_typ, builtin_num=builtin_num,
boehmes@40403
   276
      builtin_fun=builtin_fun, has_datatypes=with_datatypes}
boehmes@40403
   277
    val translate' = {prefixes=prefixes, header=header, strict=strict,
boehmes@40403
   278
      builtins=builtins', serialize=serialize}
boehmes@40403
   279
  in (with_datatypes', translate') end
boehmes@40403
   280
boehmes@40405
   281
fun trace_assumptions ctxt irules idxs =
boehmes@40405
   282
  let
boehmes@40405
   283
    val thms = filter (fn i => i >= 0) idxs
boehmes@40405
   284
      |> map_filter (AList.lookup (op =) irules)
boehmes@40405
   285
  in
boehmes@40405
   286
    if Config.get ctxt trace_used_facts andalso length thms > 0
boehmes@40405
   287
    then
boehmes@40405
   288
      tracing (Pretty.string_of (Pretty.big_list "SMT used facts:"
boehmes@40405
   289
        (map (Display.pretty_thm ctxt) thms)))
boehmes@40405
   290
    else ()
boehmes@40405
   291
  end
boehmes@40405
   292
boehmes@40437
   293
fun gen_solver name info rm ctxt irules =
boehmes@40403
   294
  let
boehmes@40403
   295
    val {env_var, is_remote, options, more_options, interface, reconstruct} =
boehmes@40403
   296
      info
boehmes@36890
   297
    val {extra_norm, translate} = interface
boehmes@40403
   298
    val (with_datatypes, translate') =
boehmes@40403
   299
      set_has_datatypes (Config.get ctxt datatypes) translate
boehmes@40437
   300
    val cmd = (rm, env_var, is_remote, name)
boehmes@40524
   301
    val keep = Config.get ctxt keep_assms
boehmes@36890
   302
  in
boehmes@40403
   303
    (irules, ctxt)
boehmes@40524
   304
    |-> SMT_Normalize.normalize trace_msg keep extra_norm with_datatypes
boehmes@40407
   305
    |-> invoke translate' name cmd more_options options
boehmes@36890
   306
    |-> reconstruct
boehmes@40402
   307
    |-> (fn (idxs, thm) => fn ctxt' => thm
boehmes@36890
   308
    |> singleton (ProofContext.export ctxt' ctxt)
boehmes@40402
   309
    |> discharge_definitions
boehmes@40405
   310
    |> tap (fn _ => trace_assumptions ctxt irules idxs)
boehmes@40402
   311
    |> pair idxs)
boehmes@36890
   312
  end
boehmes@36890
   313
boehmes@36890
   314
boehmes@36890
   315
boehmes@36890
   316
(* solver store *)
boehmes@36890
   317
boehmes@40437
   318
type solver = bool option -> Proof.context -> (int * thm) list -> int list * thm
boehmes@40403
   319
boehmes@40403
   320
type solver_info = {
boehmes@40403
   321
  env_var: string,
boehmes@40403
   322
  is_remote: bool,
boehmes@40403
   323
  options: Proof.context -> string list,
boehmes@40403
   324
  more_options: string list,
boehmes@40403
   325
  interface: interface,
boehmes@40403
   326
  reconstruct: string list * SMT_Translate.recon -> Proof.context ->
boehmes@40403
   327
    (int list * thm) * Proof.context }
boehmes@36890
   328
boehmes@36891
   329
structure Solvers = Generic_Data
boehmes@36890
   330
(
boehmes@40403
   331
  type T = solver_info Symtab.table
boehmes@36890
   332
  val empty = Symtab.empty
boehmes@36890
   333
  val extend = I
boehmes@36890
   334
  fun merge data = Symtab.merge (K true) data
boehmes@36890
   335
    handle Symtab.DUP name => error ("Duplicate SMT solver: " ^ quote name)
boehmes@36890
   336
)
boehmes@36890
   337
boehmes@36890
   338
val no_solver = "(none)"
boehmes@40403
   339
boehmes@40403
   340
fun set_solver_options name opts = Solvers.map (Symtab.map_entry name (fn
boehmes@40403
   341
  {env_var, is_remote, options, interface, reconstruct, ...} =>
boehmes@40403
   342
  {env_var=env_var, is_remote=is_remote, options=options,
boehmes@40403
   343
   more_options=String.tokens (Symbol.is_ascii_blank o str) opts,
boehmes@40403
   344
   interface=interface, reconstruct=reconstruct}))
boehmes@40403
   345
boehmes@40403
   346
local
boehmes@40403
   347
  fun finish outcome cex_parser reconstruct ocl (output, recon) ctxt =
boehmes@40403
   348
    (case outcome output of
boehmes@40403
   349
      (Unsat, ls) =>
boehmes@40403
   350
        if not (Config.get ctxt oracle) andalso is_some reconstruct
boehmes@40403
   351
        then the reconstruct ctxt recon ls
boehmes@40403
   352
        else (([], ocl ()), ctxt)
boehmes@40403
   353
    | (result, ls) =>
boehmes@40403
   354
        let val ts = (case cex_parser of SOME f => f ctxt recon ls | _ => [])
boehmes@40403
   355
        in raise SMT (Counterexample (result = Sat, ts)) end)
boehmes@40403
   356
in
boehmes@40403
   357
boehmes@40403
   358
fun add_solver cfg =
boehmes@40403
   359
  let
boehmes@40403
   360
    val {name, env_var, is_remote, options, interface, outcome, cex_parser,
boehmes@40403
   361
      reconstruct} = cfg
boehmes@40403
   362
boehmes@40403
   363
    fun core_oracle () = @{cprop False}
boehmes@40403
   364
boehmes@40403
   365
    fun solver ocl = { env_var=env_var, is_remote=is_remote, options=options,
boehmes@40403
   366
      more_options=[], interface=interface,
boehmes@40403
   367
      reconstruct=finish (outcome name) cex_parser reconstruct ocl }
boehmes@40403
   368
  in
boehmes@40403
   369
    Thm.add_oracle (Binding.name name, core_oracle) #-> (fn (_, ocl) =>
boehmes@40403
   370
    Context.theory_map (Solvers.map (Symtab.update_new (name, solver ocl)))) #>
boehmes@40403
   371
    Attrib.setup (Binding.name (name ^ "_options"))
boehmes@40403
   372
      (Scan.lift (Parse.$$$ "=" |-- Args.name) >>
boehmes@40403
   373
        (Thm.declaration_attribute o K o set_solver_options name))
boehmes@40403
   374
      ("Additional command line options for SMT solver " ^ quote name)
boehmes@40403
   375
  end
boehmes@40403
   376
boehmes@40403
   377
end
boehmes@40403
   378
boehmes@36890
   379
val all_solver_names_of = Symtab.keys o Solvers.get
boehmes@36890
   380
val lookup_solver = Symtab.lookup o Solvers.get
boehmes@36890
   381
boehmes@36890
   382
boehmes@36890
   383
boehmes@36890
   384
(* selected solver *)
boehmes@36890
   385
boehmes@36890
   386
structure Selected_Solver = Generic_Data
boehmes@36890
   387
(
boehmes@36890
   388
  type T = string
boehmes@36890
   389
  val empty = no_solver
boehmes@36890
   390
  val extend = I
boehmes@36890
   391
  fun merge (s, _) = s
boehmes@36890
   392
)
boehmes@36890
   393
boehmes@36890
   394
val solver_name_of = Selected_Solver.get
boehmes@36890
   395
boehmes@36890
   396
fun select_solver name context =
boehmes@36891
   397
  if is_none (lookup_solver context name)
boehmes@36890
   398
  then error ("SMT solver not registered: " ^ quote name)
boehmes@36890
   399
  else Selected_Solver.map (K name) context
boehmes@36890
   400
boehmes@36890
   401
fun raw_solver_of context name =
boehmes@36891
   402
  (case lookup_solver context name of
boehmes@36890
   403
    NONE => error "No SMT solver selected"
boehmes@40403
   404
  | SOME s => s)
boehmes@36890
   405
boehmes@36890
   406
fun solver_of context =
boehmes@36890
   407
  let val name = solver_name_of context
boehmes@36890
   408
  in gen_solver name (raw_solver_of context name) end
boehmes@36890
   409
boehmes@40407
   410
fun is_locally_installed ctxt =
boehmes@40407
   411
  let
boehmes@40407
   412
    val name = solver_name_of (Context.Proof ctxt)
boehmes@40407
   413
    val {env_var, ...} = raw_solver_of (Context.Proof ctxt) name
boehmes@40407
   414
  in is_some (get_local_solver env_var) end
boehmes@40407
   415
boehmes@36890
   416
boehmes@36890
   417
boehmes@40403
   418
(* SMT filter *)
boehmes@40402
   419
boehmes@40402
   420
val has_topsort = Term.exists_type (Term.exists_subtype (fn
boehmes@40402
   421
    TFree (_, []) => true
boehmes@40402
   422
  | TVar (_, []) => true
boehmes@40402
   423
  | _ => false))
boehmes@40402
   424
boehmes@40437
   425
fun smt_solver rm ctxt irules =
boehmes@40402
   426
  (* without this test, we would run into problems when atomizing the rules: *)
boehmes@40403
   427
  if exists (has_topsort o Thm.prop_of o snd) irules
boehmes@40403
   428
  then raise SMT (Other_Failure "proof state contains the universal sort {}")
boehmes@40437
   429
  else solver_of (Context.Proof ctxt) rm ctxt irules
boehmes@40403
   430
boehmes@40437
   431
fun smt_filter run_remote time_limit st xrules i =
boehmes@40403
   432
  let
boehmes@40440
   433
    val {facts, goal, ...} = Proof.goal st
boehmes@40405
   434
    val ctxt =
boehmes@40440
   435
      Proof.context_of st
boehmes@40578
   436
      |> Config.put timeout (Real.fromInt (Time.toSeconds time_limit))
boehmes@40405
   437
      |> Config.put oracle false
boehmes@40405
   438
      |> Config.put filter_only true
boehmes@40524
   439
      |> Config.put keep_assms false
boehmes@40603
   440
    val ({context=ctxt', prems, concl, ...}, _) = Subgoal.focus ctxt i goal
boehmes@40403
   441
    val cprop =
boehmes@40603
   442
      concl
boehmes@40603
   443
      |> Thm.rhs_of o SMT_Normalize.atomize_conv ctxt'
boehmes@40403
   444
      |> Thm.capply @{cterm Trueprop} o Thm.capply @{cterm Not} o Thm.dest_arg
boehmes@40603
   445
    val irs = map (pair ~1) (Thm.assume cprop :: prems @ facts)
boehmes@40437
   446
    val rm = SOME run_remote
boehmes@40403
   447
  in
boehmes@40402
   448
    split_list xrules
boehmes@40603
   449
    ||> distinct (op =) o fst o smt_solver rm ctxt' o append irs o map_index I
boehmes@40405
   450
    |-> map_filter o try o nth
boehmes@40438
   451
    |> (fn xs => {outcome=NONE, used_facts=xs, run_time_in_msecs=NONE})
boehmes@40403
   452
  end
boehmes@40438
   453
  handle SMT fail => {outcome=SOME fail, used_facts=[], run_time_in_msecs=NONE}
boehmes@40407
   454
  (* FIXME: measure runtime *)
boehmes@40402
   455
boehmes@40402
   456
boehmes@40402
   457
boehmes@36890
   458
(* SMT tactic *)
boehmes@36890
   459
boehmes@36890
   460
fun smt_tac' pass_exns ctxt rules =
boehmes@36891
   461
  CONVERSION (SMT_Normalize.atomize_conv ctxt)
boehmes@36891
   462
  THEN' Tactic.rtac @{thm ccontr}
boehmes@40407
   463
  THEN' SUBPROOF (fn {context=cx, prems, ...} =>
boehmes@40406
   464
    let
boehmes@40437
   465
      fun solve () = snd (smt_solver NONE cx (map (pair ~1) (rules @ prems)))
boehmes@40438
   466
      val name = solver_name_of (Context.Proof cx)
boehmes@40438
   467
      val trace = trace_msg cx (prefix "SMT: " o string_of_failure cx name)
boehmes@40406
   468
    in
boehmes@40407
   469
      (if pass_exns then SOME (solve ())
boehmes@40407
   470
       else (SOME (solve ()) handle SMT fail => (trace fail; NONE)))
boehmes@40406
   471
      |> (fn SOME thm => Tactic.rtac thm 1 | _ => Tactical.no_tac)
boehmes@40407
   472
    end) ctxt
boehmes@36890
   473
boehmes@36890
   474
val smt_tac = smt_tac' false
boehmes@40403
   475
boehmes@36890
   476
val smt_method =
boehmes@36890
   477
  Scan.optional Attrib.thms [] >>
boehmes@36890
   478
  (fn thms => fn ctxt => METHOD (fn facts =>
boehmes@36890
   479
    HEADGOAL (smt_tac ctxt (thms @ facts))))
boehmes@36890
   480
boehmes@36890
   481
boehmes@36890
   482
boehmes@36890
   483
(* setup *)
boehmes@36890
   484
boehmes@36890
   485
val setup =
wenzelm@39075
   486
  Attrib.setup @{binding smt_solver}
wenzelm@36970
   487
    (Scan.lift (Parse.$$$ "=" |-- Args.name) >>
boehmes@36890
   488
      (Thm.declaration_attribute o K o select_solver))
boehmes@36890
   489
    "SMT solver configuration" #>
boehmes@40403
   490
  setup_oracle #>
boehmes@40405
   491
  setup_filter_only #>
boehmes@40403
   492
  setup_datatypes #>
boehmes@40524
   493
  setup_keep_assms #>
boehmes@36890
   494
  setup_timeout #>
boehmes@36890
   495
  setup_trace #>
boehmes@40403
   496
  setup_trace_used_facts #>
boehmes@36890
   497
  setup_fixed_certificates #>
wenzelm@39075
   498
  Attrib.setup @{binding smt_certificates}
wenzelm@36970
   499
    (Scan.lift (Parse.$$$ "=" |-- Args.name) >>
boehmes@36890
   500
      (Thm.declaration_attribute o K o select_certificates))
boehmes@36890
   501
    "SMT certificates" #>
wenzelm@39075
   502
  Method.setup @{binding smt} smt_method
boehmes@36890
   503
    "Applies an SMT solver to the current goal."
boehmes@36890
   504
boehmes@36890
   505
boehmes@36891
   506
fun print_setup context =
boehmes@36890
   507
  let
boehmes@40578
   508
    val t = Time.toString (seconds (Config.get_generic context timeout))
boehmes@36891
   509
    val names = sort_strings (all_solver_names_of context)
boehmes@36890
   510
    val ns = if null names then [no_solver] else names
boehmes@40403
   511
    val n = solver_name_of context
boehmes@40403
   512
    val opts =
boehmes@40403
   513
      (case Symtab.lookup (Solvers.get context) n of
boehmes@40403
   514
        NONE => []
boehmes@40403
   515
      | SOME {more_options, options, ...} =>
boehmes@40403
   516
          more_options @ options (Context.proof_of context))
boehmes@36890
   517
    val certs_filename =
boehmes@36891
   518
      (case get_certificates_path context of
boehmes@36890
   519
        SOME path => Path.implode path
boehmes@36890
   520
      | NONE => "(disabled)")
boehmes@36891
   521
    val fixed = if Config.get_generic context fixed_certificates then "true"
boehmes@36890
   522
      else "false"
boehmes@36890
   523
  in
boehmes@36890
   524
    Pretty.writeln (Pretty.big_list "SMT setup:" [
boehmes@40403
   525
      Pretty.str ("Current SMT solver: " ^ n),
boehmes@40403
   526
      Pretty.str ("Current SMT solver options: " ^ space_implode " " opts),
boehmes@36890
   527
      Pretty.str_list "Available SMT solvers: "  "" ns,
boehmes@36890
   528
      Pretty.str ("Current timeout: " ^ t ^ " seconds"),
boehmes@40403
   529
      Pretty.str ("With proofs: " ^
boehmes@40403
   530
        (if Config.get_generic context oracle then "false" else "true")),
boehmes@36890
   531
      Pretty.str ("Certificates cache: " ^ certs_filename),
boehmes@40403
   532
      Pretty.str ("Fixed certificates: " ^ fixed)])
boehmes@36890
   533
  end
boehmes@36890
   534
wenzelm@36970
   535
val _ =
wenzelm@36970
   536
  Outer_Syntax.improper_command "smt_status"
boehmes@40402
   537
    "show the available SMT solvers and the currently selected solver"
boehmes@40402
   538
    Keyword.diag
boehmes@36890
   539
    (Scan.succeed (Toplevel.no_timing o Toplevel.keep (fn state =>
boehmes@36890
   540
      print_setup (Context.Proof (Toplevel.context_of state)))))
boehmes@36890
   541
boehmes@36890
   542
end