| author | wenzelm | 
| Thu, 04 Apr 2013 18:44:22 +0200 | |
| changeset 51620 | 7c39677f9ea0 | 
| parent 41413 | 64cd30d6b0b8 | 
| child 53062 | 3af1a6020014 | 
| permissions | -rw-r--r-- | 
| 37655 | 1 | (* Title: HOL/Word/Bit_Operations.thy | 
| 2 | Author: Author: Brian Huffman, PSU and Gerwin Klein, NICTA | |
| 24334 | 3 | *) | 
| 4 | ||
| 26559 | 5 | header {* Syntactic classes for bitwise operations *}
 | 
| 24333 | 6 | |
| 37655 | 7 | theory Bit_Operations | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
37655diff
changeset | 8 | imports "~~/src/HOL/Library/Bit" | 
| 24333 | 9 | begin | 
| 10 | ||
| 29608 | 11 | class bit = | 
| 26559 | 12 |   fixes bitNOT :: "'a \<Rightarrow> 'a"       ("NOT _" [70] 71)
 | 
| 13 | and bitAND :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "AND" 64) | |
| 14 | and bitOR :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "OR" 59) | |
| 15 | and bitXOR :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "XOR" 59) | |
| 24333 | 16 | |
| 17 | text {*
 | |
| 18 | We want the bitwise operations to bind slightly weaker | |
| 19 |   than @{text "+"} and @{text "-"}, but @{text "~~"} to 
 | |
| 20 |   bind slightly stronger than @{text "*"}.
 | |
| 21 | *} | |
| 24352 | 22 | |
| 24333 | 23 | text {*
 | 
| 24 | Testing and shifting operations. | |
| 25 | *} | |
| 26559 | 26 | |
| 27 | class bits = bit + | |
| 28 | fixes test_bit :: "'a \<Rightarrow> nat \<Rightarrow> bool" (infixl "!!" 100) | |
| 29 | and lsb :: "'a \<Rightarrow> bool" | |
| 30 | and set_bit :: "'a \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> 'a" | |
| 31 | and set_bits :: "(nat \<Rightarrow> bool) \<Rightarrow> 'a" (binder "BITS " 10) | |
| 32 | and shiftl :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infixl "<<" 55) | |
| 33 | and shiftr :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infixl ">>" 55) | |
| 34 | ||
| 35 | class bitss = bits + | |
| 36 | fixes msb :: "'a \<Rightarrow> bool" | |
| 24333 | 37 | |
| 38 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25762diff
changeset | 39 | subsection {* Bitwise operations on @{typ bit} *}
 | 
| 24333 | 40 | |
| 25762 | 41 | instantiation bit :: bit | 
| 42 | begin | |
| 43 | ||
| 26559 | 44 | primrec bitNOT_bit where | 
| 37655 | 45 | "NOT 0 = (1::bit)" | 
| 46 | | "NOT 1 = (0::bit)" | |
| 25762 | 47 | |
| 26559 | 48 | primrec bitAND_bit where | 
| 37655 | 49 | "0 AND y = (0::bit)" | 
| 50 | | "1 AND y = (y::bit)" | |
| 25762 | 51 | |
| 26559 | 52 | primrec bitOR_bit where | 
| 37655 | 53 | "0 OR y = (y::bit)" | 
| 54 | | "1 OR y = (1::bit)" | |
| 25762 | 55 | |
| 26559 | 56 | primrec bitXOR_bit where | 
| 37655 | 57 | "0 XOR y = (y::bit)" | 
| 58 | | "1 XOR y = (NOT y :: bit)" | |
| 25762 | 59 | |
| 60 | instance .. | |
| 61 | ||
| 62 | end | |
| 24333 | 63 | |
| 26559 | 64 | lemmas bit_simps = | 
| 65 | bitNOT_bit.simps bitAND_bit.simps bitOR_bit.simps bitXOR_bit.simps | |
| 24333 | 66 | |
| 24366 | 67 | lemma bit_extra_simps [simp]: | 
| 37655 | 68 | "x AND 0 = (0::bit)" | 
| 69 | "x AND 1 = (x::bit)" | |
| 70 | "x OR 1 = (1::bit)" | |
| 71 | "x OR 0 = (x::bit)" | |
| 72 | "x XOR 1 = NOT (x::bit)" | |
| 73 | "x XOR 0 = (x::bit)" | |
| 24366 | 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" | |
| 37655 | 85 | "(x::bit) XOR x = 0" | 
| 24366 | 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 |