src/HOL/SMT/Examples/SMT_Examples.thy
changeset 32618 42865636d006
child 32619 02f45a09a9f2
equal deleted inserted replaced
32608:c0056c2c1d17 32618:42865636d006
       
     1 (*  Title:       SMT_Examples.thy
       
     2     Author:      Sascha Böhme, 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=false]]
       
    12 declare [[smt_trace=true]] (*FIXME*)
       
    13 
       
    14 
       
    15 section {* Propositional and first-order logic *}
       
    16 
       
    17 lemma "True" by smt
       
    18 lemma "p \<or> \<not>p" by smt
       
    19 lemma "(p \<and> True) = p" by smt
       
    20 lemma "(p \<or> q) \<and> \<not>p \<Longrightarrow> q" by smt
       
    21 lemma "(a \<and> b) \<or> (c \<and> d) \<Longrightarrow> (a \<and> b) \<or> (c \<and> d)" by smt
       
    22 lemma "P=P=P=P=P=P=P=P=P=P" by smt
       
    23 
       
    24 axiomatization symm_f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
       
    25   symm_f: "symm_f x y = symm_f y x"
       
    26 lemma "a = a \<and> symm_f a b = symm_f b a" by (smt add: symm_f)
       
    27 
       
    28 
       
    29 section {* Linear arithmetic *}
       
    30 
       
    31 lemma "(3::int) = 3" by smt
       
    32 lemma "(3::real) = 3" by smt
       
    33 lemma "(3 :: int) + 1 = 4" by smt
       
    34 lemma "max (3::int) 8 > 5" by smt
       
    35 lemma "abs (x :: real) + abs y \<ge> abs (x + y)" by smt
       
    36 lemma "let x = (2 :: int) in x + x \<noteq> 5" by smt
       
    37 
       
    38 text{* 
       
    39 The following example was taken from HOL/ex/PresburgerEx.thy, where it says:
       
    40 
       
    41   This following theorem proves that all solutions to the
       
    42   recurrence relation $x_{i+2} = |x_{i+1}| - x_i$ are periodic with
       
    43   period 9.  The example was brought to our attention by John
       
    44   Harrison. It does does not require Presburger arithmetic but merely
       
    45   quantifier-free linear arithmetic and holds for the rationals as well.
       
    46 
       
    47   Warning: it takes (in 2006) over 4.2 minutes! 
       
    48 
       
    49 There, it is proved by "arith". SMT is able to prove this within a fraction
       
    50 of one second.
       
    51 *}
       
    52 
       
    53 lemma "\<lbrakk> x3 = abs x2 - x1; x4 = abs x3 - x2; x5 = abs x4 - x3;
       
    54          x6 = abs x5 - x4; x7 = abs x6 - x5; x8 = abs x7 - x6;
       
    55          x9 = abs x8 - x7; x10 = abs x9 - x8; x11 = abs x10 - x9 \<rbrakk>
       
    56  \<Longrightarrow> x1 = x10 & x2 = (x11::int)"
       
    57   by smt
       
    58 
       
    59 lemma "\<exists>x::int. 0 < x" by smt
       
    60 lemma "\<exists>x::real. 0 < x" by smt
       
    61 lemma "\<forall>x y::int. x < y \<longrightarrow> (2 * x + 1) < (2 * y)" by smt
       
    62 lemma "\<forall>x y::int. (2 * x + 1) \<noteq> (2 * y)" by smt
       
    63 lemma "~ (\<exists>x y z::int. 4 * x + -6 * y = (1::int))" by smt
       
    64 lemma "~ (\<exists>x::int. False)" by smt
       
    65 
       
    66 
       
    67 section {* Non-linear arithmetic *}
       
    68 
       
    69 lemma "((x::int) * (1 + y) - x * (1 - y)) = (2 * x * y)" by smt
       
    70 lemma
       
    71   "(U::int) + (1 + p) * (b + e) + p * d =
       
    72    U + (2 * (1 + p) * (b + e) + (1 + p) * d + d * p) - (1 + p) * (b + d + e)"
       
    73   by smt
       
    74 
       
    75 
       
    76 section {* Linear arithmetic for natural numbers *}
       
    77 
       
    78 lemma "a < 3 \<Longrightarrow> (7::nat) > 2 * a" by smt
       
    79 lemma "let x = (1::nat) + y in x - y > 0 * x" by smt
       
    80 lemma
       
    81   "let x = (1::nat) + y in
       
    82    let P = (if x > 0 then True else False) in
       
    83    False \<or> P = (x - 1 = y) \<or> (\<not>P \<longrightarrow> False)"
       
    84   by smt
       
    85 
       
    86 
       
    87 section {* Bitvectors *}
       
    88 
       
    89 locale bv
       
    90 begin
       
    91 
       
    92 declare [[smt_solver=z3]]
       
    93 
       
    94 lemma "(27 :: 4 word) = -5" by smt
       
    95 lemma "(27 :: 4 word) = 11" by smt
       
    96 lemma "23 < (27::8 word)" by smt
       
    97 lemma "27 + 11 = (6::5 word)" by smt
       
    98 lemma "7 * 3 = (21::8 word)" by smt
       
    99 lemma "11 - 27 = (-16::8 word)" by smt
       
   100 lemma "- -11 = (11::5 word)" by smt
       
   101 lemma "-40 + 1 = (-39::7 word)" by smt
       
   102 lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by smt
       
   103 
       
   104 lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by smt
       
   105 lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by smt
       
   106 lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)" by smt
       
   107 lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by smt
       
   108 
       
   109 lemma "word_cat (27::4 word) (27::8 word) = (2843::12 word)" by smt
       
   110 lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" 
       
   111   by smt
       
   112 
       
   113 lemma "slice 1 (0b10110 :: 4 word) = (0b11 :: 2 word)" by smt
       
   114 
       
   115 lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by smt
       
   116 lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by smt
       
   117 
       
   118 lemma "bv_lshr 0b10011 2 = (0b100::8 word)" by smt
       
   119 lemma "bv_ashr 0b10011 2 = (0b100::8 word)" by smt
       
   120 
       
   121 lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by smt
       
   122 lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by smt
       
   123 
       
   124 lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" by smt
       
   125 
       
   126 lemma "w < 256 \<Longrightarrow> (w :: 16 word) AND 0x00FF = w" by smt
       
   127 
       
   128 end
       
   129 
       
   130 
       
   131 section {* Pairs *}
       
   132 
       
   133 lemma "fst (x, y) = a \<Longrightarrow> x = a" by smt
       
   134 lemma "p1 = (x, y) \<and> p2 = (y, x) \<Longrightarrow> fst p1 = snd p2" by smt
       
   135 
       
   136 
       
   137 section {* Higher-order problems and recursion *}
       
   138 
       
   139 lemma "(f g x = (g x \<and> True)) \<or> (f g x = True) \<or> (g x = True)" by smt
       
   140 lemma "P ((2::int) < 3) = P True" by smt
       
   141 lemma "P ((2::int) < 3) = (P True :: bool)" by smt
       
   142 lemma "P (0 \<le> (a :: 4 word)) = P True" using [[smt_solver=z3]] by smt
       
   143 lemma "id 3 = 3 \<and> id True = True" by (smt add: id_def)
       
   144 lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> ((f (i1 := v1)) (i2 := v2)) i = f i" by smt
       
   145 lemma "map (\<lambda>i::nat. i + 1) [0, 1] = [1, 2]" by (smt add: map.simps)
       
   146 lemma "(ALL x. P x) | ~ All P" by smt
       
   147 
       
   148 fun dec_10 :: "nat \<Rightarrow> nat" where
       
   149   "dec_10 n = (if n < 10 then n else dec_10 (n - 10))"
       
   150 lemma "dec_10 (4 * dec_10 4) = 6" by (smt add: dec_10.simps)
       
   151 
       
   152 axiomatization
       
   153   eval_dioph :: "int list \<Rightarrow> nat list \<Rightarrow> int"
       
   154   where
       
   155   eval_dioph_mod:
       
   156   "eval_dioph ks xs mod int n = eval_dioph ks (map (\<lambda>x. x mod n) xs) mod int n"
       
   157   and
       
   158   eval_dioph_div_mult:
       
   159   "eval_dioph ks (map (\<lambda>x. x div n) xs) * int n +
       
   160    eval_dioph ks (map (\<lambda>x. x mod n) xs) = eval_dioph ks xs"
       
   161 lemma
       
   162   "(eval_dioph ks xs = l) =
       
   163    (eval_dioph ks (map (\<lambda>x. x mod 2) xs) mod 2 = l mod 2 \<and>
       
   164     eval_dioph ks (map (\<lambda>x. x div 2) xs) =
       
   165       (l - eval_dioph ks (map (\<lambda>x. x mod 2) xs)) div 2)"
       
   166   by (smt add: eval_dioph_mod[where n=2] eval_dioph_div_mult[where n=2])
       
   167 
       
   168 
       
   169 section {* Monomorphization examples *}
       
   170 
       
   171 definition P :: "'a \<Rightarrow> bool" where "P x = True"
       
   172 lemma poly_P: "P x \<and> (P [x] \<or> \<not>P[x])" by (simp add: P_def)
       
   173 lemma "P (1::int)" by (smt add: poly_P)
       
   174 
       
   175 consts g :: "'a \<Rightarrow> nat"
       
   176 axioms
       
   177   g1: "g (Some x) = g [x]"
       
   178   g2: "g None = g []"
       
   179   g3: "g xs = length xs"
       
   180 lemma "g (Some (3::int)) = g (Some True)" by (smt add: g1 g2 g3 list.size)
       
   181 
       
   182 end