| author | wenzelm | 
| Fri, 05 Jul 2024 12:53:45 +0200 | |
| changeset 80509 | 2a9abd6a164e | 
| parent 79712 | 658f17274845 | 
| child 82074 | 22d521925afc | 
| 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 | 
| 78936 | 9 | imports Numeral_Simprocs | 
| 66551 | 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 | |
| 79712 
658f17274845
new less ad hoc implementation of the 'moura' tactic for skolemization
 blanchet parents: 
79623diff
changeset | 15 | lemma ex_iff_push: "(\<exists>y. P \<longleftrightarrow> Q y) \<longleftrightarrow> (P \<longrightarrow> (\<exists>y. Q y)) \<and> ((\<forall>y. Q y) \<longrightarrow> P)" | 
| 
658f17274845
new less ad hoc implementation of the 'moura' tactic for skolemization
 blanchet parents: 
79623diff
changeset | 16 | by metis | 
| 58481 | 17 | |
| 60758 | 18 | ML \<open> | 
| 58481 | 19 | fun moura_tac ctxt = | 
| 79712 
658f17274845
new less ad hoc implementation of the 'moura' tactic for skolemization
 blanchet parents: 
79623diff
changeset | 20 | TRY o Atomize_Elim.atomize_elim_tac ctxt THEN' | 
| 
658f17274845
new less ad hoc implementation of the 'moura' tactic for skolemization
 blanchet parents: 
79623diff
changeset | 21 | REPEAT o EqSubst.eqsubst_tac ctxt [0] | 
| 
658f17274845
new less ad hoc implementation of the 'moura' tactic for skolemization
 blanchet parents: 
79623diff
changeset | 22 |     @{thms choice_iff[symmetric] bchoice_iff[symmetric]} THEN'
 | 
| 
658f17274845
new less ad hoc implementation of the 'moura' tactic for skolemization
 blanchet parents: 
79623diff
changeset | 23 | TRY o Simplifier.asm_full_simp_tac | 
| 
658f17274845
new less ad hoc implementation of the 'moura' tactic for skolemization
 blanchet parents: 
79623diff
changeset | 24 |     (clear_simpset ctxt addsimps @{thms all_simps ex_simps ex_iff_push}) THEN_ALL_NEW
 | 
| 
658f17274845
new less ad hoc implementation of the 'moura' tactic for skolemization
 blanchet parents: 
79623diff
changeset | 25 | Metis_Tactic.metis_tac (take 1 ATP_Proof_Reconstruct.partial_type_encs) | 
| 
658f17274845
new less ad hoc implementation of the 'moura' tactic for skolemization
 blanchet parents: 
79623diff
changeset | 26 | ATP_Proof_Reconstruct.default_metis_lam_trans ctxt [] | 
| 60758 | 27 | \<close> | 
| 58481 | 28 | |
| 60758 | 29 | method_setup moura = \<open> | 
| 60201 | 30 | Scan.succeed (SIMPLE_METHOD' o moura_tac) | 
| 60758 | 31 | \<close> "solve skolemization goals, especially those arising from Z3 proofs" | 
| 58481 | 32 | |
| 79712 
658f17274845
new less ad hoc implementation of the 'moura' tactic for skolemization
 blanchet parents: 
79623diff
changeset | 33 | hide_fact (open) ex_iff_push | 
| 58481 | 34 | |
| 35 | ||
| 60758 | 36 | subsection \<open>Triggers for quantifier instantiation\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 37 | |
| 60758 | 38 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 39 | 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 | 40 | 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 | 41 | triggering quantifier instantiations -- when the solver finds a | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 42 | term matching a positive pattern, it instantiates the corresponding | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 43 | 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 | 44 | inhibiting quantifier instantiations. A list of patterns | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 45 | 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 | 46 | multipattern are considered conjunctively for quantifier instantiation. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 47 | 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 | 48 | act disjunctively during quantifier instantiation. Each multipattern | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 49 | 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 | 50 | quantifier block. | 
| 60758 | 51 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 52 | |
| 57230 | 53 | typedecl 'a symb_list | 
| 54 | ||
| 55 | consts | |
| 56 | Symb_Nil :: "'a symb_list" | |
| 57 | Symb_Cons :: "'a \<Rightarrow> 'a symb_list \<Rightarrow> 'a symb_list" | |
| 58 | ||
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 59 | typedecl pattern | 
| 
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 | consts | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 62 | pat :: "'a \<Rightarrow> pattern" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 63 | nopat :: "'a \<Rightarrow> pattern" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 64 | |
| 57230 | 65 | definition trigger :: "pattern symb_list symb_list \<Rightarrow> bool \<Rightarrow> bool" where | 
| 66 | "trigger _ P = P" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 67 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 68 | |
| 60758 | 69 | subsection \<open>Higher-order encoding\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 70 | |
| 60758 | 71 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 72 | 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 | 73 | 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 | 74 | following constant. | 
| 60758 | 75 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 76 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 77 | 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 | 78 | |
| 60758 | 79 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 80 | 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 | 81 | 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 | 82 | properties of such (extensional) arrays. | 
| 60758 | 83 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 84 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 85 | 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 | 86 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 87 | |
| 60758 | 88 | subsection \<open>Normalization\<close> | 
| 56103 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 89 | |
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 90 | 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 | 91 | by simp | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 92 | |
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 93 | lemmas Ex1_def_raw = Ex1_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 94 | lemmas Ball_def_raw = Ball_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 95 | lemmas Bex_def_raw = Bex_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 96 | lemmas abs_if_raw = abs_if[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 97 | lemmas min_def_raw = min_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 98 | lemmas max_def_raw = max_def[abs_def] | 
| 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 99 | |
| 66817 | 100 | lemma nat_zero_as_int: | 
| 101 | "0 = nat 0" | |
| 102 | by simp | |
| 103 | ||
| 104 | lemma nat_one_as_int: | |
| 105 | "1 = nat 1" | |
| 106 | by simp | |
| 107 | ||
| 66298 | 108 | lemma nat_numeral_as_int: "numeral = (\<lambda>i. nat (numeral i))" by simp | 
| 67399 | 109 | lemma nat_less_as_int: "(<) = (\<lambda>a b. int a < int b)" by simp | 
| 110 | lemma nat_leq_as_int: "(\<le>) = (\<lambda>a b. int a \<le> int b)" by simp | |
| 66298 | 111 | lemma Suc_as_int: "Suc = (\<lambda>a. nat (int a + 1))" by (rule ext) simp | 
| 67399 | 112 | lemma nat_plus_as_int: "(+) = (\<lambda>a b. nat (int a + int b))" by (rule ext)+ simp | 
| 113 | 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 | 114 | lemma nat_times_as_int: "(*) = (\<lambda>a b. nat (int a * int b))" by (simp add: nat_mult_distrib) | 
| 67399 | 115 | lemma nat_div_as_int: "(div) = (\<lambda>a b. nat (int a div int b))" by (simp add: nat_div_distrib) | 
| 116 | lemma nat_mod_as_int: "(mod) = (\<lambda>a b. nat (int a mod int b))" by (simp add: nat_mod_distrib) | |
| 66298 | 117 | |
| 118 | lemma int_Suc: "int (Suc n) = int n + 1" by simp | |
| 119 | lemma int_plus: "int (n + m) = int n + int m" by (rule of_nat_add) | |
| 120 | lemma int_minus: "int (n - m) = int (nat (int n - int m))" by auto | |
| 121 | ||
| 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 | 122 | 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 | 123 | 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 | 124 | 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 | 125 | 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 | 126 | 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 | 127 | 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 | 128 | |
| 
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 | 129 | 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 | 130 | 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 | 131 | 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 | 132 | 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 | 133 | 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 | 134 | 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 | 135 | 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 | 136 | 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 | 137 | 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 | 138 | 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 | 139 | 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 | 140 | 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 | 141 | |
| 
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 | 142 | 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 | 143 | 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 | 144 | 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 | 145 | 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 | 146 | |
| 56103 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 blanchet parents: 
56102diff
changeset | 147 | |
| 60758 | 148 | 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 | 149 | |
| 60758 | 150 | text \<open> | 
| 61799 | 151 | The following Z3-inspired definitions are overspecified for the case where \<open>l = 0\<close>. This | 
| 152 | Schönheitsfehler is corrected in the \<open>div_as_z3div\<close> and \<open>mod_as_z3mod\<close> theorems. | |
| 60758 | 153 | \<close> | 
| 56102 | 154 | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 155 | definition z3div :: "int \<Rightarrow> int \<Rightarrow> int" where | 
| 56102 | 156 | "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 | 157 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 158 | definition z3mod :: "int \<Rightarrow> int \<Rightarrow> int" where | 
| 56102 | 159 | "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 | 160 | |
| 56101 | 161 | lemma div_as_z3div: | 
| 56102 | 162 | "\<forall>k l. k div l = (if l = 0 then 0 else if l > 0 then z3div k l else z3div (- k) (- l))" | 
| 56101 | 163 | by (simp add: z3div_def) | 
| 164 | ||
| 165 | lemma mod_as_z3mod: | |
| 56102 | 166 | "\<forall>k l. k mod l = (if l = 0 then k else if l > 0 then z3mod k l else - z3mod (- k) (- l))" | 
| 56101 | 167 | by (simp add: z3mod_def) | 
| 168 | ||
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 169 | |
| 69205 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 170 | 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 | 171 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 172 | 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 | 173 | 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 | 174 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 175 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 176 | 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 | 177 | by (subst verit_sko_forall) | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 178 | |
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 179 | lemma verit_sko_forall'': \<open>B = A \<Longrightarrow> (SOME x. P x) = A \<equiv> (SOME x. P x) = B\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 180 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 181 | |
| 69205 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 182 | 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 | 183 | 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 | 184 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 185 | |
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 186 | lemma verit_sko_forall_indirect2: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 187 | \<open>x = (SOME x. \<not>P x) \<Longrightarrow> (\<And>x :: 'a. (P x = P' x)) \<Longrightarrow>(\<forall>x. P' x) \<longleftrightarrow> P x\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 188 | using someI[of \<open>\<lambda>x. \<not>P x\<close>] | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 189 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 190 | |
| 69205 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 191 | 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 | 192 | 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 | 193 | by auto | 
| 
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_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 | 196 | by (subst verit_sko_ex) | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 197 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 198 | 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 | 199 | 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 | 200 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 201 | |
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 202 | lemma verit_sko_ex_indirect2: \<open>x = (SOME x. P x) \<Longrightarrow> (\<And>x. P x = P' x) \<Longrightarrow> (\<exists>x. P' x) \<longleftrightarrow> P x\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 203 | using someI[of \<open>\<lambda>x. P x\<close>] | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 204 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 205 | |
| 69205 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 206 | lemma verit_Pure_trans: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 207 | \<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 | 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_if_cong: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 211 | assumes \<open>b \<equiv> c\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 212 | 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 | 213 | 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 | 214 | 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 | 215 | 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 | 216 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 217 | lemma verit_if_weak_cong': | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 218 | \<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 | 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_or_neg: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 222 | \<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 | 223 | \<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 | 224 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 225 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 226 | 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 | 227 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 228 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 229 | lemma verit_and_pos: | 
| 72513 
75f5c63f6cfa
better handling of skolemization for Isar reconstruction in Sledgehammer for veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72458diff
changeset | 230 | \<open>(a \<Longrightarrow> \<not>(b \<and> c) \<or> A) \<Longrightarrow> \<not>(a \<and> b \<and> c) \<or> A\<close> | 
| 
75f5c63f6cfa
better handling of skolemization for Isar reconstruction in Sledgehammer for veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72458diff
changeset | 231 | \<open>(a \<Longrightarrow> b \<Longrightarrow> A) \<Longrightarrow> \<not>(a \<and> b) \<or> A\<close> | 
| 79576 
157de27b0863
fix reconstruction of Alethe's and_pos rule
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
78936diff
changeset | 232 | by blast+ | 
| 
157de27b0863
fix reconstruction of Alethe's and_pos rule
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
78936diff
changeset | 233 | |
| 
157de27b0863
fix reconstruction of Alethe's and_pos rule
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
78936diff
changeset | 234 | lemma verit_farkas: | 
| 69205 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 235 | \<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 | 236 | \<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 | 237 | by blast+ | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 238 | |
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 239 | lemma verit_or_pos: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 240 | \<open>A \<and> A' \<Longrightarrow> (c \<and> A) \<or> (\<not>c \<and> A')\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 241 | \<open>A \<and> A' \<Longrightarrow> (\<not>c \<and> A) \<or> (c \<and> A')\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 242 | by blast+ | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 243 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 244 | |
| 69205 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 245 | lemma verit_la_generic: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 246 | \<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 | 247 | by linarith | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 248 | |
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 249 | lemma verit_bfun_elim: | 
| 69205 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 250 | \<open>(if b then P True else P False) = P b\<close> | 
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 251 | \<open>(\<forall>b. P' b) = (P' False \<and> P' True)\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 252 | \<open>(\<exists>b. P' b) = (P' False \<or> P' True)\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 253 | by (cases b) (auto simp: all_bool_eq ex_bool_eq) | 
| 69205 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 254 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 255 | lemma verit_eq_true_simplify: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 256 | \<open>(P = True) \<equiv> P\<close> | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 257 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 258 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 259 | lemma verit_and_neg: | 
| 72513 
75f5c63f6cfa
better handling of skolemization for Isar reconstruction in Sledgehammer for veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72458diff
changeset | 260 | \<open>(a \<Longrightarrow> \<not>b \<or> A) \<Longrightarrow> \<not>(a \<and> b) \<or> A\<close> | 
| 
75f5c63f6cfa
better handling of skolemization for Isar reconstruction in Sledgehammer for veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72458diff
changeset | 261 | \<open>(a \<Longrightarrow> A) \<Longrightarrow> \<not>a \<or> A\<close> | 
| 
75f5c63f6cfa
better handling of skolemization for Isar reconstruction in Sledgehammer for veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72458diff
changeset | 262 | \<open>(\<not>a \<Longrightarrow> A) \<Longrightarrow> a \<or> A\<close> | 
| 
75f5c63f6cfa
better handling of skolemization for Isar reconstruction in Sledgehammer for veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72458diff
changeset | 263 | by blast+ | 
| 69205 
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_forall_inst: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 266 | \<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 | 267 | \<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 | 268 | \<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 | 269 | \<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 | 270 | \<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 | 271 | \<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 | 272 | by blast+ | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 273 | |
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 274 | lemma verit_eq_transitive: | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 275 | \<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 | 276 | \<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 | 277 | \<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 | 278 | \<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 | 279 | by auto | 
| 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 280 | |
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 281 | lemma verit_bool_simplify: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 282 | \<open>\<not>(P \<longrightarrow> Q) \<longleftrightarrow> P \<and> \<not>Q\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 283 | \<open>\<not>(P \<or> Q) \<longleftrightarrow> \<not>P \<and> \<not>Q\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 284 | \<open>\<not>(P \<and> Q) \<longleftrightarrow> \<not>P \<or> \<not>Q\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 285 | \<open>(P \<longrightarrow> (Q \<longrightarrow> R)) \<longleftrightarrow> ((P \<and> Q) \<longrightarrow> R)\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 286 | \<open>((P \<longrightarrow> Q) \<longrightarrow> Q) \<longleftrightarrow> P \<or> Q\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 287 | \<open>(Q \<longleftrightarrow> (P \<or> Q)) \<longleftrightarrow> (P \<longrightarrow> Q)\<close> \<comment> \<open>This rule was inverted\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 288 | \<open>P \<and> (P \<longrightarrow> Q) \<longleftrightarrow> P \<and> Q\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 289 | \<open>(P \<longrightarrow> Q) \<and> P \<longleftrightarrow> P \<and> Q\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 290 | (* \<comment>\<open>Additional rules:\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 291 | * \<open>((P \<longrightarrow> Q) \<longrightarrow> P) \<longleftrightarrow> P\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 292 | * \<open>((P \<longrightarrow> Q) \<longrightarrow> Q) \<longleftrightarrow> P \<or> Q\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 293 | * \<open>(P \<longrightarrow> Q) \<or> P\<close> *) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 294 | unfolding not_imp imp_conjL | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 295 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 296 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 297 | text \<open>We need the last equation for \<^term>\<open>\<not>(\<forall>a b. \<not>P a b)\<close>\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 298 | lemma verit_connective_def: \<comment> \<open>the definition of XOR is missing | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 299 | as the operator is not generated by Isabelle\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 300 | \<open>(A = B) \<longleftrightarrow> ((A \<longrightarrow> B) \<and> (B \<longrightarrow> A))\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 301 | \<open>(If A B C) = ((A \<longrightarrow> B) \<and> (\<not>A \<longrightarrow> C))\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 302 | \<open>(\<exists>x. P x) \<longleftrightarrow> \<not>(\<forall>x. \<not>P x)\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 303 | \<open>\<not>(\<exists>x. P x) \<longleftrightarrow> (\<forall>x. \<not>P x)\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 304 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 305 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 306 | lemma verit_ite_simplify: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 307 | \<open>(If True B C) = B\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 308 | \<open>(If False B C) = C\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 309 | \<open>(If A' B B) = B\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 310 | \<open>(If (\<not>A') B C) = (If A' C B)\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 311 | \<open>(If c (If c A B) C) = (If c A C)\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 312 | \<open>(If c C (If c A B)) = (If c C B)\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 313 | \<open>(If A' True False) = A'\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 314 | \<open>(If A' False True) \<longleftrightarrow> \<not>A'\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 315 | \<open>(If A' True B') \<longleftrightarrow> A'\<or>B'\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 316 | \<open>(If A' B' False) \<longleftrightarrow> A'\<and>B'\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 317 | \<open>(If A' False B') \<longleftrightarrow> \<not>A'\<and>B'\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 318 | \<open>(If A' B' True) \<longleftrightarrow> \<not>A'\<or>B'\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 319 | \<open>x \<and> True \<longleftrightarrow> x\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 320 | \<open>x \<or> False \<longleftrightarrow> x\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 321 | for B C :: 'a and A' B' C' :: bool | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 322 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 323 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 324 | lemma verit_and_simplify1: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 325 | \<open>True \<and> b \<longleftrightarrow> b\<close> \<open>b \<and> True \<longleftrightarrow> b\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 326 | \<open>False \<and> b \<longleftrightarrow> False\<close> \<open>b \<and> False \<longleftrightarrow> False\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 327 | \<open>(c \<and> \<not>c) \<longleftrightarrow> False\<close> \<open>(\<not>c \<and> c) \<longleftrightarrow> False\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 328 | \<open>\<not>\<not>a = a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 329 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 330 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 331 | lemmas verit_and_simplify = conj_ac de_Morgan_conj disj_not1 | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 332 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 333 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 334 | lemma verit_or_simplify_1: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 335 | \<open>False \<or> b \<longleftrightarrow> b\<close> \<open>b \<or> False \<longleftrightarrow> b\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 336 | \<open>b \<or> \<not>b\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 337 | \<open>\<not>b \<or> b\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 338 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 339 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 340 | lemmas verit_or_simplify = disj_ac | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 341 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 342 | lemma verit_not_simplify: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 343 | \<open>\<not> \<not>b \<longleftrightarrow> b\<close> \<open>\<not>True \<longleftrightarrow> False\<close> \<open>\<not>False \<longleftrightarrow> True\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 344 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 345 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 346 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 347 | lemma verit_implies_simplify: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 348 | \<open>(\<not>a \<longrightarrow> \<not>b) \<longleftrightarrow> (b \<longrightarrow> a)\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 349 | \<open>(False \<longrightarrow> a) \<longleftrightarrow> True\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 350 | \<open>(a \<longrightarrow> True) \<longleftrightarrow> True\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 351 | \<open>(True \<longrightarrow> a) \<longleftrightarrow> a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 352 | \<open>(a \<longrightarrow> False) \<longleftrightarrow> \<not>a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 353 | \<open>(a \<longrightarrow> a) \<longleftrightarrow> True\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 354 | \<open>(\<not>a \<longrightarrow> a) \<longleftrightarrow> a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 355 | \<open>(a \<longrightarrow> \<not>a) \<longleftrightarrow> \<not>a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 356 | \<open>((a \<longrightarrow> b) \<longrightarrow> b) \<longleftrightarrow> a \<or> b\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 357 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 358 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 359 | lemma verit_equiv_simplify: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 360 | \<open>((\<not>a) = (\<not>b)) \<longleftrightarrow> (a = b)\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 361 | \<open>(a = a) \<longleftrightarrow> True\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 362 | \<open>(a = (\<not>a)) \<longleftrightarrow> False\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 363 | \<open>((\<not>a) = a) \<longleftrightarrow> False\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 364 | \<open>(True = a) \<longleftrightarrow> a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 365 | \<open>(a = True) \<longleftrightarrow> a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 366 | \<open>(False = a) \<longleftrightarrow> \<not>a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 367 | \<open>(a = False) \<longleftrightarrow> \<not>a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 368 | \<open>\<not>\<not>a \<longleftrightarrow> a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 369 | \<open>(\<not> False) = True\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 370 | for a b :: bool | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 371 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 372 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 373 | lemmas verit_eq_simplify = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 374 | semiring_char_0_class.eq_numeral_simps eq_refl zero_neq_one num.simps | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 375 | neg_equal_zero equal_neg_zero one_neq_zero neg_equal_iff_equal | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 376 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 377 | lemma verit_minus_simplify: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 378 | \<open>(a :: 'a :: cancel_comm_monoid_add) - a = 0\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 379 | \<open>(a :: 'a :: cancel_comm_monoid_add) - 0 = a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 380 |   \<open>0 - (b :: 'b :: {group_add}) = -b\<close>
 | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 381 | \<open>- (- (b :: 'b :: group_add)) = b\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 382 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 383 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 384 | lemma verit_sum_simplify: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 385 | \<open>(a :: 'a :: cancel_comm_monoid_add) + 0 = a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 386 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 387 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 388 | lemmas verit_prod_simplify = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 389 | (* already included: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 390 | mult_zero_class.mult_zero_right | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 391 | mult_zero_class.mult_zero_left *) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 392 | mult_1 | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 393 | mult_1_right | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 394 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 395 | lemma verit_comp_simplify1: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 396 | \<open>(a :: 'a ::order) < a \<longleftrightarrow> False\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 397 | \<open>a \<le> a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 398 | \<open>\<not>(b' \<le> a') \<longleftrightarrow> (a' :: 'b :: linorder) < b'\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 399 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 400 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 401 | lemmas verit_comp_simplify = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 402 | verit_comp_simplify1 | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 403 | le_numeral_simps | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 404 | le_num_simps | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 405 | less_numeral_simps | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 406 | less_num_simps | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 407 | zero_less_one | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 408 | zero_le_one | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 409 | less_neg_numeral_simps | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 410 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 411 | lemma verit_la_disequality: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 412 | \<open>(a :: 'a ::linorder) = b \<or> \<not>a \<le> b \<or> \<not>b \<le> a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 413 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 414 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 415 | context | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 416 | begin | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 417 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 418 | text \<open>For the reconstruction, we need to keep the order of the arguments.\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 419 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 420 | named_theorems smt_arith_multiplication \<open>Theorems to reconstruct arithmetic theorems.\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 421 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 422 | named_theorems smt_arith_combine \<open>Theorems to reconstruct arithmetic theorems.\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 423 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 424 | named_theorems smt_arith_simplify \<open>Theorems to combine theorems in the LA procedure\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 425 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 426 | lemmas [smt_arith_simplify] = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 427 | div_add dvd_numeral_simp divmod_steps less_num_simps le_num_simps if_True if_False divmod_cancel | 
| 75936 | 428 | dvd_mult dvd_mult2 less_irrefl prod.case numeral_plus_one divmod_step_def order.refl le_zero_eq | 
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 429 | le_numeral_simps less_numeral_simps mult.right_neutral simp_thms divides_aux_eq | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 430 | mult_nonneg_nonneg dvd_imp_mod_0 dvd_add zero_less_one mod_mult_self4 numeral_mod_numeral | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 431 | divmod_trivial prod.sel mult.left_neutral div_pos_pos_trivial arith_simps div_add div_mult_self1 | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 432 | add_le_cancel_left add_le_same_cancel2 not_one_le_zero le_numeral_simps add_le_same_cancel1 | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 433 | zero_neq_one zero_le_one le_num_simps add_Suc mod_div_trivial nat.distinct mult_minus_right | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 434 | add.inverse_inverse distrib_left_numeral mult_num_simps numeral_times_numeral add_num_simps | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 435 | divmod_steps rel_simps if_True if_False numeral_div_numeral divmod_cancel prod.case | 
| 75936 | 436 | add_num_simps one_plus_numeral fst_conv arith_simps sub_num_simps dbl_inc_simps | 
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 437 | dbl_simps mult_1 add_le_cancel_right left_diff_distrib_numeral add_uminus_conv_diff zero_neq_one | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 438 | zero_le_one One_nat_def add_Suc mod_div_trivial nat.distinct of_int_1 numerals numeral_One | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 439 | of_int_numeral add_uminus_conv_diff zle_diff1_eq add_less_same_cancel2 minus_add_distrib | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 440 | add_uminus_conv_diff mult.left_neutral semiring_class.distrib_right | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 441 | add_diff_cancel_left' add_diff_eq ring_distribs mult_minus_left minus_diff_eq | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 442 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 443 | lemma [smt_arith_simplify]: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 444 | \<open>\<not> (a' :: 'a :: linorder) < b' \<longleftrightarrow> b' \<le> a'\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 445 | \<open>\<not> (a' :: 'a :: linorder) \<le> b' \<longleftrightarrow> b' < a'\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 446 | \<open>(c::int) mod Numeral1 = 0\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 447 | \<open>(a::nat) mod Numeral1 = 0\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 448 | \<open>(c::int) div Numeral1 = c\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 449 | \<open>a div Numeral1 = a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 450 | \<open>(c::int) mod 1 = 0\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 451 | \<open>a mod 1 = 0\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 452 | \<open>(c::int) div 1 = c\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 453 | \<open>a div 1 = a\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 454 | \<open>\<not>(a' \<noteq> b') \<longleftrightarrow> a' = b'\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 455 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 456 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 457 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 458 | lemma div_mod_decomp: "A = (A div n) * n + (A mod n)" for A :: nat | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 459 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 460 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 461 | lemma div_less_mono: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 462 | fixes A B :: nat | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 463 | assumes "A < B" "0 < n" and | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 464 | mod: "A mod n = 0""B mod n = 0" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 465 | shows "(A div n) < (B div n)" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 466 | proof - | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 467 | show ?thesis | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 468 | using assms(1) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 469 | apply (subst (asm) div_mod_decomp[of "A" n]) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 470 | apply (subst (asm) div_mod_decomp[of "B" n]) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 471 | unfolding mod | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 472 | by (use assms(2,3) in \<open>auto simp: ac_simps\<close>) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 473 | qed | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 474 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 475 | lemma verit_le_mono_div: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 476 | fixes A B :: nat | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 477 | assumes "A < B" "0 < n" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 478 | shows "(A div n) + (if B mod n = 0 then 1 else 0) \<le> (B div n)" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 479 | by (auto simp: ac_simps Suc_leI assms less_mult_imp_div_less div_le_mono less_imp_le_nat) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 480 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 481 | lemmas [smt_arith_multiplication] = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 482 | verit_le_mono_div[THEN mult_le_mono1, unfolded add_mult_distrib] | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 483 | div_le_mono[THEN mult_le_mono2, unfolded add_mult_distrib] | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 484 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 485 | lemma div_mod_decomp_int: "A = (A div n) * n + (A mod n)" for A :: int | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 486 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 487 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 488 | lemma zdiv_mono_strict: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 489 | fixes A B :: int | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 490 | assumes "A < B" "0 < n" and | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 491 | mod: "A mod n = 0""B mod n = 0" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 492 | shows "(A div n) < (B div n)" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 493 | proof - | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 494 | show ?thesis | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 495 | using assms(1) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 496 | apply (subst (asm) div_mod_decomp_int[of A n]) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 497 | apply (subst (asm) div_mod_decomp_int[of B n]) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 498 | unfolding mod | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 499 | by (use assms(2,3) in \<open>auto simp: ac_simps\<close>) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 500 | qed | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 501 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 502 | lemma verit_le_mono_div_int: | 
| 76106 | 503 | \<open>A div n + (if B mod n = 0 then 1 else 0) \<le> B div n\<close> | 
| 504 | if \<open>A < B\<close> \<open>0 < n\<close> | |
| 505 | for A B n :: int | |
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 506 | proof - | 
| 76106 | 507 | from \<open>A < B\<close> \<open>0 < n\<close> have \<open>A div n \<le> B div n\<close> | 
| 508 | by (auto intro: zdiv_mono1) | |
| 509 | show ?thesis | |
| 510 | proof (cases \<open>n dvd B\<close>) | |
| 511 | case False | |
| 512 | with \<open>A div n \<le> B div n\<close> show ?thesis | |
| 513 | by auto | |
| 514 | next | |
| 515 | case True | |
| 516 | then obtain C where \<open>B = n * C\<close> .. | |
| 517 | then have \<open>B div n = C\<close> | |
| 518 | using \<open>0 < n\<close> by simp | |
| 519 | from \<open>0 < n\<close> have \<open>A mod n \<ge> 0\<close> | |
| 520 | by simp | |
| 521 | have \<open>A div n < C\<close> | |
| 522 | proof (rule ccontr) | |
| 523 | assume \<open>\<not> A div n < C\<close> | |
| 524 | then have \<open>C \<le> A div n\<close> | |
| 525 | by simp | |
| 526 | with \<open>B div n = C\<close> \<open>A div n \<le> B div n\<close> | |
| 527 | have \<open>A div n = C\<close> | |
| 528 | by simp | |
| 529 | moreover from \<open>A < B\<close> have \<open>n * (A div n) + A mod n < B\<close> | |
| 530 | by simp | |
| 531 | ultimately have \<open>n * C + A mod n < n * C\<close> | |
| 532 | using \<open>B = n * C\<close> by simp | |
| 533 | moreover have \<open>A mod n \<ge> 0\<close> | |
| 534 | using \<open>0 < n\<close> by simp | |
| 535 | ultimately show False | |
| 536 | by simp | |
| 537 | qed | |
| 538 | with \<open>n dvd B\<close> \<open>B div n = C\<close> show ?thesis | |
| 539 | by simp | |
| 540 | qed | |
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 541 | qed | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 542 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 543 | lemma verit_less_mono_div_int2: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 544 | fixes A B :: int | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 545 | assumes "A \<le> B" "0 < -n" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 546 | shows "(A div n) \<ge> (B div n)" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 547 | using assms(1) assms(2) zdiv_mono1_neg by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 548 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 549 | lemmas [smt_arith_multiplication] = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 550 | verit_le_mono_div_int[THEN mult_left_mono, unfolded int_distrib] | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 551 | zdiv_mono1[THEN mult_left_mono, unfolded int_distrib] | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 552 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 553 | lemmas [smt_arith_multiplication] = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 554 | arg_cong[of _ _ \<open>\<lambda>a :: nat. a div n * p\<close> for n p :: nat, THEN sym] | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 555 | arg_cong[of _ _ \<open>\<lambda>a :: int. a div n * p\<close> for n p :: int, THEN sym] | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 556 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 557 | lemma [smt_arith_combine]: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 558 | "a < b \<Longrightarrow> c < d \<Longrightarrow> a + c + 2 \<le> b + d" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 559 | "a < b \<Longrightarrow> c \<le> d \<Longrightarrow> a + c + 1 \<le> b + d" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 560 | "a \<le> b \<Longrightarrow> c < d \<Longrightarrow> a + c + 1 \<le> b + d" for a b c :: int | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 561 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 562 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 563 | lemma [smt_arith_combine]: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 564 | "a < b \<Longrightarrow> c < d \<Longrightarrow> a + c + 2 \<le> b + d" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 565 | "a < b \<Longrightarrow> c \<le> d \<Longrightarrow> a + c + 1 \<le> b + d" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 566 | "a \<le> b \<Longrightarrow> c < d \<Longrightarrow> a + c + 1 \<le> b + d" for a b c :: nat | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 567 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 568 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 569 | lemmas [smt_arith_combine] = | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 570 | add_strict_mono | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 571 | add_less_le_mono | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 572 | add_mono | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 573 | add_le_less_mono | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 574 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 575 | lemma [smt_arith_combine]: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 576 | \<open>m < n \<Longrightarrow> c = d \<Longrightarrow> m + c < n + d\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 577 | \<open>m \<le> n \<Longrightarrow> c = d \<Longrightarrow> m + c \<le> n + d\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 578 | \<open>c = d \<Longrightarrow> m < n \<Longrightarrow> m + c < n + d\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 579 | \<open>c = d \<Longrightarrow> m \<le> n \<Longrightarrow> m + c \<le> n + d\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 580 | for m :: \<open>'a :: ordered_cancel_ab_semigroup_add\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 581 | by (auto intro: ordered_cancel_ab_semigroup_add_class.add_strict_right_mono | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 582 | ordered_ab_semigroup_add_class.add_right_mono) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 583 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 584 | lemma verit_negate_coefficient: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 585 |   \<open>a \<le> (b :: 'a :: {ordered_ab_group_add}) \<Longrightarrow> -a \<ge> -b\<close>
 | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 586 | \<open>a < b \<Longrightarrow> -a > -b\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 587 | \<open>a = b \<Longrightarrow> -a = -b\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 588 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 589 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 590 | end | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 591 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 592 | lemma verit_ite_intro: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 593 | \<open>(if a then P (if a then a' else b') else Q) \<longleftrightarrow> (if a then P a' else Q)\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 594 | \<open>(if a then P' else Q' (if a then a' else b')) \<longleftrightarrow> (if a then P' else Q' b')\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 595 | \<open>A = f (if a then R else S) \<longleftrightarrow> (if a then A = f R else A = f S)\<close> | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 596 | by auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 597 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 598 | lemma verit_ite_if_cong: | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 599 | fixes x y :: bool | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 600 | assumes "b=c" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 601 | and "c \<equiv> True \<Longrightarrow> x = u" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 602 | and "c \<equiv> False \<Longrightarrow> y = v" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 603 | shows "(if b then x else y) \<equiv> (if c then u else v)" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 604 | proof - | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 605 | have H: "(if b then x else y) = (if c then u else v)" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 606 | using assms by (auto split: if_splits) | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 607 | |
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 608 | show "(if b then x else y) \<equiv> (if c then u else v)" | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 609 | by (subst H) auto | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 610 | qed | 
| 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 611 | |
| 69205 
8050734eee3e
add reconstruction by veriT in method smt
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
69204diff
changeset | 612 | |
| 60758 | 613 | subsection \<open>Setup\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 614 | |
| 69605 | 615 | ML_file \<open>Tools/SMT/smt_util.ML\<close> | 
| 616 | ML_file \<open>Tools/SMT/smt_failure.ML\<close> | |
| 617 | ML_file \<open>Tools/SMT/smt_config.ML\<close> | |
| 618 | ML_file \<open>Tools/SMT/smt_builtin.ML\<close> | |
| 619 | ML_file \<open>Tools/SMT/smt_datatypes.ML\<close> | |
| 620 | ML_file \<open>Tools/SMT/smt_normalize.ML\<close> | |
| 621 | ML_file \<open>Tools/SMT/smt_translate.ML\<close> | |
| 622 | ML_file \<open>Tools/SMT/smtlib.ML\<close> | |
| 623 | ML_file \<open>Tools/SMT/smtlib_interface.ML\<close> | |
| 624 | ML_file \<open>Tools/SMT/smtlib_proof.ML\<close> | |
| 625 | ML_file \<open>Tools/SMT/smtlib_isar.ML\<close> | |
| 626 | ML_file \<open>Tools/SMT/z3_proof.ML\<close> | |
| 627 | ML_file \<open>Tools/SMT/z3_isar.ML\<close> | |
| 628 | ML_file \<open>Tools/SMT/smt_solver.ML\<close> | |
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75299diff
changeset | 629 | ML_file \<open>Tools/SMT/cvc_interface.ML\<close> | 
| 75299 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75063diff
changeset | 630 | ML_file \<open>Tools/SMT/lethe_proof.ML\<close> | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75063diff
changeset | 631 | ML_file \<open>Tools/SMT/lethe_isar.ML\<close> | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75063diff
changeset | 632 | ML_file \<open>Tools/SMT/lethe_proof_parse.ML\<close> | 
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75299diff
changeset | 633 | ML_file \<open>Tools/SMT/cvc_proof_parse.ML\<close> | 
| 69605 | 634 | ML_file \<open>Tools/SMT/conj_disj_perm.ML\<close> | 
| 635 | ML_file \<open>Tools/SMT/smt_replay_methods.ML\<close> | |
| 636 | ML_file \<open>Tools/SMT/smt_replay.ML\<close> | |
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 637 | ML_file \<open>Tools/SMT/smt_replay_arith.ML\<close> | 
| 69605 | 638 | ML_file \<open>Tools/SMT/z3_interface.ML\<close> | 
| 639 | ML_file \<open>Tools/SMT/z3_replay_rules.ML\<close> | |
| 640 | ML_file \<open>Tools/SMT/z3_replay_methods.ML\<close> | |
| 641 | ML_file \<open>Tools/SMT/z3_replay.ML\<close> | |
| 75299 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75063diff
changeset | 642 | ML_file \<open>Tools/SMT/lethe_replay_methods.ML\<close> | 
| 78177 
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
76106diff
changeset | 643 | ML_file \<open>Tools/SMT/cvc5_replay_methods.ML\<close> | 
| 69605 | 644 | ML_file \<open>Tools/SMT/verit_replay_methods.ML\<close> | 
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75936diff
changeset | 645 | ML_file \<open>Tools/SMT/verit_strategies.ML\<close> | 
| 69605 | 646 | ML_file \<open>Tools/SMT/verit_replay.ML\<close> | 
| 78177 
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
76106diff
changeset | 647 | ML_file \<open>Tools/SMT/cvc5_replay.ML\<close> | 
| 69605 | 648 | 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 | 649 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 650 | |
| 60758 | 651 | subsection \<open>Configuration\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 652 | |
| 60758 | 653 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 654 | The current configuration can be printed by the command | 
| 61799 | 655 | \<open>smt_status\<close>, which shows the values of most options. | 
| 60758 | 656 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 657 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 658 | |
| 60758 | 659 | subsection \<open>General configuration options\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 660 | |
| 60758 | 661 | text \<open> | 
| 61799 | 662 | The option \<open>smt_solver\<close> can be used to change the target SMT | 
| 663 | 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 | 664 | command. | 
| 60758 | 665 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 666 | |
| 58061 | 667 | declare [[smt_solver = z3]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 668 | |
| 60758 | 669 | 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 | 670 | 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 | 671 | (given in seconds) to restrict their runtime. | 
| 60758 | 672 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 673 | |
| 73389 | 674 | declare [[smt_timeout = 0]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 675 | |
| 60758 | 676 | 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 | 677 | 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 | 678 | solvable by an SMT solver, changing the following option might help. | 
| 60758 | 679 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 680 | |
| 58061 | 681 | declare [[smt_random_seed = 1]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 682 | |
| 60758 | 683 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 684 | 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 | 685 | 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 | 686 | option can cause the SMT solver to run in proof-producing mode, giving | 
| 75063 
7ff39293e265
added possibility of extra options to SMT slices
 blanchet parents: 
74740diff
changeset | 687 | a checkable certificate. This is currently implemented only for veriT and | 
| 
7ff39293e265
added possibility of extra options to SMT slices
 blanchet parents: 
74740diff
changeset | 688 | Z3. | 
| 60758 | 689 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 690 | |
| 58061 | 691 | declare [[smt_oracle = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 692 | |
| 60758 | 693 | text \<open> | 
| 75063 
7ff39293e265
added possibility of extra options to SMT slices
 blanchet parents: 
74740diff
changeset | 694 | Each SMT solver provides several command-line 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 | 695 | 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 | 696 | options. | 
| 60758 | 697 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 698 | |
| 75063 
7ff39293e265
added possibility of extra options to SMT slices
 blanchet parents: 
74740diff
changeset | 699 | declare [[cvc4_options = ""]] | 
| 79623 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
79576diff
changeset | 700 | declare [[cvc5_options = ""]] | 
| 
e905fb37467f
Distinguish two versions of cvc5 -- one used for Sledgehammer, one for proof reconstruction in SMT. provided by Mathias Fleury
 blanchet parents: 
79576diff
changeset | 701 | declare [[cvc5_proof_options = "--proof-format-mode=alethe --proof-granularity=dsl-rewrite"]] | 
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 702 | declare [[verit_options = ""]] | 
| 58061 | 703 | declare [[z3_options = ""]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 704 | |
| 60758 | 705 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 706 | 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 | 707 | 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 | 708 | 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 | 709 | in the SMT solver). To turn it on, set the following option. | 
| 60758 | 710 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 711 | |
| 58061 | 712 | declare [[smt_infer_triggers = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 713 | |
| 60758 | 714 | text \<open> | 
| 58360 | 715 | Enable the following option to use built-in support for datatypes, | 
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75299diff
changeset | 716 | codatatypes, and records in CVC4 and cvc5. Currently, this is implemented | 
| 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75299diff
changeset | 717 | only in oracle mode. | 
| 60758 | 718 | \<close> | 
| 58360 | 719 | |
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75299diff
changeset | 720 | declare [[cvc_extensions = false]] | 
| 58360 | 721 | |
| 60758 | 722 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 723 | 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 | 724 | and records in Z3. Currently, this is implemented only in oracle mode. | 
| 60758 | 725 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 726 | |
| 58061 | 727 | declare [[z3_extensions = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 728 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 729 | |
| 60758 | 730 | subsection \<open>Certificates\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 731 | |
| 60758 | 732 | text \<open> | 
| 61799 | 733 | 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 | 734 | 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 | 735 | 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 | 736 | 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 | 737 | solver. An empty string disables caching certificates. | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 738 | |
| 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 | 739 | 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 | 740 | practice to use the name of the current theory (with ending | 
| 61799 | 741 | \<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 | 742 | 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 | 743 | to avoid race conditions with other concurrent accesses. | 
| 60758 | 744 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 745 | |
| 58061 | 746 | declare [[smt_certificates = ""]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 747 | |
| 60758 | 748 | text \<open> | 
| 61799 | 749 | The option \<open>smt_read_only_certificates\<close> controls whether only | 
| 74740 | 750 | stored certificates should be used or invocation of an SMT solver | 
| 61799 | 751 | 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 | 752 | invoked and only the existing certificates found in the configured | 
| 61799 | 753 | 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 | 754 | 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 | 755 | invoked. | 
| 60758 | 756 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 757 | |
| 58061 | 758 | declare [[smt_read_only_certificates = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 759 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 760 | |
| 60758 | 761 | subsection \<open>Tracing\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 762 | |
| 60758 | 763 | 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 | 764 | The SMT method, when applied, traces important information. To | 
| 61799 | 765 | make it entirely silent, set the following option to \<open>false\<close>. | 
| 60758 | 766 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 767 | |
| 58061 | 768 | declare [[smt_verbose = true]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 769 | |
| 60758 | 770 | text \<open> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 771 | 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 | 772 | well as the returned result of the solver, the option | 
| 61799 | 773 | \<open>smt_trace\<close> should be set to \<open>true\<close>. | 
| 60758 | 774 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 775 | |
| 58061 | 776 | declare [[smt_trace = false]] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 777 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 778 | |
| 60758 | 779 | 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 | 780 | |
| 60758 | 781 | 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 | 782 | 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 | 783 | lemma groups which can turn failing Z3 proof reconstruction attempts | 
| 61799 | 784 | 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 | 785 | any implemented reconstruction procedure for all uncertain Z3 proof | 
| 61799 | 786 | 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 | 787 | the simplifier when reconstructing theory-specific proof steps. | 
| 60758 | 788 | \<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 789 | |
| 58061 | 790 | lemmas [z3_rule] = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 791 | 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 | 792 | 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 | 793 | 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 | 794 | NO_MATCH_def | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 795 | |
| 58061 | 796 | lemma [z3_rule]: | 
| 57169 | 797 | "(P \<and> Q) = (\<not> (\<not> P \<or> \<not> Q))" | 
| 798 | "(P \<and> Q) = (\<not> (\<not> Q \<or> \<not> P))" | |
| 799 | "(\<not> P \<and> Q) = (\<not> (P \<or> \<not> Q))" | |
| 800 | "(\<not> P \<and> Q) = (\<not> (\<not> Q \<or> P))" | |
| 801 | "(P \<and> \<not> Q) = (\<not> (\<not> P \<or> Q))" | |
| 802 | "(P \<and> \<not> Q) = (\<not> (Q \<or> \<not> P))" | |
| 803 | "(\<not> P \<and> \<not> Q) = (\<not> (P \<or> Q))" | |
| 804 | "(\<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 | 805 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 806 | |
| 58061 | 807 | lemma [z3_rule]: | 
| 57169 | 808 | "(P \<longrightarrow> Q) = (Q \<or> \<not> P)" | 
| 809 | "(\<not> P \<longrightarrow> Q) = (P \<or> Q)" | |
| 810 | "(\<not> P \<longrightarrow> Q) = (Q \<or> P)" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 811 | "(True \<longrightarrow> P) = P" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 812 | "(P \<longrightarrow> True) = True" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 813 | "(False \<longrightarrow> P) = True" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 814 | "(P \<longrightarrow> P) = True" | 
| 59037 
650dcf624729
added Z3 reconstruction rule suggested by F. Maric
 blanchet parents: 
59036diff
changeset | 815 | "(\<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 | 816 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 817 | |
| 58061 | 818 | lemma [z3_rule]: | 
| 67091 | 819 | "((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 | 820 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 821 | |
| 58061 | 822 | lemma [z3_rule]: | 
| 57169 | 823 | "(\<not> True) = False" | 
| 824 | "(\<not> False) = True" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 825 | "(x = x) = True" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 826 | "(P = True) = P" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 827 | "(True = P) = P" | 
| 57169 | 828 | "(P = False) = (\<not> P)" | 
| 829 | "(False = P) = (\<not> P)" | |
| 830 | "((\<not> P) = P) = False" | |
| 831 | "(P = (\<not> P)) = False" | |
| 832 | "((\<not> P) = (\<not> Q)) = (P = Q)" | |
| 833 | "\<not> (P = (\<not> Q)) = (P = Q)" | |
| 834 | "\<not> ((\<not> P) = Q) = (P = Q)" | |
| 835 | "(P \<noteq> Q) = (Q = (\<not> P))" | |
| 836 | "(P = Q) = ((\<not> P \<or> Q) \<and> (P \<or> \<not> Q))" | |
| 837 | "(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 | 838 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 839 | |
| 58061 | 840 | lemma [z3_rule]: | 
| 57169 | 841 | "(if P then P else \<not> P) = True" | 
| 842 | "(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 | 843 | "(if P then True else False) = P" | 
| 57169 | 844 | "(if P then False else True) = (\<not> P)" | 
| 845 | "(if P then Q else True) = ((\<not> P) \<or> Q)" | |
| 846 | "(if P then Q else True) = (Q \<or> (\<not> P))" | |
| 847 | "(if P then Q else \<not> Q) = (P = Q)" | |
| 848 | "(if P then Q else \<not> Q) = (Q = P)" | |
| 849 | "(if P then \<not> Q else Q) = (P = (\<not> Q))" | |
| 850 | "(if P then \<not> Q else Q) = ((\<not> Q) = P)" | |
| 851 | "(if \<not> P then x else y) = (if P then y else x)" | |
| 852 | "(if P then (if Q then x else y) else x) = (if P \<and> (\<not> Q) then y else x)" | |
| 853 | "(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 | 854 | "(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 | 855 | "(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 | 856 | "(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 | 857 | "(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 | 858 | "(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 | 859 | "(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 | 860 | "(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 | 861 | "(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 | 862 | by auto | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 863 | |
| 58061 | 864 | lemma [z3_rule]: | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 865 | "0 + (x::int) = x" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 866 | "x + 0 = x" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 867 | "x + x = 2 * x" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 868 | "0 * x = 0" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 869 | "1 * x = x" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 870 | "x + y = y + x" | 
| 72458 
b44e894796d5
add reconstruction for the SMT solver veriT
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
72343diff
changeset | 871 | by auto | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 872 | |
| 58061 | 873 | lemma [z3_rule]: (* for def-axiom *) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 874 | "P = Q \<or> P \<or> Q" | 
| 57169 | 875 | "P = Q \<or> \<not> P \<or> \<not> Q" | 
| 876 | "(\<not> P) = Q \<or> \<not> P \<or> Q" | |
| 877 | "(\<not> P) = Q \<or> P \<or> \<not> Q" | |
| 878 | "P = (\<not> Q) \<or> \<not> P \<or> Q" | |
| 879 | "P = (\<not> Q) \<or> P \<or> \<not> Q" | |
| 880 | "P \<noteq> Q \<or> P \<or> \<not> Q" | |
| 881 | "P \<noteq> Q \<or> \<not> P \<or> Q" | |
| 882 | "P \<noteq> (\<not> Q) \<or> P \<or> Q" | |
| 883 | "(\<not> P) \<noteq> Q \<or> P \<or> Q" | |
| 884 | "P \<or> Q \<or> P \<noteq> (\<not> Q)" | |
| 885 | "P \<or> Q \<or> (\<not> P) \<noteq> Q" | |
| 886 | "P \<or> \<not> Q \<or> P \<noteq> Q" | |
| 887 | "\<not> P \<or> Q \<or> P \<noteq> Q" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 888 | "P \<or> y = (if P then x else y)" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 889 | "P \<or> (if P then x else y) = y" | 
| 57169 | 890 | "\<not> P \<or> x = (if P then x else y)" | 
| 891 | "\<not> P \<or> (if P then x else y) = x" | |
| 892 | "P \<or> R \<or> \<not> (if P then Q else R)" | |
| 893 | "\<not> P \<or> Q \<or> \<not> (if P then Q else R)" | |
| 894 | "\<not> (if P then Q else R) \<or> \<not> P \<or> Q" | |
| 895 | "\<not> (if P then Q else R) \<or> P \<or> R" | |
| 896 | "(if P then Q else R) \<or> \<not> P \<or> \<not> Q" | |
| 897 | "(if P then Q else R) \<or> P \<or> \<not> R" | |
| 898 | "(if P then \<not> Q else R) \<or> \<not> P \<or> Q" | |
| 899 | "(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 | 900 | by auto | 
| 75806 
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
 blanchet parents: 
75299diff
changeset | 901 | |
| 57230 | 902 | hide_type (open) symb_list pattern | 
| 903 | 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 | 904 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 905 | end |