| author | wenzelm | 
| Sun, 25 Oct 2009 00:00:53 +0200 | |
| changeset 33102 | e3463e6db704 | 
| parent 33011 | ab599f7f2639 | 
| child 33299 | 73af7831ba1e | 
| permissions | -rw-r--r-- | 
| 33011 | 1 | (* Title: HOL/SMT/SMT_Examples.thy | 
| 2 | Author: Sascha Boehme, TU Muenchen | |
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 3 | *) | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 4 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 5 | header {* Examples for the 'smt' tactic. *}
 | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 6 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 7 | theory SMT_Examples | 
| 33010 | 8 | imports SMT | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 9 | begin | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 10 | |
| 33010 | 11 | declare [[smt_solver=z3, z3_proofs=true]] | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 12 | |
| 33011 | 13 | text {*
 | 
| 14 | To re-generate the certificates, replace the option 'smt_cert' with 'smt_keep' | |
| 15 | (while keeping the paths as they are) and let Isabelle process this theory. | |
| 16 | *} | |
| 17 | ||
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 18 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 19 | section {* Propositional and first-order logic *}
 | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 20 | |
| 33010 | 21 | lemma "True" | 
| 33011 | 22 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_01"]] | 
| 33010 | 23 | by smt | 
| 24 | ||
| 25 | lemma "p \<or> \<not>p" | |
| 33011 | 26 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_02"]] | 
| 33010 | 27 | by smt | 
| 28 | ||
| 29 | lemma "(p \<and> True) = p" | |
| 33011 | 30 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_03"]] | 
| 33010 | 31 | by smt | 
| 32 | ||
| 33 | lemma "(p \<or> q) \<and> \<not>p \<Longrightarrow> q" | |
| 33011 | 34 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_04"]] | 
| 33010 | 35 | by smt | 
| 36 | ||
| 37 | lemma "(a \<and> b) \<or> (c \<and> d) \<Longrightarrow> (a \<and> b) \<or> (c \<and> d)" | |
| 33011 | 38 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_05"]] | 
| 33010 | 39 | using [[z3_proofs=false]] (* no Z3 proof *) | 
| 40 | by smt | |
| 41 | ||
| 42 | lemma "(p1 \<and> p2) \<or> p3 \<longrightarrow> (p1 \<longrightarrow> (p3 \<and> p2) \<or> (p1 \<and> p3)) \<or> p1" | |
| 33011 | 43 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_06"]] | 
| 33010 | 44 | by smt | 
| 45 | ||
| 46 | lemma "P=P=P=P=P=P=P=P=P=P" | |
| 33011 | 47 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_07"]] | 
| 33010 | 48 | by smt | 
| 49 | ||
| 50 | lemma | |
| 51 | assumes "a | b | c | d" | |
| 52 | and "e | f | (a & d)" | |
| 53 | and "~(a | (c & ~c)) | b" | |
| 54 | and "~(b & (x | ~x)) | c" | |
| 55 | and "~(d | False) | c" | |
| 56 | and "~(c | (~p & (p | (q & ~q))))" | |
| 57 | shows False | |
| 33011 | 58 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_08"]] | 
| 33010 | 59 | using assms by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 60 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 61 | axiomatization symm_f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 62 | symm_f: "symm_f x y = symm_f y x" | 
| 33010 | 63 | lemma "a = a \<and> symm_f a b = symm_f b a" | 
| 33011 | 64 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_09"]] | 
| 33010 | 65 | by (smt add: symm_f) | 
| 66 | ||
| 67 | (* | |
| 68 | Taken from ~~/src/HOL/ex/SAT_Examples.thy. | |
| 69 | Translated from TPTP problem library: PUZ015-2.006.dimacs | |
| 70 | *) | |
| 71 | lemma | |
| 72 | assumes "~x0" | |
| 73 | and "~x30" | |
| 74 | and "~x29" | |
| 75 | and "~x59" | |
| 76 | and "x1 | x31 | x0" | |
| 77 | and "x2 | x32 | x1" | |
| 78 | and "x3 | x33 | x2" | |
| 79 | and "x4 | x34 | x3" | |
| 80 | and "x35 | x4" | |
| 81 | and "x5 | x36 | x30" | |
| 82 | and "x6 | x37 | x5 | x31" | |
| 83 | and "x7 | x38 | x6 | x32" | |
| 84 | and "x8 | x39 | x7 | x33" | |
| 85 | and "x9 | x40 | x8 | x34" | |
| 86 | and "x41 | x9 | x35" | |
| 87 | and "x10 | x42 | x36" | |
| 88 | and "x11 | x43 | x10 | x37" | |
| 89 | and "x12 | x44 | x11 | x38" | |
| 90 | and "x13 | x45 | x12 | x39" | |
| 91 | and "x14 | x46 | x13 | x40" | |
| 92 | and "x47 | x14 | x41" | |
| 93 | and "x15 | x48 | x42" | |
| 94 | and "x16 | x49 | x15 | x43" | |
| 95 | and "x17 | x50 | x16 | x44" | |
| 96 | and "x18 | x51 | x17 | x45" | |
| 97 | and "x19 | x52 | x18 | x46" | |
| 98 | and "x53 | x19 | x47" | |
| 99 | and "x20 | x54 | x48" | |
| 100 | and "x21 | x55 | x20 | x49" | |
| 101 | and "x22 | x56 | x21 | x50" | |
| 102 | and "x23 | x57 | x22 | x51" | |
| 103 | and "x24 | x58 | x23 | x52" | |
| 104 | and "x59 | x24 | x53" | |
| 105 | and "x25 | x54" | |
| 106 | and "x26 | x25 | x55" | |
| 107 | and "x27 | x26 | x56" | |
| 108 | and "x28 | x27 | x57" | |
| 109 | and "x29 | x28 | x58" | |
| 110 | and "~x1 | ~x31" | |
| 111 | and "~x1 | ~x0" | |
| 112 | and "~x31 | ~x0" | |
| 113 | and "~x2 | ~x32" | |
| 114 | and "~x2 | ~x1" | |
| 115 | and "~x32 | ~x1" | |
| 116 | and "~x3 | ~x33" | |
| 117 | and "~x3 | ~x2" | |
| 118 | and "~x33 | ~x2" | |
| 119 | and "~x4 | ~x34" | |
| 120 | and "~x4 | ~x3" | |
| 121 | and "~x34 | ~x3" | |
| 122 | and "~x35 | ~x4" | |
| 123 | and "~x5 | ~x36" | |
| 124 | and "~x5 | ~x30" | |
| 125 | and "~x36 | ~x30" | |
| 126 | and "~x6 | ~x37" | |
| 127 | and "~x6 | ~x5" | |
| 128 | and "~x6 | ~x31" | |
| 129 | and "~x37 | ~x5" | |
| 130 | and "~x37 | ~x31" | |
| 131 | and "~x5 | ~x31" | |
| 132 | and "~x7 | ~x38" | |
| 133 | and "~x7 | ~x6" | |
| 134 | and "~x7 | ~x32" | |
| 135 | and "~x38 | ~x6" | |
| 136 | and "~x38 | ~x32" | |
| 137 | and "~x6 | ~x32" | |
| 138 | and "~x8 | ~x39" | |
| 139 | and "~x8 | ~x7" | |
| 140 | and "~x8 | ~x33" | |
| 141 | and "~x39 | ~x7" | |
| 142 | and "~x39 | ~x33" | |
| 143 | and "~x7 | ~x33" | |
| 144 | and "~x9 | ~x40" | |
| 145 | and "~x9 | ~x8" | |
| 146 | and "~x9 | ~x34" | |
| 147 | and "~x40 | ~x8" | |
| 148 | and "~x40 | ~x34" | |
| 149 | and "~x8 | ~x34" | |
| 150 | and "~x41 | ~x9" | |
| 151 | and "~x41 | ~x35" | |
| 152 | and "~x9 | ~x35" | |
| 153 | and "~x10 | ~x42" | |
| 154 | and "~x10 | ~x36" | |
| 155 | and "~x42 | ~x36" | |
| 156 | and "~x11 | ~x43" | |
| 157 | and "~x11 | ~x10" | |
| 158 | and "~x11 | ~x37" | |
| 159 | and "~x43 | ~x10" | |
| 160 | and "~x43 | ~x37" | |
| 161 | and "~x10 | ~x37" | |
| 162 | and "~x12 | ~x44" | |
| 163 | and "~x12 | ~x11" | |
| 164 | and "~x12 | ~x38" | |
| 165 | and "~x44 | ~x11" | |
| 166 | and "~x44 | ~x38" | |
| 167 | and "~x11 | ~x38" | |
| 168 | and "~x13 | ~x45" | |
| 169 | and "~x13 | ~x12" | |
| 170 | and "~x13 | ~x39" | |
| 171 | and "~x45 | ~x12" | |
| 172 | and "~x45 | ~x39" | |
| 173 | and "~x12 | ~x39" | |
| 174 | and "~x14 | ~x46" | |
| 175 | and "~x14 | ~x13" | |
| 176 | and "~x14 | ~x40" | |
| 177 | and "~x46 | ~x13" | |
| 178 | and "~x46 | ~x40" | |
| 179 | and "~x13 | ~x40" | |
| 180 | and "~x47 | ~x14" | |
| 181 | and "~x47 | ~x41" | |
| 182 | and "~x14 | ~x41" | |
| 183 | and "~x15 | ~x48" | |
| 184 | and "~x15 | ~x42" | |
| 185 | and "~x48 | ~x42" | |
| 186 | and "~x16 | ~x49" | |
| 187 | and "~x16 | ~x15" | |
| 188 | and "~x16 | ~x43" | |
| 189 | and "~x49 | ~x15" | |
| 190 | and "~x49 | ~x43" | |
| 191 | and "~x15 | ~x43" | |
| 192 | and "~x17 | ~x50" | |
| 193 | and "~x17 | ~x16" | |
| 194 | and "~x17 | ~x44" | |
| 195 | and "~x50 | ~x16" | |
| 196 | and "~x50 | ~x44" | |
| 197 | and "~x16 | ~x44" | |
| 198 | and "~x18 | ~x51" | |
| 199 | and "~x18 | ~x17" | |
| 200 | and "~x18 | ~x45" | |
| 201 | and "~x51 | ~x17" | |
| 202 | and "~x51 | ~x45" | |
| 203 | and "~x17 | ~x45" | |
| 204 | and "~x19 | ~x52" | |
| 205 | and "~x19 | ~x18" | |
| 206 | and "~x19 | ~x46" | |
| 207 | and "~x52 | ~x18" | |
| 208 | and "~x52 | ~x46" | |
| 209 | and "~x18 | ~x46" | |
| 210 | and "~x53 | ~x19" | |
| 211 | and "~x53 | ~x47" | |
| 212 | and "~x19 | ~x47" | |
| 213 | and "~x20 | ~x54" | |
| 214 | and "~x20 | ~x48" | |
| 215 | and "~x54 | ~x48" | |
| 216 | and "~x21 | ~x55" | |
| 217 | and "~x21 | ~x20" | |
| 218 | and "~x21 | ~x49" | |
| 219 | and "~x55 | ~x20" | |
| 220 | and "~x55 | ~x49" | |
| 221 | and "~x20 | ~x49" | |
| 222 | and "~x22 | ~x56" | |
| 223 | and "~x22 | ~x21" | |
| 224 | and "~x22 | ~x50" | |
| 225 | and "~x56 | ~x21" | |
| 226 | and "~x56 | ~x50" | |
| 227 | and "~x21 | ~x50" | |
| 228 | and "~x23 | ~x57" | |
| 229 | and "~x23 | ~x22" | |
| 230 | and "~x23 | ~x51" | |
| 231 | and "~x57 | ~x22" | |
| 232 | and "~x57 | ~x51" | |
| 233 | and "~x22 | ~x51" | |
| 234 | and "~x24 | ~x58" | |
| 235 | and "~x24 | ~x23" | |
| 236 | and "~x24 | ~x52" | |
| 237 | and "~x58 | ~x23" | |
| 238 | and "~x58 | ~x52" | |
| 239 | and "~x23 | ~x52" | |
| 240 | and "~x59 | ~x24" | |
| 241 | and "~x59 | ~x53" | |
| 242 | and "~x24 | ~x53" | |
| 243 | and "~x25 | ~x54" | |
| 244 | and "~x26 | ~x25" | |
| 245 | and "~x26 | ~x55" | |
| 246 | and "~x25 | ~x55" | |
| 247 | and "~x27 | ~x26" | |
| 248 | and "~x27 | ~x56" | |
| 249 | and "~x26 | ~x56" | |
| 250 | and "~x28 | ~x27" | |
| 251 | and "~x28 | ~x57" | |
| 252 | and "~x27 | ~x57" | |
| 253 | and "~x29 | ~x28" | |
| 254 | and "~x29 | ~x58" | |
| 255 | and "~x28 | ~x58" | |
| 256 | shows False | |
| 257 | using assms | |
| 33011 | 258 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_10"]] | 
| 33010 | 259 | by smt | 
| 260 | ||
| 261 | lemma "\<forall>x::int. P x \<longrightarrow> (\<forall>y::int. P x \<or> P y)" | |
| 33011 | 262 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_fol_01"]] | 
| 33010 | 263 | by smt | 
| 264 | ||
| 265 | lemma | |
| 266 | assumes "(\<forall>x y. P x y = x)" | |
| 267 | shows "(\<exists>y. P x y) = P x c" | |
| 268 | using assms | |
| 33011 | 269 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_fol_02"]] | 
| 33010 | 270 | by smt | 
| 271 | ||
| 272 | lemma | |
| 273 | assumes "(\<forall>x y. P x y = x)" | |
| 274 | and "(\<forall>x. \<exists>y. P x y) = (\<forall>x. P x c)" | |
| 275 | shows "(EX y. P x y) = P x c" | |
| 276 | using assms | |
| 33011 | 277 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_fol_03"]] | 
| 33010 | 278 | by smt | 
| 279 | ||
| 280 | lemma | |
| 281 | assumes "if P x then \<not>(\<exists>y. P y) else (\<forall>y. \<not>P y)" | |
| 282 | shows "P x \<longrightarrow> P y" | |
| 283 | using assms | |
| 33011 | 284 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_fol_04"]] | 
| 33010 | 285 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 286 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 287 | |
| 33010 | 288 | section {* Arithmetic *}
 | 
| 289 | ||
| 290 | subsection {* Linear arithmetic over integers and reals *}
 | |
| 291 | ||
| 292 | lemma "(3::int) = 3" | |
| 33011 | 293 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_01"]] | 
| 33010 | 294 | by smt | 
| 295 | ||
| 296 | lemma "(3::real) = 3" | |
| 33011 | 297 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_02"]] | 
| 33010 | 298 | by smt | 
| 299 | ||
| 300 | lemma "(3 :: int) + 1 = 4" | |
| 33011 | 301 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_03"]] | 
| 33010 | 302 | by smt | 
| 303 | ||
| 304 | lemma "x + (y + z) = y + (z + (x::int))" | |
| 33011 | 305 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_04"]] | 
| 33010 | 306 | by smt | 
| 307 | ||
| 308 | lemma "max (3::int) 8 > 5" | |
| 33011 | 309 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_05"]] | 
| 33010 | 310 | by smt | 
| 311 | ||
| 312 | lemma "abs (x :: real) + abs y \<ge> abs (x + y)" | |
| 33011 | 313 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_06"]] | 
| 33010 | 314 | by smt | 
| 315 | ||
| 316 | lemma "P ((2::int) < 3) = P True" | |
| 33011 | 317 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_07"]] | 
| 33010 | 318 | by smt | 
| 319 | ||
| 320 | lemma "x + 3 \<ge> 4 \<or> x < (1::int)" | |
| 33011 | 321 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_08"]] | 
| 33010 | 322 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 323 | |
| 33010 | 324 | lemma | 
| 325 | assumes "x \<ge> (3::int)" and "y = x + 4" | |
| 326 | shows "y - x > 0" | |
| 327 | using assms | |
| 33011 | 328 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_09"]] | 
| 33010 | 329 | by smt | 
| 330 | ||
| 331 | lemma "let x = (2 :: int) in x + x \<noteq> 5" | |
| 33011 | 332 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_10"]] | 
| 33010 | 333 | by smt | 
| 334 | ||
| 335 | lemma | |
| 336 | fixes x :: real | |
| 337 | assumes "3 * x + 7 * a < 4" and "3 < 2 * x" | |
| 338 | shows "a < 0" | |
| 339 | using assms | |
| 33011 | 340 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_11"]] | 
| 33010 | 341 | by smt | 
| 342 | ||
| 343 | lemma "(0 \<le> y + -1 * x \<or> \<not> 0 \<le> x \<or> 0 \<le> (x::int)) = (\<not> False)" | |
| 33011 | 344 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_12"]] | 
| 33010 | 345 | by smt | 
| 346 | ||
| 347 | lemma "distinct [x < (3::int), 3 \<le> x]" | |
| 33011 | 348 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_13"]] | 
| 33010 | 349 | by smt | 
| 350 | ||
| 351 | lemma | |
| 352 | assumes "a > (0::int)" | |
| 353 | shows "distinct [a, a * 2, a - a]" | |
| 354 | using assms | |
| 33011 | 355 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_14"]] | 
| 33010 | 356 | by smt | 
| 357 | ||
| 358 | lemma " | |
| 359 | (n < m & m < n') | (n < m & m = n') | (n < n' & n' < m) | | |
| 360 | (n = n' & n' < m) | (n = m & m < n') | | |
| 361 | (n' < m & m < n) | (n' < m & m = n) | | |
| 362 | (n' < n & n < m) | (n' = n & n < m) | (n' = m & m < n) | | |
| 363 | (m < n & n < n') | (m < n & n' = n) | (m < n' & n' < n) | | |
| 364 | (m = n & n < n') | (m = n' & n' < n) | | |
| 365 | (n' = m & m = (n::int))" | |
| 33011 | 366 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_15"]] | 
| 33010 | 367 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 368 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 369 | text{* 
 | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 370 | The following example was taken from HOL/ex/PresburgerEx.thy, where it says: | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 371 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 372 | This following theorem proves that all solutions to the | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 373 |   recurrence relation $x_{i+2} = |x_{i+1}| - x_i$ are periodic with
 | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 374 | period 9. The example was brought to our attention by John | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 375 | Harrison. It does does not require Presburger arithmetic but merely | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 376 | quantifier-free linear arithmetic and holds for the rationals as well. | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 377 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 378 | Warning: it takes (in 2006) over 4.2 minutes! | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 379 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 380 | There, it is proved by "arith". SMT is able to prove this within a fraction | 
| 33010 | 381 | of one second. With proof reconstruction, it takes about 13 seconds on a Core2 | 
| 382 | processor. | |
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 383 | *} | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 384 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 385 | lemma "\<lbrakk> x3 = abs x2 - x1; x4 = abs x3 - x2; x5 = abs x4 - x3; | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 386 | x6 = abs x5 - x4; x7 = abs x6 - x5; x8 = abs x7 - x6; | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 387 | x9 = abs x8 - x7; x10 = abs x9 - x8; x11 = abs x10 - x9 \<rbrakk> | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 388 | \<Longrightarrow> x1 = x10 & x2 = (x11::int)" | 
| 33011 | 389 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_16"]] | 
| 33010 | 390 | by smt | 
| 391 | ||
| 392 | ||
| 393 | subsection {* Linear arithmetic with quantifiers *}
 | |
| 394 | ||
| 395 | lemma "~ (\<exists>x::int. False)" | |
| 33011 | 396 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_01"]] | 
| 33010 | 397 | by smt | 
| 398 | ||
| 399 | lemma "~ (\<exists>x::real. False)" | |
| 33011 | 400 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_02"]] | 
| 33010 | 401 | by smt | 
| 402 | ||
| 403 | lemma "\<exists>x::int. 0 < x" | |
| 33011 | 404 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_03"]] | 
| 33010 | 405 | using [[z3_proofs=false]] (* no Z3 proof *) | 
| 406 | by smt | |
| 407 | ||
| 408 | lemma "\<exists>x::real. 0 < x" | |
| 33011 | 409 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_04"]] | 
| 33010 | 410 | using [[z3_proofs=false]] (* no Z3 proof *) | 
| 411 | by smt | |
| 412 | ||
| 413 | lemma "\<forall>x::int. \<exists>y. y > x" | |
| 33011 | 414 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_05"]] | 
| 33010 | 415 | using [[z3_proofs=false]] (* no Z3 proof *) | 
| 416 | by smt | |
| 417 | ||
| 418 | lemma "\<forall>x y::int. (x = 0 \<and> y = 1) \<longrightarrow> x \<noteq> y" | |
| 33011 | 419 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_06"]] | 
| 33010 | 420 | by smt | 
| 421 | ||
| 422 | lemma "\<exists>x::int. \<forall>y. x < y \<longrightarrow> y < 0 \<or> y >= 0" | |
| 33011 | 423 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_07"]] | 
| 33010 | 424 | by smt | 
| 425 | ||
| 426 | lemma "\<forall>x y::int. x < y \<longrightarrow> (2 * x + 1) < (2 * y)" | |
| 33011 | 427 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_08"]] | 
| 33010 | 428 | by smt | 
| 429 | ||
| 430 | lemma "\<forall>x y::int. (2 * x + 1) \<noteq> (2 * y)" | |
| 33011 | 431 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_09"]] | 
| 33010 | 432 | by smt | 
| 433 | ||
| 434 | lemma "\<forall>x y::int. x + y > 2 \<or> x + y = 2 \<or> x + y < 2" | |
| 33011 | 435 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_10"]] | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 436 | by smt | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 437 | |
| 33010 | 438 | lemma "\<forall>x::int. if x > 0 then x + 1 > 0 else 1 > x" | 
| 33011 | 439 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_11"]] | 
| 33010 | 440 | by smt | 
| 441 | ||
| 442 | lemma "if (ALL x::int. x < 0 \<or> x > 0) then False else True" | |
| 33011 | 443 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_12"]] | 
| 33010 | 444 | by smt | 
| 445 | ||
| 446 | lemma "(if (ALL x::int. x < 0 \<or> x > 0) then -1 else 3) > (0::int)" | |
| 33011 | 447 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_13"]] | 
| 33010 | 448 | by smt | 
| 449 | ||
| 450 | lemma "~ (\<exists>x y z::int. 4 * x + -6 * y = (1::int))" | |
| 33011 | 451 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_14"]] | 
| 33010 | 452 | by smt | 
| 453 | ||
| 454 | lemma "\<exists>x::int. \<forall>x y. 0 < x \<and> 0 < y \<longrightarrow> (0::int) < x + y" | |
| 33011 | 455 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_15"]] | 
| 33010 | 456 | by smt | 
| 457 | ||
| 458 | lemma "\<exists>u::int. \<forall>(x::int) y::real. 0 < x \<and> 0 < y \<longrightarrow> -1 < x" | |
| 33011 | 459 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_16"]] | 
| 33010 | 460 | by smt | 
| 461 | ||
| 462 | lemma "\<exists>x::int. (\<forall>y. y \<ge> x \<longrightarrow> y > 0) \<longrightarrow> x > 0" | |
| 33011 | 463 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_17"]] | 
| 33010 | 464 | by smt | 
| 465 | ||
| 466 | lemma "\<forall>x::int. trigger [pat x] (x < a \<longrightarrow> 2 * x < 2 * a)" | |
| 33011 | 467 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_18"]] | 
| 33010 | 468 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 469 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 470 | |
| 33010 | 471 | subsection {* Non-linear arithmetic over integers and reals *}
 | 
| 472 | ||
| 473 | lemma "a > (0::int) \<Longrightarrow> a*b > 0 \<Longrightarrow> b > 0" | |
| 33011 | 474 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nlarith_01"]] | 
| 33010 | 475 |   using [[z3_proofs=false]]  -- {* Isabelle's arithmetic decision procedures
 | 
| 476 |     are too weak to automatically prove @{thm zero_less_mult_pos}. *}
 | |
| 477 | by smt | |
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 478 | |
| 33010 | 479 | lemma "(a::int) * (x + 1 + y) = a * x + a * (y + 1)" | 
| 33011 | 480 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nlarith_02"]] | 
| 33010 | 481 | by smt | 
| 482 | ||
| 483 | lemma "((x::real) * (1 + y) - x * (1 - y)) = (2 * x * y)" | |
| 33011 | 484 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nlarith_03"]] | 
| 33010 | 485 | by smt | 
| 486 | ||
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 487 | lemma | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 488 | "(U::int) + (1 + p) * (b + e) + p * d = | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 489 | U + (2 * (1 + p) * (b + e) + (1 + p) * d + d * p) - (1 + p) * (b + d + e)" | 
| 33011 | 490 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nlarith_04"]] | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 491 | by smt | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 492 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 493 | |
| 33010 | 494 | subsection {* Linear arithmetic for natural numbers *}
 | 
| 495 | ||
| 496 | lemma "2 * (x::nat) ~= 1" | |
| 33011 | 497 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_01"]] | 
| 33010 | 498 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 499 | |
| 33010 | 500 | lemma "a < 3 \<Longrightarrow> (7::nat) > 2 * a" | 
| 33011 | 501 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_02"]] | 
| 33010 | 502 | by smt | 
| 503 | ||
| 504 | lemma "let x = (1::nat) + y in x - y > 0 * x" | |
| 33011 | 505 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_03"]] | 
| 33010 | 506 | by smt | 
| 507 | ||
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 508 | lemma | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 509 | "let x = (1::nat) + y in | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 510 | let P = (if x > 0 then True else False) in | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 511 | False \<or> P = (x - 1 = y) \<or> (\<not>P \<longrightarrow> False)" | 
| 33011 | 512 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_04"]] | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 513 | by smt | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 514 | |
| 33010 | 515 | lemma "distinct [a + (1::nat), a * 2 + 3, a - a]" | 
| 33011 | 516 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_05"]] | 
| 33010 | 517 | by smt | 
| 518 | ||
| 519 | lemma "int (nat \<bar>x::int\<bar>) = \<bar>x\<bar>" | |
| 33011 | 520 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_06"]] | 
| 33010 | 521 | by smt | 
| 522 | ||
| 523 | definition prime_nat :: "nat \<Rightarrow> bool" where | |
| 524 | "prime_nat p = (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))" | |
| 525 | lemma "prime_nat (4*m + 1) \<Longrightarrow> m \<ge> (1::nat)" | |
| 33011 | 526 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_07"]] | 
| 33010 | 527 | by (smt add: prime_nat_def) | 
| 528 | ||
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 529 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 530 | section {* Bitvectors *}
 | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 531 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 532 | locale bv | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 533 | begin | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 534 | |
| 33010 | 535 | text {*
 | 
| 536 | The following examples only work for Z3, and only without proof reconstruction. | |
| 537 | *} | |
| 538 | ||
| 539 | declare [[smt_solver=z3, z3_proofs=false]] | |
| 540 | ||
| 541 | ||
| 542 | subsection {* Bitvector arithmetic *}
 | |
| 543 | ||
| 544 | lemma "(27 :: 4 word) = -5" | |
| 33011 | 545 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_01"]] | 
| 33010 | 546 | by smt | 
| 547 | ||
| 548 | lemma "(27 :: 4 word) = 11" | |
| 33011 | 549 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_02"]] | 
| 33010 | 550 | by smt | 
| 551 | ||
| 552 | lemma "23 < (27::8 word)" | |
| 33011 | 553 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_03"]] | 
| 33010 | 554 | by smt | 
| 555 | ||
| 556 | lemma "27 + 11 = (6::5 word)" | |
| 33011 | 557 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_04"]] | 
| 33010 | 558 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 559 | |
| 33010 | 560 | lemma "7 * 3 = (21::8 word)" | 
| 33011 | 561 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_05"]] | 
| 33010 | 562 | by smt | 
| 563 | lemma "11 - 27 = (-16::8 word)" | |
| 33011 | 564 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_06"]] | 
| 33010 | 565 | by smt | 
| 566 | ||
| 567 | lemma "- -11 = (11::5 word)" | |
| 33011 | 568 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_07"]] | 
| 33010 | 569 | by smt | 
| 570 | ||
| 571 | lemma "-40 + 1 = (-39::7 word)" | |
| 33011 | 572 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_08"]] | 
| 33010 | 573 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 574 | |
| 33010 | 575 | lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" | 
| 33011 | 576 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_09"]] | 
| 33010 | 577 | by smt | 
| 578 | ||
| 579 | lemma "x = (5 :: 4 word) \<Longrightarrow> 4 * x = 4" | |
| 33011 | 580 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_10"]] | 
| 33010 | 581 | by smt | 
| 582 | ||
| 583 | ||
| 584 | subsection {* Bit-level logic *}
 | |
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 585 | |
| 33010 | 586 | lemma "0b110 AND 0b101 = (0b100 :: 32 word)" | 
| 33011 | 587 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_01"]] | 
| 33010 | 588 | by smt | 
| 589 | ||
| 590 | lemma "0b110 OR 0b011 = (0b111 :: 8 word)" | |
| 33011 | 591 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_02"]] | 
| 33010 | 592 | by smt | 
| 593 | ||
| 594 | lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)" | |
| 33011 | 595 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_03"]] | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 596 | by smt | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 597 | |
| 33010 | 598 | lemma "NOT (0xF0 :: 16 word) = 0xFF0F" | 
| 33011 | 599 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_04"]] | 
| 33010 | 600 | by smt | 
| 601 | ||
| 602 | lemma "word_cat (27::4 word) (27::8 word) = (2843::12 word)" | |
| 33011 | 603 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_05"]] | 
| 33010 | 604 | by smt | 
| 605 | ||
| 606 | lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" | |
| 33011 | 607 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_06"]] | 
| 33010 | 608 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 609 | |
| 33010 | 610 | lemma "slice 1 (0b10110 :: 4 word) = (0b11 :: 2 word)" | 
| 33011 | 611 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_07"]] | 
| 33010 | 612 | by smt | 
| 613 | ||
| 614 | lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" | |
| 33011 | 615 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_08"]] | 
| 33010 | 616 | by smt | 
| 617 | ||
| 618 | lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" | |
| 33011 | 619 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_09"]] | 
| 33010 | 620 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 621 | |
| 33010 | 622 | lemma "bv_lshr 0b10011 2 = (0b100::8 word)" | 
| 33011 | 623 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_10"]] | 
| 33010 | 624 | by smt | 
| 625 | ||
| 626 | lemma "bv_ashr 0b10011 2 = (0b100::8 word)" | |
| 33011 | 627 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_11"]] | 
| 33010 | 628 | by smt | 
| 629 | ||
| 630 | lemma "word_rotr 2 0b0110 = (0b1001::4 word)" | |
| 33011 | 631 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_12"]] | 
| 33010 | 632 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 633 | |
| 33010 | 634 | lemma "word_rotl 1 0b1110 = (0b1101::4 word)" | 
| 33011 | 635 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_13"]] | 
| 33010 | 636 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 637 | |
| 33010 | 638 | lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" | 
| 33011 | 639 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_14"]] | 
| 33010 | 640 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 641 | |
| 33010 | 642 | lemma "w < 256 \<Longrightarrow> (w :: 16 word) AND 0x00FF = w" | 
| 33011 | 643 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_15"]] | 
| 33010 | 644 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 645 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 646 | end | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 647 | |
| 33010 | 648 | lemma | 
| 649 | assumes "bv2int 0 = 0" | |
| 650 | and "bv2int 1 = 1" | |
| 651 | and "bv2int 2 = 2" | |
| 652 | and "bv2int 3 = 3" | |
| 653 | and "\<forall>x::2 word. bv2int x > 0" | |
| 654 | shows "\<forall>i::int. i < 0 \<longrightarrow> (\<forall>x::2 word. bv2int x > i)" | |
| 655 | using assms | |
| 656 | using [[smt_solver=z3]] | |
| 33011 | 657 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_01"]] | 
| 33010 | 658 | by smt | 
| 659 | ||
| 660 | lemma "P (0 \<le> (a :: 4 word)) = P True" | |
| 661 | using [[smt_solver=z3, z3_proofs=false]] | |
| 33011 | 662 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_02"]] | 
| 33010 | 663 | by smt | 
| 664 | ||
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 665 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 666 | section {* Pairs *}
 | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 667 | |
| 33010 | 668 | lemma "fst (x, y) = a \<Longrightarrow> x = a" | 
| 33011 | 669 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_pair_01"]] | 
| 33010 | 670 | by smt | 
| 671 | ||
| 672 | lemma "p1 = (x, y) \<and> p2 = (y, x) \<Longrightarrow> fst p1 = snd p2" | |
| 33011 | 673 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_pair_02"]] | 
| 33010 | 674 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 675 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 676 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 677 | section {* Higher-order problems and recursion *}
 | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 678 | |
| 33010 | 679 | lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> (f (i1 := v1, i2 := v2)) i = f i" | 
| 33011 | 680 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_01"]] | 
| 33010 | 681 | by smt | 
| 682 | ||
| 683 | lemma "(f g x = (g x \<and> True)) \<or> (f g x = True) \<or> (g x = True)" | |
| 33011 | 684 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_02"]] | 
| 33010 | 685 | by smt | 
| 686 | ||
| 687 | lemma "id 3 = 3 \<and> id True = True" | |
| 33011 | 688 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_03"]] | 
| 33010 | 689 | by (smt add: id_def) | 
| 690 | ||
| 691 | lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> ((f (i1 := v1)) (i2 := v2)) i = f i" | |
| 33011 | 692 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_04"]] | 
| 33010 | 693 | by smt | 
| 694 | ||
| 695 | lemma "map (\<lambda>i::nat. i + 1) [0, 1] = [1, 2]" | |
| 33011 | 696 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_05"]] | 
| 33010 | 697 | by (smt add: map.simps) | 
| 698 | ||
| 699 | lemma "(ALL x. P x) | ~ All P" | |
| 33011 | 700 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_06"]] | 
| 33010 | 701 | by smt | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 702 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 703 | fun dec_10 :: "nat \<Rightarrow> nat" where | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 704 | "dec_10 n = (if n < 10 then n else dec_10 (n - 10))" | 
| 33010 | 705 | lemma "dec_10 (4 * dec_10 4) = 6" | 
| 33011 | 706 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_07"]] | 
| 33010 | 707 | by (smt add: dec_10.simps) | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 708 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 709 | axiomatization | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 710 | eval_dioph :: "int list \<Rightarrow> nat list \<Rightarrow> int" | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 711 | where | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 712 | eval_dioph_mod: | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 713 | "eval_dioph ks xs mod int n = eval_dioph ks (map (\<lambda>x. x mod n) xs) mod int n" | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 714 | and | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 715 | eval_dioph_div_mult: | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 716 | "eval_dioph ks (map (\<lambda>x. x div n) xs) * int n + | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 717 | eval_dioph ks (map (\<lambda>x. x mod n) xs) = eval_dioph ks xs" | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 718 | lemma | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 719 | "(eval_dioph ks xs = l) = | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 720 | (eval_dioph ks (map (\<lambda>x. x mod 2) xs) mod 2 = l mod 2 \<and> | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 721 | eval_dioph ks (map (\<lambda>x. x div 2) xs) = | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 722 | (l - eval_dioph ks (map (\<lambda>x. x mod 2) xs)) div 2)" | 
| 33011 | 723 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_08"]] | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 724 | by (smt add: eval_dioph_mod[where n=2] eval_dioph_div_mult[where n=2]) | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 725 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 726 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 727 | section {* Monomorphization examples *}
 | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 728 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 729 | definition P :: "'a \<Rightarrow> bool" where "P x = True" | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 730 | lemma poly_P: "P x \<and> (P [x] \<or> \<not>P[x])" by (simp add: P_def) | 
| 33010 | 731 | lemma "P (1::int)" | 
| 33011 | 732 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_mono_01"]] | 
| 33010 | 733 | by (smt add: poly_P) | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 734 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 735 | consts g :: "'a \<Rightarrow> nat" | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 736 | axioms | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 737 | g1: "g (Some x) = g [x]" | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 738 | g2: "g None = g []" | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 739 | g3: "g xs = length xs" | 
| 33010 | 740 | lemma "g (Some (3::int)) = g (Some True)" | 
| 33011 | 741 | using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_mono_02"]] | 
| 33010 | 742 | by (smt add: g1 g2 g3 list.size) | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 743 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 744 | end |