src/HOL/Tools/SMT/smt_setup_solvers.ML
author wenzelm
Sun, 30 Jan 2011 13:02:18 +0100
changeset 41648 6d736d983d5c
parent 41601 fda8511006f9
child 41761 2dc75bae5226
permissions -rw-r--r--
clarified example settings for Proof General;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/Tools/SMT/smt_setup_solvers.ML
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
     3
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
     4
Setup SMT solvers.
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
     5
*)
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
     6
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
     7
signature SMT_SETUP_SOLVERS =
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
     8
sig
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
     9
  val setup: theory -> theory
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    10
end
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    11
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    12
structure SMT_Setup_Solvers: SMT_SETUP_SOLVERS =
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    13
struct
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    14
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    15
(* helper functions *)
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    16
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    17
val remote_prefix = "remote_"
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    18
fun make_name is_remote name = name |> is_remote ? prefix remote_prefix
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    19
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    20
fun make_local_avail name () = getenv (name ^ "_INSTALLED") = "yes"
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    21
fun make_remote_avail name () = getenv (name ^ "_REMOTE_SOLVER") <> ""
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    22
fun make_avail is_remote name =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    23
  if is_remote then make_remote_avail name else make_local_avail name
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    24
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    25
fun make_local_command name () = [getenv (name ^ "_SOLVER")]
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    26
fun make_remote_command name () =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    27
  [getenv "ISABELLE_SMT_REMOTE", getenv (name ^ "_REMOTE_SOLVER")]
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    28
fun make_command is_remote name =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    29
  if is_remote then make_remote_command name
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    30
  else make_local_command name
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    31
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    32
fun outcome_of unsat sat unknown solver_name line =
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    33
  if String.isPrefix unsat line then SMT_Solver.Unsat
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    34
  else if String.isPrefix sat line then SMT_Solver.Sat
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    35
  else if String.isPrefix unknown line then SMT_Solver.Unknown
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40276
diff changeset
    36
  else raise SMT_Failure.SMT (SMT_Failure.Other_Failure ("Solver " ^
40276
6efa052b9213 clarified error message
boehmes
parents: 40208
diff changeset
    37
    quote solver_name ^ " failed. Enable SMT tracing by setting the " ^
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40276
diff changeset
    38
    "configuration option " ^ quote SMT_Config.traceN ^ " and " ^
40981
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40424
diff changeset
    39
    "see the trace for details."))
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    40
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    41
fun on_first_line test_outcome solver_name lines =
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    42
  let
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    43
    val empty_line = (fn "" => true | _ => false)
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    44
    val split_first = (fn [] => ("", []) | l :: ls => (l, ls))
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    45
    val (l, ls) = split_first (snd (chop_while empty_line lines))
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    46
  in (test_outcome solver_name l, ls) end
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    47
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    48
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    49
(* CVC3 *)
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    50
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    51
local
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    52
  fun cvc3_options ctxt = [
41121
5c5d05963f93 added option to modify the random seed of SMT solvers
boehmes
parents: 40981
diff changeset
    53
    "-seed", string_of_int (Config.get ctxt SMT_Config.random_seed),
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    54
    "-lang", "smtlib", "-output-lang", "presentation"]
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    55
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    56
  fun mk is_remote = {
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    57
    name = make_name is_remote "cvc3",
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    58
    class = SMTLIB_Interface.smtlibC,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    59
    avail = make_avail is_remote "CVC3",
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    60
    command = make_command is_remote "CVC3",
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    61
    options = cvc3_options,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    62
    default_max_relevant = 250,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    63
    supports_filter = false,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    64
    outcome =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    65
      on_first_line (outcome_of "Unsatisfiable." "Satisfiable." "Unknown."),
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    66
    cex_parser = NONE,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    67
    reconstruct = NONE }
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    68
in
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    69
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    70
fun cvc3 () = mk false
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    71
fun remote_cvc3 () = mk true
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    72
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    73
end
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    74
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    75
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    76
(* Yices *)
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    77
40208
54e5e8e0b2a4 made SML/NJ happy
boehmes
parents: 40197
diff changeset
    78
fun yices () = {
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    79
  name = "yices",
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    80
  class = SMTLIB_Interface.smtlibC,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    81
  avail = make_local_avail "YICES",
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    82
  command = make_local_command "YICES",
41121
5c5d05963f93 added option to modify the random seed of SMT solvers
boehmes
parents: 40981
diff changeset
    83
  options = (fn ctxt => [
41235
975db7bd23e3 fixed the command-line syntax for setting Yices' random seed
boehmes
parents: 41130
diff changeset
    84
    "--rand-seed=" ^ string_of_int (Config.get ctxt SMT_Config.random_seed),
41121
5c5d05963f93 added option to modify the random seed of SMT solvers
boehmes
parents: 40981
diff changeset
    85
    "--smtlib"]),
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    86
  default_max_relevant = 275,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    87
  supports_filter = false,
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    88
  outcome = on_first_line (outcome_of "unsat" "sat" "unknown"),
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    89
  cex_parser = NONE,
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    90
  reconstruct = NONE }
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    91
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    92
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    93
(* Z3 *)
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
    94
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    95
local
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
    96
  val flagN = "Z3_NON_COMMERCIAL"
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
    97
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
    98
  fun z3_make_command is_remote name () =
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
    99
    if getenv flagN = "yes" then make_command is_remote name ()
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   100
    else
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   101
      error ("The SMT solver Z3 is not enabled. To enable it, set " ^
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   102
        "the environment variable " ^ quote flagN ^ " to " ^ quote ("yes") ^
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   103
        ".")
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   104
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   105
  fun z3_options ctxt =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   106
    ["-rs:" ^ string_of_int (Config.get ctxt SMT_Config.random_seed),
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   107
      "MODEL=true", "-smt"]
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   108
    |> not (Config.get ctxt SMT_Config.oracle) ?
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   109
         append ["DISPLAY_PROOF=true","PROOF_MODE=2"]
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
   110
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   111
  fun z3_on_last_line solver_name lines =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   112
    let
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   113
      fun junk l =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   114
        if String.isPrefix "WARNING: Out of allocated virtual memory" l
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   115
        then raise SMT_Failure.SMT SMT_Failure.Out_Of_Memory
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   116
        else
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   117
          String.isPrefix "WARNING" l orelse
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   118
          String.isPrefix "ERROR" l orelse
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   119
          forall Symbol.is_ascii_blank (Symbol.explode l)
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   120
    in
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   121
      the_default ("", []) (try (swap o split_last) (filter_out junk lines))
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   122
      |>> outcome_of "unsat" "sat" "unknown" solver_name
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   123
    end
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
   124
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   125
  fun mk is_remote = {
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   126
    name = make_name is_remote "z3",
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   127
    class = Z3_Interface.smtlib_z3C,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   128
    avail = make_avail is_remote "Z3",
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   129
    command = z3_make_command is_remote "Z3",
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   130
    options = z3_options,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   131
    default_max_relevant = 225,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   132
    supports_filter = true,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   133
    outcome = z3_on_last_line,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   134
    cex_parser = SOME Z3_Model.parse_counterex,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   135
    reconstruct = SOME Z3_Proof_Reconstruction.reconstruct }
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   136
in
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   137
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   138
fun z3 () = mk false
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   139
fun remote_z3 () = mk true
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   140
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   141
end
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
   142
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
   143
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
   144
(* overall setup *)
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
   145
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
   146
val setup =
40208
54e5e8e0b2a4 made SML/NJ happy
boehmes
parents: 40197
diff changeset
   147
  SMT_Solver.add_solver (cvc3 ()) #>
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   148
  SMT_Solver.add_solver (remote_cvc3 ()) #>
40208
54e5e8e0b2a4 made SML/NJ happy
boehmes
parents: 40197
diff changeset
   149
  SMT_Solver.add_solver (yices ()) #>
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   150
  SMT_Solver.add_solver (z3 ()) #>
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   151
  SMT_Solver.add_solver (remote_z3 ())
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
   152
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents:
diff changeset
   153
end