src/HOL/Tools/SMT/smt_setup_solvers.ML
author wenzelm
Sat, 12 Jan 2013 23:07:21 +0100
changeset 50856 c091e46ec3c5
parent 50661 acea12b85315
child 55007 0c07990363a3
permissions -rw-r--r--
tuned message;
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
42075
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
     9
  datatype z3_non_commercial =
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
    10
    Z3_Non_Commercial_Unknown |
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
    11
    Z3_Non_Commercial_Accepted |
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
    12
    Z3_Non_Commercial_Declined
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
    13
  val z3_non_commercial: unit -> z3_non_commercial
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 48043
diff changeset
    14
  val z3_with_extensions: bool Config.T
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
    15
  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
    16
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
    17
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
    18
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
    19
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
    20
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    21
(* helper functions *)
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    22
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    23
val remote_prefix = "remote_"
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    24
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
    25
48391
480746f1012c get rid of redundant "xxx_INSTALLED" environment variabl
blanchet
parents: 48069
diff changeset
    26
fun make_local_avail name () = getenv (name ^ "_SOLVER") <> ""
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    27
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
    28
fun make_avail is_remote name =
41761
2dc75bae5226 more explicit errors to inform users about problems related to SMT solvers;
boehmes
parents: 41601
diff changeset
    29
  if is_remote then make_remote_avail name
2dc75bae5226 more explicit errors to inform users about problems related to SMT solvers;
boehmes
parents: 41601
diff changeset
    30
  else make_local_avail name orf make_remote_avail name
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    31
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    32
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
    33
fun make_remote_command name () =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    34
  [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
    35
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
    36
  if is_remote then make_remote_command name
41761
2dc75bae5226 more explicit errors to inform users about problems related to SMT solvers;
boehmes
parents: 41601
diff changeset
    37
  else (fn () =>
2dc75bae5226 more explicit errors to inform users about problems related to SMT solvers;
boehmes
parents: 41601
diff changeset
    38
    if make_local_avail name () then make_local_command name ()
2dc75bae5226 more explicit errors to inform users about problems related to SMT solvers;
boehmes
parents: 41601
diff changeset
    39
    else make_remote_command name ())
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    40
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
    41
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
    42
  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
    43
  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
    44
  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
    45
  else raise SMT_Failure.SMT (SMT_Failure.Other_Failure ("Solver " ^
40276
6efa052b9213 clarified error message
boehmes
parents: 40208
diff changeset
    46
    quote solver_name ^ " failed. Enable SMT tracing by setting the " ^
42616
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42075
diff changeset
    47
    "configuration option " ^ quote (Config.name_of SMT_Config.trace) ^ " 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
    48
    "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
    49
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
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
    51
  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
    52
    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
    53
    val split_first = (fn [] => ("", []) | l :: ls => (l, ls))
48902
44a6967240b7 prefer classic take_prefix/take_suffix over chop_while (cf. 0659e84bdc5f);
wenzelm
parents: 48391
diff changeset
    54
    val (l, ls) = split_first (snd (take_prefix empty_line 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
    55
  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
    56
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
    57
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
    58
(* 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
    59
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    60
local
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    61
  fun cvc3_options ctxt = [
41121
5c5d05963f93 added option to modify the random seed of SMT solvers
boehmes
parents: 40981
diff changeset
    62
    "-seed", string_of_int (Config.get ctxt SMT_Config.random_seed),
49980
34b0464d7eef soft SMT timeout
blanchet
parents: 48902
diff changeset
    63
    "-lang", "smtlib", "-output-lang", "presentation",
34b0464d7eef soft SMT timeout
blanchet
parents: 48902
diff changeset
    64
    "-timeout", string_of_int (Real.ceil (Config.get ctxt SMT_Config.timeout))]
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    65
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    66
  fun mk is_remote = {
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    67
    name = make_name is_remote "cvc3",
48043
3ff2c76c9f64 introduced option "z3_with_extensions" to control whether Z3's support for nonlinear arithmetic and datatypes should be enabled (including potential proof reconstruction failures)
boehmes
parents: 47940
diff changeset
    68
    class = K SMTLIB_Interface.smtlibC,
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    69
    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
    70
    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
    71
    options = cvc3_options,
47645
8ca67d2b21c2 tuned "max_relevant" defaults for SMT solvers based on Judgment Day
blanchet
parents: 45025
diff changeset
    72
    default_max_relevant = 400 (* FUDGE *),
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    73
    supports_filter = false,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    74
    outcome =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    75
      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
    76
    cex_parser = NONE,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    77
    reconstruct = NONE }
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    78
in
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    79
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    80
fun cvc3 () = mk false
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    81
fun remote_cvc3 () = mk true
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    82
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    83
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
    84
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
    85
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
    86
(* 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
    87
40208
54e5e8e0b2a4 made SML/NJ happy
boehmes
parents: 40197
diff changeset
    88
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
    89
  name = "yices",
48043
3ff2c76c9f64 introduced option "z3_with_extensions" to control whether Z3's support for nonlinear arithmetic and datatypes should be enabled (including potential proof reconstruction failures)
boehmes
parents: 47940
diff changeset
    90
  class = K SMTLIB_Interface.smtlibC,
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    91
  avail = make_local_avail "YICES",
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    92
  command = make_local_command "YICES",
41121
5c5d05963f93 added option to modify the random seed of SMT solvers
boehmes
parents: 40981
diff changeset
    93
  options = (fn ctxt => [
41235
975db7bd23e3 fixed the command-line syntax for setting Yices' random seed
boehmes
parents: 41130
diff changeset
    94
    "--rand-seed=" ^ string_of_int (Config.get ctxt SMT_Config.random_seed),
49980
34b0464d7eef soft SMT timeout
blanchet
parents: 48902
diff changeset
    95
    "--timeout=" ^
34b0464d7eef soft SMT timeout
blanchet
parents: 48902
diff changeset
    96
      string_of_int (Real.ceil (Config.get ctxt SMT_Config.timeout)),
41121
5c5d05963f93 added option to modify the random seed of SMT solvers
boehmes
parents: 40981
diff changeset
    97
    "--smtlib"]),
47645
8ca67d2b21c2 tuned "max_relevant" defaults for SMT solvers based on Judgment Day
blanchet
parents: 45025
diff changeset
    98
  default_max_relevant = 350 (* FUDGE *),
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    99
  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
   100
  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
   101
  cex_parser = NONE,
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   102
  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
   103
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
   104
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
   105
(* 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
   106
42075
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   107
datatype z3_non_commercial =
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   108
  Z3_Non_Commercial_Unknown |
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   109
  Z3_Non_Commercial_Accepted |
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   110
  Z3_Non_Commercial_Declined
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   111
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   112
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   113
local
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   114
  val flagN = "Z3_NON_COMMERCIAL"
42075
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   115
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   116
  val accepted = member (op =) ["yes", "Yes", "YES"]
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   117
  val declined = member (op =) ["no", "No", "NO"]
42074
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   118
in
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   119
42075
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   120
fun z3_non_commercial () =
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   121
  if accepted (getenv flagN) then Z3_Non_Commercial_Accepted
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   122
  else if declined (getenv flagN) then Z3_Non_Commercial_Declined
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   123
  else Z3_Non_Commercial_Unknown
42074
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   124
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   125
fun if_z3_non_commercial f =
42075
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   126
  (case z3_non_commercial () of
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   127
    Z3_Non_Commercial_Accepted => f ()
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   128
  | Z3_Non_Commercial_Declined =>
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   129
      error ("The SMT solver Z3 may only be used for non-commercial " ^
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   130
        "applications.")
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   131
  | Z3_Non_Commercial_Unknown =>
45025
33a1af99b3a2 more hints on Z3 configuration;
wenzelm
parents: 42616
diff changeset
   132
      error ("The SMT solver Z3 is not activated. To activate it, set\n" ^
50856
c091e46ec3c5 tuned message;
wenzelm
parents: 50661
diff changeset
   133
        "the environment variable " ^ quote flagN ^ " to " ^ quote "yes" ^ ",\n" ^
c091e46ec3c5 tuned message;
wenzelm
parents: 50661
diff changeset
   134
        "and restart the Isabelle system." ^
45025
33a1af99b3a2 more hints on Z3 configuration;
wenzelm
parents: 42616
diff changeset
   135
        (if getenv "Z3_COMPONENT" = "" then ""
33a1af99b3a2 more hints on Z3 configuration;
wenzelm
parents: 42616
diff changeset
   136
         else "\nSee also " ^ Path.print (Path.expand (Path.explode "$Z3_COMPONENT/etc/settings")))))
42074
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   137
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   138
end
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   139
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   140
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 48043
diff changeset
   141
val z3_with_extensions =
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 48043
diff changeset
   142
  Attrib.setup_config_bool @{binding z3_with_extensions} (K false)
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 48043
diff changeset
   143
42074
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   144
local
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   145
  fun z3_make_command is_remote name () =
42074
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   146
    if_z3_non_commercial (make_command is_remote name)
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   147
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   148
  fun z3_options ctxt =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   149
    ["-rs:" ^ string_of_int (Config.get ctxt SMT_Config.random_seed),
49980
34b0464d7eef soft SMT timeout
blanchet
parents: 48902
diff changeset
   150
      "MODEL=true",
34b0464d7eef soft SMT timeout
blanchet
parents: 48902
diff changeset
   151
      "SOFT_TIMEOUT=" ^
50661
acea12b85315 Z3's soft timeout is expressed in ms, not in s -- this explains the frequenty "error code 112" failures we had recently
blanchet
parents: 49980
diff changeset
   152
        string_of_int (Real.ceil (1000.0 * Config.get ctxt SMT_Config.timeout)),
49980
34b0464d7eef soft SMT timeout
blanchet
parents: 48902
diff changeset
   153
      "-smt"]
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   154
    |> not (Config.get ctxt SMT_Config.oracle) ?
49980
34b0464d7eef soft SMT timeout
blanchet
parents: 48902
diff changeset
   155
         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
   156
47940
4ef90b51641e robustly parse Z3 4.0's output (with outcome appearing on first rather than last line)
blanchet
parents: 47645
diff changeset
   157
  fun z3_on_first_or_last_line solver_name lines =
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   158
    let
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   159
      fun junk l =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   160
        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
   161
        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
   162
        else
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   163
          String.isPrefix "WARNING" l orelse
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   164
          String.isPrefix "ERROR" l orelse
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   165
          forall Symbol.is_ascii_blank (Symbol.explode l)
47940
4ef90b51641e robustly parse Z3 4.0's output (with outcome appearing on first rather than last line)
blanchet
parents: 47645
diff changeset
   166
      val lines = filter_out junk lines
4ef90b51641e robustly parse Z3 4.0's output (with outcome appearing on first rather than last line)
blanchet
parents: 47645
diff changeset
   167
      fun outcome split =
4ef90b51641e robustly parse Z3 4.0's output (with outcome appearing on first rather than last line)
blanchet
parents: 47645
diff changeset
   168
        the_default ("", []) (try split lines)
4ef90b51641e robustly parse Z3 4.0's output (with outcome appearing on first rather than last line)
blanchet
parents: 47645
diff changeset
   169
        |>> outcome_of "unsat" "sat" "unknown" solver_name
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   170
    in
47940
4ef90b51641e robustly parse Z3 4.0's output (with outcome appearing on first rather than last line)
blanchet
parents: 47645
diff changeset
   171
      (* Starting with version 4.0, Z3 puts the outcome on the first line of the
4ef90b51641e robustly parse Z3 4.0's output (with outcome appearing on first rather than last line)
blanchet
parents: 47645
diff changeset
   172
         output rather than on the last line. *)
4ef90b51641e robustly parse Z3 4.0's output (with outcome appearing on first rather than last line)
blanchet
parents: 47645
diff changeset
   173
      outcome (fn lines => (hd lines, tl lines))
4ef90b51641e robustly parse Z3 4.0's output (with outcome appearing on first rather than last line)
blanchet
parents: 47645
diff changeset
   174
      handle SMT_Failure.SMT _ => outcome (swap o split_last)
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   175
    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
   176
48043
3ff2c76c9f64 introduced option "z3_with_extensions" to control whether Z3's support for nonlinear arithmetic and datatypes should be enabled (including potential proof reconstruction failures)
boehmes
parents: 47940
diff changeset
   177
  fun select_class ctxt =
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 48043
diff changeset
   178
    if Config.get ctxt z3_with_extensions then Z3_Interface.smtlib_z3C
48043
3ff2c76c9f64 introduced option "z3_with_extensions" to control whether Z3's support for nonlinear arithmetic and datatypes should be enabled (including potential proof reconstruction failures)
boehmes
parents: 47940
diff changeset
   179
    else SMTLIB_Interface.smtlibC
3ff2c76c9f64 introduced option "z3_with_extensions" to control whether Z3's support for nonlinear arithmetic and datatypes should be enabled (including potential proof reconstruction failures)
boehmes
parents: 47940
diff changeset
   180
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   181
  fun mk is_remote = {
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   182
    name = make_name is_remote "z3",
48043
3ff2c76c9f64 introduced option "z3_with_extensions" to control whether Z3's support for nonlinear arithmetic and datatypes should be enabled (including potential proof reconstruction failures)
boehmes
parents: 47940
diff changeset
   183
    class = select_class,
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   184
    avail = make_avail is_remote "Z3",
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   185
    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
   186
    options = z3_options,
47645
8ca67d2b21c2 tuned "max_relevant" defaults for SMT solvers based on Judgment Day
blanchet
parents: 45025
diff changeset
   187
    default_max_relevant = 350 (* FUDGE *),
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   188
    supports_filter = true,
47940
4ef90b51641e robustly parse Z3 4.0's output (with outcome appearing on first rather than last line)
blanchet
parents: 47645
diff changeset
   189
    outcome = z3_on_first_or_last_line,
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   190
    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
   191
    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
   192
in
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   193
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   194
fun z3 () = mk false
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   195
fun remote_z3 () = mk true
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   196
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   197
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
   198
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
   199
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
   200
(* 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
   201
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
   202
val setup =
40208
54e5e8e0b2a4 made SML/NJ happy
boehmes
parents: 40197
diff changeset
   203
  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
   204
  SMT_Solver.add_solver (remote_cvc3 ()) #>
40208
54e5e8e0b2a4 made SML/NJ happy
boehmes
parents: 40197
diff changeset
   205
  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
   206
  SMT_Solver.add_solver (z3 ()) #>
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   207
  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
   208
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
   209
end