src/HOL/SMT.thy
changeset 58061 3d060f43accb
parent 57704 c0da3fc313e3
child 58072 a86c962de77f
equal deleted inserted replaced
58060:835b5443b978 58061:3d060f43accb
       
     1 (*  Title:      HOL/SMT.thy
       
     2     Author:     Sascha Boehme, TU Muenchen
       
     3 *)
       
     4 
       
     5 header {* Bindings to Satisfiability Modulo Theories (SMT) solvers based on SMT-LIB 2 *}
       
     6 
       
     7 theory SMT
       
     8 imports Divides
       
     9 keywords "smt_status" :: diag
       
    10 begin
       
    11 
       
    12 subsection {* Triggers for quantifier instantiation *}
       
    13 
       
    14 text {*
       
    15 Some SMT solvers support patterns as a quantifier instantiation
       
    16 heuristics. Patterns may either be positive terms (tagged by "pat")
       
    17 triggering quantifier instantiations -- when the solver finds a
       
    18 term matching a positive pattern, it instantiates the corresponding
       
    19 quantifier accordingly -- or negative terms (tagged by "nopat")
       
    20 inhibiting quantifier instantiations. A list of patterns
       
    21 of the same kind is called a multipattern, and all patterns in a
       
    22 multipattern are considered conjunctively for quantifier instantiation.
       
    23 A list of multipatterns is called a trigger, and their multipatterns
       
    24 act disjunctively during quantifier instantiation. Each multipattern
       
    25 should mention at least all quantified variables of the preceding
       
    26 quantifier block.
       
    27 *}
       
    28 
       
    29 typedecl 'a symb_list
       
    30 
       
    31 consts
       
    32   Symb_Nil :: "'a symb_list"
       
    33   Symb_Cons :: "'a \<Rightarrow> 'a symb_list \<Rightarrow> 'a symb_list"
       
    34 
       
    35 typedecl pattern
       
    36 
       
    37 consts
       
    38   pat :: "'a \<Rightarrow> pattern"
       
    39   nopat :: "'a \<Rightarrow> pattern"
       
    40 
       
    41 definition trigger :: "pattern symb_list symb_list \<Rightarrow> bool \<Rightarrow> bool" where
       
    42   "trigger _ P = P"
       
    43 
       
    44 
       
    45 subsection {* Higher-order encoding *}
       
    46 
       
    47 text {*
       
    48 Application is made explicit for constants occurring with varying
       
    49 numbers of arguments. This is achieved by the introduction of the
       
    50 following constant.
       
    51 *}
       
    52 
       
    53 definition fun_app :: "'a \<Rightarrow> 'a" where "fun_app f = f"
       
    54 
       
    55 text {*
       
    56 Some solvers support a theory of arrays which can be used to encode
       
    57 higher-order functions. The following set of lemmas specifies the
       
    58 properties of such (extensional) arrays.
       
    59 *}
       
    60 
       
    61 lemmas array_rules = ext fun_upd_apply fun_upd_same fun_upd_other  fun_upd_upd fun_app_def
       
    62 
       
    63 
       
    64 subsection {* Normalization *}
       
    65 
       
    66 lemma case_bool_if[abs_def]: "case_bool x y P = (if P then x else y)"
       
    67   by simp
       
    68 
       
    69 lemmas Ex1_def_raw = Ex1_def[abs_def]
       
    70 lemmas Ball_def_raw = Ball_def[abs_def]
       
    71 lemmas Bex_def_raw = Bex_def[abs_def]
       
    72 lemmas abs_if_raw = abs_if[abs_def]
       
    73 lemmas min_def_raw = min_def[abs_def]
       
    74 lemmas max_def_raw = max_def[abs_def]
       
    75 
       
    76 
       
    77 subsection {* Integer division and modulo for Z3 *}
       
    78 
       
    79 text {*
       
    80 The following Z3-inspired definitions are overspecified for the case where @{text "l = 0"}. This
       
    81 Schönheitsfehler is corrected in the @{text div_as_z3div} and @{text mod_as_z3mod} theorems.
       
    82 *}
       
    83 
       
    84 definition z3div :: "int \<Rightarrow> int \<Rightarrow> int" where
       
    85   "z3div k l = (if l \<ge> 0 then k div l else - (k div - l))"
       
    86 
       
    87 definition z3mod :: "int \<Rightarrow> int \<Rightarrow> int" where
       
    88   "z3mod k l = k mod (if l \<ge> 0 then l else - l)"
       
    89 
       
    90 lemma div_as_z3div:
       
    91   "\<forall>k l. k div l = (if l = 0 then 0 else if l > 0 then z3div k l else z3div (- k) (- l))"
       
    92   by (simp add: z3div_def)
       
    93 
       
    94 lemma mod_as_z3mod:
       
    95   "\<forall>k l. k mod l = (if l = 0 then k else if l > 0 then z3mod k l else - z3mod (- k) (- l))"
       
    96   by (simp add: z3mod_def)
       
    97 
       
    98 
       
    99 subsection {* Setup *}
       
   100 
       
   101 ML_file "Tools/SMT/smt_util.ML"
       
   102 ML_file "Tools/SMT/smt_failure.ML"
       
   103 ML_file "Tools/SMT/smt_config.ML"
       
   104 ML_file "Tools/SMT/smt_builtin.ML"
       
   105 ML_file "Tools/SMT/smt_datatypes.ML"
       
   106 ML_file "Tools/SMT/smt_normalize.ML"
       
   107 ML_file "Tools/SMT/smt_translate.ML"
       
   108 ML_file "Tools/SMT/smtlib.ML"
       
   109 ML_file "Tools/SMT/smtlib_interface.ML"
       
   110 ML_file "Tools/SMT/smtlib_proof.ML"
       
   111 ML_file "Tools/SMT/smtlib_isar.ML"
       
   112 ML_file "Tools/SMT/z3_proof.ML"
       
   113 ML_file "Tools/SMT/z3_isar.ML"
       
   114 ML_file "Tools/SMT/smt_solver.ML"
       
   115 ML_file "Tools/SMT/z3_interface.ML"
       
   116 ML_file "Tools/SMT/z3_replay_util.ML"
       
   117 ML_file "Tools/SMT/z3_replay_literals.ML"
       
   118 ML_file "Tools/SMT/z3_replay_rules.ML"
       
   119 ML_file "Tools/SMT/z3_replay_methods.ML"
       
   120 ML_file "Tools/SMT/z3_replay.ML"
       
   121 ML_file "Tools/SMT/verit_proof.ML"
       
   122 ML_file "Tools/SMT/verit_isar.ML"
       
   123 ML_file "Tools/SMT/verit_proof_parse.ML"
       
   124 ML_file "Tools/SMT/smt_systems.ML"
       
   125 
       
   126 method_setup smt = {*
       
   127   Scan.optional Attrib.thms [] >>
       
   128     (fn thms => fn ctxt =>
       
   129       METHOD (fn facts => HEADGOAL (SMT_Solver.smt_tac ctxt (thms @ facts))))
       
   130 *} "apply an SMT solver to the current goal (based on SMT-LIB 2)"
       
   131 
       
   132 
       
   133 subsection {* Configuration *}
       
   134 
       
   135 text {*
       
   136 The current configuration can be printed by the command
       
   137 @{text smt_status}, which shows the values of most options.
       
   138 *}
       
   139 
       
   140 
       
   141 subsection {* General configuration options *}
       
   142 
       
   143 text {*
       
   144 The option @{text smt_solver} can be used to change the target SMT
       
   145 solver. The possible values can be obtained from the @{text smt_status}
       
   146 command.
       
   147 
       
   148 Due to licensing restrictions, Z3 is not enabled by default. Z3 is free
       
   149 for non-commercial applications and can be enabled by setting Isabelle
       
   150 system option @{text z3_non_commercial} to @{text yes}.
       
   151 *}
       
   152 
       
   153 declare [[smt_solver = z3]]
       
   154 
       
   155 text {*
       
   156 Since SMT solvers are potentially nonterminating, there is a timeout
       
   157 (given in seconds) to restrict their runtime.
       
   158 *}
       
   159 
       
   160 declare [[smt_timeout = 20]]
       
   161 
       
   162 text {*
       
   163 SMT solvers apply randomized heuristics. In case a problem is not
       
   164 solvable by an SMT solver, changing the following option might help.
       
   165 *}
       
   166 
       
   167 declare [[smt_random_seed = 1]]
       
   168 
       
   169 text {*
       
   170 In general, the binding to SMT solvers runs as an oracle, i.e, the SMT
       
   171 solvers are fully trusted without additional checks. The following
       
   172 option can cause the SMT solver to run in proof-producing mode, giving
       
   173 a checkable certificate. This is currently only implemented for Z3.
       
   174 *}
       
   175 
       
   176 declare [[smt_oracle = false]]
       
   177 
       
   178 text {*
       
   179 Each SMT solver provides several commandline options to tweak its
       
   180 behaviour. They can be passed to the solver by setting the following
       
   181 options.
       
   182 *}
       
   183 
       
   184 declare [[cvc3_options = ""]]
       
   185 declare [[cvc4_options = ""]]
       
   186 declare [[veriT_options = ""]]
       
   187 declare [[z3_options = ""]]
       
   188 
       
   189 text {*
       
   190 The SMT method provides an inference mechanism to detect simple triggers
       
   191 in quantified formulas, which might increase the number of problems
       
   192 solvable by SMT solvers (note: triggers guide quantifier instantiations
       
   193 in the SMT solver). To turn it on, set the following option.
       
   194 *}
       
   195 
       
   196 declare [[smt_infer_triggers = false]]
       
   197 
       
   198 text {*
       
   199 Enable the following option to use built-in support for div/mod, datatypes,
       
   200 and records in Z3. Currently, this is implemented only in oracle mode.
       
   201 *}
       
   202 
       
   203 declare [[z3_extensions = false]]
       
   204 
       
   205 
       
   206 subsection {* Certificates *}
       
   207 
       
   208 text {*
       
   209 By setting the option @{text smt_certificates} to the name of a file,
       
   210 all following applications of an SMT solver a cached in that file.
       
   211 Any further application of the same SMT solver (using the very same
       
   212 configuration) re-uses the cached certificate instead of invoking the
       
   213 solver. An empty string disables caching certificates.
       
   214 
       
   215 The filename should be given as an explicit path. It is good
       
   216 practice to use the name of the current theory (with ending
       
   217 @{text ".certs"} instead of @{text ".thy"}) as the certificates file.
       
   218 Certificate files should be used at most once in a certain theory context,
       
   219 to avoid race conditions with other concurrent accesses.
       
   220 *}
       
   221 
       
   222 declare [[smt_certificates = ""]]
       
   223 
       
   224 text {*
       
   225 The option @{text smt_read_only_certificates} controls whether only
       
   226 stored certificates are should be used or invocation of an SMT solver
       
   227 is allowed. When set to @{text true}, no SMT solver will ever be
       
   228 invoked and only the existing certificates found in the configured
       
   229 cache are used;  when set to @{text false} and there is no cached
       
   230 certificate for some proposition, then the configured SMT solver is
       
   231 invoked.
       
   232 *}
       
   233 
       
   234 declare [[smt_read_only_certificates = false]]
       
   235 
       
   236 
       
   237 subsection {* Tracing *}
       
   238 
       
   239 text {*
       
   240 The SMT method, when applied, traces important information. To
       
   241 make it entirely silent, set the following option to @{text false}.
       
   242 *}
       
   243 
       
   244 declare [[smt_verbose = true]]
       
   245 
       
   246 text {*
       
   247 For tracing the generated problem file given to the SMT solver as
       
   248 well as the returned result of the solver, the option
       
   249 @{text smt_trace} should be set to @{text true}.
       
   250 *}
       
   251 
       
   252 declare [[smt_trace = false]]
       
   253 
       
   254 
       
   255 subsection {* Schematic rules for Z3 proof reconstruction *}
       
   256 
       
   257 text {*
       
   258 Several prof rules of Z3 are not very well documented. There are two
       
   259 lemma groups which can turn failing Z3 proof reconstruction attempts
       
   260 into succeeding ones: the facts in @{text z3_rule} are tried prior to
       
   261 any implemented reconstruction procedure for all uncertain Z3 proof
       
   262 rules;  the facts in @{text z3_simp} are only fed to invocations of
       
   263 the simplifier when reconstructing theory-specific proof steps.
       
   264 *}
       
   265 
       
   266 lemmas [z3_rule] =
       
   267   refl eq_commute conj_commute disj_commute simp_thms nnf_simps
       
   268   ring_distribs field_simps times_divide_eq_right times_divide_eq_left
       
   269   if_True if_False not_not
       
   270 
       
   271 lemma [z3_rule]:
       
   272   "(P \<and> Q) = (\<not> (\<not> P \<or> \<not> Q))"
       
   273   "(P \<and> Q) = (\<not> (\<not> Q \<or> \<not> P))"
       
   274   "(\<not> P \<and> Q) = (\<not> (P \<or> \<not> Q))"
       
   275   "(\<not> P \<and> Q) = (\<not> (\<not> Q \<or> P))"
       
   276   "(P \<and> \<not> Q) = (\<not> (\<not> P \<or> Q))"
       
   277   "(P \<and> \<not> Q) = (\<not> (Q \<or> \<not> P))"
       
   278   "(\<not> P \<and> \<not> Q) = (\<not> (P \<or> Q))"
       
   279   "(\<not> P \<and> \<not> Q) = (\<not> (Q \<or> P))"
       
   280   by auto
       
   281 
       
   282 lemma [z3_rule]:
       
   283   "(P \<longrightarrow> Q) = (Q \<or> \<not> P)"
       
   284   "(\<not> P \<longrightarrow> Q) = (P \<or> Q)"
       
   285   "(\<not> P \<longrightarrow> Q) = (Q \<or> P)"
       
   286   "(True \<longrightarrow> P) = P"
       
   287   "(P \<longrightarrow> True) = True"
       
   288   "(False \<longrightarrow> P) = True"
       
   289   "(P \<longrightarrow> P) = True"
       
   290   by auto
       
   291 
       
   292 lemma [z3_rule]:
       
   293   "((P = Q) \<longrightarrow> R) = (R | (Q = (\<not> P)))"
       
   294   by auto
       
   295 
       
   296 lemma [z3_rule]:
       
   297   "(\<not> True) = False"
       
   298   "(\<not> False) = True"
       
   299   "(x = x) = True"
       
   300   "(P = True) = P"
       
   301   "(True = P) = P"
       
   302   "(P = False) = (\<not> P)"
       
   303   "(False = P) = (\<not> P)"
       
   304   "((\<not> P) = P) = False"
       
   305   "(P = (\<not> P)) = False"
       
   306   "((\<not> P) = (\<not> Q)) = (P = Q)"
       
   307   "\<not> (P = (\<not> Q)) = (P = Q)"
       
   308   "\<not> ((\<not> P) = Q) = (P = Q)"
       
   309   "(P \<noteq> Q) = (Q = (\<not> P))"
       
   310   "(P = Q) = ((\<not> P \<or> Q) \<and> (P \<or> \<not> Q))"
       
   311   "(P \<noteq> Q) = ((\<not> P \<or> \<not> Q) \<and> (P \<or> Q))"
       
   312   by auto
       
   313 
       
   314 lemma [z3_rule]:
       
   315   "(if P then P else \<not> P) = True"
       
   316   "(if \<not> P then \<not> P else P) = True"
       
   317   "(if P then True else False) = P"
       
   318   "(if P then False else True) = (\<not> P)"
       
   319   "(if P then Q else True) = ((\<not> P) \<or> Q)"
       
   320   "(if P then Q else True) = (Q \<or> (\<not> P))"
       
   321   "(if P then Q else \<not> Q) = (P = Q)"
       
   322   "(if P then Q else \<not> Q) = (Q = P)"
       
   323   "(if P then \<not> Q else Q) = (P = (\<not> Q))"
       
   324   "(if P then \<not> Q else Q) = ((\<not> Q) = P)"
       
   325   "(if \<not> P then x else y) = (if P then y else x)"
       
   326   "(if P then (if Q then x else y) else x) = (if P \<and> (\<not> Q) then y else x)"
       
   327   "(if P then (if Q then x else y) else x) = (if (\<not> Q) \<and> P then y else x)"
       
   328   "(if P then (if Q then x else y) else y) = (if P \<and> Q then x else y)"
       
   329   "(if P then (if Q then x else y) else y) = (if Q \<and> P then x else y)"
       
   330   "(if P then x else if P then y else z) = (if P then x else z)"
       
   331   "(if P then x else if Q then x else y) = (if P \<or> Q then x else y)"
       
   332   "(if P then x else if Q then x else y) = (if Q \<or> P then x else y)"
       
   333   "(if P then x = y else x = z) = (x = (if P then y else z))"
       
   334   "(if P then x = y else y = z) = (y = (if P then x else z))"
       
   335   "(if P then x = y else z = y) = (y = (if P then x else z))"
       
   336   by auto
       
   337 
       
   338 lemma [z3_rule]:
       
   339   "0 + (x::int) = x"
       
   340   "x + 0 = x"
       
   341   "x + x = 2 * x"
       
   342   "0 * x = 0"
       
   343   "1 * x = x"
       
   344   "x + y = y + x"
       
   345   by (auto simp add: mult_2)
       
   346 
       
   347 lemma [z3_rule]:  (* for def-axiom *)
       
   348   "P = Q \<or> P \<or> Q"
       
   349   "P = Q \<or> \<not> P \<or> \<not> Q"
       
   350   "(\<not> P) = Q \<or> \<not> P \<or> Q"
       
   351   "(\<not> P) = Q \<or> P \<or> \<not> Q"
       
   352   "P = (\<not> Q) \<or> \<not> P \<or> Q"
       
   353   "P = (\<not> Q) \<or> P \<or> \<not> Q"
       
   354   "P \<noteq> Q \<or> P \<or> \<not> Q"
       
   355   "P \<noteq> Q \<or> \<not> P \<or> Q"
       
   356   "P \<noteq> (\<not> Q) \<or> P \<or> Q"
       
   357   "(\<not> P) \<noteq> Q \<or> P \<or> Q"
       
   358   "P \<or> Q \<or> P \<noteq> (\<not> Q)"
       
   359   "P \<or> Q \<or> (\<not> P) \<noteq> Q"
       
   360   "P \<or> \<not> Q \<or> P \<noteq> Q"
       
   361   "\<not> P \<or> Q \<or> P \<noteq> Q"
       
   362   "P \<or> y = (if P then x else y)"
       
   363   "P \<or> (if P then x else y) = y"
       
   364   "\<not> P \<or> x = (if P then x else y)"
       
   365   "\<not> P \<or> (if P then x else y) = x"
       
   366   "P \<or> R \<or> \<not> (if P then Q else R)"
       
   367   "\<not> P \<or> Q \<or> \<not> (if P then Q else R)"
       
   368   "\<not> (if P then Q else R) \<or> \<not> P \<or> Q"
       
   369   "\<not> (if P then Q else R) \<or> P \<or> R"
       
   370   "(if P then Q else R) \<or> \<not> P \<or> \<not> Q"
       
   371   "(if P then Q else R) \<or> P \<or> \<not> R"
       
   372   "(if P then \<not> Q else R) \<or> \<not> P \<or> Q"
       
   373   "(if P then Q else \<not> R) \<or> P \<or> R"
       
   374   by auto
       
   375 
       
   376 hide_type (open) symb_list pattern
       
   377 hide_const (open) Symb_Nil Symb_Cons trigger pat nopat fun_app z3div z3mod
       
   378 
       
   379 end