| author | paulson | 
| Wed, 21 Aug 2024 14:09:44 +0100 | |
| changeset 80733 | 17d8b3f6d744 | 
| parent 79623 | e905fb37467f | 
| child 82090 | 023845c29d48 | 
| 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 | 
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 9 | val cvc_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 | |
| 75029 | 16 | val mashN = "mash" | 
| 17 | val mepoN = "mepo" | |
| 18 | val meshN = "mesh" | |
| 19 | ||
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 20 | (* helper functions *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 21 | |
| 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 | 22 | 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 | 23 | (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 | 24 | "" => NONE | 
| 72479 | 25 | | s => | 
| 26 | if File.is_file (Path.variable var |> Path.expand |> Path.platform_exe) | |
| 27 | 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 | 28 | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 29 | fun make_avail name () = getenv (name ^ "_SOLVER") <> "" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 30 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 31 | fun make_command name () = [getenv (name ^ "_SOLVER")] | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 32 | |
| 70327 
c04d4951a155
handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
 blanchet parents: 
69593diff
changeset | 33 | fun outcome_of unsat sat unknown timeout solver_name line = | 
| 58061 | 34 | if String.isPrefix unsat line then SMT_Solver.Unsat | 
| 35 | else if String.isPrefix sat line then SMT_Solver.Sat | |
| 36 | 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 | 37 | else if String.isPrefix timeout line then SMT_Solver.Time_Out | 
| 58061 | 38 |   else raise SMT_Failure.SMT (SMT_Failure.Other_Failure ("Solver " ^ quote solver_name ^
 | 
| 39 | " failed -- enable tracing using the " ^ quote (Config.name_of SMT_Config.trace) ^ | |
| 56094 | 40 | " option for details")) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 41 | |
| 73104 
6520d59fbdd7
ignore error messages produced by CVC4 when generating BV
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72667diff
changeset | 42 | (* 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 | 43 | |
| 
6520d59fbdd7
ignore error messages produced by CVC4 when generating BV
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72667diff
changeset | 44 | $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 | 45 | |
| 
6520d59fbdd7
ignore error messages produced by CVC4 when generating BV
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72667diff
changeset | 46 | 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 | 47 | *) | 
| 60201 | 48 | 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 | 49 | | 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 | 50 | String.isPrefix "(error " s orelse String.isPrefix (getenv "ISABELLE_TMP_PREFIX") s | 
| 60201 | 51 | |
| 57239 | 52 | 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 | 53 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 54 |     val split_first = (fn [] => ("", []) | l :: ls => (l, ls))
 | 
| 67522 | 55 | 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 | 56 | in (test_outcome solver_name l, ls) end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 57 | |
| 57704 | 58 | 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 | 59 | 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 | 60 | |
| 66551 | 61 | |
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 62 | (* CVC4 and cvc5 *) | 
| 57240 | 63 | |
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 64 | val cvc_extensions = Attrib.setup_config_bool \<^binding>\<open>cvc_extensions\<close> (K false) | 
| 58360 | 65 | |
| 57240 | 66 | local | 
| 73388 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 67 | fun cvc4_options ctxt = | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 68 | ["--no-stats", | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 69 | "--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 | 70 | "--lang=smt2"] @ | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 71 | (case SMT_Config.get_timeout ctxt of | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 72 | NONE => [] | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 73 | | SOME t => ["--tlimit", string_of_int (Time.toMilliseconds t)]) | 
| 58360 | 74 | |
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 75 | fun cvc5_options ctxt = | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 76 | ["--no-stats", | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 77 | "--sat-random-seed=" ^ string_of_int (Config.get ctxt SMT_Config.random_seed), | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 78 | "--lang=smt2"] @ | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 79 | (case SMT_Config.get_timeout ctxt of | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 80 | NONE => [] | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 81 | | SOME t => ["--tlimit", string_of_int (Time.toMilliseconds t)]) | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 82 | |
| 58360 | 83 | fun select_class ctxt = | 
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 84 | if Config.get ctxt cvc_extensions then | 
| 66551 | 85 | if Config.get ctxt SMT_Config.higher_order then | 
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 86 | CVC_Interface.hosmtlib_cvcC | 
| 66551 | 87 | else | 
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 88 | CVC_Interface.smtlib_cvcC | 
| 66551 | 89 | else | 
| 90 | if Config.get ctxt SMT_Config.higher_order then | |
| 91 | SMTLIB_Interface.hosmtlibC | |
| 74817 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74691diff
changeset | 92 | else if Config.get ctxt SMT_Config.native_bv then | 
| 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74691diff
changeset | 93 | SMTLIB_Interface.bvsmlibC | 
| 66551 | 94 | else | 
| 95 | SMTLIB_Interface.smtlibC | |
| 57240 | 96 | in | 
| 97 | ||
| 58061 | 98 | val cvc4: SMT_Solver.solver_config = {
 | 
| 57240 | 99 | name = "cvc4", | 
| 58360 | 100 | class = select_class, | 
| 57240 | 101 | avail = make_avail "CVC4", | 
| 102 | command = make_command "CVC4", | |
| 103 | options = cvc4_options, | |
| 59015 | 104 |   smt_options = [(":produce-unsat-cores", "true")],
 | 
| 75029 | 105 | good_slices = | 
| 106 | (* FUDGE *) | |
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 107 | [((2, false, false, 512, meshN), ["--full-saturate-quant", "--inst-when=full-last-call", "--inst-no-entail", "--term-db-mode=relevant", "--multi-trigger-linear"]), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 108 | ((2, false, false, 64, meshN), ["--decision=internal", "--simplification=none", "--full-saturate-quant"]), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 109 | ((2, false, true, 256, mepoN), ["--trigger-sel=max", "--full-saturate-quant"]), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 110 | ((2, false, false, 1024, meshN), ["--relevant-triggers", "--full-saturate-quant"]), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 111 | ((2, false, false, 32, meshN), ["--term-db-mode=relevant", "--full-saturate-quant"]), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 112 | ((2, false, false, 128, meshN), ["--no-e-matching", "--full-saturate-quant"]), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 113 | ((2, false, false, 256, meshN), ["--finite-model-find", "--fmf-inst-engine"])], | 
| 70327 
c04d4951a155
handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
 blanchet parents: 
69593diff
changeset | 114 | outcome = on_first_line (outcome_of "unsat" "sat" "unknown" "timeout"), | 
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 115 | parse_proof = SOME (K CVC_Proof_Parse.parse_proof), | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 116 | replay = NONE } | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 117 | |
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 118 | val cvc5: SMT_Solver.solver_config = {
 | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 119 | name = "cvc5", | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 120 | class = select_class, | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 121 | avail = make_avail "CVC5", | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 122 | command = make_command "CVC5", | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 123 | options = cvc5_options, | 
| 79623 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 124 |   smt_options = [(":produce-unsat-cores", "true")],
 | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 125 | good_slices = | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 126 | (* FUDGE *) | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 127 | [((2, false, false, 512, meshN), ["--full-saturate-quant", "--inst-when=full-last-call", "--inst-no-entail", "--term-db-mode=relevant", "--multi-trigger-linear"]), | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 128 | ((2, false, false, 64, meshN), ["--decision=internal", "--simplification=none", "--full-saturate-quant"]), | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 129 | ((2, false, true, 256, mepoN), ["--trigger-sel=max", "--full-saturate-quant"]), | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 130 | ((2, false, false, 1024, meshN), ["--relevant-triggers", "--full-saturate-quant"]), | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 131 | ((2, false, false, 32, meshN), ["--term-db-mode=relevant", "--full-saturate-quant"]), | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 132 | ((2, false, false, 128, meshN), ["--no-e-matching", "--full-saturate-quant"]), | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 133 | ((2, false, false, 256, meshN), ["--finite-model-find", "--fmf-inst-engine"])], | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 134 | outcome = on_first_line (outcome_of "unsat" "sat" "unknown" "timeout"), | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 135 | parse_proof = SOME (K CVC_Proof_Parse.parse_proof), | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 136 | replay = NONE } | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 137 | |
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 138 | (* | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 139 | We need different options for alethe proof production by cvc5 and the smt_options | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 140 | cannot be changed, so different configuration. | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 141 | *) | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 142 | val cvc5_proof: SMT_Solver.solver_config = {
 | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 143 | name = "cvc5_proof", | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 144 | class = select_class, | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 145 | avail = make_avail "CVC5", | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 146 | command = make_command "CVC5", | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 147 | options = cvc5_options, | 
| 78177 
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
77269diff
changeset | 148 |   smt_options = [(":produce-proofs", "true")],
 | 
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 149 | good_slices = | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 150 | (* FUDGE *) | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 151 | [((2, false, false, 512, meshN), ["--full-saturate-quant", "--inst-when=full-last-call", "--inst-no-entail", "--term-db-mode=relevant", "--multi-trigger-linear"]), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 152 | ((2, false, false, 64, meshN), ["--decision=internal", "--simplification=none", "--full-saturate-quant"]), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 153 | ((2, false, true, 256, mepoN), ["--trigger-sel=max", "--full-saturate-quant"]), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 154 | ((2, false, false, 1024, meshN), ["--relevant-triggers", "--full-saturate-quant"]), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 155 | ((2, false, false, 32, meshN), ["--term-db-mode=relevant", "--full-saturate-quant"]), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 156 | ((2, false, false, 128, meshN), ["--no-e-matching", "--full-saturate-quant"]), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 157 | ((2, false, false, 256, meshN), ["--finite-model-find", "--fmf-inst-engine"])], | 
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 158 | outcome = on_first_line (outcome_of "unsat" "sat" "unknown" "timeout"), | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 159 | parse_proof = SOME (K CVC_Proof_Parse.parse_proof), | 
| 78177 
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
77269diff
changeset | 160 | replay = SOME CVC5_Replay.replay } | 
| 57240 | 161 | |
| 162 | end | |
| 163 | ||
| 66551 | 164 | |
| 57704 | 165 | (* veriT *) | 
| 166 | ||
| 66551 | 167 | local | 
| 168 | fun select_class ctxt = | |
| 169 | if Config.get ctxt SMT_Config.higher_order then | |
| 170 | SMTLIB_Interface.hosmtlibC | |
| 171 | else | |
| 172 | SMTLIB_Interface.smtlibC | |
| 73388 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 173 | |
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 174 | fun veriT_options ctxt = | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 175 | ["--proof-with-sharing", | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 176 | "--proof-define-skolems", | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 177 | "--proof-prune", | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 178 | "--proof-merge", | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 179 | "--disable-print-success", | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 180 | "--disable-banner"] @ | 
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75872diff
changeset | 181 | Verit_Strategies.veriT_current_strategy (Context.Proof ctxt) @ | 
| 74476 | 182 | (case SMT_Config.get_timeout ctxt of | 
| 183 | NONE => [] | |
| 74553 | 184 | | SOME t => ["--max-time=" ^ string_of_int (Time.toMilliseconds t)]) | 
| 66551 | 185 | in | 
| 186 | ||
| 58061 | 187 | 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 | 188 | name = "verit", | 
| 66551 | 189 | 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 | 190 | 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 | 191 | command = the o check_tool "ISABELLE_VERIT", | 
| 73388 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 192 | options = veriT_options, | 
| 61587 
c3974cd2d381
updating options to verit
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
60201diff
changeset | 193 |   smt_options = [(":produce-proofs", "true")],
 | 
| 75029 | 194 | good_slices = | 
| 195 | (* FUDGE *) | |
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 196 | [((2, false, false, 1024, meshN), []), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 197 | ((2, false, false, 512, mashN), []), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 198 | ((2, false, true, 128, meshN), []), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 199 | ((2, false, false, 64, meshN), []), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 200 | ((2, false, false, 256, mepoN), []), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 201 | ((2, false, false, 32, meshN), [])], | 
| 74691 
634e2323b6cf
proper support of verit's return code for timeout
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74553diff
changeset | 202 | outcome = on_first_non_unsupported_line (outcome_of "unsat" "sat" "unknown" "Time limit exceeded"), | 
| 75299 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75063diff
changeset | 203 | parse_proof = SOME (K Lethe_Proof_Parse.parse_proof), | 
| 69205 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
67522diff
changeset | 204 | replay = SOME Verit_Replay.replay } | 
| 57240 | 205 | |
| 66551 | 206 | end | 
| 207 | ||
| 208 | ||
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 209 | (* Z3 *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 210 | |
| 69593 | 211 | 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 | 212 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 213 | local | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 214 | fun z3_options ctxt = | 
| 58061 | 215 | ["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 | 216 | "smt.refine_inj_axioms=false"] @ | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 217 | (case SMT_Config.get_timeout ctxt of | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 218 | NONE => [] | 
| 
a40e69fde2b4
clarified smt: support Timeout.ignored and Timeout.scale_time;
 wenzelm parents: 
73104diff
changeset | 219 | | 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 | 220 | ["-smt2"] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 221 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 222 | fun select_class ctxt = | 
| 74817 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74691diff
changeset | 223 | if Config.get ctxt z3_extensions then Z3_Interface.smtlib_z3C | 
| 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74691diff
changeset | 224 | else if Config.get ctxt SMT_Config.native_bv then SMTLIB_Interface.bvsmlibC | 
| 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74691diff
changeset | 225 | else SMTLIB_Interface.smtlibC | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 226 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 227 | |
| 58061 | 228 | val z3: SMT_Solver.solver_config = {
 | 
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57168diff
changeset | 229 | name = "z3", | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 230 | class = select_class, | 
| 59960 
372ddff01244
updated SMT module and Sledgehammer to fully open source Z3
 blanchet parents: 
59035diff
changeset | 231 | avail = make_avail "Z3", | 
| 
372ddff01244
updated SMT module and Sledgehammer to fully open source Z3
 blanchet parents: 
59035diff
changeset | 232 | command = make_command "Z3", | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 233 | options = z3_options, | 
| 57239 | 234 |   smt_options = [(":produce-proofs", "true")],
 | 
| 75029 | 235 | good_slices = | 
| 236 | (* FUDGE *) | |
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 237 | [((2, false, false, 1024, meshN), []), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 238 | ((2, false, false, 512, mepoN), []), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 239 | ((2, false, false, 64, meshN), []), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 240 | ((2, false, true, 256, meshN), []), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 241 | ((2, false, false, 128, mashN), []), | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 242 | ((2, false, false, 32, meshN), [])], | 
| 70327 
c04d4951a155
handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
 blanchet parents: 
69593diff
changeset | 243 | outcome = on_first_line (outcome_of "unsat" "sat" "unknown" "timeout"), | 
| 58061 | 244 | parse_proof = SOME Z3_Replay.parse_proof, | 
| 245 | replay = SOME Z3_Replay.replay } | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 246 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 247 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 248 | |
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 249 | (* smt tactic *) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 250 | val parse_smt_options = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 251 | Scan.optional | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 252 | (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 | 253 | (NONE, NONE) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 254 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 255 | 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 | 256 | let | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 257 | val default_solver = SMT_Config.solver_of ctxt | 
| 79623 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 258 | val solver = | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 259 | (case solver of | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 260 | NONE => default_solver | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 261 | | SOME "cvc5" => "cvc5_proof" | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 262 | | SOME a => a) | 
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 263 | val _ = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 264 | 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 | 265 |       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 | 266 | else () | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 267 | val ctxt = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 268 | ctxt | 
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75872diff
changeset | 269 | |> (if stgy <> NONE then Context.proof_map (Verit_Strategies.select_veriT_stgy (the stgy)) else I) | 
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 270 | |> Context.Proof | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 271 | |> SMT_Config.select_solver solver | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 272 | |> Context.proof_of | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 273 | in | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 274 | 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 | 275 | end | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 276 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 277 | val _ = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 278 | Theory.setup | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
70327diff
changeset | 279 | (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 | 280 | (Scan.lift parse_smt_options -- Attrib.thms >> (METHOD oo smt_method)) | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
75956diff
changeset | 281 | "Call to the SMT solver veriT or z3") | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 282 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 283 | (* overall setup *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 284 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 285 | val _ = Theory.setup ( | 
| 58061 | 286 | SMT_Solver.add_solver cvc4 #> | 
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75339diff
changeset | 287 | SMT_Solver.add_solver cvc5 #> | 
| 79623 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
78177diff
changeset | 288 | SMT_Solver.add_solver cvc5_proof #> | 
| 58061 | 289 | SMT_Solver.add_solver veriT #> | 
| 290 | SMT_Solver.add_solver z3) | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 291 | |
| 57229 | 292 | end; |