--- a/src/HOL/Word/BitSyntax.thy Fri Apr 04 13:40:25 2008 +0200
+++ b/src/HOL/Word/BitSyntax.thy Fri Apr 04 13:40:26 2008 +0200
@@ -6,17 +6,17 @@
*)
-header {* Syntactic class for bitwise operations *}
+header {* Syntactic classes for bitwise operations *}
theory BitSyntax
-imports Main Num_Lemmas
+imports BinGeneral
begin
class bit = type +
- fixes bitNOT :: "'a \<Rightarrow> 'a"
- and bitAND :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
- and bitOR :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
- and bitXOR :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+ fixes bitNOT :: "'a \<Rightarrow> 'a" ("NOT _" [70] 71)
+ and bitAND :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "AND" 64)
+ and bitOR :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "OR" 59)
+ and bitXOR :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "XOR" 59)
text {*
We want the bitwise operations to bind slightly weaker
@@ -24,23 +24,20 @@
bind slightly stronger than @{text "*"}.
*}
-notation
- bitNOT ("NOT _" [70] 71) and
- bitAND (infixr "AND" 64) and
- bitOR (infixr "OR" 59) and
- bitXOR (infixr "XOR" 59)
-
text {*
Testing and shifting operations.
*}
-consts
- test_bit :: "'a::bit \<Rightarrow> nat \<Rightarrow> bool" (infixl "!!" 100)
- lsb :: "'a::bit \<Rightarrow> bool"
- msb :: "'a::bit \<Rightarrow> bool"
- set_bit :: "'a::bit \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> 'a"
- set_bits :: "(nat \<Rightarrow> bool) \<Rightarrow> 'a::bit" (binder "BITS " 10)
- shiftl :: "'a::bit \<Rightarrow> nat \<Rightarrow> 'a" (infixl "<<" 55)
- shiftr :: "'a::bit \<Rightarrow> nat \<Rightarrow> 'a" (infixl ">>" 55)
+
+class bits = bit +
+ fixes test_bit :: "'a \<Rightarrow> nat \<Rightarrow> bool" (infixl "!!" 100)
+ and lsb :: "'a \<Rightarrow> bool"
+ and set_bit :: "'a \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> 'a"
+ and set_bits :: "(nat \<Rightarrow> bool) \<Rightarrow> 'a" (binder "BITS " 10)
+ and shiftl :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infixl "<<" 55)
+ and shiftr :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infixl ">>" 55)
+
+class bitss = bits +
+ fixes msb :: "'a \<Rightarrow> bool"
subsection {* Bitwise operations on @{typ bit} *}
@@ -48,33 +45,28 @@
instantiation bit :: bit
begin
-definition
- NOT_bit_def: "NOT x = (case x of bit.B0 \<Rightarrow> bit.B1 | bit.B1 \<Rightarrow> bit.B0)"
+primrec bitNOT_bit where
+ "NOT bit.B0 = bit.B1"
+ | "NOT bit.B1 = bit.B0"
-definition
- AND_bit_def: "x AND y = (case x of bit.B0 \<Rightarrow> bit.B0 | bit.B1 \<Rightarrow> y)"
+primrec bitAND_bit where
+ "bit.B0 AND y = bit.B0"
+ | "bit.B1 AND y = y"
-definition
- OR_bit_def: "(x OR y \<Colon> bit) = NOT (NOT x AND NOT y)"
+primrec bitOR_bit where
+ "bit.B0 OR y = y"
+ | "bit.B1 OR y = bit.B1"
-definition
- XOR_bit_def: "(x XOR y \<Colon> bit) = (x AND NOT y) OR (NOT x AND y)"
+primrec bitXOR_bit where
+ "bit.B0 XOR y = y"
+ | "bit.B1 XOR y = NOT y"
instance ..
end
-lemma bit_simps [simp]:
- "NOT bit.B0 = bit.B1"
- "NOT bit.B1 = bit.B0"
- "bit.B0 AND y = bit.B0"
- "bit.B1 AND y = y"
- "bit.B0 OR y = y"
- "bit.B1 OR y = bit.B1"
- "bit.B0 XOR y = y"
- "bit.B1 XOR y = NOT y"
- by (simp_all add: NOT_bit_def AND_bit_def OR_bit_def
- XOR_bit_def split: bit.split)
+lemmas bit_simps =
+ bitNOT_bit.simps bitAND_bit.simps bitOR_bit.simps bitXOR_bit.simps
lemma bit_extra_simps [simp]:
"x AND bit.B0 = bit.B0"