| author | wenzelm | 
| Sat, 13 Mar 2021 12:45:31 +0100 | |
| changeset 73420 | 2c5d58e58fd2 | 
| parent 73388 | a40e69fde2b4 | 
| child 74048 | a0c9fc9c7dbe | 
| permissions | -rw-r--r-- | 
| 58061 | 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 | 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: 
59035diff
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 | 11 | end; | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 12 | |
| 58061 | 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: 
72458diff
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: 
72458diff
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: 
72458diff
changeset | 20 | "" => NONE | 
| 72479 | 21 | | s => | 
| 22 | if File.is_file (Path.variable var |> Path.expand |> Path.platform_exe) | |
| 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: 
72458diff
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: 
69593diff
changeset | 29 | fun outcome_of unsat sat unknown timeout solver_name line = | 
| 58061 | 30 | if String.isPrefix unsat line then SMT_Solver.Unsat | 
| 31 | else if String.isPrefix sat line then SMT_Solver.Sat | |
| 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: 
69593diff
changeset | 33 | else if String.isPrefix timeout line then SMT_Solver.Time_Out | 
| 58061 | 34 |   else raise SMT_Failure.SMT (SMT_Failure.Other_Failure ("Solver " ^ quote solver_name ^
 | 
| 35 | " failed -- enable tracing using the " ^ quote (Config.name_of SMT_Config.trace) ^ | |
| 56094 | 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: 
72667diff
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: 
72667diff
changeset | 39 | |
| 
6520d59fbdd7
ignore error messages produced by CVC4 when generating BV
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72667diff
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: 
72667diff
changeset | 41 | |
| 
6520d59fbdd7
ignore error messages produced by CVC4 when generating BV
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72667diff
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: 
72667diff
changeset | 43 | *) | 
| 60201 | 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: 
72667diff
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: 
72667diff
changeset | 46 | String.isPrefix "(error " s orelse String.isPrefix (getenv "ISABELLE_TMP_PREFIX") s | 
| 60201 | 47 | |
| 57239 | 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 | 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 | 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: 
67399diff
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 | 
| 73388 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 60 | fun cvc3_options ctxt = | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 61 | ["-seed", string_of_int (Config.get ctxt SMT_Config.random_seed), | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 62 | "-lang", "smt2"] @ | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 63 | (case SMT_Config.get_timeout ctxt of | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 64 | NONE => [] | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 65 | | SOME t => ["-timeout", string_of_int (Real.ceil (Time.toReal t))]) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 66 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 67 | |
| 58061 | 68 | val cvc3: SMT_Solver.solver_config = {
 | 
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57168diff
changeset | 69 | name = "cvc3", | 
| 58061 | 70 | 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: 
57209diff
changeset | 71 | 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: 
57209diff
changeset | 72 | command = make_command "CVC3", | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 73 | options = cvc3_options, | 
| 57239 | 74 | smt_options = [], | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 75 | default_max_relevant = 400 (* FUDGE *), | 
| 70327 
c04d4951a155
handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
 blanchet parents: 
69593diff
changeset | 76 | outcome = on_first_line (outcome_of "unsat" "sat" "unknown" "timeout"), | 
| 57157 | 77 | parse_proof = NONE, | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 78 | replay = NONE } | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 79 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 80 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 81 | |
| 66551 | 82 | |
| 57240 | 83 | (* CVC4 *) | 
| 84 | ||
| 69593 | 85 | val cvc4_extensions = Attrib.setup_config_bool \<^binding>\<open>cvc4_extensions\<close> (K false) | 
| 58360 | 86 | |
| 57240 | 87 | local | 
| 73388 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 88 | fun cvc4_options ctxt = | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 89 | ["--no-stats", | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 90 | "--random-seed=" ^ string_of_int (Config.get ctxt SMT_Config.random_seed), | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 91 | "--lang=smt2"] @ | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 92 | (case SMT_Config.get_timeout ctxt of | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 93 | NONE => [] | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 94 | | SOME t => ["--tlimit", string_of_int (Time.toMilliseconds t)]) | 
| 58360 | 95 | |
| 96 | fun select_class ctxt = | |
| 66551 | 97 | if Config.get ctxt cvc4_extensions then | 
| 98 | if Config.get ctxt SMT_Config.higher_order then | |
| 99 | CVC4_Interface.hosmtlib_cvc4C | |
| 100 | else | |
| 101 | CVC4_Interface.smtlib_cvc4C | |
| 102 | else | |
| 103 | if Config.get ctxt SMT_Config.higher_order then | |
| 104 | SMTLIB_Interface.hosmtlibC | |
| 105 | else | |
| 106 | SMTLIB_Interface.smtlibC | |
| 57240 | 107 | in | 
| 108 | ||
| 58061 | 109 | val cvc4: SMT_Solver.solver_config = {
 | 
| 57240 | 110 | name = "cvc4", | 
| 58360 | 111 | class = select_class, | 
| 57240 | 112 | avail = make_avail "CVC4", | 
| 113 | command = make_command "CVC4", | |
| 114 | options = cvc4_options, | |
| 59015 | 115 |   smt_options = [(":produce-unsat-cores", "true")],
 | 
| 57240 | 116 | default_max_relevant = 400 (* FUDGE *), | 
| 70327 
c04d4951a155
handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
 blanchet parents: 
69593diff
changeset | 117 | outcome = on_first_line (outcome_of "unsat" "sat" "unknown" "timeout"), | 
| 59015 | 118 | parse_proof = SOME (K CVC4_Proof_Parse.parse_proof), | 
| 57240 | 119 | replay = NONE } | 
| 120 | ||
| 121 | end | |
| 122 | ||
| 66551 | 123 | |
| 57704 | 124 | (* veriT *) | 
| 125 | ||
| 66551 | 126 | local | 
| 127 | fun select_class ctxt = | |
| 128 | if Config.get ctxt SMT_Config.higher_order then | |
| 129 | SMTLIB_Interface.hosmtlibC | |
| 130 | else | |
| 131 | SMTLIB_Interface.smtlibC | |
| 73388 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 132 | |
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 133 | fun veriT_options ctxt = | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 134 | ["--proof-with-sharing", | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 135 | "--proof-define-skolems", | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 136 | "--proof-prune", | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 137 | "--proof-merge", | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 138 | "--disable-print-success", | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 139 | "--disable-banner"] @ | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 140 | Verit_Proof.veriT_current_strategy (Context.Proof ctxt) | 
| 66551 | 141 | in | 
| 142 | ||
| 58061 | 143 | val veriT: SMT_Solver.solver_config = {
 | 
| 59035 
3a2153676705
renamed 'veriT' to 'verit', to stick to all-lowercase rule for prover names
 blanchet parents: 
59015diff
changeset | 144 | name = "verit", | 
| 66551 | 145 | 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: 
72458diff
changeset | 146 | 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: 
72458diff
changeset | 147 | command = the o check_tool "ISABELLE_VERIT", | 
| 73388 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 148 | options = veriT_options, | 
| 61587 
c3974cd2d381
updating options to verit
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
60201diff
changeset | 149 |   smt_options = [(":produce-proofs", "true")],
 | 
| 58496 
2ba52ecc4996
give more facts to veriT -- it seems to be able to cope with them
 blanchet parents: 
58491diff
changeset | 150 | default_max_relevant = 200 (* FUDGE *), | 
| 70327 
c04d4951a155
handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
 blanchet parents: 
69593diff
changeset | 151 | outcome = on_first_non_unsupported_line (outcome_of "unsat" "sat" "unknown" "timeout"), | 
| 58491 | 152 | 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: 
67522diff
changeset | 153 | replay = SOME Verit_Replay.replay } | 
| 57240 | 154 | |
| 66551 | 155 | end | 
| 156 | ||
| 157 | ||
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 158 | (* Z3 *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 159 | |
| 69593 | 160 | 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 | 161 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 162 | local | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 163 | fun z3_options ctxt = | 
| 58061 | 164 | ["smt.random_seed=" ^ string_of_int (Config.get ctxt SMT_Config.random_seed), | 
| 73388 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 165 | "smt.refine_inj_axioms=false"] @ | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 166 | (case SMT_Config.get_timeout ctxt of | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 167 | NONE => [] | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 168 | | SOME t => ["-T:" ^ string_of_int (Real.ceil (Time.toReal t))]) @ | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 169 | ["-smt2"] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 170 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 171 | fun select_class ctxt = | 
| 58061 | 172 | 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 | 173 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 174 | |
| 58061 | 175 | val z3: SMT_Solver.solver_config = {
 | 
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57168diff
changeset | 176 | name = "z3", | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 177 | class = select_class, | 
| 59960 
372ddff01244
updated SMT module and Sledgehammer to fully open source Z3
 blanchet parents: 
59035diff
changeset | 178 | avail = make_avail "Z3", | 
| 
372ddff01244
updated SMT module and Sledgehammer to fully open source Z3
 blanchet parents: 
59035diff
changeset | 179 | command = make_command "Z3", | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 180 | options = z3_options, | 
| 57239 | 181 |   smt_options = [(":produce-proofs", "true")],
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 182 | default_max_relevant = 350 (* FUDGE *), | 
| 70327 
c04d4951a155
handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
 blanchet parents: 
69593diff
changeset | 183 | outcome = on_first_line (outcome_of "unsat" "sat" "unknown" "timeout"), | 
| 58061 | 184 | parse_proof = SOME Z3_Replay.parse_proof, | 
| 185 | replay = SOME Z3_Replay.replay } | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 186 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 187 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 188 | |
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 189 | (* smt tactic *) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 190 | val parse_smt_options = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 191 | Scan.optional | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 192 | (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: 
70327diff
changeset | 193 | (NONE, NONE) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 194 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 195 | 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: 
70327diff
changeset | 196 | let | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 197 | 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: 
70327diff
changeset | 198 | val solver = the_default default_solver solver | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 199 | val _ = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 200 | if solver = "z3" andalso stgy <> NONE | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 201 |       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: 
70327diff
changeset | 202 | else () | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 203 | val ctxt = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 204 | ctxt | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 205 | |> (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: 
70327diff
changeset | 206 | |> Context.Proof | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 207 | |> SMT_Config.select_solver solver | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 208 | |> Context.proof_of | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 209 | in | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 210 | 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: 
70327diff
changeset | 211 | end | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 212 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 213 | val _ = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 214 | Theory.setup | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 215 | (Method.setup \<^binding>\<open>smt\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 216 | (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: 
70327diff
changeset | 217 | "Call to the SMT solvers veriT or z3") | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 218 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 219 | (* overall setup *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 220 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 221 | val _ = Theory.setup ( | 
| 58061 | 222 | SMT_Solver.add_solver cvc3 #> | 
| 223 | SMT_Solver.add_solver cvc4 #> | |
| 224 | SMT_Solver.add_solver veriT #> | |
| 225 | SMT_Solver.add_solver z3) | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 226 | |
| 57229 | 227 | end; |