src/HOL/Word/Examples/WordExamples.thy
author Thomas Sewell <thomas.sewell@nicta.com.au>
Tue, 17 Apr 2012 16:21:47 +1000
changeset 47567 407cabf66f21
parent 46173 5cc700033194
child 58410 6d46ad54a2ab
permissions -rw-r--r--
New tactic "word_bitwise" expands word equalities/inequalities into logic.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
     1
(* 
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
     2
  Authors: Gerwin Klein and Thomas Sewell, NICTA
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
     3
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
     4
  Examples demonstrating and testing various word operations.
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
     5
*)
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
     6
25262
d0928156e326 Added reference to Jeremy Dawson's paper on the word library.
kleing
parents: 24465
diff changeset
     7
header "Examples of word operations"
d0928156e326 Added reference to Jeremy Dawson's paper on the word library.
kleing
parents: 24465
diff changeset
     8
29629
5111ce425e7a tuned header
haftmann
parents: 26086
diff changeset
     9
theory WordExamples
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
    10
imports "../Word" "../WordBitwise"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    11
begin
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    12
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 29640
diff changeset
    13
type_synonym word32 = "32 word"
f270e3e18be5 modernized specifications;
wenzelm
parents: 29640
diff changeset
    14
type_synonym word8 = "8 word"
f270e3e18be5 modernized specifications;
wenzelm
parents: 29640
diff changeset
    15
type_synonym byte = word8
29640
741f26190c96 recovered example types from WordMain.thy;
wenzelm
parents: 29629
diff changeset
    16
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    17
text "modulus"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    18
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    19
lemma "(27 :: 4 word) = -5" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    20
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    21
lemma "(27 :: 4 word) = 11" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    22
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    23
lemma "27 \<noteq> (11 :: 6 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    24
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    25
text "signed"
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    26
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    27
lemma "(127 :: 6 word) = -1" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    28
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    29
text "number ring simps"
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    30
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    31
lemma 
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
    32
  "27 + 11 = (38::'a::len word)"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    33
  "27 + 11 = (6::5 word)"
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
    34
  "7 * 3 = (21::'a::len word)"
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
    35
  "11 - 27 = (-16::'a::len word)"
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
    36
  "- -11 = (11::'a::len word)"
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
    37
  "-40 + 1 = (-39::'a::len word)"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    38
  by simp_all
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    39
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    40
lemma "word_pred 2 = 1" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    41
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    42
lemma "word_succ -3 = -2" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    43
  
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    44
lemma "23 < (27::8 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    45
lemma "23 \<le> (27::8 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    46
lemma "\<not> 23 < (27::2 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    47
lemma "0 < (4::3 word)" by simp
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    48
lemma "1 < (4::3 word)" by simp
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    49
lemma "0 < (1::3 word)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    50
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    51
text "ring operations"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    52
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    53
lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    54
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    55
text "casting"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    56
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    57
lemma "uint (234567 :: 10 word) = 71" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    58
lemma "uint (-234567 :: 10 word) = 953" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    59
lemma "sint (234567 :: 10 word) = 71" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    60
lemma "sint (-234567 :: 10 word) = -71" by simp
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    61
lemma "uint (1 :: 10 word) = 1" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    62
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    63
lemma "unat (-234567 :: 10 word) = 953" by simp
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    64
lemma "unat (1 :: 10 word) = 1" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    65
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    66
lemma "ucast (0b1010 :: 4 word) = (0b10 :: 2 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    67
lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    68
lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by simp
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    69
lemma "ucast (1 :: 4 word) = (1 :: 2 word)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    70
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    71
text "reducing goals to nat or int and arith:"
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
    72
lemma "i < x ==> i < (i + 1 :: 'a :: len word)" by unat_arith
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
    73
lemma "i < x ==> i < (i + 1 :: 'a :: len word)" by uint_arith
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    74
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    75
text "bool lists"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    76
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
    77
lemma "of_bl [True, False, True, True] = (0b1011::'a::len word)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    78
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    79
lemma "to_bl (0b110::4 word) = [False, True, True, False]" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    80
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    81
text "this is not exactly fast, but bearable"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    82
lemma "of_bl (replicate 32 True) = (0xFFFFFFFF::32 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    83
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    84
text "this works only for replicate n True"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    85
lemma "of_bl (replicate 32 True) = (0xFFFFFFFF::32 word)"
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    86
  by (unfold mask_bl [symmetric]) (simp add: mask_def)
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    87
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    88
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    89
text "bit operations"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    90
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    91
lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    92
lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    93
lemma "0xF0 XOR 0xFF = (0x0F :: byte)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    94
lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by simp
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    95
lemma "0 AND 5 = (0 :: byte)" by simp
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    96
lemma "1 AND 1 = (1 :: byte)" by simp
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    97
lemma "1 AND 0 = (0 :: byte)" by simp
46064
88ef116e0522 add simp rules for bitwise word operations with 1
huffman
parents: 46015
diff changeset
    98
lemma "1 AND 5 = (1 :: byte)" by simp
88ef116e0522 add simp rules for bitwise word operations with 1
huffman
parents: 46015
diff changeset
    99
lemma "1 OR 6 = (7 :: byte)" by simp
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
   100
lemma "1 OR 1 = (1 :: byte)" by simp
46064
88ef116e0522 add simp rules for bitwise word operations with 1
huffman
parents: 46015
diff changeset
   101
lemma "1 XOR 7 = (6 :: byte)" by simp
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
   102
lemma "1 XOR 1 = (0 :: byte)" by simp
46064
88ef116e0522 add simp rules for bitwise word operations with 1
huffman
parents: 46015
diff changeset
   103
lemma "NOT 1 = (254 :: byte)" by simp
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
   104
lemma "NOT 0 = (255 :: byte)" apply simp oops
46064
88ef116e0522 add simp rules for bitwise word operations with 1
huffman
parents: 46015
diff changeset
   105
(* FIXME: "NOT 0" rewrites to "max_word" instead of "-1" *)
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   106
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   107
lemma "(-1 :: 32 word) = 0xFFFFFFFF" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   108
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   109
lemma "(0b0010 :: 4 word) !! 1" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   110
lemma "\<not> (0b0010 :: 4 word) !! 0" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   111
lemma "\<not> (0b1000 :: 3 word) !! 4" by simp
46172
c06e868dc339 add simp rule test_bit_1
huffman
parents: 46064
diff changeset
   112
lemma "\<not> (1 :: 3 word) !! 2" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   113
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   114
lemma "(0b11000 :: 10 word) !! n = (n = 4 \<or> n = 3)" 
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25262
diff changeset
   115
  by (auto simp add: bin_nth_Bit0 bin_nth_Bit1)
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   116
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
   117
lemma "set_bit 55 7 True = (183::'a::len0 word)" by simp
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
   118
lemma "set_bit 0b0010 7 True = (0b10000010::'a::len0 word)" by simp
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
   119
lemma "set_bit 0b0010 1 False = (0::'a::len0 word)" by simp
46173
5cc700033194 add simp rules for set_bit and msb applied to 0 and 1
huffman
parents: 46172
diff changeset
   120
lemma "set_bit 1 3 True = (0b1001::'a::len0 word)" by simp
5cc700033194 add simp rules for set_bit and msb applied to 0 and 1
huffman
parents: 46172
diff changeset
   121
lemma "set_bit 1 0 False = (0::'a::len0 word)" by simp
5cc700033194 add simp rules for set_bit and msb applied to 0 and 1
huffman
parents: 46172
diff changeset
   122
lemma "set_bit 0 3 True = (0b1000::'a::len0 word)" by simp
5cc700033194 add simp rules for set_bit and msb applied to 0 and 1
huffman
parents: 46172
diff changeset
   123
lemma "set_bit 0 3 False = (0::'a::len0 word)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   124
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
   125
lemma "lsb (0b0101::'a::len word)" by simp
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
   126
lemma "\<not> lsb (0b1000::'a::len word)" by simp
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
   127
lemma "lsb (1::'a::len word)" by simp
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
   128
lemma "\<not> lsb (0::'a::len word)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   129
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   130
lemma "\<not> msb (0b0101::4 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   131
lemma   "msb (0b1000::4 word)" by simp
46173
5cc700033194 add simp rules for set_bit and msb applied to 0 and 1
huffman
parents: 46172
diff changeset
   132
lemma "\<not> msb (1::4 word)" by simp
5cc700033194 add simp rules for set_bit and msb applied to 0 and 1
huffman
parents: 46172
diff changeset
   133
lemma "\<not> msb (0::4 word)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   134
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
   135
lemma "word_cat (27::4 word) (27::8 word) = (2843::'a::len word)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   136
lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" 
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   137
  by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   138
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
   139
lemma "0b1011 << 2 = (0b101100::'a::len0 word)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   140
lemma "0b1011 >> 2 = (0b10::8 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   141
lemma "0b1011 >>> 2 = (0b10::8 word)" by simp
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
   142
lemma "1 << 2 = (0b100::'a::len0 word)" apply simp? oops
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   143
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   144
lemma "slice 3 (0b101111::6 word) = (0b101::3 word)" by simp
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
   145
lemma "slice 3 (1::6 word) = (0::3 word)" apply simp? oops
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   146
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   147
lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   148
lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   149
lemma "word_roti 2 0b1110 = (0b1011::4 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   150
lemma "word_roti -2 0b0110 = (0b1001::4 word)" by simp
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
   151
lemma "word_rotr 2 0 = (0::4 word)" by simp
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
   152
lemma "word_rotr 2 1 = (0b0100::4 word)" apply simp? oops
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
   153
lemma "word_rotl 2 1 = (0b0100::4 word)" apply simp? oops
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
   154
lemma "word_roti -2 1 = (0b0100::4 word)" apply simp? oops
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   155
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   156
lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)"
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   157
proof -
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   158
  have "(x AND 0xff00) OR (x AND 0x00ff) = x AND (0xff00 OR 0x00ff)"
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   159
    by (simp only: word_ao_dist2)
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   160
  also have "0xff00 OR 0x00ff = (-1::16 word)"
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   161
    by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   162
  also have "x AND -1 = x"
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   163
    by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   164
  finally show ?thesis .
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   165
qed
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   166
47567
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   167
text "alternative proof using bitwise expansion"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   168
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   169
lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   170
  by word_bitwise
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   171
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   172
text "more proofs using bitwise expansion"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   173
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   174
lemma "((x :: 10 word) AND NOT 3) >> 4 << 2 = ((x >> 2) AND NOT 3)"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   175
  by word_bitwise
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   176
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   177
lemma "(((x :: 12 word) AND -8) >> 3) AND 7 = (x AND 56) >> 3"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   178
  by word_bitwise
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   179
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   180
text "some problems require further reasoning after bit expansion"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   181
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   182
lemma "x \<le> (42 :: 8 word) \<Longrightarrow> x \<le> 89"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   183
  apply word_bitwise
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   184
  apply blast
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   185
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   186
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   187
lemma "((x :: word32) AND 1023) = 0 \<Longrightarrow> x \<le> -1024"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   188
  apply word_bitwise
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   189
  apply clarsimp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   190
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   191
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   192
text "operations like shifts by non-numerals will expose some internal list
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   193
 representations but may still be easy to solve"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   194
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   195
lemma shiftr_overflow:
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   196
  "32 \<le> a \<Longrightarrow> (b::word32) >> a = 0"
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   197
  apply (word_bitwise)
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   198
  apply simp
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   199
  done
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   200
407cabf66f21 New tactic "word_bitwise" expands word equalities/inequalities into logic.
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 46173
diff changeset
   201
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   202
end