src/HOL/Tools/SMT/smt_systems.ML
author wenzelm
Sat, 20 Feb 2021 13:42:37 +0100
changeset 73255 7e2a9a8c2b85
parent 73104 6520d59fbdd7
child 73388 a40e69fde2b4
permissions -rw-r--r--
provide naproche-755224402e36;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
     1
(*  Title:      HOL/Tools/SMT/smt_systems.ML
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     3
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     4
Setup SMT solvers.
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     5
*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     6
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
     7
signature SMT_SYSTEMS =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     8
sig
59960
372ddff01244 updated SMT module and Sledgehammer to fully open source Z3
blanchet
parents: 59035
diff changeset
     9
  val cvc4_extensions: bool Config.T
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    10
  val z3_extensions: bool Config.T
57229
blanchet
parents: 57210
diff changeset
    11
end;
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    12
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
    13
structure SMT_Systems: SMT_SYSTEMS =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    14
struct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    15
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    16
(* helper functions *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    17
72478
b452242dce36 proper Isabelle component settings: prefer standard terminology "ISABELLE_VERIT", avoid conflict of "VERIT_VERSION" with processing of implicit options by veriT;
wenzelm
parents: 72458
diff changeset
    18
fun check_tool var () =
b452242dce36 proper Isabelle component settings: prefer standard terminology "ISABELLE_VERIT", avoid conflict of "VERIT_VERSION" with processing of implicit options by veriT;
wenzelm
parents: 72458
diff changeset
    19
  (case getenv var of
b452242dce36 proper Isabelle component settings: prefer standard terminology "ISABELLE_VERIT", avoid conflict of "VERIT_VERSION" with processing of implicit options by veriT;
wenzelm
parents: 72458
diff changeset
    20
    "" => NONE
72479
7d0861af3cb0 proper support for Windows exe;
wenzelm
parents: 72478
diff changeset
    21
  | s =>
7d0861af3cb0 proper support for Windows exe;
wenzelm
parents: 72478
diff changeset
    22
      if File.is_file (Path.variable var |> Path.expand |> Path.platform_exe)
7d0861af3cb0 proper support for Windows exe;
wenzelm
parents: 72478
diff changeset
    23
      then SOME [s] else NONE);
72478
b452242dce36 proper Isabelle component settings: prefer standard terminology "ISABELLE_VERIT", avoid conflict of "VERIT_VERSION" with processing of implicit options by veriT;
wenzelm
parents: 72458
diff changeset
    24
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    25
fun make_avail name () = getenv (name ^ "_SOLVER") <> ""
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    26
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    27
fun make_command name () = [getenv (name ^ "_SOLVER")]
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    28
70327
c04d4951a155 handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
blanchet
parents: 69593
diff changeset
    29
fun outcome_of unsat sat unknown timeout solver_name line =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
    30
  if String.isPrefix unsat line then SMT_Solver.Unsat
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
    31
  else if String.isPrefix sat line then SMT_Solver.Sat
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
    32
  else if String.isPrefix unknown line then SMT_Solver.Unknown
70327
c04d4951a155 handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
blanchet
parents: 69593
diff changeset
    33
  else if String.isPrefix timeout line then SMT_Solver.Time_Out
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
    34
  else raise SMT_Failure.SMT (SMT_Failure.Other_Failure ("Solver " ^ quote solver_name ^
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
    35
    " failed -- enable tracing using the " ^ quote (Config.name_of SMT_Config.trace) ^
56094
2adbc6e4cd8f let exception pass through in debug mode
blanchet
parents: 56091
diff changeset
    36
    " option for details"))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    37
73104
6520d59fbdd7 ignore error messages produced by CVC4 when generating BV
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72667
diff changeset
    38
(* When used with bitvectors, CVC4 can produce error messages like:
6520d59fbdd7 ignore error messages produced by CVC4 when generating BV
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72667
diff changeset
    39
6520d59fbdd7 ignore error messages produced by CVC4 when generating BV
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72667
diff changeset
    40
$ISABELLE_TMP_PREFIX/... No set-logic command was given before this point.
6520d59fbdd7 ignore error messages produced by CVC4 when generating BV
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72667
diff changeset
    41
6520d59fbdd7 ignore error messages produced by CVC4 when generating BV
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72667
diff changeset
    42
These message should be ignored.
6520d59fbdd7 ignore error messages produced by CVC4 when generating BV
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72667
diff changeset
    43
*)
60201
90e88e521e0e made CVC4 support work also without unsat cores
blanchet
parents: 59960
diff changeset
    44
fun is_blank_or_error_line "" = true
73104
6520d59fbdd7 ignore error messages produced by CVC4 when generating BV
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72667
diff changeset
    45
  | is_blank_or_error_line s =
6520d59fbdd7 ignore error messages produced by CVC4 when generating BV
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72667
diff changeset
    46
  String.isPrefix "(error " s orelse String.isPrefix (getenv "ISABELLE_TMP_PREFIX") s
60201
90e88e521e0e made CVC4 support work also without unsat cores
blanchet
parents: 59960
diff changeset
    47
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57237
diff changeset
    48
fun on_first_line test_outcome solver_name lines =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    49
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    50
    val split_first = (fn [] => ("", []) | l :: ls => (l, ls))
67522
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 67405
diff changeset
    51
    val (l, ls) = split_first (drop_prefix is_blank_or_error_line lines)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    52
  in (test_outcome solver_name l, ls) end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    53
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents: 57240
diff changeset
    54
fun on_first_non_unsupported_line test_outcome solver_name lines =
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67399
diff changeset
    55
  on_first_line test_outcome solver_name (filter (curry (op <>) "unsupported") lines)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    56
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    57
(* CVC3 *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    58
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    59
local
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    60
  fun cvc3_options ctxt = [
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
    61
    "-seed", string_of_int (Config.get ctxt SMT_Config.random_seed),
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57237
diff changeset
    62
    "-lang", "smt2",
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
    63
    "-timeout", string_of_int (Real.ceil (Config.get ctxt SMT_Config.timeout))]
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    64
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    65
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
    66
val cvc3: SMT_Solver.solver_config = {
57209
7ffa0f7e2775 removed '_new' sufffix in SMT2 solver names (in some cases)
blanchet
parents: 57168
diff changeset
    67
  name = "cvc3",
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
    68
  class = K SMTLIB_Interface.smtlibC,
57210
5d61d875076a rationalized CVC3 and Yices environment variable -- no need (unlike for Z3) to distinguish between old and new versions
blanchet
parents: 57209
diff changeset
    69
  avail = make_avail "CVC3",
5d61d875076a rationalized CVC3 and Yices environment variable -- no need (unlike for Z3) to distinguish between old and new versions
blanchet
parents: 57209
diff changeset
    70
  command = make_command "CVC3",
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    71
  options = cvc3_options,
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57237
diff changeset
    72
  smt_options = [],
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    73
  default_max_relevant = 400 (* FUDGE *),
70327
c04d4951a155 handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
blanchet
parents: 69593
diff changeset
    74
  outcome = on_first_line (outcome_of "unsat" "sat" "unknown" "timeout"),
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
    75
  parse_proof = NONE,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    76
  replay = NONE }
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    77
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    78
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    79
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
    80
57240
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
    81
(* CVC4 *)
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
    82
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69205
diff changeset
    83
val cvc4_extensions = Attrib.setup_config_bool \<^binding>\<open>cvc4_extensions\<close> (K false)
58360
dee1fd1cc631 added interface for CVC4 extensions
blanchet
parents: 58061
diff changeset
    84
57240
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
    85
local
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
    86
  fun cvc4_options ctxt = [
72667
b83988b436dc Add support for CVC4 1.8 to Sledgehammer
desharna
parents: 72482
diff changeset
    87
    "--no-stats",
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
    88
    "--random-seed=" ^ string_of_int (Config.get ctxt SMT_Config.random_seed),
57240
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
    89
    "--lang=smt2",
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
    90
    "--tlimit", string_of_int (Real.ceil (1000.0 * Config.get ctxt SMT_Config.timeout))]
58360
dee1fd1cc631 added interface for CVC4 extensions
blanchet
parents: 58061
diff changeset
    91
dee1fd1cc631 added interface for CVC4 extensions
blanchet
parents: 58061
diff changeset
    92
  fun select_class ctxt =
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
    93
    if Config.get ctxt cvc4_extensions then
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
    94
      if Config.get ctxt SMT_Config.higher_order then
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
    95
        CVC4_Interface.hosmtlib_cvc4C
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
    96
      else
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
    97
        CVC4_Interface.smtlib_cvc4C
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
    98
    else
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
    99
      if Config.get ctxt SMT_Config.higher_order then
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   100
        SMTLIB_Interface.hosmtlibC
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   101
      else
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   102
        SMTLIB_Interface.smtlibC
57240
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
   103
in
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
   104
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
   105
val cvc4: SMT_Solver.solver_config = {
57240
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
   106
  name = "cvc4",
58360
dee1fd1cc631 added interface for CVC4 extensions
blanchet
parents: 58061
diff changeset
   107
  class = select_class,
57240
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
   108
  avail = make_avail "CVC4",
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
   109
  command = make_command "CVC4",
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
   110
  options = cvc4_options,
59015
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58496
diff changeset
   111
  smt_options = [(":produce-unsat-cores", "true")],
57240
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
   112
  default_max_relevant = 400 (* FUDGE *),
70327
c04d4951a155 handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
blanchet
parents: 69593
diff changeset
   113
  outcome = on_first_line (outcome_of "unsat" "sat" "unknown" "timeout"),
59015
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58496
diff changeset
   114
  parse_proof = SOME (K CVC4_Proof_Parse.parse_proof),
57240
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
   115
  replay = NONE }
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
   116
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
   117
end
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
   118
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   119
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents: 57240
diff changeset
   120
(* veriT *)
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents: 57240
diff changeset
   121
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   122
local
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   123
  fun select_class ctxt =
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   124
    if Config.get ctxt SMT_Config.higher_order then
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   125
      SMTLIB_Interface.hosmtlibC
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   126
    else
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   127
      SMTLIB_Interface.smtlibC
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   128
in
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   129
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
   130
val veriT: SMT_Solver.solver_config = {
59035
3a2153676705 renamed 'veriT' to 'verit', to stick to all-lowercase rule for prover names
blanchet
parents: 59015
diff changeset
   131
  name = "verit",
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   132
  class = select_class,
72478
b452242dce36 proper Isabelle component settings: prefer standard terminology "ISABELLE_VERIT", avoid conflict of "VERIT_VERSION" with processing of implicit options by veriT;
wenzelm
parents: 72458
diff changeset
   133
  avail = is_some o check_tool "ISABELLE_VERIT",
b452242dce36 proper Isabelle component settings: prefer standard terminology "ISABELLE_VERIT", avoid conflict of "VERIT_VERSION" with processing of implicit options by veriT;
wenzelm
parents: 72458
diff changeset
   134
  command = the o check_tool "ISABELLE_VERIT",
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents: 57240
diff changeset
   135
  options = (fn ctxt => [
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   136
    "--proof-with-sharing",
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   137
    "--proof-define-skolems",
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents: 57240
diff changeset
   138
    "--proof-prune",
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents: 57240
diff changeset
   139
    "--proof-merge",
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents: 57240
diff changeset
   140
    "--disable-print-success",
72482
11f645d25498 remove unsupported max-time option from veriT calls
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72479
diff changeset
   141
    "--disable-banner"] @
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   142
    Verit_Proof.veriT_current_strategy (Context.Proof ctxt)),
61587
c3974cd2d381 updating options to verit
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 60201
diff changeset
   143
  smt_options = [(":produce-proofs", "true")],
58496
2ba52ecc4996 give more facts to veriT -- it seems to be able to cope with them
blanchet
parents: 58491
diff changeset
   144
  default_max_relevant = 200 (* FUDGE *),
70327
c04d4951a155 handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
blanchet
parents: 69593
diff changeset
   145
  outcome = on_first_non_unsupported_line (outcome_of "unsat" "sat" "unknown" "timeout"),
58491
blanchet
parents: 58360
diff changeset
   146
  parse_proof = SOME (K VeriT_Proof_Parse.parse_proof),
69205
8050734eee3e add reconstruction by veriT in method smt
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67522
diff changeset
   147
  replay = SOME Verit_Replay.replay }
57240
9a5729600ba9 added support for CVC4 in SMT2
blanchet
parents: 57239
diff changeset
   148
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   149
end
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   150
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 64461
diff changeset
   151
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   152
(* Z3 *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   153
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69205
diff changeset
   154
val z3_extensions = Attrib.setup_config_bool \<^binding>\<open>z3_extensions\<close> (K false)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   155
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   156
local
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   157
  fun z3_options ctxt =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
   158
    ["smt.random_seed=" ^ string_of_int (Config.get ctxt SMT_Config.random_seed),
57168
af95a414136a disable hard-to-reconstruct Z3 feature
blanchet
parents: 57164
diff changeset
   159
     "smt.refine_inj_axioms=false",
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
   160
     "-T:" ^ string_of_int (Real.ceil (Config.get ctxt SMT_Config.timeout)),
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   161
     "-smt2"]
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   162
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   163
  fun select_class ctxt =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
   164
    if Config.get ctxt z3_extensions then Z3_Interface.smtlib_z3C else SMTLIB_Interface.smtlibC
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   165
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   166
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
   167
val z3: SMT_Solver.solver_config = {
57209
7ffa0f7e2775 removed '_new' sufffix in SMT2 solver names (in some cases)
blanchet
parents: 57168
diff changeset
   168
  name = "z3",
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   169
  class = select_class,
59960
372ddff01244 updated SMT module and Sledgehammer to fully open source Z3
blanchet
parents: 59035
diff changeset
   170
  avail = make_avail "Z3",
372ddff01244 updated SMT module and Sledgehammer to fully open source Z3
blanchet
parents: 59035
diff changeset
   171
  command = make_command "Z3",
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   172
  options = z3_options,
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57237
diff changeset
   173
  smt_options = [(":produce-proofs", "true")],
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   174
  default_max_relevant = 350 (* FUDGE *),
70327
c04d4951a155 handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
blanchet
parents: 69593
diff changeset
   175
  outcome = on_first_line (outcome_of "unsat" "sat" "unknown" "timeout"),
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
   176
  parse_proof = SOME Z3_Replay.parse_proof,
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
   177
  replay = SOME Z3_Replay.replay }
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   178
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   179
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   180
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   181
(* smt tactic *)
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   182
val parse_smt_options =
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   183
  Scan.optional
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   184
    (Args.parens (Args.name -- Scan.option (\<^keyword>\<open>,\<close> |-- Args.name)) >> apfst SOME)
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   185
    (NONE, NONE)
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   186
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   187
fun smt_method ((solver, stgy), thms) ctxt facts =
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   188
  let
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   189
    val default_solver = SMT_Config.solver_of ctxt
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   190
    val solver = the_default default_solver solver
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   191
    val _ = 
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   192
      if solver = "z3" andalso stgy <> NONE
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   193
      then warning ("No strategy is available for z3. Ignoring " ^ quote (the stgy)) 
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   194
      else ()
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   195
    val ctxt =
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   196
      ctxt
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   197
      |> (if stgy <> NONE then Context.proof_map (Verit_Proof.select_veriT_stgy (the stgy)) else I)
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   198
      |> Context.Proof
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   199
      |> SMT_Config.select_solver solver
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   200
      |> Context.proof_of
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   201
  in
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   202
    HEADGOAL (SOLVED' (SMT_Solver.smt_tac ctxt (thms @ facts)))
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   203
  end
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   204
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   205
val _ =
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   206
  Theory.setup
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   207
    (Method.setup \<^binding>\<open>smt\<close>
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   208
      (Scan.lift parse_smt_options -- Attrib.thms >> (METHOD oo smt_method))
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 70327
diff changeset
   209
      "Call to the SMT solvers veriT or z3")
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   210
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   211
(* overall setup *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   212
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   213
val _ = Theory.setup (
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
   214
  SMT_Solver.add_solver cvc3 #>
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
   215
  SMT_Solver.add_solver cvc4 #>
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
   216
  SMT_Solver.add_solver veriT #>
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57714
diff changeset
   217
  SMT_Solver.add_solver z3)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   218
57229
blanchet
parents: 57210
diff changeset
   219
end;