| author | wenzelm |
| Thu, 28 Feb 2008 15:54:37 +0100 | |
| changeset 26179 | bc5d582d6cfe |
| parent 26086 | 3c243098b64a |
| child 26559 | 799983936aad |
| permissions | -rw-r--r-- |
| 24334 | 1 |
(* |
2 |
ID: $Id$ |
|
3 |
Author: Brian Huffman, PSU and Gerwin Klein, NICTA |
|
4 |
||
5 |
Syntactic class for bitwise operations. |
|
6 |
*) |
|
7 |
||
8 |
||
| 24333 | 9 |
header {* Syntactic class for bitwise operations *}
|
10 |
||
11 |
theory BitSyntax |
|
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25762
diff
changeset
|
12 |
imports Main Num_Lemmas |
| 24333 | 13 |
begin |
14 |
||
| 24352 | 15 |
class bit = type + |
16 |
fixes bitNOT :: "'a \<Rightarrow> 'a" |
|
17 |
and bitAND :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" |
|
18 |
and bitOR :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" |
|
19 |
and bitXOR :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" |
|
| 24333 | 20 |
|
21 |
text {*
|
|
22 |
We want the bitwise operations to bind slightly weaker |
|
23 |
than @{text "+"} and @{text "-"}, but @{text "~~"} to
|
|
24 |
bind slightly stronger than @{text "*"}.
|
|
25 |
*} |
|
| 24352 | 26 |
|
27 |
notation |
|
| 25382 | 28 |
bitNOT ("NOT _" [70] 71) and
|
29 |
bitAND (infixr "AND" 64) and |
|
30 |
bitOR (infixr "OR" 59) and |
|
| 24352 | 31 |
bitXOR (infixr "XOR" 59) |
| 24333 | 32 |
|
33 |
text {*
|
|
34 |
Testing and shifting operations. |
|
35 |
*} |
|
36 |
consts |
|
37 |
test_bit :: "'a::bit \<Rightarrow> nat \<Rightarrow> bool" (infixl "!!" 100) |
|
38 |
lsb :: "'a::bit \<Rightarrow> bool" |
|
39 |
msb :: "'a::bit \<Rightarrow> bool" |
|
40 |
set_bit :: "'a::bit \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> 'a" |
|
41 |
set_bits :: "(nat \<Rightarrow> bool) \<Rightarrow> 'a::bit" (binder "BITS " 10) |
|
42 |
shiftl :: "'a::bit \<Rightarrow> nat \<Rightarrow> 'a" (infixl "<<" 55) |
|
43 |
shiftr :: "'a::bit \<Rightarrow> nat \<Rightarrow> 'a" (infixl ">>" 55) |
|
44 |
||
45 |
||
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25762
diff
changeset
|
46 |
subsection {* Bitwise operations on @{typ bit} *}
|
| 24333 | 47 |
|
| 25762 | 48 |
instantiation bit :: bit |
49 |
begin |
|
50 |
||
51 |
definition |
|
52 |
NOT_bit_def: "NOT x = (case x of bit.B0 \<Rightarrow> bit.B1 | bit.B1 \<Rightarrow> bit.B0)" |
|
53 |
||
54 |
definition |
|
55 |
AND_bit_def: "x AND y = (case x of bit.B0 \<Rightarrow> bit.B0 | bit.B1 \<Rightarrow> y)" |
|
56 |
||
57 |
definition |
|
58 |
OR_bit_def: "(x OR y \<Colon> bit) = NOT (NOT x AND NOT y)" |
|
59 |
||
60 |
definition |
|
61 |
XOR_bit_def: "(x XOR y \<Colon> bit) = (x AND NOT y) OR (NOT x AND y)" |
|
62 |
||
63 |
instance .. |
|
64 |
||
65 |
end |
|
| 24333 | 66 |
|
67 |
lemma bit_simps [simp]: |
|
68 |
"NOT bit.B0 = bit.B1" |
|
69 |
"NOT bit.B1 = bit.B0" |
|
70 |
"bit.B0 AND y = bit.B0" |
|
71 |
"bit.B1 AND y = y" |
|
72 |
"bit.B0 OR y = y" |
|
73 |
"bit.B1 OR y = bit.B1" |
|
74 |
"bit.B0 XOR y = y" |
|
75 |
"bit.B1 XOR y = NOT y" |
|
| 24334 | 76 |
by (simp_all add: NOT_bit_def AND_bit_def OR_bit_def |
77 |
XOR_bit_def split: bit.split) |
|
| 24333 | 78 |
|
| 24366 | 79 |
lemma bit_extra_simps [simp]: |
80 |
"x AND bit.B0 = bit.B0" |
|
81 |
"x AND bit.B1 = x" |
|
82 |
"x OR bit.B1 = bit.B1" |
|
83 |
"x OR bit.B0 = x" |
|
84 |
"x XOR bit.B1 = NOT x" |
|
85 |
"x XOR bit.B0 = x" |
|
86 |
by (cases x, auto)+ |
|
87 |
||
88 |
lemma bit_ops_comm: |
|
89 |
"(x::bit) AND y = y AND x" |
|
90 |
"(x::bit) OR y = y OR x" |
|
91 |
"(x::bit) XOR y = y XOR x" |
|
92 |
by (cases y, auto)+ |
|
93 |
||
94 |
lemma bit_ops_same [simp]: |
|
95 |
"(x::bit) AND x = x" |
|
96 |
"(x::bit) OR x = x" |
|
97 |
"(x::bit) XOR x = bit.B0" |
|
98 |
by (cases x, auto)+ |
|
99 |
||
100 |
lemma bit_not_not [simp]: "NOT (NOT (x::bit)) = x" |
|
101 |
by (cases x) auto |
|
| 24333 | 102 |
|
103 |
end |