src/HOL/Word/Word_Examples.thy
author haftmann
Wed, 01 Jul 2020 17:32:11 +0000
changeset 71985 a1cf296a7786
parent 71954 13bb3f5cdc5b
child 71997 4a013c92a091
permissions -rw-r--r--
moved to Word_Lib
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
70173
c2786fe88064 tuned theory names
haftmann
parents: 67122
diff changeset
     1
(*  Title:      HOL/Word/Word_Examples.thy
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 58874
diff changeset
     2
    Authors:    Gerwin Klein and Thomas Sewell, NICTA
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
     3
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 58874
diff changeset
     4
Examples demonstrating and testing various word operations.
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
     5
*)
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
     6
58874
7172c7ffb047 modernized header;
wenzelm
parents: 58410
diff changeset
     7
section "Examples of word operations"
25262
d0928156e326 Added reference to Jeremy Dawson's paper on the word library.
kleing
parents: 24465
diff changeset
     8
70173
c2786fe88064 tuned theory names
haftmann
parents: 67122
diff changeset
     9
theory Word_Examples
71985
a1cf296a7786 moved to Word_Lib
haftmann
parents: 71954
diff changeset
    10
  imports Word
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
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 58874
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)"
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 47567
diff changeset
    36
  "- (- 11) = (11::'a::len word)"
24465
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
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 47567
diff changeset
    42
lemma "word_succ (- 3) = -2" by simp
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 58874
diff changeset
    43
24441
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:"
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 58874
diff changeset
    72
lemma "i < x \<Longrightarrow> i < i + 1" for i x :: "'a::len word"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 58874
diff changeset
    73
  by unat_arith
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 58874
diff changeset
    74
lemma "i < x \<Longrightarrow> i < i + 1" for i x :: "'a::len word"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 58874
diff changeset
    75
  by unat_arith
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    76
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    77
text "bool lists"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    78
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
    79
lemma "of_bl [True, False, True, True] = (0b1011::'a::len word)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    80
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    81
lemma "to_bl (0b110::4 word) = [False, True, True, False]" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    82
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    83
text "this is not exactly fast, but bearable"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    84
lemma "of_bl (replicate 32 True) = (0xFFFFFFFF::32 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    85
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    86
text "this works only for replicate n True"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    87
lemma "of_bl (replicate 32 True) = (0xFFFFFFFF::32 word)"
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    88
  by (unfold mask_bl [symmetric]) (simp add: mask_def)
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    89
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    90
46015
713c1f396e33 add several new tests, most of which don't work yet
huffman
parents: 44762
diff changeset
    91
text "bit operations"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    92
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    93
lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    94
lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    95
lemma "0xF0 XOR 0xFF = (0x0F :: byte)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
    96
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
    97
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
    98
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
    99
lemma "1 AND 0 = (0 :: byte)" by simp
46064
88ef116e0522 add simp rules for bitwise word operations with 1
huffman
parents: 46015
diff changeset
   100
lemma "1 AND 5 = (1 :: byte)" by simp
88ef116e0522 add simp rules for bitwise word operations with 1
huffman
parents: 46015
diff changeset
   101
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
   102
lemma "1 OR 1 = (1 :: byte)" by simp
46064
88ef116e0522 add simp rules for bitwise word operations with 1
huffman
parents: 46015
diff changeset
   103
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
   104
lemma "1 XOR 1 = (0 :: byte)" by simp
46064
88ef116e0522 add simp rules for bitwise word operations with 1
huffman
parents: 46015
diff changeset
   105
lemma "NOT 1 = (254 :: byte)" by simp
71946
4d4079159be7 replaced mere alias by abbreviation
haftmann
parents: 70173
diff changeset
   106
lemma "NOT 0 = (255 :: byte)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   107
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   108
lemma "(-1 :: 32 word) = 0xFFFFFFFF" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   109
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   110
lemma "(0b0010 :: 4 word) !! 1" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   111
lemma "\<not> (0b0010 :: 4 word) !! 0" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   112
lemma "\<not> (0b1000 :: 3 word) !! 4" by simp
46172
c06e868dc339 add simp rule test_bit_1
huffman
parents: 46064
diff changeset
   113
lemma "\<not> (1 :: 3 word) !! 2" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   114
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 58874
diff changeset
   115
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
   116
  by (auto simp add: bin_nth_Bit0 bin_nth_Bit1)
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   117
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71946
diff changeset
   118
lemma "set_bit 55 7 True = (183::'a::len word)" by simp
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71946
diff changeset
   119
lemma "set_bit 0b0010 7 True = (0b10000010::'a::len word)" by simp
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71946
diff changeset
   120
lemma "set_bit 0b0010 1 False = (0::'a::len word)" by simp
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71946
diff changeset
   121
lemma "set_bit 1 3 True = (0b1001::'a::len word)" by simp
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71946
diff changeset
   122
lemma "set_bit 1 0 False = (0::'a::len word)" by simp
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71946
diff changeset
   123
lemma "set_bit 0 3 True = (0b1000::'a::len word)" by simp
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71946
diff changeset
   124
lemma "set_bit 0 3 False = (0::'a::len word)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   125
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
   126
lemma "lsb (0b0101::'a::len word)" by simp
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
   127
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
   128
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
   129
lemma "\<not> lsb (0::'a::len word)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   130
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   131
lemma "\<not> msb (0b0101::4 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   132
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
   133
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
   134
lemma "\<not> msb (0::4 word)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   135
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents: 24441
diff changeset
   136
lemma "word_cat (27::4 word) (27::8 word) = (2843::'a::len word)" by simp
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 58874
diff changeset
   137
lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)"
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   138
  by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   139
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71946
diff changeset
   140
lemma "0b1011 << 2 = (0b101100::'a::len word)" by simp
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   141
lemma "0b1011 >> 2 = (0b10::8 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   142
lemma "0b1011 >>> 2 = (0b10::8 word)" by simp
71954
13bb3f5cdc5b pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
haftmann
parents: 71946
diff changeset
   143
lemma "1 << 2 = (0b100::'a::len word)" apply simp? oops
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   144
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   145
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
   146
lemma "slice 3 (1::6 word) = (0::3 word)" apply simp? oops
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   147
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   148
lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   149
lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   150
lemma "word_roti 2 0b1110 = (0b1011::4 word)" by simp
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 47567
diff changeset
   151
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
   152
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
   153
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
   154
lemma "word_rotl 2 1 = (0b0100::4 word)" apply simp? oops
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 47567
diff changeset
   155
lemma "word_roti (- 2) 1 = (0b0100::4 word)" apply simp? oops
24441
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   156
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   157
lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)"
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   158
proof -
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   159
  have "(x AND 0xff00) OR (x AND 0x00ff) = x AND (0xff00 OR 0x00ff)"
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   160
    by (simp only: word_ao_dist2)
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   161
  also have "0xff00 OR 0x00ff = (-1::16 word)"
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   162
    by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   163
  also have "x AND -1 = x"
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   164
    by simp
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   165
  finally show ?thesis .
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   166
qed
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   167
d2a5295570d0 Word Examples directory
huffman
parents:
diff changeset
   168
end