| author | wenzelm | 
| Sun, 20 Jul 2014 17:54:01 +0200 | |
| changeset 57581 | 74bbe9317aa4 | 
| parent 57246 | 62746a41cc0c | 
| child 57696 | fb71c6f100f8 | 
| permissions | -rw-r--r-- | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 1 | (* Title: HOL/SMT2.thy | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 2 | Author: Sascha Boehme, TU Muenchen | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 3 | *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 4 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 5 | header {* Bindings to Satisfiability Modulo Theories (SMT) solvers based on SMT-LIB 2 *}
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 6 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 7 | theory SMT2 | 
| 57230 | 8 | imports Divides | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 9 | keywords "smt2_status" :: diag | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 10 | begin | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 11 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 12 | subsection {* Triggers for quantifier instantiation *}
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 13 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 14 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 15 | Some SMT solvers support patterns as a quantifier instantiation | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 16 | heuristics. Patterns may either be positive terms (tagged by "pat") | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 17 | triggering quantifier instantiations -- when the solver finds a | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 18 | term matching a positive pattern, it instantiates the corresponding | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 19 | quantifier accordingly -- or negative terms (tagged by "nopat") | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 20 | inhibiting quantifier instantiations. A list of patterns | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 21 | of the same kind is called a multipattern, and all patterns in a | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 22 | multipattern are considered conjunctively for quantifier instantiation. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 23 | A list of multipatterns is called a trigger, and their multipatterns | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 24 | act disjunctively during quantifier instantiation. Each multipattern | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 25 | should mention at least all quantified variables of the preceding | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 26 | quantifier block. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 27 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 28 | |
| 57230 | 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 | ||
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 35 | typedecl pattern | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 36 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 37 | consts | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 38 | pat :: "'a \<Rightarrow> pattern" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 39 | nopat :: "'a \<Rightarrow> pattern" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 40 | |
| 57230 | 41 | definition trigger :: "pattern symb_list symb_list \<Rightarrow> bool \<Rightarrow> bool" where | 
| 42 | "trigger _ P = P" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 43 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 44 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 45 | subsection {* Higher-order encoding *}
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 46 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 47 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 48 | Application is made explicit for constants occurring with varying | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 49 | numbers of arguments. This is achieved by the introduction of the | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 50 | following constant. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 51 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 52 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 53 | definition fun_app :: "'a \<Rightarrow> 'a" where "fun_app f = f" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 54 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 55 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 56 | Some solvers support a theory of arrays which can be used to encode | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 57 | higher-order functions. The following set of lemmas specifies the | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 58 | properties of such (extensional) arrays. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 59 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 60 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 61 | lemmas array_rules = ext fun_upd_apply fun_upd_same fun_upd_other fun_upd_upd fun_app_def | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 62 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 63 | |
| 56103 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 64 | subsection {* Normalization *}
 | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 65 | |
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 66 | lemma case_bool_if[abs_def]: "case_bool x y P = (if P then x else y)" | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 67 | by simp | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 68 | |
| 56107 | 69 | lemma nat_int': "\<forall>n. nat (int n) = n" by simp | 
| 56103 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 70 | lemma int_nat_nneg: "\<forall>i. i \<ge> 0 \<longrightarrow> int (nat i) = i" by simp | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 71 | lemma int_nat_neg: "\<forall>i. i < 0 \<longrightarrow> int (nat i) = 0" by simp | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 72 | |
| 57230 | 73 | lemma nat_zero_as_int: "0 = nat 0" by simp | 
| 74 | lemma nat_one_as_int: "1 = nat 1" by simp | |
| 56103 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 75 | lemma nat_numeral_as_int: "numeral = (\<lambda>i. nat (numeral i))" by simp | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 76 | lemma nat_less_as_int: "op < = (\<lambda>a b. int a < int b)" by simp | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 77 | lemma nat_leq_as_int: "op \<le> = (\<lambda>a b. int a <= int b)" by simp | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 78 | lemma Suc_as_int: "Suc = (\<lambda>a. nat (int a + 1))" by (rule ext) simp | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 79 | lemma nat_plus_as_int: "op + = (\<lambda>a b. nat (int a + int b))" by (rule ext)+ simp | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 80 | lemma nat_minus_as_int: "op - = (\<lambda>a b. nat (int a - int b))" by (rule ext)+ simp | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 81 | lemma nat_times_as_int: "op * = (\<lambda>a b. nat (int a * int b))" by (simp add: nat_mult_distrib) | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 82 | lemma nat_div_as_int: "op div = (\<lambda>a b. nat (int a div int b))" by (simp add: nat_div_distrib) | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 83 | lemma nat_mod_as_int: "op mod = (\<lambda>a b. nat (int a mod int b))" by (simp add: nat_mod_distrib) | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 84 | |
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 85 | lemma int_Suc: "int (Suc n) = int n + 1" by simp | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 86 | lemma int_plus: "int (n + m) = int n + int m" by (rule of_nat_add) | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 87 | lemma int_minus: "int (n - m) = int (nat (int n - int m))" by auto | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 88 | |
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 89 | lemmas Ex1_def_raw = Ex1_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 90 | lemmas Ball_def_raw = Ball_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 91 | lemmas Bex_def_raw = Bex_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 92 | lemmas abs_if_raw = abs_if[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 93 | lemmas min_def_raw = min_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 94 | lemmas max_def_raw = max_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 95 | |
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 96 | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 97 | subsection {* Integer division and modulo for Z3 *}
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 98 | |
| 56102 | 99 | text {*
 | 
| 100 | The following Z3-inspired definitions are overspecified for the case where @{text "l = 0"}. This
 | |
| 101 | Schönheitsfehler is corrected in the @{text div_as_z3div} and @{text mod_as_z3mod} theorems.
 | |
| 102 | *} | |
| 103 | ||
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 104 | definition z3div :: "int \<Rightarrow> int \<Rightarrow> int" where | 
| 56102 | 105 | "z3div k l = (if l \<ge> 0 then k div l else - (k div - l))" | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 106 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 107 | definition z3mod :: "int \<Rightarrow> int \<Rightarrow> int" where | 
| 56102 | 108 | "z3mod k l = k mod (if l \<ge> 0 then l else - l)" | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 109 | |
| 56101 | 110 | lemma div_as_z3div: | 
| 56102 | 111 | "\<forall>k l. k div l = (if l = 0 then 0 else if l > 0 then z3div k l else z3div (- k) (- l))" | 
| 56101 | 112 | by (simp add: z3div_def) | 
| 113 | ||
| 114 | lemma mod_as_z3mod: | |
| 56102 | 115 | "\<forall>k l. k mod l = (if l = 0 then k else if l > 0 then z3mod k l else - z3mod (- k) (- l))" | 
| 56101 | 116 | by (simp add: z3mod_def) | 
| 117 | ||
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 118 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 119 | subsection {* Setup *}
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 120 | |
| 57230 | 121 | ML_file "Tools/SMT2/smt2_util.ML" | 
| 122 | ML_file "Tools/SMT2/smt2_failure.ML" | |
| 123 | ML_file "Tools/SMT2/smt2_config.ML" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 124 | ML_file "Tools/SMT2/smt2_builtin.ML" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 125 | ML_file "Tools/SMT2/smt2_datatypes.ML" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 126 | ML_file "Tools/SMT2/smt2_normalize.ML" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 127 | ML_file "Tools/SMT2/smt2_translate.ML" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 128 | ML_file "Tools/SMT2/smtlib2.ML" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 129 | ML_file "Tools/SMT2/smtlib2_interface.ML" | 
| 57219 
34018603e0d0
factor out SMT-LIB 2 type/term parsing from Z3-specific code
 blanchet parents: 
57209diff
changeset | 130 | ML_file "Tools/SMT2/smtlib2_proof.ML" | 
| 56083 
b5d1d9c60341
have Sledgehammer generate Isar proofs from Z3 proofs
 blanchet parents: 
56078diff
changeset | 131 | ML_file "Tools/SMT2/z3_new_proof.ML" | 
| 57159 | 132 | ML_file "Tools/SMT2/z3_new_isar.ML" | 
| 56083 
b5d1d9c60341
have Sledgehammer generate Isar proofs from Z3 proofs
 blanchet parents: 
56078diff
changeset | 133 | ML_file "Tools/SMT2/smt2_solver.ML" | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 134 | ML_file "Tools/SMT2/z3_new_interface.ML" | 
| 56090 | 135 | ML_file "Tools/SMT2/z3_new_replay_util.ML" | 
| 136 | ML_file "Tools/SMT2/z3_new_replay_literals.ML" | |
| 137 | ML_file "Tools/SMT2/z3_new_replay_rules.ML" | |
| 138 | ML_file "Tools/SMT2/z3_new_replay_methods.ML" | |
| 139 | ML_file "Tools/SMT2/z3_new_replay.ML" | |
| 140 | ML_file "Tools/SMT2/smt2_systems.ML" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 141 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 142 | method_setup smt2 = {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 143 | Scan.optional Attrib.thms [] >> | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 144 | (fn thms => fn ctxt => | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 145 | METHOD (fn facts => HEADGOAL (SMT2_Solver.smt2_tac ctxt (thms @ facts)))) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 146 | *} "apply an SMT solver to the current goal (based on SMT-LIB 2)" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 147 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 148 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 149 | subsection {* Configuration *}
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 150 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 151 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 152 | The current configuration can be printed by the command | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 153 | @{text smt2_status}, which shows the values of most options.
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 154 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 155 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 156 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 157 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 158 | subsection {* General configuration options *}
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 159 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 160 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 161 | The option @{text smt2_solver} can be used to change the target SMT
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 162 | solver.  The possible values can be obtained from the @{text smt2_status}
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 163 | command. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 164 | |
| 57237 
bc51864c2ac4
took out broken support for Yices from SMT2 stack -- see 'NEWS' for rationale
 blanchet parents: 
57230diff
changeset | 165 | Due to licensing restrictions, Z3 is not enabled by default. Z3 is free | 
| 
bc51864c2ac4
took out broken support for Yices from SMT2 stack -- see 'NEWS' for rationale
 blanchet parents: 
57230diff
changeset | 166 | for non-commercial applications and can be enabled by setting Isabelle | 
| 
bc51864c2ac4
took out broken support for Yices from SMT2 stack -- see 'NEWS' for rationale
 blanchet parents: 
57230diff
changeset | 167 | system option @{text z3_non_commercial} to @{text yes}.
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 168 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 169 | |
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57169diff
changeset | 170 | declare [[smt2_solver = z3]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 171 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 172 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 173 | Since SMT solvers are potentially non-terminating, there is a timeout | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 174 | (given in seconds) to restrict their runtime. A value greater than | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 175 | 120 (seconds) is in most cases not advisable. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 176 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 177 | |
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57169diff
changeset | 178 | declare [[smt2_timeout = 20]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 179 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 180 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 181 | SMT solvers apply randomized heuristics. In case a problem is not | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 182 | solvable by an SMT solver, changing the following option might help. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 183 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 184 | |
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57169diff
changeset | 185 | declare [[smt2_random_seed = 1]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 186 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 187 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 188 | In general, the binding to SMT solvers runs as an oracle, i.e, the SMT | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 189 | solvers are fully trusted without additional checks. The following | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 190 | option can cause the SMT solver to run in proof-producing mode, giving | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 191 | a checkable certificate. This is currently only implemented for Z3. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 192 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 193 | |
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57169diff
changeset | 194 | declare [[smt2_oracle = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 195 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 196 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 197 | Each SMT solver provides several commandline options to tweak its | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 198 | behaviour. They can be passed to the solver by setting the following | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 199 | options. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 200 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 201 | |
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57169diff
changeset | 202 | declare [[cvc3_new_options = ""]] | 
| 57246 | 203 | declare [[cvc4_new_options = ""]] | 
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57169diff
changeset | 204 | declare [[z3_new_options = ""]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 205 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 206 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 207 | The SMT method provides an inference mechanism to detect simple triggers | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 208 | in quantified formulas, which might increase the number of problems | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 209 | solvable by SMT solvers (note: triggers guide quantifier instantiations | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 210 | in the SMT solver). To turn it on, set the following option. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 211 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 212 | |
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57169diff
changeset | 213 | declare [[smt2_infer_triggers = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 214 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 215 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 216 | Enable the following option to use built-in support for div/mod, datatypes, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 217 | and records in Z3. Currently, this is implemented only in oracle mode. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 218 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 219 | |
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57169diff
changeset | 220 | declare [[z3_new_extensions = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 221 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 222 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 223 | subsection {* Certificates *}
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 224 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 225 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 226 | By setting the option @{text smt2_certificates} to the name of a file,
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 227 | all following applications of an SMT solver a cached in that file. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 228 | Any further application of the same SMT solver (using the very same | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 229 | configuration) re-uses the cached certificate instead of invoking the | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 230 | solver. An empty string disables caching certificates. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 231 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 232 | The filename should be given as an explicit path. It is good | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 233 | practice to use the name of the current theory (with ending | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 234 | @{text ".certs"} instead of @{text ".thy"}) as the certificates file.
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 235 | Certificate files should be used at most once in a certain theory context, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 236 | to avoid race conditions with other concurrent accesses. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 237 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 238 | |
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57169diff
changeset | 239 | declare [[smt2_certificates = ""]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 240 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 241 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 242 | The option @{text smt2_read_only_certificates} controls whether only
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 243 | stored certificates are should be used or invocation of an SMT solver | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 244 | is allowed.  When set to @{text true}, no SMT solver will ever be
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 245 | invoked and only the existing certificates found in the configured | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 246 | cache are used;  when set to @{text false} and there is no cached
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 247 | certificate for some proposition, then the configured SMT solver is | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 248 | invoked. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 249 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 250 | |
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57169diff
changeset | 251 | declare [[smt2_read_only_certificates = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 252 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 253 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 254 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 255 | subsection {* Tracing *}
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 256 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 257 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 258 | The SMT method, when applied, traces important information. To | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 259 | make it entirely silent, set the following option to @{text false}.
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 260 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 261 | |
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57169diff
changeset | 262 | declare [[smt2_verbose = true]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 263 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 264 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 265 | For tracing the generated problem file given to the SMT solver as | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 266 | well as the returned result of the solver, the option | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 267 | @{text smt2_trace} should be set to @{text true}.
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 268 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 269 | |
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
57169diff
changeset | 270 | declare [[smt2_trace = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 271 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 272 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 273 | subsection {* Schematic rules for Z3 proof reconstruction *}
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 274 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 275 | text {*
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 276 | Several prof rules of Z3 are not very well documented. There are two | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 277 | lemma groups which can turn failing Z3 proof reconstruction attempts | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 278 | into succeeding ones: the facts in @{text z3_rule} are tried prior to
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 279 | any implemented reconstruction procedure for all uncertain Z3 proof | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 280 | rules;  the facts in @{text z3_simp} are only fed to invocations of
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 281 | the simplifier when reconstructing theory-specific proof steps. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 282 | *} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 283 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 284 | lemmas [z3_new_rule] = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 285 | refl eq_commute conj_commute disj_commute simp_thms nnf_simps | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 286 | ring_distribs field_simps times_divide_eq_right times_divide_eq_left | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 287 | if_True if_False not_not | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 288 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 289 | lemma [z3_new_rule]: | 
| 57169 | 290 | "(P \<and> Q) = (\<not> (\<not> P \<or> \<not> Q))" | 
| 291 | "(P \<and> Q) = (\<not> (\<not> Q \<or> \<not> P))" | |
| 292 | "(\<not> P \<and> Q) = (\<not> (P \<or> \<not> Q))" | |
| 293 | "(\<not> P \<and> Q) = (\<not> (\<not> Q \<or> P))" | |
| 294 | "(P \<and> \<not> Q) = (\<not> (\<not> P \<or> Q))" | |
| 295 | "(P \<and> \<not> Q) = (\<not> (Q \<or> \<not> P))" | |
| 296 | "(\<not> P \<and> \<not> Q) = (\<not> (P \<or> Q))" | |
| 297 | "(\<not> P \<and> \<not> Q) = (\<not> (Q \<or> P))" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 298 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 299 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 300 | lemma [z3_new_rule]: | 
| 57169 | 301 | "(P \<longrightarrow> Q) = (Q \<or> \<not> P)" | 
| 302 | "(\<not> P \<longrightarrow> Q) = (P \<or> Q)" | |
| 303 | "(\<not> P \<longrightarrow> Q) = (Q \<or> P)" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 304 | "(True \<longrightarrow> P) = P" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 305 | "(P \<longrightarrow> True) = True" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 306 | "(False \<longrightarrow> P) = True" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 307 | "(P \<longrightarrow> P) = True" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 308 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 309 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 310 | lemma [z3_new_rule]: | 
| 57169 | 311 | "((P = Q) \<longrightarrow> R) = (R | (Q = (\<not> P)))" | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 312 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 313 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 314 | lemma [z3_new_rule]: | 
| 57169 | 315 | "(\<not> True) = False" | 
| 316 | "(\<not> False) = True" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 317 | "(x = x) = True" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 318 | "(P = True) = P" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 319 | "(True = P) = P" | 
| 57169 | 320 | "(P = False) = (\<not> P)" | 
| 321 | "(False = P) = (\<not> P)" | |
| 322 | "((\<not> P) = P) = False" | |
| 323 | "(P = (\<not> P)) = False" | |
| 324 | "((\<not> P) = (\<not> Q)) = (P = Q)" | |
| 325 | "\<not> (P = (\<not> Q)) = (P = Q)" | |
| 326 | "\<not> ((\<not> P) = Q) = (P = Q)" | |
| 327 | "(P \<noteq> Q) = (Q = (\<not> P))" | |
| 328 | "(P = Q) = ((\<not> P \<or> Q) \<and> (P \<or> \<not> Q))" | |
| 329 | "(P \<noteq> Q) = ((\<not> P \<or> \<not> Q) \<and> (P \<or> Q))" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 330 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 331 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 332 | lemma [z3_new_rule]: | 
| 57169 | 333 | "(if P then P else \<not> P) = True" | 
| 334 | "(if \<not> P then \<not> P else P) = True" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 335 | "(if P then True else False) = P" | 
| 57169 | 336 | "(if P then False else True) = (\<not> P)" | 
| 337 | "(if P then Q else True) = ((\<not> P) \<or> Q)" | |
| 338 | "(if P then Q else True) = (Q \<or> (\<not> P))" | |
| 339 | "(if P then Q else \<not> Q) = (P = Q)" | |
| 340 | "(if P then Q else \<not> Q) = (Q = P)" | |
| 341 | "(if P then \<not> Q else Q) = (P = (\<not> Q))" | |
| 342 | "(if P then \<not> Q else Q) = ((\<not> Q) = P)" | |
| 343 | "(if \<not> P then x else y) = (if P then y else x)" | |
| 344 | "(if P then (if Q then x else y) else x) = (if P \<and> (\<not> Q) then y else x)" | |
| 345 | "(if P then (if Q then x else y) else x) = (if (\<not> Q) \<and> P then y else x)" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 346 | "(if P then (if Q then x else y) else y) = (if P \<and> Q then x else y)" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 347 | "(if P then (if Q then x else y) else y) = (if Q \<and> P then x else y)" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 348 | "(if P then x else if P then y else z) = (if P then x else z)" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 349 | "(if P then x else if Q then x else y) = (if P \<or> Q then x else y)" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 350 | "(if P then x else if Q then x else y) = (if Q \<or> P then x else y)" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 351 | "(if P then x = y else x = z) = (x = (if P then y else z))" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 352 | "(if P then x = y else y = z) = (y = (if P then x else z))" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 353 | "(if P then x = y else z = y) = (y = (if P then x else z))" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 354 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 355 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 356 | lemma [z3_new_rule]: | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 357 | "0 + (x::int) = x" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 358 | "x + 0 = x" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 359 | "x + x = 2 * x" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 360 | "0 * x = 0" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 361 | "1 * x = x" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 362 | "x + y = y + x" | 
| 57230 | 363 | by (auto simp add: mult_2) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 364 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 365 | lemma [z3_new_rule]: (* for def-axiom *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 366 | "P = Q \<or> P \<or> Q" | 
| 57169 | 367 | "P = Q \<or> \<not> P \<or> \<not> Q" | 
| 368 | "(\<not> P) = Q \<or> \<not> P \<or> Q" | |
| 369 | "(\<not> P) = Q \<or> P \<or> \<not> Q" | |
| 370 | "P = (\<not> Q) \<or> \<not> P \<or> Q" | |
| 371 | "P = (\<not> Q) \<or> P \<or> \<not> Q" | |
| 372 | "P \<noteq> Q \<or> P \<or> \<not> Q" | |
| 373 | "P \<noteq> Q \<or> \<not> P \<or> Q" | |
| 374 | "P \<noteq> (\<not> Q) \<or> P \<or> Q" | |
| 375 | "(\<not> P) \<noteq> Q \<or> P \<or> Q" | |
| 376 | "P \<or> Q \<or> P \<noteq> (\<not> Q)" | |
| 377 | "P \<or> Q \<or> (\<not> P) \<noteq> Q" | |
| 378 | "P \<or> \<not> Q \<or> P \<noteq> Q" | |
| 379 | "\<not> P \<or> Q \<or> P \<noteq> Q" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 380 | "P \<or> y = (if P then x else y)" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 381 | "P \<or> (if P then x else y) = y" | 
| 57169 | 382 | "\<not> P \<or> x = (if P then x else y)" | 
| 383 | "\<not> P \<or> (if P then x else y) = x" | |
| 384 | "P \<or> R \<or> \<not> (if P then Q else R)" | |
| 385 | "\<not> P \<or> Q \<or> \<not> (if P then Q else R)" | |
| 386 | "\<not> (if P then Q else R) \<or> \<not> P \<or> Q" | |
| 387 | "\<not> (if P then Q else R) \<or> P \<or> R" | |
| 388 | "(if P then Q else R) \<or> \<not> P \<or> \<not> Q" | |
| 389 | "(if P then Q else R) \<or> P \<or> \<not> R" | |
| 390 | "(if P then \<not> Q else R) \<or> \<not> P \<or> Q" | |
| 391 | "(if P then Q else \<not> R) \<or> P \<or> R" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 392 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 393 | |
| 57230 | 394 | hide_type (open) symb_list pattern | 
| 395 | hide_const (open) Symb_Nil Symb_Cons trigger pat nopat fun_app z3div z3mod | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 396 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 397 | end |