| author | wenzelm | 
| Fri, 16 Sep 2016 15:54:50 +0200 | |
| changeset 63888 | 5a9a1985e9fb | 
| parent 61799 | 4cf66f21b764 | 
| child 65363 | 5eb619751b14 | 
| permissions | -rw-r--r-- | 
| 55210 | 1  | 
(* Title: HOL/Word/Bits.thy  | 
| 37655 | 2  | 
Author: Author: Brian Huffman, PSU and Gerwin Klein, NICTA  | 
| 24334 | 3  | 
*)  | 
4  | 
||
| 61799 | 5  | 
section \<open>Syntactic classes for bitwise operations\<close>  | 
| 24333 | 6  | 
|
| 
54854
 
3324a0078636
prefer "Bits" as theory name for abstract bit operations, similar to "Orderings", "Lattices", "Groups" etc.
 
haftmann 
parents: 
54847 
diff
changeset
 | 
7  | 
theory Bits  | 
| 
54847
 
d6cf9a5b9be9
prefer plain bool over dedicated type for binary digits
 
haftmann 
parents: 
54224 
diff
changeset
 | 
8  | 
imports Main  | 
| 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  | 
|
| 61799 | 17  | 
text \<open>  | 
| 24333 | 18  | 
We want the bitwise operations to bind slightly weaker  | 
| 61799 | 19  | 
than \<open>+\<close> and \<open>-\<close>, but \<open>~~\<close> to  | 
20  | 
bind slightly stronger than \<open>*\<close>.  | 
|
21  | 
\<close>  | 
|
| 24352 | 22  | 
|
| 61799 | 23  | 
text \<open>  | 
| 24333 | 24  | 
Testing and shifting operations.  | 
| 61799 | 25  | 
\<close>  | 
| 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  |