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