src/HOL/Tools/SMT/smt_setup_solvers.ML
author nipkow
Mon, 01 Aug 2011 12:08:53 +0200
changeset 44045 2814ff2a6e3e
parent 42616 92715b528e78
child 45025 33a1af99b3a2
permissions -rw-r--r--
infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
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
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
    14
  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
    15
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
    16
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
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
    18
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
    19
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    20
(* helper functions *)
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    21
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    22
val remote_prefix = "remote_"
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    23
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
    24
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    25
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
    26
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
    27
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
    28
  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
    29
  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
    30
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    31
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
    32
fun make_remote_command name () =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    33
  [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
    34
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
    35
  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
    36
  else (fn () =>
2dc75bae5226 more explicit errors to inform users about problems related to SMT solvers;
boehmes
parents: 41601
diff changeset
    37
    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
    38
    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
    39
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
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
    41
  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
    42
  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
    43
  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
    44
  else raise SMT_Failure.SMT (SMT_Failure.Other_Failure ("Solver " ^
40276
6efa052b9213 clarified error message
boehmes
parents: 40208
diff changeset
    45
    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
    46
    "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
    47
    "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
    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
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
    50
  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
    51
    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
    52
    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
    53
    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
    54
  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
    55
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
(* 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
    58
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    59
local
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    60
  fun cvc3_options ctxt = [
41121
5c5d05963f93 added option to modify the random seed of SMT solvers
boehmes
parents: 40981
diff changeset
    61
    "-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
    62
    "-lang", "smtlib", "-output-lang", "presentation"]
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    63
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    64
  fun mk is_remote = {
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    65
    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
    66
    class = SMTLIB_Interface.smtlibC,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    67
    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
    68
    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
    69
    options = cvc3_options,
41800
7f333b59d5fb better fudge factors for CVC3 and Yices based on Judgment Day
blanchet
parents: 41787
diff changeset
    70
    default_max_relevant = 150 (* FUDGE *),
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    71
    supports_filter = false,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    72
    outcome =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    73
      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
    74
    cex_parser = NONE,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    75
    reconstruct = NONE }
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    76
in
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    77
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    78
fun cvc3 () = mk false
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    79
fun remote_cvc3 () = mk true
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    80
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    81
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
    82
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
    83
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
(* 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
    85
40208
54e5e8e0b2a4 made SML/NJ happy
boehmes
parents: 40197
diff changeset
    86
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
    87
  name = "yices",
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    88
  class = SMTLIB_Interface.smtlibC,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    89
  avail = make_local_avail "YICES",
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    90
  command = make_local_command "YICES",
41121
5c5d05963f93 added option to modify the random seed of SMT solvers
boehmes
parents: 40981
diff changeset
    91
  options = (fn ctxt => [
41235
975db7bd23e3 fixed the command-line syntax for setting Yices' random seed
boehmes
parents: 41130
diff changeset
    92
    "--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
    93
    "--smtlib"]),
41800
7f333b59d5fb better fudge factors for CVC3 and Yices based on Judgment Day
blanchet
parents: 41787
diff changeset
    94
  default_max_relevant = 150 (* FUDGE *),
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    95
  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
    96
  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
    97
  cex_parser = NONE,
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
    98
  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
    99
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
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
(* 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
   102
42075
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   103
datatype z3_non_commercial =
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   104
  Z3_Non_Commercial_Unknown |
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   105
  Z3_Non_Commercial_Accepted |
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   106
  Z3_Non_Commercial_Declined
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   107
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   108
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   109
local
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   110
  val flagN = "Z3_NON_COMMERCIAL"
42075
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
  val accepted = member (op =) ["yes", "Yes", "YES"]
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   113
  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
   114
in
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   115
42075
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   116
fun z3_non_commercial () =
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   117
  if accepted (getenv flagN) then Z3_Non_Commercial_Accepted
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   118
  else if declined (getenv flagN) then Z3_Non_Commercial_Declined
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   119
  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
   120
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   121
fun if_z3_non_commercial f =
42075
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   122
  (case z3_non_commercial () of
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   123
    Z3_Non_Commercial_Accepted => f ()
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   124
  | Z3_Non_Commercial_Declined =>
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   125
      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
   126
        "applications.")
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   127
  | Z3_Non_Commercial_Unknown =>
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   128
      error ("The SMT solver Z3 is not activated. To activate it, set " ^
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   129
        "the environment variable " ^ quote flagN ^ " to " ^ quote ("yes") ^
c8be98f12b1c Z3 non-commercial usage may explicitly be declined
boehmes
parents: 42074
diff changeset
   130
        "."))
42074
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   131
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   132
end
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   133
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   134
621321627d0f export status function to query whether Z3 has been activated for usage within Isabelle
boehmes
parents: 41800
diff changeset
   135
local
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   136
  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
   137
    if_z3_non_commercial (make_command is_remote name)
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   138
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   139
  fun z3_options ctxt =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   140
    ["-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
   141
      "MODEL=true", "-smt"]
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   142
    |> 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
   143
         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
   144
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   145
  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
   146
    let
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   147
      fun junk l =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   148
        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
   149
        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
   150
        else
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   151
          String.isPrefix "WARNING" l orelse
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   152
          String.isPrefix "ERROR" l orelse
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   153
          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
   154
    in
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   155
      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
   156
      |>> 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
   157
    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
   158
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   159
  fun mk is_remote = {
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   160
    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
   161
    class = Z3_Interface.smtlib_z3C,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   162
    avail = make_avail is_remote "Z3",
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41432
diff changeset
   163
    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
   164
    options = z3_options,
41787
5acde4abb07b comments to find fudge factors easily
blanchet
parents: 41771
diff changeset
   165
    default_max_relevant = 250 (* FUDGE *),
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   166
    supports_filter = true,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   167
    outcome = z3_on_last_line,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   168
    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
   169
    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
   170
in
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   171
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   172
fun z3 () = mk false
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   173
fun remote_z3 () = mk true
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   174
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
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
   177
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
   178
(* 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
   179
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
   180
val setup =
40208
54e5e8e0b2a4 made SML/NJ happy
boehmes
parents: 40197
diff changeset
   181
  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
   182
  SMT_Solver.add_solver (remote_cvc3 ()) #>
40208
54e5e8e0b2a4 made SML/NJ happy
boehmes
parents: 40197
diff changeset
   183
  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
   184
  SMT_Solver.add_solver (z3 ()) #>
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41235
diff changeset
   185
  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
   186
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
   187
end