| author | wenzelm | 
| Mon, 08 Nov 2010 11:28:22 +0100 | |
| changeset 40407 | 2ff10e613689 | 
| parent 40277 | 4e3a3461c1a6 | 
| child 40424 | 7550b2cba1cb | 
| permissions | -rw-r--r-- | 
| 36898 | 1 | (* Title: HOL/SMT.thy | 
| 2 | Author: Sascha Boehme, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 5 | header {* Bindings to Satisfiability Modulo Theories (SMT) solvers *}
 | |
| 6 | ||
| 7 | theory SMT | |
| 8 | imports List | |
| 9 | uses | |
| 39483 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 boehmes parents: 
39298diff
changeset | 10 | "Tools/Datatype/datatype_selectors.ML" | 
| 36898 | 11 |   ("Tools/SMT/smt_monomorph.ML")
 | 
| 40277 | 12 |   ("Tools/SMT/smt_builtin.ML")
 | 
| 36898 | 13 |   ("Tools/SMT/smt_normalize.ML")
 | 
| 14 |   ("Tools/SMT/smt_translate.ML")
 | |
| 15 |   ("Tools/SMT/smt_solver.ML")
 | |
| 16 |   ("Tools/SMT/smtlib_interface.ML")
 | |
| 17 |   ("Tools/SMT/z3_proof_parser.ML")
 | |
| 18 |   ("Tools/SMT/z3_proof_tools.ML")
 | |
| 19 |   ("Tools/SMT/z3_proof_literals.ML")
 | |
| 20 |   ("Tools/SMT/z3_proof_reconstruction.ML")
 | |
| 21 |   ("Tools/SMT/z3_model.ML")
 | |
| 22 |   ("Tools/SMT/z3_interface.ML")
 | |
| 40162 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 23 |   ("Tools/SMT/smt_setup_solvers.ML")
 | 
| 36898 | 24 | begin | 
| 25 | ||
| 26 | ||
| 27 | ||
| 36902 
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
 huffman parents: 
36899diff
changeset | 28 | subsection {* Triggers for quantifier instantiation *}
 | 
| 36898 | 29 | |
| 30 | text {*
 | |
| 31 | Some SMT solvers support triggers for quantifier instantiation. | |
| 32 | Each trigger consists of one ore more patterns. A pattern may either | |
| 37124 | 33 | be a list of positive subterms (each being tagged by "pat"), or a | 
| 34 | list of negative subterms (each being tagged by "nopat"). | |
| 35 | ||
| 36 | When an SMT solver finds a term matching a positive pattern (a | |
| 37 | pattern with positive subterms only), it instantiates the | |
| 38 | corresponding quantifier accordingly. Negative patterns inhibit | |
| 39 | quantifier instantiations. Each pattern should mention all preceding | |
| 40 | bound variables. | |
| 36898 | 41 | *} | 
| 42 | ||
| 43 | datatype pattern = Pattern | |
| 44 | ||
| 37124 | 45 | definition pat :: "'a \<Rightarrow> pattern" where "pat _ = Pattern" | 
| 46 | definition nopat :: "'a \<Rightarrow> pattern" where "nopat _ = Pattern" | |
| 36898 | 47 | |
| 37124 | 48 | definition trigger :: "pattern list list \<Rightarrow> bool \<Rightarrow> bool" | 
| 36898 | 49 | where "trigger _ P = P" | 
| 50 | ||
| 51 | ||
| 52 | ||
| 40274 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 53 | subsection {* Distinctness *}
 | 
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 54 | |
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 55 | text {*
 | 
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 56 | As an abbreviation for a quadratic number of inequalities, SMT solvers | 
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 57 | provide a built-in @{text distinct}.  To avoid confusion with the
 | 
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 58 | already defined (and more general) @{term List.distinct}, a separate
 | 
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 59 | constant is defined. | 
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 60 | *} | 
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 61 | |
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 62 | definition distinct :: "'a list \<Rightarrow> bool" | 
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 63 | where "distinct xs = List.distinct xs" | 
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 64 | |
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 65 | |
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 66 | |
| 36902 
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
 huffman parents: 
36899diff
changeset | 67 | subsection {* Higher-order encoding *}
 | 
| 36898 | 68 | |
| 69 | text {*
 | |
| 70 | Application is made explicit for constants occurring with varying | |
| 71 | numbers of arguments. This is achieved by the introduction of the | |
| 72 | following constant. | |
| 73 | *} | |
| 74 | ||
| 37153 
8feed34275ce
renamed constant "apply" to "fun_app" (which is closer to the related "fun_upd")
 boehmes parents: 
37151diff
changeset | 75 | definition fun_app where "fun_app f x = f x" | 
| 36898 | 76 | |
| 77 | text {*
 | |
| 78 | Some solvers support a theory of arrays which can be used to encode | |
| 79 | higher-order functions. The following set of lemmas specifies the | |
| 80 | properties of such (extensional) arrays. | |
| 81 | *} | |
| 82 | ||
| 83 | lemmas array_rules = ext fun_upd_apply fun_upd_same fun_upd_other | |
| 37157 | 84 | fun_upd_upd fun_app_def | 
| 36898 | 85 | |
| 86 | ||
| 87 | ||
| 36902 
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
 huffman parents: 
36899diff
changeset | 88 | subsection {* First-order logic *}
 | 
| 36898 | 89 | |
| 90 | text {*
 | |
| 91 | Some SMT solvers require a strict separation between formulas and | |
| 92 | terms. When translating higher-order into first-order problems, | |
| 93 | all uninterpreted constants (those not builtin in the target solver) | |
| 94 | are treated as function symbols in the first-order sense. Their | |
| 95 | occurrences as head symbols in atoms (i.e., as predicate symbols) is | |
| 96 | turned into terms by equating such atoms with @{term True} using the
 | |
| 97 | following term-level equation symbol. | |
| 98 | *} | |
| 99 | ||
| 37124 | 100 | definition term_eq :: "bool \<Rightarrow> bool \<Rightarrow> bool" where "term_eq x y = (x = y)" | 
| 36898 | 101 | |
| 102 | ||
| 103 | ||
| 37151 | 104 | subsection {* Integer division and modulo for Z3 *}
 | 
| 105 | ||
| 106 | definition z3div :: "int \<Rightarrow> int \<Rightarrow> int" where | |
| 107 | "z3div k l = (if 0 \<le> l then k div l else -(k div (-l)))" | |
| 108 | ||
| 109 | definition z3mod :: "int \<Rightarrow> int \<Rightarrow> int" where | |
| 110 | "z3mod k l = (if 0 \<le> l then k mod l else k mod (-l))" | |
| 111 | ||
| 112 | lemma div_by_z3div: "k div l = ( | |
| 113 | if k = 0 \<or> l = 0 then 0 | |
| 114 | else if (0 < k \<and> 0 < l) \<or> (k < 0 \<and> 0 < l) then z3div k l | |
| 115 | else z3div (-k) (-l))" | |
| 116 | by (auto simp add: z3div_def) | |
| 117 | ||
| 118 | lemma mod_by_z3mod: "k mod l = ( | |
| 119 | if l = 0 then k | |
| 120 | else if k = 0 then 0 | |
| 121 | else if (0 < k \<and> 0 < l) \<or> (k < 0 \<and> 0 < l) then z3mod k l | |
| 122 | else - z3mod (-k) (-l))" | |
| 123 | by (auto simp add: z3mod_def) | |
| 124 | ||
| 125 | ||
| 126 | ||
| 36902 
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
 huffman parents: 
36899diff
changeset | 127 | subsection {* Setup *}
 | 
| 36898 | 128 | |
| 129 | use "Tools/SMT/smt_monomorph.ML" | |
| 40277 | 130 | use "Tools/SMT/smt_builtin.ML" | 
| 36898 | 131 | use "Tools/SMT/smt_normalize.ML" | 
| 132 | use "Tools/SMT/smt_translate.ML" | |
| 133 | use "Tools/SMT/smt_solver.ML" | |
| 134 | use "Tools/SMT/smtlib_interface.ML" | |
| 135 | use "Tools/SMT/z3_interface.ML" | |
| 136 | use "Tools/SMT/z3_proof_parser.ML" | |
| 137 | use "Tools/SMT/z3_proof_tools.ML" | |
| 138 | use "Tools/SMT/z3_proof_literals.ML" | |
| 139 | use "Tools/SMT/z3_proof_reconstruction.ML" | |
| 140 | use "Tools/SMT/z3_model.ML" | |
| 40162 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 141 | use "Tools/SMT/smt_setup_solvers.ML" | 
| 36898 | 142 | |
| 143 | setup {*
 | |
| 144 | SMT_Solver.setup #> | |
| 145 | Z3_Proof_Reconstruction.setup #> | |
| 40162 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 146 | SMT_Setup_Solvers.setup | 
| 36898 | 147 | *} | 
| 148 | ||
| 149 | ||
| 150 | ||
| 36902 
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
 huffman parents: 
36899diff
changeset | 151 | subsection {* Configuration *}
 | 
| 36898 | 152 | |
| 153 | text {*
 | |
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 154 | The current configuration can be printed by the command | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 155 | @{text smt_status}, which shows the values of most options.
 | 
| 36898 | 156 | *} | 
| 157 | ||
| 158 | ||
| 159 | ||
| 160 | subsection {* General configuration options *}
 | |
| 161 | ||
| 162 | text {*
 | |
| 163 | The option @{text smt_solver} can be used to change the target SMT
 | |
| 164 | solver.  The possible values are @{text cvc3}, @{text yices}, and
 | |
| 165 | @{text z3}.  It is advisable to locally install the selected solver,
 | |
| 166 | although this is not necessary for @{text cvc3} and @{text z3}, which
 | |
| 167 | can also be used over an Internet-based service. | |
| 168 | ||
| 169 | When using local SMT solvers, the path to their binaries should be | |
| 170 | declared by setting the following environment variables: | |
| 171 | @{text CVC3_SOLVER}, @{text YICES_SOLVER}, and @{text Z3_SOLVER}.
 | |
| 172 | *} | |
| 173 | ||
| 174 | declare [[ smt_solver = z3 ]] | |
| 175 | ||
| 176 | text {*
 | |
| 177 | Since SMT solvers are potentially non-terminating, there is a timeout | |
| 178 | (given in seconds) to restrict their runtime. A value greater than | |
| 179 | 120 (seconds) is in most cases not advisable. | |
| 180 | *} | |
| 181 | ||
| 182 | declare [[ smt_timeout = 20 ]] | |
| 183 | ||
| 40162 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 184 | text {*
 | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 185 | In general, the binding to SMT solvers runs as an oracle, i.e, the SMT | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 186 | solvers are fully trusted without additional checks. The following | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 187 | option can cause the SMT solver to run in proof-producing mode, giving | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 188 | a checkable certificate. This is currently only implemented for Z3. | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 189 | *} | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 190 | |
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 191 | declare [[ smt_oracle = false ]] | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 192 | |
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 193 | text {*
 | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 194 | Each SMT solver provides several commandline options to tweak its | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 195 | behaviour. They can be passed to the solver by setting the following | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 196 | options. | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 197 | *} | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 198 | |
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 199 | declare [[ cvc3_options = "", yices_options = "", z3_options = "" ]] | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 200 | |
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 201 | text {*
 | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 202 | Enable the following option to use built-in support for datatypes and | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 203 | records. Currently, this is only implemented for Z3 running in oracle | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 204 | mode. | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 205 | *} | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 206 | |
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 207 | declare [[ smt_datatypes = false ]] | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 208 | |
| 36898 | 209 | |
| 210 | ||
| 211 | subsection {* Certificates *}
 | |
| 212 | ||
| 213 | text {*
 | |
| 214 | By setting the option @{text smt_certificates} to the name of a file,
 | |
| 215 | all following applications of an SMT solver a cached in that file. | |
| 216 | Any further application of the same SMT solver (using the very same | |
| 217 | configuration) re-uses the cached certificate instead of invoking the | |
| 218 | solver. An empty string disables caching certificates. | |
| 219 | ||
| 220 | The filename should be given as an explicit path. It is good | |
| 221 | practice to use the name of the current theory (with ending | |
| 222 | @{text ".certs"} instead of @{text ".thy"}) as the certificates file.
 | |
| 223 | *} | |
| 224 | ||
| 225 | declare [[ smt_certificates = "" ]] | |
| 226 | ||
| 227 | text {*
 | |
| 228 | The option @{text smt_fixed} controls whether only stored
 | |
| 229 | certificates are should be used or invocation of an SMT solver is | |
| 230 | allowed.  When set to @{text true}, no SMT solver will ever be
 | |
| 231 | invoked and only the existing certificates found in the configured | |
| 232 | cache are used;  when set to @{text false} and there is no cached
 | |
| 233 | certificate for some proposition, then the configured SMT solver is | |
| 234 | invoked. | |
| 235 | *} | |
| 236 | ||
| 237 | declare [[ smt_fixed = false ]] | |
| 238 | ||
| 239 | ||
| 240 | ||
| 241 | subsection {* Tracing *}
 | |
| 242 | ||
| 243 | text {*
 | |
| 244 | For tracing the generated problem file given to the SMT solver as | |
| 245 | well as the returned result of the solver, the option | |
| 246 | @{text smt_trace} should be set to @{text true}.
 | |
| 247 | *} | |
| 248 | ||
| 249 | declare [[ smt_trace = false ]] | |
| 250 | ||
| 251 | text {*
 | |
| 40162 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 252 | From the set of assumptions given to the SMT solver, those assumptions | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 253 | used in the proof are traced when the following option is set to | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 254 | @{term true}.  This only works for Z3 when it runs in non-oracle mode
 | 
| 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 255 | (see options @{text smt_solver} and @{text smt_oracle} above).
 | 
| 36898 | 256 | *} | 
| 257 | ||
| 40162 
7f58a9a843c2
joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
 boehmes parents: 
39483diff
changeset | 258 | declare [[ smt_trace_used_facts = false ]] | 
| 39298 
5aefb5bc8a93
added preliminary support for SMT datatypes (for now restricted to tuples and lists); only the Z3 interface (in oracle mode) makes use of it, there is especially no Z3 proof reconstruction support for datatypes yet
 boehmes parents: 
37818diff
changeset | 259 | |
| 36898 | 260 | |
| 261 | ||
| 36902 
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
 huffman parents: 
36899diff
changeset | 262 | subsection {* Schematic rules for Z3 proof reconstruction *}
 | 
| 36898 | 263 | |
| 264 | text {*
 | |
| 265 | Several prof rules of Z3 are not very well documented. There are two | |
| 266 | lemma groups which can turn failing Z3 proof reconstruction attempts | |
| 267 | into succeeding ones: the facts in @{text z3_rule} are tried prior to
 | |
| 268 | any implemented reconstruction procedure for all uncertain Z3 proof | |
| 269 | rules;  the facts in @{text z3_simp} are only fed to invocations of
 | |
| 270 | the simplifier when reconstructing theory-specific proof steps. | |
| 271 | *} | |
| 272 | ||
| 273 | lemmas [z3_rule] = | |
| 274 | refl eq_commute conj_commute disj_commute simp_thms nnf_simps | |
| 275 | ring_distribs field_simps times_divide_eq_right times_divide_eq_left | |
| 276 | if_True if_False not_not | |
| 277 | ||
| 278 | lemma [z3_rule]: | |
| 279 | "(P \<longrightarrow> Q) = (Q \<or> \<not>P)" | |
| 280 | "(\<not>P \<longrightarrow> Q) = (P \<or> Q)" | |
| 281 | "(\<not>P \<longrightarrow> Q) = (Q \<or> P)" | |
| 282 | by auto | |
| 283 | ||
| 284 | lemma [z3_rule]: | |
| 285 | "((P = Q) \<longrightarrow> R) = (R | (Q = (\<not>P)))" | |
| 286 | by auto | |
| 287 | ||
| 288 | lemma [z3_rule]: | |
| 289 | "((\<not>P) = P) = False" | |
| 290 | "(P = (\<not>P)) = False" | |
| 291 | "(P \<noteq> Q) = (Q = (\<not>P))" | |
| 292 | "(P = Q) = ((\<not>P \<or> Q) \<and> (P \<or> \<not>Q))" | |
| 293 | "(P \<noteq> Q) = ((\<not>P \<or> \<not>Q) \<and> (P \<or> Q))" | |
| 294 | by auto | |
| 295 | ||
| 296 | lemma [z3_rule]: | |
| 297 | "(if P then P else \<not>P) = True" | |
| 298 | "(if \<not>P then \<not>P else P) = True" | |
| 299 | "(if P then True else False) = P" | |
| 300 | "(if P then False else True) = (\<not>P)" | |
| 301 | "(if \<not>P then x else y) = (if P then y else x)" | |
| 302 | by auto | |
| 303 | ||
| 304 | lemma [z3_rule]: | |
| 305 | "P = Q \<or> P \<or> Q" | |
| 306 | "P = Q \<or> \<not>P \<or> \<not>Q" | |
| 307 | "(\<not>P) = Q \<or> \<not>P \<or> Q" | |
| 308 | "(\<not>P) = Q \<or> P \<or> \<not>Q" | |
| 309 | "P = (\<not>Q) \<or> \<not>P \<or> Q" | |
| 310 | "P = (\<not>Q) \<or> P \<or> \<not>Q" | |
| 311 | "P \<noteq> Q \<or> P \<or> \<not>Q" | |
| 312 | "P \<noteq> Q \<or> \<not>P \<or> Q" | |
| 313 | "P \<noteq> (\<not>Q) \<or> P \<or> Q" | |
| 314 | "(\<not>P) \<noteq> Q \<or> P \<or> Q" | |
| 315 | "P \<or> Q \<or> P \<noteq> (\<not>Q)" | |
| 316 | "P \<or> Q \<or> (\<not>P) \<noteq> Q" | |
| 317 | "P \<or> \<not>Q \<or> P \<noteq> Q" | |
| 318 | "\<not>P \<or> Q \<or> P \<noteq> Q" | |
| 319 | by auto | |
| 320 | ||
| 321 | lemma [z3_rule]: | |
| 322 | "0 + (x::int) = x" | |
| 323 | "x + 0 = x" | |
| 324 | "0 * x = 0" | |
| 325 | "1 * x = x" | |
| 326 | "x + y = y + x" | |
| 327 | by auto | |
| 328 | ||
| 37124 | 329 | |
| 330 | ||
| 331 | hide_type (open) pattern | |
| 37153 
8feed34275ce
renamed constant "apply" to "fun_app" (which is closer to the related "fun_upd")
 boehmes parents: 
37151diff
changeset | 332 | hide_const Pattern term_eq | 
| 40274 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40162diff
changeset | 333 | hide_const (open) trigger pat nopat distinct fun_app z3div z3mod | 
| 37124 | 334 | |
| 39483 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 boehmes parents: 
39298diff
changeset | 335 | |
| 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 boehmes parents: 
39298diff
changeset | 336 | |
| 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 boehmes parents: 
39298diff
changeset | 337 | subsection {* Selectors for datatypes *}
 | 
| 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 boehmes parents: 
39298diff
changeset | 338 | |
| 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 boehmes parents: 
39298diff
changeset | 339 | setup {* Datatype_Selectors.setup *}
 | 
| 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 boehmes parents: 
39298diff
changeset | 340 | |
| 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 boehmes parents: 
39298diff
changeset | 341 | declare [[ selector Pair 1 = fst, selector Pair 2 = snd ]] | 
| 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 boehmes parents: 
39298diff
changeset | 342 | declare [[ selector Cons 1 = hd, selector Cons 2 = tl ]] | 
| 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 boehmes parents: 
39298diff
changeset | 343 | |
| 36898 | 344 | end |