blanchet@56078: (* Title: HOL/SMT2.thy blanchet@56078: Author: Sascha Boehme, TU Muenchen blanchet@56078: *) blanchet@56078: blanchet@56078: header {* Bindings to Satisfiability Modulo Theories (SMT) solvers based on SMT-LIB 2 *} blanchet@56078: blanchet@56078: theory SMT2 blanchet@57230: imports Divides blanchet@56078: keywords "smt2_status" :: diag blanchet@56078: begin blanchet@56078: blanchet@56078: subsection {* Triggers for quantifier instantiation *} blanchet@56078: blanchet@56078: text {* blanchet@56078: Some SMT solvers support patterns as a quantifier instantiation blanchet@56078: heuristics. Patterns may either be positive terms (tagged by "pat") blanchet@56078: triggering quantifier instantiations -- when the solver finds a blanchet@56078: term matching a positive pattern, it instantiates the corresponding blanchet@56078: quantifier accordingly -- or negative terms (tagged by "nopat") blanchet@56078: inhibiting quantifier instantiations. A list of patterns blanchet@56078: of the same kind is called a multipattern, and all patterns in a blanchet@56078: multipattern are considered conjunctively for quantifier instantiation. blanchet@56078: A list of multipatterns is called a trigger, and their multipatterns blanchet@56078: act disjunctively during quantifier instantiation. Each multipattern blanchet@56078: should mention at least all quantified variables of the preceding blanchet@56078: quantifier block. blanchet@56078: *} blanchet@56078: blanchet@57230: typedecl 'a symb_list blanchet@57230: blanchet@57230: consts blanchet@57230: Symb_Nil :: "'a symb_list" blanchet@57230: Symb_Cons :: "'a \ 'a symb_list \ 'a symb_list" blanchet@57230: blanchet@56078: typedecl pattern blanchet@56078: blanchet@56078: consts blanchet@56078: pat :: "'a \ pattern" blanchet@56078: nopat :: "'a \ pattern" blanchet@56078: blanchet@57230: definition trigger :: "pattern symb_list symb_list \ bool \ bool" where blanchet@57230: "trigger _ P = P" blanchet@56078: blanchet@56078: blanchet@56078: subsection {* Higher-order encoding *} blanchet@56078: blanchet@56078: text {* blanchet@56078: Application is made explicit for constants occurring with varying blanchet@56078: numbers of arguments. This is achieved by the introduction of the blanchet@56078: following constant. blanchet@56078: *} blanchet@56078: blanchet@56078: definition fun_app :: "'a \ 'a" where "fun_app f = f" blanchet@56078: blanchet@56078: text {* blanchet@56078: Some solvers support a theory of arrays which can be used to encode blanchet@56078: higher-order functions. The following set of lemmas specifies the blanchet@56078: properties of such (extensional) arrays. blanchet@56078: *} blanchet@56078: blanchet@56078: lemmas array_rules = ext fun_upd_apply fun_upd_same fun_upd_other fun_upd_upd fun_app_def blanchet@56078: blanchet@56078: blanchet@56103: subsection {* Normalization *} blanchet@56103: blanchet@56103: lemma case_bool_if[abs_def]: "case_bool x y P = (if P then x else y)" blanchet@56103: by simp blanchet@56103: blanchet@56107: lemma nat_int': "\n. nat (int n) = n" by simp blanchet@56103: lemma int_nat_nneg: "\i. i \ 0 \ int (nat i) = i" by simp blanchet@56103: lemma int_nat_neg: "\i. i < 0 \ int (nat i) = 0" by simp blanchet@56103: blanchet@57230: lemma nat_zero_as_int: "0 = nat 0" by simp blanchet@57230: lemma nat_one_as_int: "1 = nat 1" by simp blanchet@56103: lemma nat_numeral_as_int: "numeral = (\i. nat (numeral i))" by simp blanchet@56103: lemma nat_less_as_int: "op < = (\a b. int a < int b)" by simp blanchet@56103: lemma nat_leq_as_int: "op \ = (\a b. int a <= int b)" by simp blanchet@56103: lemma Suc_as_int: "Suc = (\a. nat (int a + 1))" by (rule ext) simp blanchet@56103: lemma nat_plus_as_int: "op + = (\a b. nat (int a + int b))" by (rule ext)+ simp blanchet@56103: lemma nat_minus_as_int: "op - = (\a b. nat (int a - int b))" by (rule ext)+ simp blanchet@56103: lemma nat_times_as_int: "op * = (\a b. nat (int a * int b))" by (simp add: nat_mult_distrib) blanchet@56103: lemma nat_div_as_int: "op div = (\a b. nat (int a div int b))" by (simp add: nat_div_distrib) blanchet@56103: lemma nat_mod_as_int: "op mod = (\a b. nat (int a mod int b))" by (simp add: nat_mod_distrib) blanchet@56103: blanchet@56103: lemma int_Suc: "int (Suc n) = int n + 1" by simp blanchet@56103: lemma int_plus: "int (n + m) = int n + int m" by (rule of_nat_add) blanchet@56103: lemma int_minus: "int (n - m) = int (nat (int n - int m))" by auto blanchet@56103: blanchet@56103: lemmas Ex1_def_raw = Ex1_def[abs_def] blanchet@56103: lemmas Ball_def_raw = Ball_def[abs_def] blanchet@56103: lemmas Bex_def_raw = Bex_def[abs_def] blanchet@56103: lemmas abs_if_raw = abs_if[abs_def] blanchet@56103: lemmas min_def_raw = min_def[abs_def] blanchet@56103: lemmas max_def_raw = max_def[abs_def] blanchet@56103: blanchet@56103: blanchet@56078: subsection {* Integer division and modulo for Z3 *} blanchet@56078: blanchet@56102: text {* blanchet@56102: The following Z3-inspired definitions are overspecified for the case where @{text "l = 0"}. This blanchet@56102: Schönheitsfehler is corrected in the @{text div_as_z3div} and @{text mod_as_z3mod} theorems. blanchet@56102: *} blanchet@56102: blanchet@56078: definition z3div :: "int \ int \ int" where blanchet@56102: "z3div k l = (if l \ 0 then k div l else - (k div - l))" blanchet@56078: blanchet@56078: definition z3mod :: "int \ int \ int" where blanchet@56102: "z3mod k l = k mod (if l \ 0 then l else - l)" blanchet@56078: blanchet@56101: lemma div_as_z3div: blanchet@56102: "\k l. k div l = (if l = 0 then 0 else if l > 0 then z3div k l else z3div (- k) (- l))" blanchet@56101: by (simp add: z3div_def) blanchet@56101: blanchet@56101: lemma mod_as_z3mod: blanchet@56102: "\k l. k mod l = (if l = 0 then k else if l > 0 then z3mod k l else - z3mod (- k) (- l))" blanchet@56101: by (simp add: z3mod_def) blanchet@56101: blanchet@56078: blanchet@56078: subsection {* Setup *} blanchet@56078: blanchet@57230: ML_file "Tools/SMT2/smt2_util.ML" blanchet@57230: ML_file "Tools/SMT2/smt2_failure.ML" blanchet@57230: ML_file "Tools/SMT2/smt2_config.ML" blanchet@56078: ML_file "Tools/SMT2/smt2_builtin.ML" blanchet@56078: ML_file "Tools/SMT2/smt2_datatypes.ML" blanchet@56078: ML_file "Tools/SMT2/smt2_normalize.ML" blanchet@56078: ML_file "Tools/SMT2/smt2_translate.ML" blanchet@56078: ML_file "Tools/SMT2/smtlib2.ML" blanchet@56078: ML_file "Tools/SMT2/smtlib2_interface.ML" blanchet@57219: ML_file "Tools/SMT2/smtlib2_proof.ML" blanchet@56083: ML_file "Tools/SMT2/z3_new_proof.ML" blanchet@57159: ML_file "Tools/SMT2/z3_new_isar.ML" blanchet@56083: ML_file "Tools/SMT2/smt2_solver.ML" blanchet@56078: ML_file "Tools/SMT2/z3_new_interface.ML" blanchet@56090: ML_file "Tools/SMT2/z3_new_replay_util.ML" blanchet@56090: ML_file "Tools/SMT2/z3_new_replay_literals.ML" blanchet@56090: ML_file "Tools/SMT2/z3_new_replay_rules.ML" blanchet@56090: ML_file "Tools/SMT2/z3_new_replay_methods.ML" blanchet@56090: ML_file "Tools/SMT2/z3_new_replay.ML" blanchet@56090: ML_file "Tools/SMT2/smt2_systems.ML" blanchet@56078: blanchet@56078: method_setup smt2 = {* blanchet@56078: Scan.optional Attrib.thms [] >> blanchet@56078: (fn thms => fn ctxt => blanchet@56078: METHOD (fn facts => HEADGOAL (SMT2_Solver.smt2_tac ctxt (thms @ facts)))) blanchet@56078: *} "apply an SMT solver to the current goal (based on SMT-LIB 2)" blanchet@56078: blanchet@56078: blanchet@56078: subsection {* Configuration *} blanchet@56078: blanchet@56078: text {* blanchet@56078: The current configuration can be printed by the command blanchet@56078: @{text smt2_status}, which shows the values of most options. blanchet@56078: *} blanchet@56078: blanchet@56078: blanchet@56078: blanchet@56078: subsection {* General configuration options *} blanchet@56078: blanchet@56078: text {* blanchet@56078: The option @{text smt2_solver} can be used to change the target SMT blanchet@56078: solver. The possible values can be obtained from the @{text smt2_status} blanchet@56078: command. blanchet@56078: blanchet@57237: Due to licensing restrictions, Z3 is not enabled by default. Z3 is free blanchet@57237: for non-commercial applications and can be enabled by setting Isabelle blanchet@57237: system option @{text z3_non_commercial} to @{text yes}. blanchet@56078: *} blanchet@56078: blanchet@57209: declare [[smt2_solver = z3]] blanchet@56078: blanchet@56078: text {* blanchet@56078: Since SMT solvers are potentially non-terminating, there is a timeout blanchet@56078: (given in seconds) to restrict their runtime. A value greater than blanchet@56078: 120 (seconds) is in most cases not advisable. blanchet@56078: *} blanchet@56078: blanchet@57209: declare [[smt2_timeout = 20]] blanchet@56078: blanchet@56078: text {* blanchet@56078: SMT solvers apply randomized heuristics. In case a problem is not blanchet@56078: solvable by an SMT solver, changing the following option might help. blanchet@56078: *} blanchet@56078: blanchet@57209: declare [[smt2_random_seed = 1]] blanchet@56078: blanchet@56078: text {* blanchet@56078: In general, the binding to SMT solvers runs as an oracle, i.e, the SMT blanchet@56078: solvers are fully trusted without additional checks. The following blanchet@56078: option can cause the SMT solver to run in proof-producing mode, giving blanchet@56078: a checkable certificate. This is currently only implemented for Z3. blanchet@56078: *} blanchet@56078: blanchet@57209: declare [[smt2_oracle = false]] blanchet@56078: blanchet@56078: text {* blanchet@56078: Each SMT solver provides several commandline options to tweak its blanchet@56078: behaviour. They can be passed to the solver by setting the following blanchet@56078: options. blanchet@56078: *} blanchet@56078: blanchet@57209: declare [[cvc3_new_options = ""]] blanchet@57246: declare [[cvc4_new_options = ""]] blanchet@57209: declare [[z3_new_options = ""]] blanchet@56078: blanchet@56078: text {* blanchet@56078: The SMT method provides an inference mechanism to detect simple triggers blanchet@56078: in quantified formulas, which might increase the number of problems blanchet@56078: solvable by SMT solvers (note: triggers guide quantifier instantiations blanchet@56078: in the SMT solver). To turn it on, set the following option. blanchet@56078: *} blanchet@56078: blanchet@57209: declare [[smt2_infer_triggers = false]] blanchet@56078: blanchet@56078: text {* blanchet@56078: Enable the following option to use built-in support for div/mod, datatypes, blanchet@56078: and records in Z3. Currently, this is implemented only in oracle mode. blanchet@56078: *} blanchet@56078: blanchet@57209: declare [[z3_new_extensions = false]] blanchet@56078: blanchet@56078: blanchet@56078: subsection {* Certificates *} blanchet@56078: blanchet@56078: text {* blanchet@56078: By setting the option @{text smt2_certificates} to the name of a file, blanchet@56078: all following applications of an SMT solver a cached in that file. blanchet@56078: Any further application of the same SMT solver (using the very same blanchet@56078: configuration) re-uses the cached certificate instead of invoking the blanchet@56078: solver. An empty string disables caching certificates. blanchet@56078: blanchet@56078: The filename should be given as an explicit path. It is good blanchet@56078: practice to use the name of the current theory (with ending blanchet@56078: @{text ".certs"} instead of @{text ".thy"}) as the certificates file. blanchet@56078: Certificate files should be used at most once in a certain theory context, blanchet@56078: to avoid race conditions with other concurrent accesses. blanchet@56078: *} blanchet@56078: blanchet@57209: declare [[smt2_certificates = ""]] blanchet@56078: blanchet@56078: text {* blanchet@56078: The option @{text smt2_read_only_certificates} controls whether only blanchet@56078: stored certificates are should be used or invocation of an SMT solver blanchet@56078: is allowed. When set to @{text true}, no SMT solver will ever be blanchet@56078: invoked and only the existing certificates found in the configured blanchet@56078: cache are used; when set to @{text false} and there is no cached blanchet@56078: certificate for some proposition, then the configured SMT solver is blanchet@56078: invoked. blanchet@56078: *} blanchet@56078: blanchet@57209: declare [[smt2_read_only_certificates = false]] blanchet@56078: blanchet@56078: blanchet@56078: blanchet@56078: subsection {* Tracing *} blanchet@56078: blanchet@56078: text {* blanchet@56078: The SMT method, when applied, traces important information. To blanchet@56078: make it entirely silent, set the following option to @{text false}. blanchet@56078: *} blanchet@56078: blanchet@57209: declare [[smt2_verbose = true]] blanchet@56078: blanchet@56078: text {* blanchet@56078: For tracing the generated problem file given to the SMT solver as blanchet@56078: well as the returned result of the solver, the option blanchet@56078: @{text smt2_trace} should be set to @{text true}. blanchet@56078: *} blanchet@56078: blanchet@57209: declare [[smt2_trace = false]] blanchet@56078: blanchet@56078: blanchet@56078: subsection {* Schematic rules for Z3 proof reconstruction *} blanchet@56078: blanchet@56078: text {* blanchet@56078: Several prof rules of Z3 are not very well documented. There are two blanchet@56078: lemma groups which can turn failing Z3 proof reconstruction attempts blanchet@56078: into succeeding ones: the facts in @{text z3_rule} are tried prior to blanchet@56078: any implemented reconstruction procedure for all uncertain Z3 proof blanchet@56078: rules; the facts in @{text z3_simp} are only fed to invocations of blanchet@56078: the simplifier when reconstructing theory-specific proof steps. blanchet@56078: *} blanchet@56078: blanchet@56078: lemmas [z3_new_rule] = blanchet@56078: refl eq_commute conj_commute disj_commute simp_thms nnf_simps blanchet@56078: ring_distribs field_simps times_divide_eq_right times_divide_eq_left blanchet@56078: if_True if_False not_not blanchet@56078: blanchet@56078: lemma [z3_new_rule]: blanchet@57169: "(P \ Q) = (\ (\ P \ \ Q))" blanchet@57169: "(P \ Q) = (\ (\ Q \ \ P))" blanchet@57169: "(\ P \ Q) = (\ (P \ \ Q))" blanchet@57169: "(\ P \ Q) = (\ (\ Q \ P))" blanchet@57169: "(P \ \ Q) = (\ (\ P \ Q))" blanchet@57169: "(P \ \ Q) = (\ (Q \ \ P))" blanchet@57169: "(\ P \ \ Q) = (\ (P \ Q))" blanchet@57169: "(\ P \ \ Q) = (\ (Q \ P))" blanchet@56078: by auto blanchet@56078: blanchet@56078: lemma [z3_new_rule]: blanchet@57169: "(P \ Q) = (Q \ \ P)" blanchet@57169: "(\ P \ Q) = (P \ Q)" blanchet@57169: "(\ P \ Q) = (Q \ P)" blanchet@56078: "(True \ P) = P" blanchet@56078: "(P \ True) = True" blanchet@56078: "(False \ P) = True" blanchet@56078: "(P \ P) = True" blanchet@56078: by auto blanchet@56078: blanchet@56078: lemma [z3_new_rule]: blanchet@57169: "((P = Q) \ R) = (R | (Q = (\ P)))" blanchet@56078: by auto blanchet@56078: blanchet@56078: lemma [z3_new_rule]: blanchet@57169: "(\ True) = False" blanchet@57169: "(\ False) = True" blanchet@56078: "(x = x) = True" blanchet@56078: "(P = True) = P" blanchet@56078: "(True = P) = P" blanchet@57169: "(P = False) = (\ P)" blanchet@57169: "(False = P) = (\ P)" blanchet@57169: "((\ P) = P) = False" blanchet@57169: "(P = (\ P)) = False" blanchet@57169: "((\ P) = (\ Q)) = (P = Q)" blanchet@57169: "\ (P = (\ Q)) = (P = Q)" blanchet@57169: "\ ((\ P) = Q) = (P = Q)" blanchet@57169: "(P \ Q) = (Q = (\ P))" blanchet@57169: "(P = Q) = ((\ P \ Q) \ (P \ \ Q))" blanchet@57169: "(P \ Q) = ((\ P \ \ Q) \ (P \ Q))" blanchet@56078: by auto blanchet@56078: blanchet@56078: lemma [z3_new_rule]: blanchet@57169: "(if P then P else \ P) = True" blanchet@57169: "(if \ P then \ P else P) = True" blanchet@56078: "(if P then True else False) = P" blanchet@57169: "(if P then False else True) = (\ P)" blanchet@57169: "(if P then Q else True) = ((\ P) \ Q)" blanchet@57169: "(if P then Q else True) = (Q \ (\ P))" blanchet@57169: "(if P then Q else \ Q) = (P = Q)" blanchet@57169: "(if P then Q else \ Q) = (Q = P)" blanchet@57169: "(if P then \ Q else Q) = (P = (\ Q))" blanchet@57169: "(if P then \ Q else Q) = ((\ Q) = P)" blanchet@57169: "(if \ P then x else y) = (if P then y else x)" blanchet@57169: "(if P then (if Q then x else y) else x) = (if P \ (\ Q) then y else x)" blanchet@57169: "(if P then (if Q then x else y) else x) = (if (\ Q) \ P then y else x)" blanchet@56078: "(if P then (if Q then x else y) else y) = (if P \ Q then x else y)" blanchet@56078: "(if P then (if Q then x else y) else y) = (if Q \ P then x else y)" blanchet@56078: "(if P then x else if P then y else z) = (if P then x else z)" blanchet@56078: "(if P then x else if Q then x else y) = (if P \ Q then x else y)" blanchet@56078: "(if P then x else if Q then x else y) = (if Q \ P then x else y)" blanchet@56078: "(if P then x = y else x = z) = (x = (if P then y else z))" blanchet@56078: "(if P then x = y else y = z) = (y = (if P then x else z))" blanchet@56078: "(if P then x = y else z = y) = (y = (if P then x else z))" blanchet@56078: by auto blanchet@56078: blanchet@56078: lemma [z3_new_rule]: blanchet@56078: "0 + (x::int) = x" blanchet@56078: "x + 0 = x" blanchet@56078: "x + x = 2 * x" blanchet@56078: "0 * x = 0" blanchet@56078: "1 * x = x" blanchet@56078: "x + y = y + x" blanchet@57230: by (auto simp add: mult_2) blanchet@56078: blanchet@56078: lemma [z3_new_rule]: (* for def-axiom *) blanchet@56078: "P = Q \ P \ Q" blanchet@57169: "P = Q \ \ P \ \ Q" blanchet@57169: "(\ P) = Q \ \ P \ Q" blanchet@57169: "(\ P) = Q \ P \ \ Q" blanchet@57169: "P = (\ Q) \ \ P \ Q" blanchet@57169: "P = (\ Q) \ P \ \ Q" blanchet@57169: "P \ Q \ P \ \ Q" blanchet@57169: "P \ Q \ \ P \ Q" blanchet@57169: "P \ (\ Q) \ P \ Q" blanchet@57169: "(\ P) \ Q \ P \ Q" blanchet@57169: "P \ Q \ P \ (\ Q)" blanchet@57169: "P \ Q \ (\ P) \ Q" blanchet@57169: "P \ \ Q \ P \ Q" blanchet@57169: "\ P \ Q \ P \ Q" blanchet@56078: "P \ y = (if P then x else y)" blanchet@56078: "P \ (if P then x else y) = y" blanchet@57169: "\ P \ x = (if P then x else y)" blanchet@57169: "\ P \ (if P then x else y) = x" blanchet@57169: "P \ R \ \ (if P then Q else R)" blanchet@57169: "\ P \ Q \ \ (if P then Q else R)" blanchet@57169: "\ (if P then Q else R) \ \ P \ Q" blanchet@57169: "\ (if P then Q else R) \ P \ R" blanchet@57169: "(if P then Q else R) \ \ P \ \ Q" blanchet@57169: "(if P then Q else R) \ P \ \ R" blanchet@57169: "(if P then \ Q else R) \ \ P \ Q" blanchet@57169: "(if P then Q else \ R) \ P \ R" blanchet@56078: by auto blanchet@56078: blanchet@57230: hide_type (open) symb_list pattern blanchet@57230: hide_const (open) Symb_Nil Symb_Cons trigger pat nopat fun_app z3div z3mod blanchet@56078: blanchet@56078: end