| author | wenzelm | 
| Fri, 06 Sep 2019 19:44:54 +0200 | |
| changeset 70665 | 94442fce40a5 | 
| parent 69605 | a96320074298 | 
| child 72343 | 478b7599a1a0 | 
| permissions | -rw-r--r-- | 
| 58061 | 1 | (* Title: HOL/SMT.thy | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 2 | Author: Sascha Boehme, TU Muenchen | 
| 66551 | 3 | Author: Jasmin Blanchette, VU Amsterdam | 
| 56078 
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 | |
| 61626 | 6 | section \<open>Bindings to Satisfiability Modulo Theories (SMT) solvers based on SMT-LIB 2\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 7 | |
| 58061 | 8 | theory SMT | 
| 66551 | 9 | imports Divides | 
| 10 | keywords "smt_status" :: diag | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 11 | begin | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 12 | |
| 60758 | 13 | subsection \<open>A skolemization tactic and proof method\<close> | 
| 58481 | 14 | |
| 15 | lemma choices: | |
| 16 | "\<And>Q. \<forall>x. \<exists>y ya. Q x y ya \<Longrightarrow> \<exists>f fa. \<forall>x. Q x (f x) (fa x)" | |
| 17 | "\<And>Q. \<forall>x. \<exists>y ya yb. Q x y ya yb \<Longrightarrow> \<exists>f fa fb. \<forall>x. Q x (f x) (fa x) (fb x)" | |
| 18 | "\<And>Q. \<forall>x. \<exists>y ya yb yc. Q x y ya yb yc \<Longrightarrow> \<exists>f fa fb fc. \<forall>x. Q x (f x) (fa x) (fb x) (fc x)" | |
| 58598 | 19 | "\<And>Q. \<forall>x. \<exists>y ya yb yc yd. Q x y ya yb yc yd \<Longrightarrow> | 
| 20 | \<exists>f fa fb fc fd. \<forall>x. Q x (f x) (fa x) (fb x) (fc x) (fd x)" | |
| 21 | "\<And>Q. \<forall>x. \<exists>y ya yb yc yd ye. Q x y ya yb yc yd ye \<Longrightarrow> | |
| 22 | \<exists>f fa fb fc fd fe. \<forall>x. Q x (f x) (fa x) (fb x) (fc x) (fd x) (fe x)" | |
| 23 | "\<And>Q. \<forall>x. \<exists>y ya yb yc yd ye yf. Q x y ya yb yc yd ye yf \<Longrightarrow> | |
| 24 | \<exists>f fa fb fc fd fe ff. \<forall>x. Q x (f x) (fa x) (fb x) (fc x) (fd x) (fe x) (ff x)" | |
| 25 | "\<And>Q. \<forall>x. \<exists>y ya yb yc yd ye yf yg. Q x y ya yb yc yd ye yf yg \<Longrightarrow> | |
| 26 | \<exists>f fa fb fc fd fe ff fg. \<forall>x. Q x (f x) (fa x) (fb x) (fc x) (fd x) (fe x) (ff x) (fg x)" | |
| 58481 | 27 | by metis+ | 
| 28 | ||
| 29 | lemma bchoices: | |
| 30 | "\<And>Q. \<forall>x \<in> S. \<exists>y ya. Q x y ya \<Longrightarrow> \<exists>f fa. \<forall>x \<in> S. Q x (f x) (fa x)" | |
| 31 | "\<And>Q. \<forall>x \<in> S. \<exists>y ya yb. Q x y ya yb \<Longrightarrow> \<exists>f fa fb. \<forall>x \<in> S. Q x (f x) (fa x) (fb x)" | |
| 32 | "\<And>Q. \<forall>x \<in> S. \<exists>y ya yb yc. Q x y ya yb yc \<Longrightarrow> \<exists>f fa fb fc. \<forall>x \<in> S. Q x (f x) (fa x) (fb x) (fc x)" | |
| 58598 | 33 | "\<And>Q. \<forall>x \<in> S. \<exists>y ya yb yc yd. Q x y ya yb yc yd \<Longrightarrow> | 
| 34 | \<exists>f fa fb fc fd. \<forall>x \<in> S. Q x (f x) (fa x) (fb x) (fc x) (fd x)" | |
| 35 | "\<And>Q. \<forall>x \<in> S. \<exists>y ya yb yc yd ye. Q x y ya yb yc yd ye \<Longrightarrow> | |
| 36 | \<exists>f fa fb fc fd fe. \<forall>x \<in> S. Q x (f x) (fa x) (fb x) (fc x) (fd x) (fe x)" | |
| 37 | "\<And>Q. \<forall>x \<in> S. \<exists>y ya yb yc yd ye yf. Q x y ya yb yc yd ye yf \<Longrightarrow> | |
| 38 | \<exists>f fa fb fc fd fe ff. \<forall>x \<in> S. Q x (f x) (fa x) (fb x) (fc x) (fd x) (fe x) (ff x)" | |
| 39 | "\<And>Q. \<forall>x \<in> S. \<exists>y ya yb yc yd ye yf yg. Q x y ya yb yc yd ye yf yg \<Longrightarrow> | |
| 40 | \<exists>f fa fb fc fd fe ff fg. \<forall>x \<in> S. Q x (f x) (fa x) (fb x) (fc x) (fd x) (fe x) (ff x) (fg x)" | |
| 58481 | 41 | by metis+ | 
| 42 | ||
| 60758 | 43 | ML \<open> | 
| 58481 | 44 | fun moura_tac ctxt = | 
| 45 | Atomize_Elim.atomize_elim_tac ctxt THEN' | |
| 46 |   SELECT_GOAL (Clasimp.auto_tac (ctxt addSIs @{thms choice choices bchoice bchoices}) THEN
 | |
| 58598 | 47 | ALLGOALS (Metis_Tactic.metis_tac (take 1 ATP_Proof_Reconstruct.partial_type_encs) | 
| 48 | ATP_Proof_Reconstruct.default_metis_lam_trans ctxt [] ORELSE' | |
| 49 | blast_tac ctxt)) | |
| 60758 | 50 | \<close> | 
| 58481 | 51 | |
| 60758 | 52 | method_setup moura = \<open> | 
| 60201 | 53 | Scan.succeed (SIMPLE_METHOD' o moura_tac) | 
| 60758 | 54 | \<close> "solve skolemization goals, especially those arising from Z3 proofs" | 
| 58481 | 55 | |
| 56 | hide_fact (open) choices bchoices | |
| 57 | ||
| 58 | ||
| 60758 | 59 | subsection \<open>Triggers for quantifier instantiation\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 60 | |
| 60758 | 61 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 62 | Some SMT solvers support patterns as a quantifier instantiation | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 63 | heuristics. Patterns may either be positive terms (tagged by "pat") | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 64 | triggering quantifier instantiations -- when the solver finds a | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 65 | term matching a positive pattern, it instantiates the corresponding | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 66 | quantifier accordingly -- or negative terms (tagged by "nopat") | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 67 | inhibiting quantifier instantiations. A list of patterns | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 68 | 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 | 69 | multipattern are considered conjunctively for quantifier instantiation. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 70 | A list of multipatterns is called a trigger, and their multipatterns | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 71 | act disjunctively during quantifier instantiation. Each multipattern | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 72 | 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 | 73 | quantifier block. | 
| 60758 | 74 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 75 | |
| 57230 | 76 | typedecl 'a symb_list | 
| 77 | ||
| 78 | consts | |
| 79 | Symb_Nil :: "'a symb_list" | |
| 80 | Symb_Cons :: "'a \<Rightarrow> 'a symb_list \<Rightarrow> 'a symb_list" | |
| 81 | ||
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 82 | typedecl pattern | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 83 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 84 | consts | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 85 | pat :: "'a \<Rightarrow> pattern" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 86 | nopat :: "'a \<Rightarrow> pattern" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 87 | |
| 57230 | 88 | definition trigger :: "pattern symb_list symb_list \<Rightarrow> bool \<Rightarrow> bool" where | 
| 89 | "trigger _ P = P" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 90 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 91 | |
| 60758 | 92 | subsection \<open>Higher-order encoding\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 93 | |
| 60758 | 94 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 95 | Application is made explicit for constants occurring with varying | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 96 | numbers of arguments. This is achieved by the introduction of the | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 97 | following constant. | 
| 60758 | 98 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 99 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 100 | 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 | 101 | |
| 60758 | 102 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 103 | Some solvers support a theory of arrays which can be used to encode | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 104 | higher-order functions. The following set of lemmas specifies the | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 105 | properties of such (extensional) arrays. | 
| 60758 | 106 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 107 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 108 | 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 | 109 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 110 | |
| 60758 | 111 | subsection \<open>Normalization\<close> | 
| 56103 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 112 | |
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 113 | 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 | 114 | by simp | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 115 | |
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 116 | lemmas Ex1_def_raw = Ex1_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 117 | lemmas Ball_def_raw = Ball_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 118 | lemmas Bex_def_raw = Bex_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 119 | lemmas abs_if_raw = abs_if[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 120 | lemmas min_def_raw = min_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 121 | lemmas max_def_raw = max_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 122 | |
| 66817 | 123 | lemma nat_zero_as_int: | 
| 124 | "0 = nat 0" | |
| 125 | by simp | |
| 126 | ||
| 127 | lemma nat_one_as_int: | |
| 128 | "1 = nat 1" | |
| 129 | by simp | |
| 130 | ||
| 66298 | 131 | lemma nat_numeral_as_int: "numeral = (\<lambda>i. nat (numeral i))" by simp | 
| 67399 | 132 | lemma nat_less_as_int: "(<) = (\<lambda>a b. int a < int b)" by simp | 
| 133 | lemma nat_leq_as_int: "(\<le>) = (\<lambda>a b. int a \<le> int b)" by simp | |
| 66298 | 134 | lemma Suc_as_int: "Suc = (\<lambda>a. nat (int a + 1))" by (rule ext) simp | 
| 67399 | 135 | lemma nat_plus_as_int: "(+) = (\<lambda>a b. nat (int a + int b))" by (rule ext)+ simp | 
| 136 | lemma nat_minus_as_int: "(-) = (\<lambda>a b. nat (int a - int b))" by (rule ext)+ simp | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
67972diff
changeset | 137 | lemma nat_times_as_int: "(*) = (\<lambda>a b. nat (int a * int b))" by (simp add: nat_mult_distrib) | 
| 67399 | 138 | lemma nat_div_as_int: "(div) = (\<lambda>a b. nat (int a div int b))" by (simp add: nat_div_distrib) | 
| 139 | lemma nat_mod_as_int: "(mod) = (\<lambda>a b. nat (int a mod int b))" by (simp add: nat_mod_distrib) | |
| 66298 | 140 | |
| 141 | lemma int_Suc: "int (Suc n) = int n + 1" by simp | |
| 142 | lemma int_plus: "int (n + m) = int n + int m" by (rule of_nat_add) | |
| 143 | lemma int_minus: "int (n - m) = int (nat (int n - int m))" by auto | |
| 144 | ||
| 67972 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 145 | lemma nat_int_comparison: | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 146 | fixes a b :: nat | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 147 | shows "(a = b) = (int a = int b)" | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 148 | and "(a < b) = (int a < int b)" | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 149 | and "(a \<le> b) = (int a \<le> int b)" | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 150 | by simp_all | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 151 | |
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 152 | lemma int_ops: | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 153 | fixes a b :: nat | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 154 | shows "int 0 = 0" | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 155 | and "int 1 = 1" | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 156 | and "int (numeral n) = numeral n" | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 157 | and "int (Suc a) = int a + 1" | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 158 | and "int (a + b) = int a + int b" | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 159 | and "int (a - b) = (if int a < int b then 0 else int a - int b)" | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 160 | and "int (a * b) = int a * int b" | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 161 | and "int (a div b) = int a div int b" | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 162 | and "int (a mod b) = int a mod int b" | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 163 | by (auto intro: zdiv_int zmod_int) | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 164 | |
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 165 | lemma int_if: | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 166 | fixes a b :: nat | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 167 | shows "int (if P then a else b) = (if P then int a else int b)" | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 168 | by simp | 
| 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 boehmes parents: 
67399diff
changeset | 169 | |
| 56103 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 170 | |
| 60758 | 171 | subsection \<open>Integer division and modulo for Z3\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 172 | |
| 60758 | 173 | text \<open> | 
| 61799 | 174 | The following Z3-inspired definitions are overspecified for the case where \<open>l = 0\<close>. This | 
| 175 | Schönheitsfehler is corrected in the \<open>div_as_z3div\<close> and \<open>mod_as_z3mod\<close> theorems. | |
| 60758 | 176 | \<close> | 
| 56102 | 177 | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 178 | definition z3div :: "int \<Rightarrow> int \<Rightarrow> int" where | 
| 56102 | 179 | "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 | 180 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 181 | definition z3mod :: "int \<Rightarrow> int \<Rightarrow> int" where | 
| 56102 | 182 | "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 | 183 | |
| 56101 | 184 | lemma div_as_z3div: | 
| 56102 | 185 | "\<forall>k l. k div l = (if l = 0 then 0 else if l > 0 then z3div k l else z3div (- k) (- l))" | 
| 56101 | 186 | by (simp add: z3div_def) | 
| 187 | ||
| 188 | lemma mod_as_z3mod: | |
| 56102 | 189 | "\<forall>k l. k mod l = (if l = 0 then k else if l > 0 then z3mod k l else - z3mod (- k) (- l))" | 
| 56101 | 190 | by (simp add: z3mod_def) | 
| 191 | ||
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 192 | |
| 69205 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 193 | subsection \<open>Extra theorems for veriT reconstruction\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 194 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 195 | lemma verit_sko_forall: \<open>(\<forall>x. P x) \<longleftrightarrow> P (SOME x. \<not>P x)\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 196 | using someI[of \<open>\<lambda>x. \<not>P x\<close>] | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 197 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 198 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 199 | lemma verit_sko_forall': \<open>P (SOME x. \<not>P x) = A \<Longrightarrow> (\<forall>x. P x) = A\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 200 | by (subst verit_sko_forall) | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 201 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 202 | lemma verit_sko_forall_indirect: \<open>x = (SOME x. \<not>P x) \<Longrightarrow> (\<forall>x. P x) \<longleftrightarrow> P x\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 203 | using someI[of \<open>\<lambda>x. \<not>P x\<close>] | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 204 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 205 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 206 | lemma verit_sko_ex: \<open>(\<exists>x. P x) \<longleftrightarrow> P (SOME x. P x)\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 207 | using someI[of \<open>\<lambda>x. P x\<close>] | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 208 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 209 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 210 | lemma verit_sko_ex': \<open>P (SOME x. P x) = A \<Longrightarrow> (\<exists>x. P x) = A\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 211 | by (subst verit_sko_ex) | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 212 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 213 | lemma verit_sko_ex_indirect: \<open>x = (SOME x. P x) \<Longrightarrow> (\<exists>x. P x) \<longleftrightarrow> P x\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 214 | using someI[of \<open>\<lambda>x. P x\<close>] | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 215 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 216 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 217 | lemma verit_Pure_trans: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 218 | \<open>P \<equiv> Q \<Longrightarrow> Q \<Longrightarrow> P\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 219 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 220 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 221 | lemma verit_if_cong: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 222 | assumes \<open>b \<equiv> c\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 223 | and \<open>c \<Longrightarrow> x \<equiv> u\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 224 | and \<open>\<not> c \<Longrightarrow> y \<equiv> v\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 225 | shows \<open>(if b then x else y) \<equiv> (if c then u else v)\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 226 | using assms if_cong[of b c x u] by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 227 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 228 | lemma verit_if_weak_cong': | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 229 | \<open>b \<equiv> c \<Longrightarrow> (if b then x else y) \<equiv> (if c then x else y)\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 230 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 231 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 232 | lemma verit_ite_intro_simp: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 233 | \<open>(if c then (a :: 'a) = (if c then P else Q') else Q) = (if c then a = P else Q)\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 234 | \<open>(if c then R else b = (if c then R' else Q')) = | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 235 | (if c then R else b = Q')\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 236 | \<open>(if c then a' = a' else b' = b')\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 237 | by (auto split: if_splits) | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 238 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 239 | lemma verit_or_neg: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 240 | \<open>(A \<Longrightarrow> B) \<Longrightarrow> B \<or> \<not>A\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 241 | \<open>(\<not>A \<Longrightarrow> B) \<Longrightarrow> B \<or> A\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 242 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 243 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 244 | lemma verit_subst_bool: \<open>P \<Longrightarrow> f True \<Longrightarrow> f P\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 245 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 246 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 247 | lemma verit_and_pos: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 248 | \<open>(a \<Longrightarrow> \<not>b \<or> A) \<Longrightarrow> \<not>(a \<and> b) \<or> A\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 249 | \<open>(a \<Longrightarrow> A) \<Longrightarrow> \<not>a \<or> A\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 250 | \<open>(\<not>a \<Longrightarrow> A) \<Longrightarrow> a \<or> A\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 251 | by blast+ | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 252 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 253 | lemma verit_la_generic: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 254 | \<open>(a::int) \<le> x \<or> a = x \<or> a \<ge> x\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 255 | by linarith | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 256 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 257 | lemma verit_tmp_bfun_elim: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 258 | \<open>(if b then P True else P False) = P b\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 259 | by (cases b) auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 260 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 261 | lemma verit_eq_true_simplify: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 262 | \<open>(P = True) \<equiv> P\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 263 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 264 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 265 | lemma verit_and_neg: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 266 | \<open>B \<or> B' \<Longrightarrow> (A \<and> B) \<or> \<not>A \<or> B'\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 267 | \<open>B \<or> B' \<Longrightarrow> (\<not>A \<and> B) \<or> A \<or> B'\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 268 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 269 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 270 | lemma verit_forall_inst: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 271 | \<open>A \<longleftrightarrow> B \<Longrightarrow> \<not>A \<or> B\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 272 | \<open>\<not>A \<longleftrightarrow> B \<Longrightarrow> A \<or> B\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 273 | \<open>A \<longleftrightarrow> B \<Longrightarrow> \<not>B \<or> A\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 274 | \<open>A \<longleftrightarrow> \<not>B \<Longrightarrow> B \<or> A\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 275 | \<open>A \<longrightarrow> B \<Longrightarrow> \<not>A \<or> B\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 276 | \<open>\<not>A \<longrightarrow> B \<Longrightarrow> A \<or> B\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 277 | by blast+ | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 278 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 279 | lemma verit_eq_transitive: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 280 | \<open>A = B \<Longrightarrow> B = C \<Longrightarrow> A = C\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 281 | \<open>A = B \<Longrightarrow> C = B \<Longrightarrow> A = C\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 282 | \<open>B = A \<Longrightarrow> B = C \<Longrightarrow> A = C\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 283 | \<open>B = A \<Longrightarrow> C = B \<Longrightarrow> A = C\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 284 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 285 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 286 | |
| 60758 | 287 | subsection \<open>Setup\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 288 | |
| 69605 | 289 | ML_file \<open>Tools/SMT/smt_util.ML\<close> | 
| 290 | ML_file \<open>Tools/SMT/smt_failure.ML\<close> | |
| 291 | ML_file \<open>Tools/SMT/smt_config.ML\<close> | |
| 292 | ML_file \<open>Tools/SMT/smt_builtin.ML\<close> | |
| 293 | ML_file \<open>Tools/SMT/smt_datatypes.ML\<close> | |
| 294 | ML_file \<open>Tools/SMT/smt_normalize.ML\<close> | |
| 295 | ML_file \<open>Tools/SMT/smt_translate.ML\<close> | |
| 296 | ML_file \<open>Tools/SMT/smtlib.ML\<close> | |
| 297 | ML_file \<open>Tools/SMT/smtlib_interface.ML\<close> | |
| 298 | ML_file \<open>Tools/SMT/smtlib_proof.ML\<close> | |
| 299 | ML_file \<open>Tools/SMT/smtlib_isar.ML\<close> | |
| 300 | ML_file \<open>Tools/SMT/z3_proof.ML\<close> | |
| 301 | ML_file \<open>Tools/SMT/z3_isar.ML\<close> | |
| 302 | ML_file \<open>Tools/SMT/smt_solver.ML\<close> | |
| 303 | ML_file \<open>Tools/SMT/cvc4_interface.ML\<close> | |
| 304 | ML_file \<open>Tools/SMT/cvc4_proof_parse.ML\<close> | |
| 305 | ML_file \<open>Tools/SMT/verit_proof.ML\<close> | |
| 306 | ML_file \<open>Tools/SMT/verit_isar.ML\<close> | |
| 307 | ML_file \<open>Tools/SMT/verit_proof_parse.ML\<close> | |
| 308 | ML_file \<open>Tools/SMT/conj_disj_perm.ML\<close> | |
| 309 | ML_file \<open>Tools/SMT/smt_replay_methods.ML\<close> | |
| 310 | ML_file \<open>Tools/SMT/smt_replay.ML\<close> | |
| 311 | ML_file \<open>Tools/SMT/z3_interface.ML\<close> | |
| 312 | ML_file \<open>Tools/SMT/z3_replay_rules.ML\<close> | |
| 313 | ML_file \<open>Tools/SMT/z3_replay_methods.ML\<close> | |
| 314 | ML_file \<open>Tools/SMT/z3_replay.ML\<close> | |
| 315 | ML_file \<open>Tools/SMT/verit_replay_methods.ML\<close> | |
| 316 | ML_file \<open>Tools/SMT/verit_replay.ML\<close> | |
| 317 | ML_file \<open>Tools/SMT/smt_systems.ML\<close> | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 318 | |
| 60758 | 319 | method_setup smt = \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 320 | Scan.optional Attrib.thms [] >> | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 321 | (fn thms => fn ctxt => | 
| 58061 | 322 | METHOD (fn facts => HEADGOAL (SMT_Solver.smt_tac ctxt (thms @ facts)))) | 
| 60758 | 323 | \<close> "apply an SMT solver to the current goal" | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 324 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 325 | |
| 60758 | 326 | subsection \<open>Configuration\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 327 | |
| 60758 | 328 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 329 | The current configuration can be printed by the command | 
| 61799 | 330 | \<open>smt_status\<close>, which shows the values of most options. | 
| 60758 | 331 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 332 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 333 | |
| 60758 | 334 | subsection \<open>General configuration options\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 335 | |
| 60758 | 336 | text \<open> | 
| 61799 | 337 | The option \<open>smt_solver\<close> can be used to change the target SMT | 
| 338 | solver. The possible values can be obtained from the \<open>smt_status\<close> | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 339 | command. | 
| 60758 | 340 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 341 | |
| 58061 | 342 | declare [[smt_solver = z3]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 343 | |
| 60758 | 344 | text \<open> | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 345 | Since SMT solvers are potentially nonterminating, there is a timeout | 
| 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 346 | (given in seconds) to restrict their runtime. | 
| 60758 | 347 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 348 | |
| 58061 | 349 | declare [[smt_timeout = 20]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 350 | |
| 60758 | 351 | text \<open> | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 352 | SMT solvers apply randomized heuristics. In case a problem is not | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 353 | solvable by an SMT solver, changing the following option might help. | 
| 60758 | 354 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 355 | |
| 58061 | 356 | declare [[smt_random_seed = 1]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 357 | |
| 60758 | 358 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 359 | In general, the binding to SMT solvers runs as an oracle, i.e, the SMT | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 360 | solvers are fully trusted without additional checks. The following | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 361 | option can cause the SMT solver to run in proof-producing mode, giving | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 362 | a checkable certificate. This is currently only implemented for Z3. | 
| 60758 | 363 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 364 | |
| 58061 | 365 | declare [[smt_oracle = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 366 | |
| 60758 | 367 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 368 | Each SMT solver provides several commandline options to tweak its | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 369 | behaviour. They can be passed to the solver by setting the following | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 370 | options. | 
| 60758 | 371 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 372 | |
| 58061 | 373 | declare [[cvc3_options = ""]] | 
| 66323 
c41642bc1ebb
pass option recommended by Andy Reynolds to CVC4 1.5 (released) or better
 blanchet parents: 
66298diff
changeset | 374 | declare [[cvc4_options = "--full-saturate-quant --inst-when=full-last-call --inst-no-entail --term-db-mode=relevant --multi-trigger-linear"]] | 
| 69205 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 375 | declare [[verit_options = "--index-fresh-sorts"]] | 
| 58061 | 376 | declare [[z3_options = ""]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 377 | |
| 60758 | 378 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 379 | 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 | 380 | 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 | 381 | solvable by SMT solvers (note: triggers guide quantifier instantiations | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 382 | in the SMT solver). To turn it on, set the following option. | 
| 60758 | 383 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 384 | |
| 58061 | 385 | declare [[smt_infer_triggers = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 386 | |
| 60758 | 387 | text \<open> | 
| 58360 | 388 | Enable the following option to use built-in support for datatypes, | 
| 389 | codatatypes, and records in CVC4. Currently, this is implemented only | |
| 390 | in oracle mode. | |
| 60758 | 391 | \<close> | 
| 58360 | 392 | |
| 393 | declare [[cvc4_extensions = false]] | |
| 394 | ||
| 60758 | 395 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 396 | Enable the following option to use built-in support for div/mod, datatypes, | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 397 | and records in Z3. Currently, this is implemented only in oracle mode. | 
| 60758 | 398 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 399 | |
| 58061 | 400 | declare [[z3_extensions = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 401 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 402 | |
| 60758 | 403 | subsection \<open>Certificates\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 404 | |
| 60758 | 405 | text \<open> | 
| 61799 | 406 | By setting the option \<open>smt_certificates\<close> to the name of a file, | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 407 | 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 | 408 | 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 | 409 | configuration) re-uses the cached certificate instead of invoking the | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 410 | solver. An empty string disables caching certificates. | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 411 | |
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 412 | The filename should be given as an explicit path. It is good | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 413 | practice to use the name of the current theory (with ending | 
| 61799 | 414 | \<open>.certs\<close> instead of \<open>.thy\<close>) as the certificates file. | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 415 | 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 | 416 | to avoid race conditions with other concurrent accesses. | 
| 60758 | 417 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 418 | |
| 58061 | 419 | declare [[smt_certificates = ""]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 420 | |
| 60758 | 421 | text \<open> | 
| 61799 | 422 | The option \<open>smt_read_only_certificates\<close> controls whether only | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 423 | stored certificates are should be used or invocation of an SMT solver | 
| 61799 | 424 | is allowed. When set to \<open>true\<close>, no SMT solver will ever be | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 425 | invoked and only the existing certificates found in the configured | 
| 61799 | 426 | cache are used; when set to \<open>false\<close> and there is no cached | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 427 | 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 | 428 | invoked. | 
| 60758 | 429 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 430 | |
| 58061 | 431 | declare [[smt_read_only_certificates = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 432 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 433 | |
| 60758 | 434 | subsection \<open>Tracing\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 435 | |
| 60758 | 436 | text \<open> | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 437 | The SMT method, when applied, traces important information. To | 
| 61799 | 438 | make it entirely silent, set the following option to \<open>false\<close>. | 
| 60758 | 439 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 440 | |
| 58061 | 441 | declare [[smt_verbose = true]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 442 | |
| 60758 | 443 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 444 | 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 | 445 | well as the returned result of the solver, the option | 
| 61799 | 446 | \<open>smt_trace\<close> should be set to \<open>true\<close>. | 
| 60758 | 447 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 448 | |
| 58061 | 449 | declare [[smt_trace = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 450 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 451 | |
| 60758 | 452 | subsection \<open>Schematic rules for Z3 proof reconstruction\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 453 | |
| 60758 | 454 | text \<open> | 
| 57696 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 blanchet parents: 
57246diff
changeset | 455 | Several prof rules of Z3 are not very well documented. There are two | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 456 | lemma groups which can turn failing Z3 proof reconstruction attempts | 
| 61799 | 457 | into succeeding ones: the facts in \<open>z3_rule\<close> are tried prior to | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 458 | any implemented reconstruction procedure for all uncertain Z3 proof | 
| 61799 | 459 | rules; the facts in \<open>z3_simp\<close> are only fed to invocations of | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 460 | the simplifier when reconstructing theory-specific proof steps. | 
| 60758 | 461 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 462 | |
| 58061 | 463 | lemmas [z3_rule] = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 464 | 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 | 465 | 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 | 466 | if_True if_False not_not | 
| 58776 
95e58e04e534
use NO_MATCH-simproc for distribution rules in field_simps, otherwise field_simps on '(a / (c + d)) * (e + f)' can be non-terminating
 hoelzl parents: 
58598diff
changeset | 467 | NO_MATCH_def | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 468 | |
| 58061 | 469 | lemma [z3_rule]: | 
| 57169 | 470 | "(P \<and> Q) = (\<not> (\<not> P \<or> \<not> Q))" | 
| 471 | "(P \<and> Q) = (\<not> (\<not> Q \<or> \<not> P))" | |
| 472 | "(\<not> P \<and> Q) = (\<not> (P \<or> \<not> Q))" | |
| 473 | "(\<not> P \<and> Q) = (\<not> (\<not> Q \<or> P))" | |
| 474 | "(P \<and> \<not> Q) = (\<not> (\<not> P \<or> Q))" | |
| 475 | "(P \<and> \<not> Q) = (\<not> (Q \<or> \<not> P))" | |
| 476 | "(\<not> P \<and> \<not> Q) = (\<not> (P \<or> Q))" | |
| 477 | "(\<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 | 478 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 479 | |
| 58061 | 480 | lemma [z3_rule]: | 
| 57169 | 481 | "(P \<longrightarrow> Q) = (Q \<or> \<not> P)" | 
| 482 | "(\<not> P \<longrightarrow> Q) = (P \<or> Q)" | |
| 483 | "(\<not> P \<longrightarrow> Q) = (Q \<or> P)" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 484 | "(True \<longrightarrow> P) = P" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 485 | "(P \<longrightarrow> True) = True" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 486 | "(False \<longrightarrow> P) = True" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 487 | "(P \<longrightarrow> P) = True" | 
| 59037 
650dcf624729
added Z3 reconstruction rule suggested by F. Maric
 blanchet parents: 
59036diff
changeset | 488 | "(\<not> (A \<longleftrightarrow> \<not> B)) \<longleftrightarrow> (A \<longleftrightarrow> B)" | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 489 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 490 | |
| 58061 | 491 | lemma [z3_rule]: | 
| 67091 | 492 | "((P = Q) \<longrightarrow> R) = (R \<or> (Q = (\<not> P)))" | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 493 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 494 | |
| 58061 | 495 | lemma [z3_rule]: | 
| 57169 | 496 | "(\<not> True) = False" | 
| 497 | "(\<not> False) = True" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 498 | "(x = x) = True" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 499 | "(P = True) = P" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 500 | "(True = P) = P" | 
| 57169 | 501 | "(P = False) = (\<not> P)" | 
| 502 | "(False = P) = (\<not> P)" | |
| 503 | "((\<not> P) = P) = False" | |
| 504 | "(P = (\<not> P)) = False" | |
| 505 | "((\<not> P) = (\<not> Q)) = (P = Q)" | |
| 506 | "\<not> (P = (\<not> Q)) = (P = Q)" | |
| 507 | "\<not> ((\<not> P) = Q) = (P = Q)" | |
| 508 | "(P \<noteq> Q) = (Q = (\<not> P))" | |
| 509 | "(P = Q) = ((\<not> P \<or> Q) \<and> (P \<or> \<not> Q))" | |
| 510 | "(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 | 511 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 512 | |
| 58061 | 513 | lemma [z3_rule]: | 
| 57169 | 514 | "(if P then P else \<not> P) = True" | 
| 515 | "(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 | 516 | "(if P then True else False) = P" | 
| 57169 | 517 | "(if P then False else True) = (\<not> P)" | 
| 518 | "(if P then Q else True) = ((\<not> P) \<or> Q)" | |
| 519 | "(if P then Q else True) = (Q \<or> (\<not> P))" | |
| 520 | "(if P then Q else \<not> Q) = (P = Q)" | |
| 521 | "(if P then Q else \<not> Q) = (Q = P)" | |
| 522 | "(if P then \<not> Q else Q) = (P = (\<not> Q))" | |
| 523 | "(if P then \<not> Q else Q) = ((\<not> Q) = P)" | |
| 524 | "(if \<not> P then x else y) = (if P then y else x)" | |
| 525 | "(if P then (if Q then x else y) else x) = (if P \<and> (\<not> Q) then y else x)" | |
| 526 | "(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 | 527 | "(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 | 528 | "(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 | 529 | "(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 | 530 | "(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 | 531 | "(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 | 532 | "(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 | 533 | "(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 | 534 | "(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 | 535 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 536 | |
| 58061 | 537 | lemma [z3_rule]: | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 538 | "0 + (x::int) = x" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 539 | "x + 0 = x" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 540 | "x + x = 2 * x" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 541 | "0 * x = 0" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 542 | "1 * x = x" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 543 | "x + y = y + x" | 
| 57230 | 544 | by (auto simp add: mult_2) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 545 | |
| 58061 | 546 | lemma [z3_rule]: (* for def-axiom *) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 547 | "P = Q \<or> P \<or> Q" | 
| 57169 | 548 | "P = Q \<or> \<not> P \<or> \<not> Q" | 
| 549 | "(\<not> P) = Q \<or> \<not> P \<or> Q" | |
| 550 | "(\<not> P) = Q \<or> P \<or> \<not> Q" | |
| 551 | "P = (\<not> Q) \<or> \<not> P \<or> Q" | |
| 552 | "P = (\<not> Q) \<or> P \<or> \<not> Q" | |
| 553 | "P \<noteq> Q \<or> P \<or> \<not> Q" | |
| 554 | "P \<noteq> Q \<or> \<not> P \<or> Q" | |
| 555 | "P \<noteq> (\<not> Q) \<or> P \<or> Q" | |
| 556 | "(\<not> P) \<noteq> Q \<or> P \<or> Q" | |
| 557 | "P \<or> Q \<or> P \<noteq> (\<not> Q)" | |
| 558 | "P \<or> Q \<or> (\<not> P) \<noteq> Q" | |
| 559 | "P \<or> \<not> Q \<or> P \<noteq> Q" | |
| 560 | "\<not> P \<or> Q \<or> P \<noteq> Q" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 561 | "P \<or> y = (if P then x else y)" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 562 | "P \<or> (if P then x else y) = y" | 
| 57169 | 563 | "\<not> P \<or> x = (if P then x else y)" | 
| 564 | "\<not> P \<or> (if P then x else y) = x" | |
| 565 | "P \<or> R \<or> \<not> (if P then Q else R)" | |
| 566 | "\<not> P \<or> Q \<or> \<not> (if P then Q else R)" | |
| 567 | "\<not> (if P then Q else R) \<or> \<not> P \<or> Q" | |
| 568 | "\<not> (if P then Q else R) \<or> P \<or> R" | |
| 569 | "(if P then Q else R) \<or> \<not> P \<or> \<not> Q" | |
| 570 | "(if P then Q else R) \<or> P \<or> \<not> R" | |
| 571 | "(if P then \<not> Q else R) \<or> \<not> P \<or> Q" | |
| 572 | "(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 | 573 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 574 | |
| 57230 | 575 | hide_type (open) symb_list pattern | 
| 576 | 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 | 577 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 578 | end |