author | wenzelm |
Tue, 22 Apr 2014 12:03:58 +0200 | |
changeset 56629 | ca302c495bca |
parent 56078 | 624faeda77b5 |
child 57213 | 9daec42f6784 |
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 |
|
41426
09615ed31f04
re-implemented support for datatypes (including records and typedefs);
boehmes
parents:
41328
diff
changeset
|
8 |
imports Record |
46950
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents:
45392
diff
changeset
|
9 |
keywords "smt_status" :: diag |
36898 | 10 |
begin |
11 |
||
48892 | 12 |
ML_file "Tools/SMT/smt_utils.ML" |
13 |
ML_file "Tools/SMT/smt_failure.ML" |
|
14 |
ML_file "Tools/SMT/smt_config.ML" |
|
36898 | 15 |
|
16 |
||
36902
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
huffman
parents:
36899
diff
changeset
|
17 |
subsection {* Triggers for quantifier instantiation *} |
36898 | 18 |
|
19 |
text {* |
|
41125 | 20 |
Some SMT solvers support patterns as a quantifier instantiation |
21 |
heuristics. Patterns may either be positive terms (tagged by "pat") |
|
22 |
triggering quantifier instantiations -- when the solver finds a |
|
23 |
term matching a positive pattern, it instantiates the corresponding |
|
24 |
quantifier accordingly -- or negative terms (tagged by "nopat") |
|
25 |
inhibiting quantifier instantiations. A list of patterns |
|
26 |
of the same kind is called a multipattern, and all patterns in a |
|
27 |
multipattern are considered conjunctively for quantifier instantiation. |
|
28 |
A list of multipatterns is called a trigger, and their multipatterns |
|
29 |
act disjunctively during quantifier instantiation. Each multipattern |
|
30 |
should mention at least all quantified variables of the preceding |
|
31 |
quantifier block. |
|
36898 | 32 |
*} |
33 |
||
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
56046
diff
changeset
|
34 |
typedecl pattern |
36898 | 35 |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
56046
diff
changeset
|
36 |
consts |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
56046
diff
changeset
|
37 |
pat :: "'a \<Rightarrow> pattern" |
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
56046
diff
changeset
|
38 |
nopat :: "'a \<Rightarrow> pattern" |
36898 | 39 |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
56046
diff
changeset
|
40 |
definition trigger :: "pattern list list \<Rightarrow> bool \<Rightarrow> bool" where "trigger _ P = P" |
36898 | 41 |
|
42 |
||
40664 | 43 |
subsection {* Quantifier weights *} |
44 |
||
45 |
text {* |
|
46 |
Weight annotations to quantifiers influence the priority of quantifier |
|
47 |
instantiations. They should be handled with care for solvers, which support |
|
48 |
them, because incorrect choices of weights might render a problem unsolvable. |
|
49 |
*} |
|
50 |
||
51 |
definition weight :: "int \<Rightarrow> bool \<Rightarrow> bool" where "weight _ P = P" |
|
52 |
||
53 |
text {* |
|
54 |
Weights must be non-negative. The value @{text 0} is equivalent to providing |
|
55 |
no weight at all. |
|
56 |
||
57 |
Weights should only be used at quantifiers and only inside triggers (if the |
|
58 |
quantifier has triggers). Valid usages of weights are as follows: |
|
59 |
||
60 |
\begin{itemize} |
|
61 |
\item |
|
62 |
@{term "\<forall>x. trigger [[pat (P x)]] (weight 2 (P x))"} |
|
63 |
\item |
|
64 |
@{term "\<forall>x. weight 3 (P x)"} |
|
65 |
\end{itemize} |
|
66 |
*} |
|
67 |
||
68 |
||
36902
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
huffman
parents:
36899
diff
changeset
|
69 |
subsection {* Higher-order encoding *} |
36898 | 70 |
|
71 |
text {* |
|
72 |
Application is made explicit for constants occurring with varying |
|
73 |
numbers of arguments. This is achieved by the introduction of the |
|
74 |
following constant. |
|
75 |
*} |
|
76 |
||
41127
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents:
41126
diff
changeset
|
77 |
definition fun_app where "fun_app f = f" |
36898 | 78 |
|
79 |
text {* |
|
80 |
Some solvers support a theory of arrays which can be used to encode |
|
81 |
higher-order functions. The following set of lemmas specifies the |
|
82 |
properties of such (extensional) arrays. |
|
83 |
*} |
|
84 |
||
85 |
lemmas array_rules = ext fun_upd_apply fun_upd_same fun_upd_other |
|
37157 | 86 |
fun_upd_upd fun_app_def |
36898 | 87 |
|
88 |
||
36902
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
huffman
parents:
36899
diff
changeset
|
89 |
subsection {* First-order logic *} |
36898 | 90 |
|
91 |
text {* |
|
41059
d2b1fc1b8e19
centralized handling of built-in types and constants;
boehmes
parents:
40806
diff
changeset
|
92 |
Some SMT solvers only accept problems in first-order logic, i.e., |
d2b1fc1b8e19
centralized handling of built-in types and constants;
boehmes
parents:
40806
diff
changeset
|
93 |
where formulas and terms are syntactically separated. When |
d2b1fc1b8e19
centralized handling of built-in types and constants;
boehmes
parents:
40806
diff
changeset
|
94 |
translating higher-order into first-order problems, all |
d2b1fc1b8e19
centralized handling of built-in types and constants;
boehmes
parents:
40806
diff
changeset
|
95 |
uninterpreted constants (those not built-in in the target solver) |
36898 | 96 |
are treated as function symbols in the first-order sense. Their |
41059
d2b1fc1b8e19
centralized handling of built-in types and constants;
boehmes
parents:
40806
diff
changeset
|
97 |
occurrences as head symbols in atoms (i.e., as predicate symbols) are |
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41280
diff
changeset
|
98 |
turned into terms by logically equating such atoms with @{term True}. |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41280
diff
changeset
|
99 |
For technical reasons, @{term True} and @{term False} occurring inside |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41280
diff
changeset
|
100 |
terms are replaced by the following constants. |
36898 | 101 |
*} |
102 |
||
41281
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41280
diff
changeset
|
103 |
definition term_true where "term_true = True" |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41280
diff
changeset
|
104 |
definition term_false where "term_false = False" |
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents:
41280
diff
changeset
|
105 |
|
36898 | 106 |
|
37151 | 107 |
subsection {* Integer division and modulo for Z3 *} |
108 |
||
109 |
definition z3div :: "int \<Rightarrow> int \<Rightarrow> int" where |
|
110 |
"z3div k l = (if 0 \<le> l then k div l else -(k div (-l)))" |
|
111 |
||
112 |
definition z3mod :: "int \<Rightarrow> int \<Rightarrow> int" where |
|
113 |
"z3mod k l = (if 0 \<le> l then k mod l else k mod (-l))" |
|
114 |
||
115 |
||
36902
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
huffman
parents:
36899
diff
changeset
|
116 |
subsection {* Setup *} |
36898 | 117 |
|
48892 | 118 |
ML_file "Tools/SMT/smt_builtin.ML" |
119 |
ML_file "Tools/SMT/smt_datatypes.ML" |
|
120 |
ML_file "Tools/SMT/smt_normalize.ML" |
|
121 |
ML_file "Tools/SMT/smt_translate.ML" |
|
122 |
ML_file "Tools/SMT/smt_solver.ML" |
|
123 |
ML_file "Tools/SMT/smtlib_interface.ML" |
|
124 |
ML_file "Tools/SMT/z3_interface.ML" |
|
125 |
ML_file "Tools/SMT/z3_proof_parser.ML" |
|
126 |
ML_file "Tools/SMT/z3_proof_tools.ML" |
|
127 |
ML_file "Tools/SMT/z3_proof_literals.ML" |
|
128 |
ML_file "Tools/SMT/z3_proof_methods.ML" |
|
129 |
ML_file "Tools/SMT/z3_proof_reconstruction.ML" |
|
130 |
ML_file "Tools/SMT/z3_model.ML" |
|
131 |
ML_file "Tools/SMT/smt_setup_solvers.ML" |
|
36898 | 132 |
|
133 |
setup {* |
|
40424
7550b2cba1cb
better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents:
40277
diff
changeset
|
134 |
SMT_Config.setup #> |
41059
d2b1fc1b8e19
centralized handling of built-in types and constants;
boehmes
parents:
40806
diff
changeset
|
135 |
SMT_Normalize.setup #> |
d2b1fc1b8e19
centralized handling of built-in types and constants;
boehmes
parents:
40806
diff
changeset
|
136 |
SMTLIB_Interface.setup #> |
d2b1fc1b8e19
centralized handling of built-in types and constants;
boehmes
parents:
40806
diff
changeset
|
137 |
Z3_Interface.setup #> |
36898 | 138 |
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:
39483
diff
changeset
|
139 |
SMT_Setup_Solvers.setup |
36898 | 140 |
*} |
141 |
||
47701 | 142 |
method_setup smt = {* |
143 |
Scan.optional Attrib.thms [] >> |
|
144 |
(fn thms => fn ctxt => |
|
145 |
METHOD (fn facts => HEADGOAL (SMT_Solver.smt_tac ctxt (thms @ facts)))) |
|
146 |
*} "apply an SMT solver to the current goal" |
|
36898 | 147 |
|
148 |
||
36902
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
huffman
parents:
36899
diff
changeset
|
149 |
subsection {* Configuration *} |
36898 | 150 |
|
151 |
text {* |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
36898
diff
changeset
|
152 |
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:
36898
diff
changeset
|
153 |
@{text smt_status}, which shows the values of most options. |
36898 | 154 |
*} |
155 |
||
156 |
||
157 |
||
158 |
subsection {* General configuration options *} |
|
159 |
||
160 |
text {* |
|
161 |
The option @{text smt_solver} can be used to change the target SMT |
|
41432
3214c39777ab
differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents:
41426
diff
changeset
|
162 |
solver. The possible values can be obtained from the @{text smt_status} |
3214c39777ab
differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents:
41426
diff
changeset
|
163 |
command. |
41459
f0db8f40d656
added hints about licensing restrictions and how to enable Z3
boehmes
parents:
41432
diff
changeset
|
164 |
|
f0db8f40d656
added hints about licensing restrictions and how to enable Z3
boehmes
parents:
41432
diff
changeset
|
165 |
Due to licensing restrictions, Yices and Z3 are not installed/enabled |
f0db8f40d656
added hints about licensing restrictions and how to enable Z3
boehmes
parents:
41432
diff
changeset
|
166 |
by default. Z3 is free for non-commercial applications and can be enabled |
55007
0c07990363a3
activation of Z3 via "z3_non_commercial" system option (without requiring restart);
wenzelm
parents:
50317
diff
changeset
|
167 |
by setting Isabelle system option @{text z3_non_commercial} to @{text yes}. |
36898 | 168 |
*} |
169 |
||
41601 | 170 |
declare [[ smt_solver = z3 ]] |
36898 | 171 |
|
172 |
text {* |
|
173 |
Since SMT solvers are potentially non-terminating, there is a timeout |
|
174 |
(given in seconds) to restrict their runtime. A value greater than |
|
175 |
120 (seconds) is in most cases not advisable. |
|
176 |
*} |
|
177 |
||
178 |
declare [[ smt_timeout = 20 ]] |
|
179 |
||
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:
39483
diff
changeset
|
180 |
text {* |
41121
5c5d05963f93
added option to modify the random seed of SMT solvers
boehmes
parents:
41059
diff
changeset
|
181 |
SMT solvers apply randomized heuristics. In case a problem is not |
5c5d05963f93
added option to modify the random seed of SMT solvers
boehmes
parents:
41059
diff
changeset
|
182 |
solvable by an SMT solver, changing the following option might help. |
5c5d05963f93
added option to modify the random seed of SMT solvers
boehmes
parents:
41059
diff
changeset
|
183 |
*} |
5c5d05963f93
added option to modify the random seed of SMT solvers
boehmes
parents:
41059
diff
changeset
|
184 |
|
5c5d05963f93
added option to modify the random seed of SMT solvers
boehmes
parents:
41059
diff
changeset
|
185 |
declare [[ smt_random_seed = 1 ]] |
5c5d05963f93
added option to modify the random seed of SMT solvers
boehmes
parents:
41059
diff
changeset
|
186 |
|
5c5d05963f93
added option to modify the random seed of SMT solvers
boehmes
parents:
41059
diff
changeset
|
187 |
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:
39483
diff
changeset
|
188 |
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:
39483
diff
changeset
|
189 |
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:
39483
diff
changeset
|
190 |
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:
39483
diff
changeset
|
191 |
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:
39483
diff
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:
39483
diff
changeset
|
193 |
|
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:
39483
diff
changeset
|
194 |
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:
39483
diff
changeset
|
195 |
|
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:
39483
diff
changeset
|
196 |
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:
39483
diff
changeset
|
197 |
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:
39483
diff
changeset
|
198 |
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:
39483
diff
changeset
|
199 |
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:
39483
diff
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:
39483
diff
changeset
|
201 |
|
55049 | 202 |
declare [[ cvc3_options = "" ]] |
41432
3214c39777ab
differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents:
41426
diff
changeset
|
203 |
declare [[ yices_options = "" ]] |
55049 | 204 |
declare [[ z3_options = "" ]] |
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:
39483
diff
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:
39483
diff
changeset
|
206 |
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:
39483
diff
changeset
|
207 |
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:
39483
diff
changeset
|
208 |
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:
39483
diff
changeset
|
209 |
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:
39483
diff
changeset
|
210 |
*} |
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:
39483
diff
changeset
|
211 |
|
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:
39483
diff
changeset
|
212 |
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:
39483
diff
changeset
|
213 |
|
41125 | 214 |
text {* |
215 |
The SMT method provides an inference mechanism to detect simple triggers |
|
216 |
in quantified formulas, which might increase the number of problems |
|
217 |
solvable by SMT solvers (note: triggers guide quantifier instantiations |
|
218 |
in the SMT solver). To turn it on, set the following option. |
|
219 |
*} |
|
220 |
||
221 |
declare [[ smt_infer_triggers = false ]] |
|
222 |
||
223 |
text {* |
|
224 |
The SMT method monomorphizes the given facts, that is, it tries to |
|
225 |
instantiate all schematic type variables with fixed types occurring |
|
226 |
in the problem. This is a (possibly nonterminating) fixed-point |
|
227 |
construction whose cycles are limited by the following option. |
|
228 |
*} |
|
229 |
||
43230
dabf6e311213
clarified meaning of monomorphization configuration option by renaming it
boehmes
parents:
41762
diff
changeset
|
230 |
declare [[ monomorph_max_rounds = 5 ]] |
41125 | 231 |
|
41762
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41601
diff
changeset
|
232 |
text {* |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41601
diff
changeset
|
233 |
In addition, the number of generated monomorphic instances is limited |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41601
diff
changeset
|
234 |
by the following option. |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41601
diff
changeset
|
235 |
*} |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41601
diff
changeset
|
236 |
|
43230
dabf6e311213
clarified meaning of monomorphization configuration option by renaming it
boehmes
parents:
41762
diff
changeset
|
237 |
declare [[ monomorph_max_new_instances = 500 ]] |
41762
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41601
diff
changeset
|
238 |
|
36898 | 239 |
|
240 |
||
241 |
subsection {* Certificates *} |
|
242 |
||
243 |
text {* |
|
244 |
By setting the option @{text smt_certificates} to the name of a file, |
|
245 |
all following applications of an SMT solver a cached in that file. |
|
246 |
Any further application of the same SMT solver (using the very same |
|
247 |
configuration) re-uses the cached certificate instead of invoking the |
|
248 |
solver. An empty string disables caching certificates. |
|
249 |
||
250 |
The filename should be given as an explicit path. It is good |
|
251 |
practice to use the name of the current theory (with ending |
|
252 |
@{text ".certs"} instead of @{text ".thy"}) as the certificates file. |
|
50317
4d1590544b91
synchronized read access to cache file -- avoid potential conflict with ongoing write (which is non-atomic);
wenzelm
parents:
48892
diff
changeset
|
253 |
Certificate files should be used at most once in a certain theory context, |
4d1590544b91
synchronized read access to cache file -- avoid potential conflict with ongoing write (which is non-atomic);
wenzelm
parents:
48892
diff
changeset
|
254 |
to avoid race conditions with other concurrent accesses. |
36898 | 255 |
*} |
256 |
||
257 |
declare [[ smt_certificates = "" ]] |
|
258 |
||
259 |
text {* |
|
47152
446cfc760ccf
renamed "smt_fixed" to "smt_read_only_certificates"
blanchet
parents:
46950
diff
changeset
|
260 |
The option @{text smt_read_only_certificates} controls whether only |
446cfc760ccf
renamed "smt_fixed" to "smt_read_only_certificates"
blanchet
parents:
46950
diff
changeset
|
261 |
stored certificates are should be used or invocation of an SMT solver |
446cfc760ccf
renamed "smt_fixed" to "smt_read_only_certificates"
blanchet
parents:
46950
diff
changeset
|
262 |
is allowed. When set to @{text true}, no SMT solver will ever be |
36898 | 263 |
invoked and only the existing certificates found in the configured |
264 |
cache are used; when set to @{text false} and there is no cached |
|
265 |
certificate for some proposition, then the configured SMT solver is |
|
266 |
invoked. |
|
267 |
*} |
|
268 |
||
47152
446cfc760ccf
renamed "smt_fixed" to "smt_read_only_certificates"
blanchet
parents:
46950
diff
changeset
|
269 |
declare [[ smt_read_only_certificates = false ]] |
36898 | 270 |
|
271 |
||
272 |
||
273 |
subsection {* Tracing *} |
|
274 |
||
275 |
text {* |
|
40424
7550b2cba1cb
better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents:
40277
diff
changeset
|
276 |
The SMT method, when applied, traces important information. To |
7550b2cba1cb
better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents:
40277
diff
changeset
|
277 |
make it entirely silent, set the following option to @{text false}. |
7550b2cba1cb
better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents:
40277
diff
changeset
|
278 |
*} |
7550b2cba1cb
better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents:
40277
diff
changeset
|
279 |
|
7550b2cba1cb
better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents:
40277
diff
changeset
|
280 |
declare [[ smt_verbose = true ]] |
7550b2cba1cb
better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents:
40277
diff
changeset
|
281 |
|
7550b2cba1cb
better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents:
40277
diff
changeset
|
282 |
text {* |
36898 | 283 |
For tracing the generated problem file given to the SMT solver as |
284 |
well as the returned result of the solver, the option |
|
285 |
@{text smt_trace} should be set to @{text true}. |
|
286 |
*} |
|
287 |
||
288 |
declare [[ smt_trace = false ]] |
|
289 |
||
290 |
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:
39483
diff
changeset
|
291 |
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:
39483
diff
changeset
|
292 |
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:
39483
diff
changeset
|
293 |
@{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:
39483
diff
changeset
|
294 |
(see options @{text smt_solver} and @{text smt_oracle} above). |
36898 | 295 |
*} |
296 |
||
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:
39483
diff
changeset
|
297 |
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:
37818
diff
changeset
|
298 |
|
36898 | 299 |
|
300 |
||
36902
c6bae4456741
use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
huffman
parents:
36899
diff
changeset
|
301 |
subsection {* Schematic rules for Z3 proof reconstruction *} |
36898 | 302 |
|
303 |
text {* |
|
304 |
Several prof rules of Z3 are not very well documented. There are two |
|
305 |
lemma groups which can turn failing Z3 proof reconstruction attempts |
|
306 |
into succeeding ones: the facts in @{text z3_rule} are tried prior to |
|
307 |
any implemented reconstruction procedure for all uncertain Z3 proof |
|
308 |
rules; the facts in @{text z3_simp} are only fed to invocations of |
|
309 |
the simplifier when reconstructing theory-specific proof steps. |
|
310 |
*} |
|
311 |
||
312 |
lemmas [z3_rule] = |
|
313 |
refl eq_commute conj_commute disj_commute simp_thms nnf_simps |
|
314 |
ring_distribs field_simps times_divide_eq_right times_divide_eq_left |
|
315 |
if_True if_False not_not |
|
316 |
||
317 |
lemma [z3_rule]: |
|
44488
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
318 |
"(P \<and> Q) = (\<not>(\<not>P \<or> \<not>Q))" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
319 |
"(P \<and> Q) = (\<not>(\<not>Q \<or> \<not>P))" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
320 |
"(\<not>P \<and> Q) = (\<not>(P \<or> \<not>Q))" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
321 |
"(\<not>P \<and> Q) = (\<not>(\<not>Q \<or> P))" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
322 |
"(P \<and> \<not>Q) = (\<not>(\<not>P \<or> Q))" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
323 |
"(P \<and> \<not>Q) = (\<not>(Q \<or> \<not>P))" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
324 |
"(\<not>P \<and> \<not>Q) = (\<not>(P \<or> Q))" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
325 |
"(\<not>P \<and> \<not>Q) = (\<not>(Q \<or> P))" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
326 |
by auto |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
327 |
|
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
328 |
lemma [z3_rule]: |
36898 | 329 |
"(P \<longrightarrow> Q) = (Q \<or> \<not>P)" |
330 |
"(\<not>P \<longrightarrow> Q) = (P \<or> Q)" |
|
331 |
"(\<not>P \<longrightarrow> Q) = (Q \<or> P)" |
|
44488
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
332 |
"(True \<longrightarrow> P) = P" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
333 |
"(P \<longrightarrow> True) = True" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
334 |
"(False \<longrightarrow> P) = True" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
335 |
"(P \<longrightarrow> P) = True" |
36898 | 336 |
by auto |
337 |
||
338 |
lemma [z3_rule]: |
|
339 |
"((P = Q) \<longrightarrow> R) = (R | (Q = (\<not>P)))" |
|
340 |
by auto |
|
341 |
||
342 |
lemma [z3_rule]: |
|
44488
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
343 |
"(\<not>True) = False" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
344 |
"(\<not>False) = True" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
345 |
"(x = x) = True" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
346 |
"(P = True) = P" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
347 |
"(True = P) = P" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
348 |
"(P = False) = (\<not>P)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
349 |
"(False = P) = (\<not>P)" |
36898 | 350 |
"((\<not>P) = P) = False" |
351 |
"(P = (\<not>P)) = False" |
|
44488
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
352 |
"((\<not>P) = (\<not>Q)) = (P = Q)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
353 |
"\<not>(P = (\<not>Q)) = (P = Q)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
354 |
"\<not>((\<not>P) = Q) = (P = Q)" |
36898 | 355 |
"(P \<noteq> Q) = (Q = (\<not>P))" |
356 |
"(P = Q) = ((\<not>P \<or> Q) \<and> (P \<or> \<not>Q))" |
|
357 |
"(P \<noteq> Q) = ((\<not>P \<or> \<not>Q) \<and> (P \<or> Q))" |
|
358 |
by auto |
|
359 |
||
360 |
lemma [z3_rule]: |
|
361 |
"(if P then P else \<not>P) = True" |
|
362 |
"(if \<not>P then \<not>P else P) = True" |
|
363 |
"(if P then True else False) = P" |
|
364 |
"(if P then False else True) = (\<not>P)" |
|
44488
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
365 |
"(if P then Q else True) = ((\<not>P) \<or> Q)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
366 |
"(if P then Q else True) = (Q \<or> (\<not>P))" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
367 |
"(if P then Q else \<not>Q) = (P = Q)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
368 |
"(if P then Q else \<not>Q) = (Q = P)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
369 |
"(if P then \<not>Q else Q) = (P = (\<not>Q))" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
370 |
"(if P then \<not>Q else Q) = ((\<not>Q) = P)" |
36898 | 371 |
"(if \<not>P then x else y) = (if P then y else x)" |
44488
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
372 |
"(if P then (if Q then x else y) else x) = (if P \<and> (\<not>Q) then y else x)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
373 |
"(if P then (if Q then x else y) else x) = (if (\<not>Q) \<and> P then y else x)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
374 |
"(if P then (if Q then x else y) else y) = (if P \<and> Q then x else y)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
375 |
"(if P then (if Q then x else y) else y) = (if Q \<and> P then x else y)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
376 |
"(if P then x else if P then y else z) = (if P then x else z)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
377 |
"(if P then x else if Q then x else y) = (if P \<or> Q then x else y)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
378 |
"(if P then x else if Q then x else y) = (if Q \<or> P then x else y)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
379 |
"(if P then x = y else x = z) = (x = (if P then y else z))" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
380 |
"(if P then x = y else y = z) = (y = (if P then x else z))" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
381 |
"(if P then x = y else z = y) = (y = (if P then x else z))" |
36898 | 382 |
by auto |
383 |
||
384 |
lemma [z3_rule]: |
|
44488
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
385 |
"0 + (x::int) = x" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
386 |
"x + 0 = x" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
387 |
"x + x = 2 * x" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
388 |
"0 * x = 0" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
389 |
"1 * x = x" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
390 |
"x + y = y + x" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
391 |
by auto |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
392 |
|
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
393 |
lemma [z3_rule]: (* for def-axiom *) |
36898 | 394 |
"P = Q \<or> P \<or> Q" |
395 |
"P = Q \<or> \<not>P \<or> \<not>Q" |
|
396 |
"(\<not>P) = Q \<or> \<not>P \<or> Q" |
|
397 |
"(\<not>P) = Q \<or> P \<or> \<not>Q" |
|
398 |
"P = (\<not>Q) \<or> \<not>P \<or> Q" |
|
399 |
"P = (\<not>Q) \<or> P \<or> \<not>Q" |
|
400 |
"P \<noteq> Q \<or> P \<or> \<not>Q" |
|
401 |
"P \<noteq> Q \<or> \<not>P \<or> Q" |
|
402 |
"P \<noteq> (\<not>Q) \<or> P \<or> Q" |
|
403 |
"(\<not>P) \<noteq> Q \<or> P \<or> Q" |
|
404 |
"P \<or> Q \<or> P \<noteq> (\<not>Q)" |
|
405 |
"P \<or> Q \<or> (\<not>P) \<noteq> Q" |
|
406 |
"P \<or> \<not>Q \<or> P \<noteq> Q" |
|
407 |
"\<not>P \<or> Q \<or> P \<noteq> Q" |
|
44488
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
408 |
"P \<or> y = (if P then x else y)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
409 |
"P \<or> (if P then x else y) = y" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
410 |
"\<not>P \<or> x = (if P then x else y)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
411 |
"\<not>P \<or> (if P then x else y) = x" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
412 |
"P \<or> R \<or> \<not>(if P then Q else R)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
413 |
"\<not>P \<or> Q \<or> \<not>(if P then Q else R)" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
414 |
"\<not>(if P then Q else R) \<or> \<not>P \<or> Q" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
415 |
"\<not>(if P then Q else R) \<or> P \<or> R" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
416 |
"(if P then Q else R) \<or> \<not>P \<or> \<not>Q" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
417 |
"(if P then Q else R) \<or> P \<or> \<not>R" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
418 |
"(if P then \<not>Q else R) \<or> \<not>P \<or> Q" |
587bf61a00a1
improved completeness and efficiency of Z3 proof reconstruction, especially by an improved handling of Skolemization
boehmes
parents:
44087
diff
changeset
|
419 |
"(if P then Q else \<not>R) \<or> P \<or> R" |
36898 | 420 |
by auto |
421 |
||
37124 | 422 |
|
423 |
hide_type (open) pattern |
|
56078
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
56046
diff
changeset
|
424 |
hide_const fun_app term_true term_false z3div z3mod |
41280
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
boehmes
parents:
41174
diff
changeset
|
425 |
hide_const (open) trigger pat nopat weight |
37124 | 426 |
|
36898 | 427 |
end |