src/HOL/SMT_Examples/SMT_Examples.thy
changeset 36890 8e55aa1306c5
child 36891 bcd6fce5bf06
equal deleted inserted replaced
36889:6d1ecdb81ff0 36890:8e55aa1306c5
       
     1 (*  Title:      HOL/SMT/SMT_Examples.thy
       
     2     Author:     Sascha Boehme, TU Muenchen
       
     3 *)
       
     4 
       
     5 header {* Examples for the 'smt' tactic. *}
       
     6 
       
     7 theory SMT_Examples
       
     8 imports SMT
       
     9 begin
       
    10 
       
    11 declare [[smt_solver=z3, z3_proofs=true]]
       
    12 
       
    13 declare [[smt_certificates="$ISABELLE_SMT/Examples/SMT_Examples.certs"]]
       
    14 
       
    15 text {*
       
    16 To avoid re-generation of certificates,
       
    17 the following option is set to "false":
       
    18 *}
       
    19 
       
    20 declare [[smt_fixed=true]]
       
    21 
       
    22 
       
    23 
       
    24 section {* Propositional and first-order logic *}
       
    25 
       
    26 lemma "True" by smt
       
    27 
       
    28 lemma "p \<or> \<not>p" by smt
       
    29 
       
    30 lemma "(p \<and> True) = p" by smt
       
    31 
       
    32 lemma "(p \<or> q) \<and> \<not>p \<Longrightarrow> q" by smt
       
    33 
       
    34 lemma "(a \<and> b) \<or> (c \<and> d) \<Longrightarrow> (a \<and> b) \<or> (c \<and> d)"
       
    35   using [[z3_proofs=false]] (* no Z3 proof *)
       
    36   by smt
       
    37 
       
    38 lemma "(p1 \<and> p2) \<or> p3 \<longrightarrow> (p1 \<longrightarrow> (p3 \<and> p2) \<or> (p1 \<and> p3)) \<or> p1" by smt
       
    39 
       
    40 lemma "P=P=P=P=P=P=P=P=P=P" by smt
       
    41 
       
    42 lemma 
       
    43   assumes "a | b | c | d"
       
    44       and "e | f | (a & d)"
       
    45       and "~(a | (c & ~c)) | b"
       
    46       and "~(b & (x | ~x)) | c"
       
    47       and "~(d | False) | c"
       
    48       and "~(c | (~p & (p | (q & ~q))))"
       
    49   shows False
       
    50   using assms by smt
       
    51 
       
    52 axiomatization symm_f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
       
    53   symm_f: "symm_f x y = symm_f y x"
       
    54 lemma "a = a \<and> symm_f a b = symm_f b a" by (smt symm_f)
       
    55 
       
    56 (* 
       
    57 Taken from ~~/src/HOL/ex/SAT_Examples.thy.
       
    58 Translated from TPTP problem library: PUZ015-2.006.dimacs
       
    59 *)
       
    60 lemma 
       
    61   assumes "~x0"
       
    62   and "~x30"
       
    63   and "~x29"
       
    64   and "~x59"
       
    65   and "x1 | x31 | x0"
       
    66   and "x2 | x32 | x1"
       
    67   and "x3 | x33 | x2"
       
    68   and "x4 | x34 | x3"
       
    69   and "x35 | x4"
       
    70   and "x5 | x36 | x30"
       
    71   and "x6 | x37 | x5 | x31"
       
    72   and "x7 | x38 | x6 | x32"
       
    73   and "x8 | x39 | x7 | x33"
       
    74   and "x9 | x40 | x8 | x34"
       
    75   and "x41 | x9 | x35"
       
    76   and "x10 | x42 | x36"
       
    77   and "x11 | x43 | x10 | x37"
       
    78   and "x12 | x44 | x11 | x38"
       
    79   and "x13 | x45 | x12 | x39"
       
    80   and "x14 | x46 | x13 | x40"
       
    81   and "x47 | x14 | x41"
       
    82   and "x15 | x48 | x42"
       
    83   and "x16 | x49 | x15 | x43"
       
    84   and "x17 | x50 | x16 | x44"
       
    85   and "x18 | x51 | x17 | x45"
       
    86   and "x19 | x52 | x18 | x46"
       
    87   and "x53 | x19 | x47"
       
    88   and "x20 | x54 | x48"
       
    89   and "x21 | x55 | x20 | x49"
       
    90   and "x22 | x56 | x21 | x50"
       
    91   and "x23 | x57 | x22 | x51"
       
    92   and "x24 | x58 | x23 | x52"
       
    93   and "x59 | x24 | x53"
       
    94   and "x25 | x54"
       
    95   and "x26 | x25 | x55"
       
    96   and "x27 | x26 | x56"
       
    97   and "x28 | x27 | x57"
       
    98   and "x29 | x28 | x58"
       
    99   and "~x1 | ~x31"
       
   100   and "~x1 | ~x0"
       
   101   and "~x31 | ~x0"
       
   102   and "~x2 | ~x32"
       
   103   and "~x2 | ~x1"
       
   104   and "~x32 | ~x1"
       
   105   and "~x3 | ~x33"
       
   106   and "~x3 | ~x2"
       
   107   and "~x33 | ~x2"
       
   108   and "~x4 | ~x34"
       
   109   and "~x4 | ~x3"
       
   110   and "~x34 | ~x3"
       
   111   and "~x35 | ~x4"
       
   112   and "~x5 | ~x36"
       
   113   and "~x5 | ~x30"
       
   114   and "~x36 | ~x30"
       
   115   and "~x6 | ~x37"
       
   116   and "~x6 | ~x5"
       
   117   and "~x6 | ~x31"
       
   118   and "~x37 | ~x5"
       
   119   and "~x37 | ~x31"
       
   120   and "~x5 | ~x31"
       
   121   and "~x7 | ~x38"
       
   122   and "~x7 | ~x6"
       
   123   and "~x7 | ~x32"
       
   124   and "~x38 | ~x6"
       
   125   and "~x38 | ~x32"
       
   126   and "~x6 | ~x32"
       
   127   and "~x8 | ~x39"
       
   128   and "~x8 | ~x7"
       
   129   and "~x8 | ~x33"
       
   130   and "~x39 | ~x7"
       
   131   and "~x39 | ~x33"
       
   132   and "~x7 | ~x33"
       
   133   and "~x9 | ~x40"
       
   134   and "~x9 | ~x8"
       
   135   and "~x9 | ~x34"
       
   136   and "~x40 | ~x8"
       
   137   and "~x40 | ~x34"
       
   138   and "~x8 | ~x34"
       
   139   and "~x41 | ~x9"
       
   140   and "~x41 | ~x35"
       
   141   and "~x9 | ~x35"
       
   142   and "~x10 | ~x42"
       
   143   and "~x10 | ~x36"
       
   144   and "~x42 | ~x36"
       
   145   and "~x11 | ~x43"
       
   146   and "~x11 | ~x10"
       
   147   and "~x11 | ~x37"
       
   148   and "~x43 | ~x10"
       
   149   and "~x43 | ~x37"
       
   150   and "~x10 | ~x37"
       
   151   and "~x12 | ~x44"
       
   152   and "~x12 | ~x11"
       
   153   and "~x12 | ~x38"
       
   154   and "~x44 | ~x11"
       
   155   and "~x44 | ~x38"
       
   156   and "~x11 | ~x38"
       
   157   and "~x13 | ~x45"
       
   158   and "~x13 | ~x12"
       
   159   and "~x13 | ~x39"
       
   160   and "~x45 | ~x12"
       
   161   and "~x45 | ~x39"
       
   162   and "~x12 | ~x39"
       
   163   and "~x14 | ~x46"
       
   164   and "~x14 | ~x13"
       
   165   and "~x14 | ~x40"
       
   166   and "~x46 | ~x13"
       
   167   and "~x46 | ~x40"
       
   168   and "~x13 | ~x40"
       
   169   and "~x47 | ~x14"
       
   170   and "~x47 | ~x41"
       
   171   and "~x14 | ~x41"
       
   172   and "~x15 | ~x48"
       
   173   and "~x15 | ~x42"
       
   174   and "~x48 | ~x42"
       
   175   and "~x16 | ~x49"
       
   176   and "~x16 | ~x15"
       
   177   and "~x16 | ~x43"
       
   178   and "~x49 | ~x15"
       
   179   and "~x49 | ~x43"
       
   180   and "~x15 | ~x43"
       
   181   and "~x17 | ~x50"
       
   182   and "~x17 | ~x16"
       
   183   and "~x17 | ~x44"
       
   184   and "~x50 | ~x16"
       
   185   and "~x50 | ~x44"
       
   186   and "~x16 | ~x44"
       
   187   and "~x18 | ~x51"
       
   188   and "~x18 | ~x17"
       
   189   and "~x18 | ~x45"
       
   190   and "~x51 | ~x17"
       
   191   and "~x51 | ~x45"
       
   192   and "~x17 | ~x45"
       
   193   and "~x19 | ~x52"
       
   194   and "~x19 | ~x18"
       
   195   and "~x19 | ~x46"
       
   196   and "~x52 | ~x18"
       
   197   and "~x52 | ~x46"
       
   198   and "~x18 | ~x46"
       
   199   and "~x53 | ~x19"
       
   200   and "~x53 | ~x47"
       
   201   and "~x19 | ~x47"
       
   202   and "~x20 | ~x54"
       
   203   and "~x20 | ~x48"
       
   204   and "~x54 | ~x48"
       
   205   and "~x21 | ~x55"
       
   206   and "~x21 | ~x20"
       
   207   and "~x21 | ~x49"
       
   208   and "~x55 | ~x20"
       
   209   and "~x55 | ~x49"
       
   210   and "~x20 | ~x49"
       
   211   and "~x22 | ~x56"
       
   212   and "~x22 | ~x21"
       
   213   and "~x22 | ~x50"
       
   214   and "~x56 | ~x21"
       
   215   and "~x56 | ~x50"
       
   216   and "~x21 | ~x50"
       
   217   and "~x23 | ~x57"
       
   218   and "~x23 | ~x22"
       
   219   and "~x23 | ~x51"
       
   220   and "~x57 | ~x22"
       
   221   and "~x57 | ~x51"
       
   222   and "~x22 | ~x51"
       
   223   and "~x24 | ~x58"
       
   224   and "~x24 | ~x23"
       
   225   and "~x24 | ~x52"
       
   226   and "~x58 | ~x23"
       
   227   and "~x58 | ~x52"
       
   228   and "~x23 | ~x52"
       
   229   and "~x59 | ~x24"
       
   230   and "~x59 | ~x53"
       
   231   and "~x24 | ~x53"
       
   232   and "~x25 | ~x54"
       
   233   and "~x26 | ~x25"
       
   234   and "~x26 | ~x55"
       
   235   and "~x25 | ~x55"
       
   236   and "~x27 | ~x26"
       
   237   and "~x27 | ~x56"
       
   238   and "~x26 | ~x56"
       
   239   and "~x28 | ~x27"
       
   240   and "~x28 | ~x57"
       
   241   and "~x27 | ~x57"
       
   242   and "~x29 | ~x28"
       
   243   and "~x29 | ~x58"
       
   244   and "~x28 | ~x58"
       
   245   shows False
       
   246   using assms by smt
       
   247 
       
   248 lemma "\<forall>x::int. P x \<longrightarrow> (\<forall>y::int. P x \<or> P y)"
       
   249   by smt
       
   250 
       
   251 lemma 
       
   252   assumes "(\<forall>x y. P x y = x)"
       
   253   shows "(\<exists>y. P x y) = P x c"
       
   254   using assms by smt
       
   255 
       
   256 lemma 
       
   257   assumes "(\<forall>x y. P x y = x)"
       
   258   and "(\<forall>x. \<exists>y. P x y) = (\<forall>x. P x c)"
       
   259   shows "(EX y. P x y) = P x c"
       
   260   using assms by smt
       
   261 
       
   262 lemma
       
   263   assumes "if P x then \<not>(\<exists>y. P y) else (\<forall>y. \<not>P y)"
       
   264   shows "P x \<longrightarrow> P y"
       
   265   using assms by smt
       
   266 
       
   267 
       
   268 section {* Arithmetic *}
       
   269 
       
   270 subsection {* Linear arithmetic over integers and reals *}
       
   271 
       
   272 lemma "(3::int) = 3" by smt
       
   273 
       
   274 lemma "(3::real) = 3" by smt
       
   275 
       
   276 lemma "(3 :: int) + 1 = 4" by smt
       
   277 
       
   278 lemma "x + (y + z) = y + (z + (x::int))" by smt
       
   279 
       
   280 lemma "max (3::int) 8 > 5" by smt
       
   281 
       
   282 lemma "abs (x :: real) + abs y \<ge> abs (x + y)" by smt
       
   283 
       
   284 lemma "P ((2::int) < 3) = P True" by smt
       
   285 
       
   286 lemma "x + 3 \<ge> 4 \<or> x < (1::int)" by smt
       
   287 
       
   288 lemma
       
   289   assumes "x \<ge> (3::int)" and "y = x + 4"
       
   290   shows "y - x > 0" 
       
   291   using assms by smt
       
   292 
       
   293 lemma "let x = (2 :: int) in x + x \<noteq> 5" by smt
       
   294 
       
   295 lemma
       
   296   fixes x :: real
       
   297   assumes "3 * x + 7 * a < 4" and "3 < 2 * x"
       
   298   shows "a < 0"
       
   299   using assms by smt
       
   300 
       
   301 lemma "(0 \<le> y + -1 * x \<or> \<not> 0 \<le> x \<or> 0 \<le> (x::int)) = (\<not> False)" by smt
       
   302 
       
   303 lemma "distinct [x < (3::int), 3 \<le> x]" by smt
       
   304 
       
   305 lemma
       
   306   assumes "a > (0::int)"
       
   307   shows "distinct [a, a * 2, a - a]"
       
   308   using assms by smt
       
   309 
       
   310 lemma "
       
   311   (n < m & m < n') | (n < m & m = n') | (n < n' & n' < m) |
       
   312   (n = n' & n' < m) | (n = m & m < n') |
       
   313   (n' < m & m < n) | (n' < m & m = n) |
       
   314   (n' < n & n < m) | (n' = n & n < m) | (n' = m & m < n) |
       
   315   (m < n & n < n') | (m < n & n' = n) | (m < n' & n' < n) |
       
   316   (m = n & n < n') | (m = n' & n' < n) |
       
   317   (n' = m & m = (n::int))"
       
   318   by smt
       
   319 
       
   320 text{* 
       
   321 The following example was taken from HOL/ex/PresburgerEx.thy, where it says:
       
   322 
       
   323   This following theorem proves that all solutions to the
       
   324   recurrence relation $x_{i+2} = |x_{i+1}| - x_i$ are periodic with
       
   325   period 9.  The example was brought to our attention by John
       
   326   Harrison. It does does not require Presburger arithmetic but merely
       
   327   quantifier-free linear arithmetic and holds for the rationals as well.
       
   328 
       
   329   Warning: it takes (in 2006) over 4.2 minutes! 
       
   330 
       
   331 There, it is proved by "arith". SMT is able to prove this within a fraction
       
   332 of one second. With proof reconstruction, it takes about 13 seconds on a Core2
       
   333 processor.
       
   334 *}
       
   335 
       
   336 lemma "\<lbrakk> x3 = abs x2 - x1; x4 = abs x3 - x2; x5 = abs x4 - x3;
       
   337          x6 = abs x5 - x4; x7 = abs x6 - x5; x8 = abs x7 - x6;
       
   338          x9 = abs x8 - x7; x10 = abs x9 - x8; x11 = abs x10 - x9 \<rbrakk>
       
   339  \<Longrightarrow> x1 = x10 & x2 = (x11::int)"
       
   340   by smt
       
   341 
       
   342 
       
   343 lemma "let P = 2 * x + 1 > x + (x::real) in P \<or> False \<or> P" by smt
       
   344 
       
   345 lemma "x + (let y = x mod 2 in 2 * y + 1) \<ge> x + (1::int)" by smt
       
   346 
       
   347 lemma "x + (let y = x mod 2 in y + y) < x + (3::int)" by smt
       
   348 
       
   349 lemma
       
   350   assumes "x \<noteq> (0::real)"
       
   351   shows "x + x \<noteq> (let P = (abs x > 1) in if P \<or> \<not>P then 4 else 2) * x"
       
   352   using assms by smt
       
   353 
       
   354 lemma                                                                         
       
   355   assumes "(n + m) mod 2 = 0" and "n mod 4 = 3"                               
       
   356   shows "n mod 2 = 1 & m mod 2 = (1::int)"      
       
   357   using assms by smt
       
   358 
       
   359 
       
   360 subsection {* Linear arithmetic with quantifiers *}
       
   361 
       
   362 lemma "~ (\<exists>x::int. False)" by smt
       
   363 
       
   364 lemma "~ (\<exists>x::real. False)" by smt
       
   365 
       
   366 lemma "\<exists>x::int. 0 < x"
       
   367   using [[z3_proofs=false]] (* no Z3 proof *)
       
   368   by smt
       
   369 
       
   370 lemma "\<exists>x::real. 0 < x"
       
   371   using [[z3_proofs=false]] (* no Z3 proof *)
       
   372   by smt
       
   373 
       
   374 lemma "\<forall>x::int. \<exists>y. y > x"
       
   375   using [[z3_proofs=false]] (* no Z3 proof *)
       
   376   by smt
       
   377 
       
   378 lemma "\<forall>x y::int. (x = 0 \<and> y = 1) \<longrightarrow> x \<noteq> y" by smt
       
   379 
       
   380 lemma "\<exists>x::int. \<forall>y. x < y \<longrightarrow> y < 0 \<or> y >= 0" by smt
       
   381 
       
   382 lemma "\<forall>x y::int. x < y \<longrightarrow> (2 * x + 1) < (2 * y)"  by smt
       
   383 
       
   384 lemma "\<forall>x y::int. (2 * x + 1) \<noteq> (2 * y)" by smt
       
   385 
       
   386 lemma "\<forall>x y::int. x + y > 2 \<or> x + y = 2 \<or> x + y < 2" by smt
       
   387 
       
   388 lemma "\<forall>x::int. if x > 0 then x + 1 > 0 else 1 > x" by smt
       
   389 
       
   390 lemma "if (ALL x::int. x < 0 \<or> x > 0) then False else True" by smt
       
   391 
       
   392 lemma "(if (ALL x::int. x < 0 \<or> x > 0) then -1 else 3) > (0::int)" by smt
       
   393 
       
   394 lemma "~ (\<exists>x y z::int. 4 * x + -6 * y = (1::int))" by smt
       
   395 
       
   396 lemma "\<exists>x::int. \<forall>x y. 0 < x \<and> 0 < y \<longrightarrow> (0::int) < x + y" by smt
       
   397 
       
   398 lemma "\<exists>u::int. \<forall>(x::int) y::real. 0 < x \<and> 0 < y \<longrightarrow> -1 < x" by smt
       
   399 
       
   400 lemma "\<exists>x::int. (\<forall>y. y \<ge> x \<longrightarrow> y > 0) \<longrightarrow> x > 0" by smt
       
   401 
       
   402 lemma "\<forall>x::int. trigger [pat x] (x < a \<longrightarrow> 2 * x < 2 * a)" by smt
       
   403 
       
   404 
       
   405 subsection {* Non-linear arithmetic over integers and reals *}
       
   406 
       
   407 lemma "a > (0::int) \<Longrightarrow> a*b > 0 \<Longrightarrow> b > 0"
       
   408   using [[z3_proofs=false]]  -- {* Isabelle's arithmetic decision procedures
       
   409     are too weak to automatically prove @{thm zero_less_mult_pos}. *}
       
   410   by smt
       
   411 
       
   412 lemma  "(a::int) * (x + 1 + y) = a * x + a * (y + 1)" by smt
       
   413 
       
   414 lemma "((x::real) * (1 + y) - x * (1 - y)) = (2 * x * y)" by smt
       
   415 
       
   416 lemma
       
   417   "(U::int) + (1 + p) * (b + e) + p * d =
       
   418    U + (2 * (1 + p) * (b + e) + (1 + p) * d + d * p) - (1 + p) * (b + d + e)"
       
   419   by smt
       
   420 
       
   421 
       
   422 subsection {* Linear arithmetic for natural numbers *}
       
   423 
       
   424 lemma "2 * (x::nat) ~= 1" by smt
       
   425 
       
   426 lemma "a < 3 \<Longrightarrow> (7::nat) > 2 * a" by smt
       
   427 
       
   428 lemma "let x = (1::nat) + y in x - y > 0 * x" by smt
       
   429 
       
   430 lemma
       
   431   "let x = (1::nat) + y in
       
   432    let P = (if x > 0 then True else False) in
       
   433    False \<or> P = (x - 1 = y) \<or> (\<not>P \<longrightarrow> False)"
       
   434   by smt
       
   435 
       
   436 lemma "distinct [a + (1::nat), a * 2 + 3, a - a]" by smt
       
   437 
       
   438 lemma "int (nat \<bar>x::int\<bar>) = \<bar>x\<bar>" by smt
       
   439 
       
   440 definition prime_nat :: "nat \<Rightarrow> bool" where
       
   441   "prime_nat p = (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))"
       
   442 lemma "prime_nat (4*m + 1) \<Longrightarrow> m \<ge> (1::nat)" by (smt prime_nat_def)
       
   443 
       
   444 
       
   445 section {* Bitvectors *}
       
   446 
       
   447 locale z3_bv_test
       
   448 begin
       
   449 
       
   450 text {*
       
   451 The following examples only work for Z3, and only without proof reconstruction.
       
   452 *}
       
   453 
       
   454 declare [[smt_solver=z3, z3_proofs=false]]
       
   455 
       
   456 
       
   457 subsection {* Bitvector arithmetic *}
       
   458 
       
   459 lemma "(27 :: 4 word) = -5" by smt
       
   460 
       
   461 lemma "(27 :: 4 word) = 11" by smt
       
   462 
       
   463 lemma "23 < (27::8 word)" by smt
       
   464 
       
   465 lemma "27 + 11 = (6::5 word)" by smt
       
   466 
       
   467 lemma "7 * 3 = (21::8 word)" by smt
       
   468 
       
   469 lemma "11 - 27 = (-16::8 word)" by smt
       
   470 
       
   471 lemma "- -11 = (11::5 word)" by smt
       
   472 
       
   473 lemma "-40 + 1 = (-39::7 word)" by smt
       
   474 
       
   475 lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by smt
       
   476 
       
   477 lemma "x = (5 :: 4 word) \<Longrightarrow> 4 * x = 4" by smt
       
   478 
       
   479 
       
   480 subsection {* Bit-level logic *}
       
   481 
       
   482 lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by smt
       
   483 
       
   484 lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by smt
       
   485 
       
   486 lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)" by smt
       
   487 
       
   488 lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by smt
       
   489 
       
   490 lemma "word_cat (27::4 word) (27::8 word) = (2843::12 word)" by smt
       
   491 
       
   492 lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)"
       
   493   by smt
       
   494 
       
   495 lemma "slice 1 (0b10110 :: 4 word) = (0b11 :: 2 word)" by smt
       
   496 
       
   497 lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by smt
       
   498 
       
   499 lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by smt
       
   500 
       
   501 lemma "bv_lshr 0b10011 2 = (0b100::8 word)" by smt
       
   502 
       
   503 lemma "bv_ashr 0b10011 2 = (0b100::8 word)" by smt
       
   504 
       
   505 lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by smt
       
   506 
       
   507 lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by smt
       
   508 
       
   509 lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" by smt
       
   510 
       
   511 lemma "w < 256 \<Longrightarrow> (w :: 16 word) AND 0x00FF = w" by smt
       
   512 
       
   513 end
       
   514 
       
   515 lemma
       
   516   assumes "bv2int 0 = 0"
       
   517       and "bv2int 1 = 1"
       
   518       and "bv2int 2 = 2"
       
   519       and "bv2int 3 = 3"
       
   520       and "\<forall>x::2 word. bv2int x > 0"
       
   521   shows "\<forall>i::int. i < 0 \<longrightarrow> (\<forall>x::2 word. bv2int x > i)"
       
   522   using assms 
       
   523   using [[smt_solver=z3]]
       
   524   by smt
       
   525 
       
   526 lemma "P (0 \<le> (a :: 4 word)) = P True"
       
   527   using [[smt_solver=z3, z3_proofs=false]]
       
   528   by smt
       
   529 
       
   530 
       
   531 section {* Pairs *}
       
   532 
       
   533 lemma "fst (x, y) = a \<Longrightarrow> x = a" by smt
       
   534 
       
   535 lemma "p1 = (x, y) \<and> p2 = (y, x) \<Longrightarrow> fst p1 = snd p2" by smt
       
   536 
       
   537 
       
   538 section {* Higher-order problems and recursion *}
       
   539 
       
   540 lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> (f (i1 := v1, i2 := v2)) i = f i" by smt
       
   541 
       
   542 lemma "(f g (x::'a::type) = (g x \<and> True)) \<or> (f g x = True) \<or> (g x = True)"
       
   543   by smt
       
   544 
       
   545 lemma "id 3 = 3 \<and> id True = True" by (smt id_def)
       
   546 
       
   547 lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> ((f (i1 := v1)) (i2 := v2)) i = f i" by smt
       
   548 
       
   549 lemma "map (\<lambda>i::nat. i + 1) [0, 1] = [1, 2]" by (smt map.simps)
       
   550 
       
   551 lemma "(ALL x. P x) | ~ All P" by smt
       
   552 
       
   553 fun dec_10 :: "nat \<Rightarrow> nat" where
       
   554   "dec_10 n = (if n < 10 then n else dec_10 (n - 10))"
       
   555 lemma "dec_10 (4 * dec_10 4) = 6" by (smt dec_10.simps)
       
   556 
       
   557 axiomatization
       
   558   eval_dioph :: "int list \<Rightarrow> nat list \<Rightarrow> int"
       
   559   where
       
   560   eval_dioph_mod:
       
   561   "eval_dioph ks xs mod int n = eval_dioph ks (map (\<lambda>x. x mod n) xs) mod int n"
       
   562   and
       
   563   eval_dioph_div_mult:
       
   564   "eval_dioph ks (map (\<lambda>x. x div n) xs) * int n +
       
   565    eval_dioph ks (map (\<lambda>x. x mod n) xs) = eval_dioph ks xs"
       
   566 lemma
       
   567   "(eval_dioph ks xs = l) =
       
   568    (eval_dioph ks (map (\<lambda>x. x mod 2) xs) mod 2 = l mod 2 \<and>
       
   569     eval_dioph ks (map (\<lambda>x. x div 2) xs) =
       
   570       (l - eval_dioph ks (map (\<lambda>x. x mod 2) xs)) div 2)"
       
   571   by (smt eval_dioph_mod[where n=2] eval_dioph_div_mult[where n=2])
       
   572 
       
   573 
       
   574 section {* Monomorphization examples *}
       
   575 
       
   576 definition P :: "'a \<Rightarrow> bool" where "P x = True"
       
   577 lemma poly_P: "P x \<and> (P [x] \<or> \<not>P[x])" by (simp add: P_def)
       
   578 lemma "P (1::int)" by (smt poly_P)
       
   579 
       
   580 consts g :: "'a \<Rightarrow> nat"
       
   581 axioms
       
   582   g1: "g (Some x) = g [x]"
       
   583   g2: "g None = g []"
       
   584   g3: "g xs = length xs"
       
   585 lemma "g (Some (3::int)) = g (Some True)" by (smt g1 g2 g3 list.size)
       
   586 
       
   587 end