| author | wenzelm | 
| Thu, 12 Dec 2013 16:56:53 +0100 | |
| changeset 54727 | a806e7251cf0 | 
| parent 54224 | 9fda41a04c32 | 
| child 54847 | d6cf9a5b9be9 | 
| 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 | |
| 25762 | 38 | end | 
| 24333 | 39 |