src/HOL/SMT.thy
author wenzelm
Thu, 07 Jul 2011 23:55:15 +0200
changeset 44572 91c4d7397f0e
parent 44071 dabf6e311213
child 44798 3a87cb597832
permissions -rw-r--r--
simplified make_option/dest_option;
added make_variant/dest_variant -- usual representation of datatypes;
     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 Record
     9 uses
    10   "Tools/SMT/smt_utils.ML"
    11   "Tools/SMT/smt_failure.ML"
    12   "Tools/SMT/smt_config.ML"
    13   ("Tools/SMT/smt_monomorph.ML")
    14   ("Tools/SMT/smt_builtin.ML")
    15   ("Tools/SMT/smt_datatypes.ML")
    16   ("Tools/SMT/smt_normalize.ML")
    17   ("Tools/SMT/smt_translate.ML")
    18   ("Tools/SMT/smt_solver.ML")
    19   ("Tools/SMT/smtlib_interface.ML")
    20   ("Tools/SMT/z3_interface.ML")
    21   ("Tools/SMT/z3_proof_parser.ML")
    22   ("Tools/SMT/z3_proof_tools.ML")
    23   ("Tools/SMT/z3_proof_literals.ML")
    24   ("Tools/SMT/z3_proof_methods.ML")
    25   ("Tools/SMT/z3_proof_reconstruction.ML")
    26   ("Tools/SMT/z3_model.ML")
    27   ("Tools/SMT/smt_setup_solvers.ML")
    28 begin
    29 
    30 
    31 
    32 subsection {* Triggers for quantifier instantiation *}
    33 
    34 text {*
    35 Some SMT solvers support patterns as a quantifier instantiation
    36 heuristics.  Patterns may either be positive terms (tagged by "pat")
    37 triggering quantifier instantiations -- when the solver finds a
    38 term matching a positive pattern, it instantiates the corresponding
    39 quantifier accordingly -- or negative terms (tagged by "nopat")
    40 inhibiting quantifier instantiations.  A list of patterns
    41 of the same kind is called a multipattern, and all patterns in a
    42 multipattern are considered conjunctively for quantifier instantiation.
    43 A list of multipatterns is called a trigger, and their multipatterns
    44 act disjunctively during quantifier instantiation.  Each multipattern
    45 should mention at least all quantified variables of the preceding
    46 quantifier block.
    47 *}
    48 
    49 datatype pattern = Pattern
    50 
    51 definition pat :: "'a \<Rightarrow> pattern" where "pat _ = Pattern"
    52 definition nopat :: "'a \<Rightarrow> pattern" where "nopat _ = Pattern"
    53 
    54 definition trigger :: "pattern list list \<Rightarrow> bool \<Rightarrow> bool"
    55 where "trigger _ P = P"
    56 
    57 
    58 
    59 subsection {* Quantifier weights *}
    60 
    61 text {*
    62 Weight annotations to quantifiers influence the priority of quantifier
    63 instantiations.  They should be handled with care for solvers, which support
    64 them, because incorrect choices of weights might render a problem unsolvable.
    65 *}
    66 
    67 definition weight :: "int \<Rightarrow> bool \<Rightarrow> bool" where "weight _ P = P"
    68 
    69 text {*
    70 Weights must be non-negative.  The value @{text 0} is equivalent to providing
    71 no weight at all.
    72 
    73 Weights should only be used at quantifiers and only inside triggers (if the
    74 quantifier has triggers).  Valid usages of weights are as follows:
    75 
    76 \begin{itemize}
    77 \item
    78 @{term "\<forall>x. trigger [[pat (P x)]] (weight 2 (P x))"}
    79 \item
    80 @{term "\<forall>x. weight 3 (P x)"}
    81 \end{itemize}
    82 *}
    83 
    84 
    85 
    86 subsection {* Higher-order encoding *}
    87 
    88 text {*
    89 Application is made explicit for constants occurring with varying
    90 numbers of arguments.  This is achieved by the introduction of the
    91 following constant.
    92 *}
    93 
    94 definition fun_app where "fun_app f = f"
    95 
    96 text {*
    97 Some solvers support a theory of arrays which can be used to encode
    98 higher-order functions.  The following set of lemmas specifies the
    99 properties of such (extensional) arrays.
   100 *}
   101 
   102 lemmas array_rules = ext fun_upd_apply fun_upd_same fun_upd_other
   103   fun_upd_upd fun_app_def
   104 
   105 
   106 
   107 subsection {* First-order logic *}
   108 
   109 text {*
   110 Some SMT solvers only accept problems in first-order logic, i.e.,
   111 where formulas and terms are syntactically separated. When
   112 translating higher-order into first-order problems, all
   113 uninterpreted constants (those not built-in in the target solver)
   114 are treated as function symbols in the first-order sense.  Their
   115 occurrences as head symbols in atoms (i.e., as predicate symbols) are
   116 turned into terms by logically equating such atoms with @{term True}.
   117 For technical reasons, @{term True} and @{term False} occurring inside
   118 terms are replaced by the following constants.
   119 *}
   120 
   121 definition term_true where "term_true = True"
   122 definition term_false where "term_false = False"
   123 
   124 
   125 
   126 subsection {* Integer division and modulo for Z3 *}
   127 
   128 definition z3div :: "int \<Rightarrow> int \<Rightarrow> int" where
   129   "z3div k l = (if 0 \<le> l then k div l else -(k div (-l)))"
   130 
   131 definition z3mod :: "int \<Rightarrow> int \<Rightarrow> int" where
   132   "z3mod k l = (if 0 \<le> l then k mod l else k mod (-l))"
   133 
   134 
   135 
   136 subsection {* Setup *}
   137 
   138 use "Tools/SMT/smt_monomorph.ML"
   139 use "Tools/SMT/smt_builtin.ML"
   140 use "Tools/SMT/smt_datatypes.ML"
   141 use "Tools/SMT/smt_normalize.ML"
   142 use "Tools/SMT/smt_translate.ML"
   143 use "Tools/SMT/smt_solver.ML"
   144 use "Tools/SMT/smtlib_interface.ML"
   145 use "Tools/SMT/z3_interface.ML"
   146 use "Tools/SMT/z3_proof_parser.ML"
   147 use "Tools/SMT/z3_proof_tools.ML"
   148 use "Tools/SMT/z3_proof_literals.ML"
   149 use "Tools/SMT/z3_proof_methods.ML"
   150 use "Tools/SMT/z3_proof_reconstruction.ML"
   151 use "Tools/SMT/z3_model.ML"
   152 use "Tools/SMT/smt_setup_solvers.ML"
   153 
   154 setup {*
   155   SMT_Config.setup #>
   156   SMT_Normalize.setup #>
   157   SMT_Solver.setup #>
   158   SMTLIB_Interface.setup #>
   159   Z3_Interface.setup #>
   160   Z3_Proof_Reconstruction.setup #>
   161   SMT_Setup_Solvers.setup
   162 *}
   163 
   164 
   165 
   166 subsection {* Configuration *}
   167 
   168 text {*
   169 The current configuration can be printed by the command
   170 @{text smt_status}, which shows the values of most options.
   171 *}
   172 
   173 
   174 
   175 subsection {* General configuration options *}
   176 
   177 text {*
   178 The option @{text smt_solver} can be used to change the target SMT
   179 solver.  The possible values can be obtained from the @{text smt_status}
   180 command.
   181 
   182 Due to licensing restrictions, Yices and Z3 are not installed/enabled
   183 by default.  Z3 is free for non-commercial applications and can be enabled
   184 by simply setting the environment variable @{text Z3_NON_COMMERCIAL} to
   185 @{text yes}.
   186 *}
   187 
   188 declare [[ smt_solver = z3 ]]
   189 
   190 text {*
   191 Since SMT solvers are potentially non-terminating, there is a timeout
   192 (given in seconds) to restrict their runtime.  A value greater than
   193 120 (seconds) is in most cases not advisable.
   194 *}
   195 
   196 declare [[ smt_timeout = 20 ]]
   197 
   198 text {*
   199 SMT solvers apply randomized heuristics.  In case a problem is not
   200 solvable by an SMT solver, changing the following option might help.
   201 *}
   202 
   203 declare [[ smt_random_seed = 1 ]]
   204 
   205 text {*
   206 In general, the binding to SMT solvers runs as an oracle, i.e, the SMT
   207 solvers are fully trusted without additional checks.  The following
   208 option can cause the SMT solver to run in proof-producing mode, giving
   209 a checkable certificate.  This is currently only implemented for Z3.
   210 *}
   211 
   212 declare [[ smt_oracle = false ]]
   213 
   214 text {*
   215 Each SMT solver provides several commandline options to tweak its
   216 behaviour.  They can be passed to the solver by setting the following
   217 options.
   218 *}
   219 
   220 declare [[ cvc3_options = "", remote_cvc3_options = "" ]]
   221 declare [[ yices_options = "" ]]
   222 declare [[ z3_options = "", remote_z3_options = "" ]]
   223 
   224 text {*
   225 Enable the following option to use built-in support for datatypes and
   226 records.  Currently, this is only implemented for Z3 running in oracle
   227 mode.
   228 *}
   229 
   230 declare [[ smt_datatypes = false ]]
   231 
   232 text {*
   233 The SMT method provides an inference mechanism to detect simple triggers
   234 in quantified formulas, which might increase the number of problems
   235 solvable by SMT solvers (note: triggers guide quantifier instantiations
   236 in the SMT solver).  To turn it on, set the following option.
   237 *}
   238 
   239 declare [[ smt_infer_triggers = false ]]
   240 
   241 text {*
   242 The SMT method monomorphizes the given facts, that is, it tries to
   243 instantiate all schematic type variables with fixed types occurring
   244 in the problem.  This is a (possibly nonterminating) fixed-point
   245 construction whose cycles are limited by the following option.
   246 *}
   247 
   248 declare [[ monomorph_max_rounds = 5 ]]
   249 
   250 text {*
   251 In addition, the number of generated monomorphic instances is limited
   252 by the following option.
   253 *}
   254 
   255 declare [[ monomorph_max_new_instances = 500 ]]
   256 
   257 
   258 
   259 subsection {* Certificates *}
   260 
   261 text {*
   262 By setting the option @{text smt_certificates} to the name of a file,
   263 all following applications of an SMT solver a cached in that file.
   264 Any further application of the same SMT solver (using the very same
   265 configuration) re-uses the cached certificate instead of invoking the
   266 solver.  An empty string disables caching certificates.
   267 
   268 The filename should be given as an explicit path.  It is good
   269 practice to use the name of the current theory (with ending
   270 @{text ".certs"} instead of @{text ".thy"}) as the certificates file.
   271 *}
   272 
   273 declare [[ smt_certificates = "" ]]
   274 
   275 text {*
   276 The option @{text smt_fixed} controls whether only stored
   277 certificates are should be used or invocation of an SMT solver is
   278 allowed.  When set to @{text true}, no SMT solver will ever be
   279 invoked and only the existing certificates found in the configured
   280 cache are used;  when set to @{text false} and there is no cached
   281 certificate for some proposition, then the configured SMT solver is
   282 invoked.
   283 *}
   284 
   285 declare [[ smt_fixed = false ]]
   286 
   287 
   288 
   289 subsection {* Tracing *}
   290 
   291 text {*
   292 The SMT method, when applied, traces important information.  To
   293 make it entirely silent, set the following option to @{text false}.
   294 *}
   295 
   296 declare [[ smt_verbose = true ]]
   297 
   298 text {*
   299 For tracing the generated problem file given to the SMT solver as
   300 well as the returned result of the solver, the option
   301 @{text smt_trace} should be set to @{text true}.
   302 *}
   303 
   304 declare [[ smt_trace = false ]]
   305 
   306 text {*
   307 From the set of assumptions given to the SMT solver, those assumptions
   308 used in the proof are traced when the following option is set to
   309 @{term true}.  This only works for Z3 when it runs in non-oracle mode
   310 (see options @{text smt_solver} and @{text smt_oracle} above).
   311 *}
   312 
   313 declare [[ smt_trace_used_facts = false ]]
   314 
   315 
   316 
   317 subsection {* Schematic rules for Z3 proof reconstruction *}
   318 
   319 text {*
   320 Several prof rules of Z3 are not very well documented.  There are two
   321 lemma groups which can turn failing Z3 proof reconstruction attempts
   322 into succeeding ones: the facts in @{text z3_rule} are tried prior to
   323 any implemented reconstruction procedure for all uncertain Z3 proof
   324 rules;  the facts in @{text z3_simp} are only fed to invocations of
   325 the simplifier when reconstructing theory-specific proof steps.
   326 *}
   327 
   328 lemmas [z3_rule] =
   329   refl eq_commute conj_commute disj_commute simp_thms nnf_simps
   330   ring_distribs field_simps times_divide_eq_right times_divide_eq_left
   331   if_True if_False not_not
   332 
   333 lemma [z3_rule]:
   334   "(P \<longrightarrow> Q) = (Q \<or> \<not>P)"
   335   "(\<not>P \<longrightarrow> Q) = (P \<or> Q)"
   336   "(\<not>P \<longrightarrow> Q) = (Q \<or> P)"
   337   by auto
   338 
   339 lemma [z3_rule]:
   340   "((P = Q) \<longrightarrow> R) = (R | (Q = (\<not>P)))"
   341   by auto
   342 
   343 lemma [z3_rule]:
   344   "((\<not>P) = P) = False"
   345   "(P = (\<not>P)) = False"
   346   "(P \<noteq> Q) = (Q = (\<not>P))"
   347   "(P = Q) = ((\<not>P \<or> Q) \<and> (P \<or> \<not>Q))"
   348   "(P \<noteq> Q) = ((\<not>P \<or> \<not>Q) \<and> (P \<or> Q))"
   349   by auto
   350 
   351 lemma [z3_rule]:
   352   "(if P then P else \<not>P) = True"
   353   "(if \<not>P then \<not>P else P) = True"
   354   "(if P then True else False) = P"
   355   "(if P then False else True) = (\<not>P)"
   356   "(if \<not>P then x else y) = (if P then y else x)"
   357   "f (if P then x else y) = (if P then f x else f y)"
   358   by auto
   359 
   360 lemma [z3_rule]:
   361   "P = Q \<or> P \<or> Q"
   362   "P = Q \<or> \<not>P \<or> \<not>Q"
   363   "(\<not>P) = Q \<or> \<not>P \<or> Q"
   364   "(\<not>P) = Q \<or> P \<or> \<not>Q"
   365   "P = (\<not>Q) \<or> \<not>P \<or> Q"
   366   "P = (\<not>Q) \<or> P \<or> \<not>Q"
   367   "P \<noteq> Q \<or> P \<or> \<not>Q"
   368   "P \<noteq> Q \<or> \<not>P \<or> Q"
   369   "P \<noteq> (\<not>Q) \<or> P \<or> Q"
   370   "(\<not>P) \<noteq> Q \<or> P \<or> Q"
   371   "P \<or> Q \<or> P \<noteq> (\<not>Q)"
   372   "P \<or> Q \<or> (\<not>P) \<noteq> Q"
   373   "P \<or> \<not>Q \<or> P \<noteq> Q"
   374   "\<not>P \<or> Q \<or> P \<noteq> Q"
   375   by auto
   376 
   377 lemma [z3_rule]:
   378   "0 + (x::int) = x"
   379   "x + 0 = x"
   380   "0 * x = 0"
   381   "1 * x = x"
   382   "x + y = y + x"
   383   by auto
   384 
   385 
   386 
   387 hide_type (open) pattern
   388 hide_const Pattern fun_app term_true term_false z3div z3mod
   389 hide_const (open) trigger pat nopat weight
   390 
   391 end