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
|
|
12 |
imports Main
|
|
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 |
|
|
46 |
subsection {* Bitwise operations on type @{typ bit} *}
|
|
47 |
|
24352
|
48 |
instance bit :: bit
|
24333
|
49 |
NOT_bit_def: "NOT x \<equiv> case x of bit.B0 \<Rightarrow> bit.B1 | bit.B1 \<Rightarrow> bit.B0"
|
|
50 |
AND_bit_def: "x AND y \<equiv> case x of bit.B0 \<Rightarrow> bit.B0 | bit.B1 \<Rightarrow> y"
|
|
51 |
OR_bit_def: "x OR y :: bit \<equiv> NOT (NOT x AND NOT y)"
|
|
52 |
XOR_bit_def: "x XOR y :: bit \<equiv> (x AND NOT y) OR (NOT x AND y)"
|
24352
|
53 |
..
|
24333
|
54 |
|
|
55 |
lemma bit_simps [simp]:
|
|
56 |
"NOT bit.B0 = bit.B1"
|
|
57 |
"NOT bit.B1 = bit.B0"
|
|
58 |
"bit.B0 AND y = bit.B0"
|
|
59 |
"bit.B1 AND y = y"
|
|
60 |
"bit.B0 OR y = y"
|
|
61 |
"bit.B1 OR y = bit.B1"
|
|
62 |
"bit.B0 XOR y = y"
|
|
63 |
"bit.B1 XOR y = NOT y"
|
24334
|
64 |
by (simp_all add: NOT_bit_def AND_bit_def OR_bit_def
|
|
65 |
XOR_bit_def split: bit.split)
|
24333
|
66 |
|
24366
|
67 |
lemma bit_extra_simps [simp]:
|
|
68 |
"x AND bit.B0 = bit.B0"
|
|
69 |
"x AND bit.B1 = x"
|
|
70 |
"x OR bit.B1 = bit.B1"
|
|
71 |
"x OR bit.B0 = x"
|
|
72 |
"x XOR bit.B1 = NOT x"
|
|
73 |
"x XOR bit.B0 = x"
|
|
74 |
by (cases x, auto)+
|
|
75 |
|
|
76 |
lemma bit_ops_comm:
|
|
77 |
"(x::bit) AND y = y AND x"
|
|
78 |
"(x::bit) OR y = y OR x"
|
|
79 |
"(x::bit) XOR y = y XOR x"
|
|
80 |
by (cases y, auto)+
|
|
81 |
|
|
82 |
lemma bit_ops_same [simp]:
|
|
83 |
"(x::bit) AND x = x"
|
|
84 |
"(x::bit) OR x = x"
|
|
85 |
"(x::bit) XOR x = bit.B0"
|
|
86 |
by (cases x, auto)+
|
|
87 |
|
|
88 |
lemma bit_not_not [simp]: "NOT (NOT (x::bit)) = x"
|
|
89 |
by (cases x) auto
|
24333
|
90 |
|
|
91 |
end
|