src/HOL/SMT/Examples/SMT_Examples.thy
author boehmes
Mon, 21 Sep 2009 08:34:56 +0200
changeset 32619 02f45a09a9f2
parent 32618 42865636d006
child 32622 8ed38c7bd21a
permissions -rw-r--r--
tuned author
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     1
(*  Title:       SMT_Examples.thy
32619
02f45a09a9f2 tuned author
boehmes
parents: 32618
diff changeset
     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
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     8
imports "../SMT"
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
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    11
declare [[smt_solver=z3, z3_proofs=false]]
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    12
declare [[smt_trace=true]] (*FIXME*)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    13
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    14
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    15
section {* Propositional and first-order logic *}
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    16
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    17
lemma "True" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    18
lemma "p \<or> \<not>p" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    19
lemma "(p \<and> True) = p" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    20
lemma "(p \<or> q) \<and> \<not>p \<Longrightarrow> q" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    21
lemma "(a \<and> b) \<or> (c \<and> d) \<Longrightarrow> (a \<and> b) \<or> (c \<and> d)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    22
lemma "P=P=P=P=P=P=P=P=P=P" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    23
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    24
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
    25
  symm_f: "symm_f x y = symm_f y x"
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    26
lemma "a = a \<and> symm_f a b = symm_f b a" by (smt add: symm_f)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    27
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    28
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    29
section {* Linear arithmetic *}
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    30
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    31
lemma "(3::int) = 3" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    32
lemma "(3::real) = 3" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    33
lemma "(3 :: int) + 1 = 4" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    34
lemma "max (3::int) 8 > 5" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    35
lemma "abs (x :: real) + abs y \<ge> abs (x + y)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    36
lemma "let x = (2 :: int) in x + x \<noteq> 5" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    37
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    38
text{* 
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    39
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
    40
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    41
  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
    42
  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
    43
  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
    44
  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
    45
  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
    46
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    47
  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
    48
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    49
There, it is proved by "arith". SMT is able to prove this within a fraction
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    50
of one second.
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    51
*}
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    52
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    53
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
    54
         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
    55
         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
    56
 \<Longrightarrow> x1 = x10 & x2 = (x11::int)"
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    57
  by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    58
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    59
lemma "\<exists>x::int. 0 < x" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    60
lemma "\<exists>x::real. 0 < x" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    61
lemma "\<forall>x y::int. x < y \<longrightarrow> (2 * x + 1) < (2 * y)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    62
lemma "\<forall>x y::int. (2 * x + 1) \<noteq> (2 * y)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    63
lemma "~ (\<exists>x y z::int. 4 * x + -6 * y = (1::int))" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    64
lemma "~ (\<exists>x::int. False)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    65
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    66
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    67
section {* Non-linear arithmetic *}
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    68
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    69
lemma "((x::int) * (1 + y) - x * (1 - y)) = (2 * x * y)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    70
lemma
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    71
  "(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
    72
   U + (2 * (1 + p) * (b + e) + (1 + p) * d + d * p) - (1 + p) * (b + d + e)"
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    73
  by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    74
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    75
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    76
section {* Linear arithmetic for natural numbers *}
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    77
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    78
lemma "a < 3 \<Longrightarrow> (7::nat) > 2 * a" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    79
lemma "let x = (1::nat) + y in x - y > 0 * x" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    80
lemma
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    81
  "let x = (1::nat) + y in
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    82
   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
    83
   False \<or> P = (x - 1 = y) \<or> (\<not>P \<longrightarrow> False)"
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    84
  by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    85
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    86
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    87
section {* Bitvectors *}
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    88
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    89
locale bv
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    90
begin
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    91
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    92
declare [[smt_solver=z3]]
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    93
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    94
lemma "(27 :: 4 word) = -5" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    95
lemma "(27 :: 4 word) = 11" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    96
lemma "23 < (27::8 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    97
lemma "27 + 11 = (6::5 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    98
lemma "7 * 3 = (21::8 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    99
lemma "11 - 27 = (-16::8 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   100
lemma "- -11 = (11::5 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   101
lemma "-40 + 1 = (-39::7 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   102
lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   103
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   104
lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   105
lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   106
lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   107
lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   108
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   109
lemma "word_cat (27::4 word) (27::8 word) = (2843::12 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   110
lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" 
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   111
  by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   112
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   113
lemma "slice 1 (0b10110 :: 4 word) = (0b11 :: 2 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   114
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   115
lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   116
lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   117
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   118
lemma "bv_lshr 0b10011 2 = (0b100::8 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   119
lemma "bv_ashr 0b10011 2 = (0b100::8 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   120
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   121
lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   122
lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   123
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   124
lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   125
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   126
lemma "w < 256 \<Longrightarrow> (w :: 16 word) AND 0x00FF = w" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   127
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   128
end
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   129
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   130
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   131
section {* Pairs *}
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   132
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   133
lemma "fst (x, y) = a \<Longrightarrow> x = a" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   134
lemma "p1 = (x, y) \<and> p2 = (y, x) \<Longrightarrow> fst p1 = snd p2" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   135
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   136
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   137
section {* Higher-order problems and recursion *}
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   138
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   139
lemma "(f g x = (g x \<and> True)) \<or> (f g x = True) \<or> (g x = True)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   140
lemma "P ((2::int) < 3) = P True" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   141
lemma "P ((2::int) < 3) = (P True :: bool)" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   142
lemma "P (0 \<le> (a :: 4 word)) = P True" using [[smt_solver=z3]] by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   143
lemma "id 3 = 3 \<and> id True = True" by (smt add: id_def)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   144
lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> ((f (i1 := v1)) (i2 := v2)) i = f i" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   145
lemma "map (\<lambda>i::nat. i + 1) [0, 1] = [1, 2]" by (smt add: map.simps)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   146
lemma "(ALL x. P x) | ~ All P" by smt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   147
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   148
fun dec_10 :: "nat \<Rightarrow> nat" where
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   149
  "dec_10 n = (if n < 10 then n else dec_10 (n - 10))"
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   150
lemma "dec_10 (4 * dec_10 4) = 6" by (smt add: dec_10.simps)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   151
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   152
axiomatization
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   153
  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
   154
  where
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   155
  eval_dioph_mod:
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   156
  "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
   157
  and
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   158
  eval_dioph_div_mult:
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   159
  "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
   160
   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
   161
lemma
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   162
  "(eval_dioph ks xs = l) =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   163
   (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
   164
    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
   165
      (l - eval_dioph ks (map (\<lambda>x. x mod 2) xs)) div 2)"
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   166
  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
   167
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   168
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   169
section {* Monomorphization examples *}
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   170
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   171
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
   172
lemma poly_P: "P x \<and> (P [x] \<or> \<not>P[x])" by (simp add: P_def)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   173
lemma "P (1::int)" by (smt add: poly_P)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   174
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   175
consts g :: "'a \<Rightarrow> nat"
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   176
axioms
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   177
  g1: "g (Some x) = g [x]"
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   178
  g2: "g None = g []"
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   179
  g3: "g xs = length xs"
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   180
lemma "g (Some (3::int)) = g (Some True)" by (smt add: g1 g2 g3 list.size)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   181
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   182
end