src/HOL/Word/WordExamples.thy
author kleing
Mon, 20 Aug 2007 04:34:31 +0200
changeset 24333 e77ea0ea7f2c
child 24408 058c5613a86f
permissions -rw-r--r--
* HOL-Word: New extensive library and type for generic, fixed size machine words, with arithemtic, bit-wise, shifting and rotating operations, reflection into int, nat, and bool lists, automation for linear arithmetic (by automatic reflection into nat or int), including lemmas on overflow and monotonicity. Instantiated to all appropriate arithmetic type classes, supporting automatic simplification of numerals on all operations. Jointly developed by NICTA, Galois, and PSU. * still to do: README.html/document + moving some of the generic lemmas to appropriate place in distribution
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     1
(* 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     2
  ID:     $Id$
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     3
  Author: Gerwin Klein, NICTA
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     4
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     5
  Examples demonstrating and testing various word operations.
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     6
*)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     7
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     8
theory WordExamples imports WordMain
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     9
begin
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    10
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    11
-- "modulus"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    12
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    13
lemma "(27 :: 4 word) = -5" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    14
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    15
lemma "(27 :: 4 word) = 11" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    16
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    17
lemma "27 \<noteq> (11 :: 6 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    18
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    19
-- "signed"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    20
lemma "(127 :: 6 word) = -1" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    21
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    22
-- "number ring simps"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    23
lemma 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    24
  "27 + 11 = (38::'a::len word)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    25
  "27 + 11 = (6::5 word)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    26
  "7 * 3 = (21::'a::len word)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    27
  "11 - 27 = (-16::'a::len word)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    28
  "- -11 = (11::'a::len word)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    29
  "-40 + 1 = (-39::'a::len word)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    30
  by simp_all
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    31
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    32
lemma "word_pred 2 = 1" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    33
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    34
lemma "word_succ -3 = -2" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    35
  
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    36
lemma "23 < (27::8 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    37
lemma "23 \<le> (27::8 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    38
lemma "\<not> 23 < (27::2 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    39
lemma "0 < (4::3 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    40
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    41
-- "ring operations"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    42
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    43
lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    44
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    45
-- "casting"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    46
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    47
lemma "uint (234567 :: 10 word) = 71" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    48
lemma "uint (-234567 :: 10 word) = 953" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    49
lemma "sint (234567 :: 10 word) = 71" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    50
lemma "sint (-234567 :: 10 word) = -71" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    51
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    52
lemma "unat (-234567 :: 10 word) = 953" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    53
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    54
lemma "ucast (0b1010 :: 4 word) = (0b10 :: 2 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    55
lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    56
lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    57
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    58
-- "reducing goals to nat or int and arith:"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    59
lemma "i < x ==> i < (i + 1 :: 'a :: len word)" by unat_arith
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    60
lemma "i < x ==> i < (i + 1 :: 'a :: len word)" by uint_arith
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    61
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    62
-- "bool lists"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    63
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    64
lemma "of_bl [True, False, True, True] = (0b1011::'a::len word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    65
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    66
lemma "to_bl (0b110::4 word) = [False, True, True, False]" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    67
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    68
-- "this is not exactly fast, but bearable"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    69
lemma "of_bl (replicate 32 True) = (0xFFFFFFFF::32 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    70
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    71
-- "this works only for replicate n True"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    72
lemma "of_bl (replicate 32 True) = (0xFFFFFFFF::32 word)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    73
  by (unfold mask_bl [symmetric]) (simp add: mask_def)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    74
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    75
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    76
-- "bit operations"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    77
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    78
lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    79
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    80
lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    81
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    82
lemma "0xF0 XOR 0xFF = (0x0F :: byte)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    83
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    84
lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    85
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    86
lemma "(-1 :: 32 word) = 0xFFFFFFFF" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    87
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    88
lemma "(0b0010 :: 4 word) !! 1" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    89
lemma "\<not> (0b0010 :: 4 word) !! 0" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    90
lemma "\<not> (0b1000 :: 3 word) !! 4" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    91
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    92
lemma "(0b11000 :: 10 word) !! n = (n = 4 \<or> n = 3)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    93
  by (auto simp add: bin_nth_Bit)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    94
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    95
lemma "set_bit 55 7 True = (183::'a::len0 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    96
lemma "set_bit 0b0010 7 True = (0b10000010::'a::len0 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    97
lemma "set_bit 0b0010 1 False = (0::'a::len0 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    98
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    99
lemma "lsb (0b0101::'a::len word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   100
lemma "\<not> lsb (0b1000::'a::len word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   101
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   102
lemma "\<not> msb (0b0101::4 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   103
lemma   "msb (0b1000::4 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   104
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   105
lemma "word_cat (27::4 word) (27::8 word) = (2843::'a::len word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   106
lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   107
  by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   108
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   109
lemma "0b1011 << 2 = (0b101100::'a::len0 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   110
lemma "0b1011 >> 2 = (0b10::8 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   111
lemma "0b1011 >>> 2 = (0b10::8 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   112
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   113
lemma "slice 3 (0b101111::6 word) = (0b101::3 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   114
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   115
lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   116
lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   117
lemma "word_roti 2 0b1110 = (0b1011::4 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   118
lemma "word_roti -2 0b0110 = (0b1001::4 word)" by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   119
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   120
lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   121
proof -
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   122
  have "(x AND 0xff00) OR (x AND 0x00ff) = x AND (0xff00 OR 0x00ff)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   123
    by (simp only: word_ao_dist2)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   124
  also have "0xff00 OR 0x00ff = (-1::16 word)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   125
    by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   126
  also have "x AND -1 = x"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   127
    by simp
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   128
  finally show ?thesis .
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   129
qed
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   130
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
   131
end