--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Word/Bit_Bit.thy Thu Oct 31 11:44:20 2013 +0100
@@ -0,0 +1,73 @@
+(* Title: HOL/Word/Bit_Bit.thy
+ Author: Author: Brian Huffman, PSU and Gerwin Klein, NICTA
+*)
+
+header {* Bit operations in \<Z>\<^sub>2*}
+
+theory Bit_Bit
+imports Bit_Operations "~~/src/HOL/Library/Bit"
+begin
+
+instantiation bit :: bit
+begin
+
+primrec bitNOT_bit where
+ "NOT 0 = (1::bit)"
+ | "NOT 1 = (0::bit)"
+
+primrec bitAND_bit where
+ "0 AND y = (0::bit)"
+ | "1 AND y = (y::bit)"
+
+primrec bitOR_bit where
+ "0 OR y = (y::bit)"
+ | "1 OR y = (1::bit)"
+
+primrec bitXOR_bit where
+ "0 XOR y = (y::bit)"
+ | "1 XOR y = (NOT y :: bit)"
+
+instance ..
+
+end
+
+lemmas bit_simps =
+ bitNOT_bit.simps bitAND_bit.simps bitOR_bit.simps bitXOR_bit.simps
+
+lemma bit_extra_simps [simp]:
+ "x AND 0 = (0::bit)"
+ "x AND 1 = (x::bit)"
+ "x OR 1 = (1::bit)"
+ "x OR 0 = (x::bit)"
+ "x XOR 1 = NOT (x::bit)"
+ "x XOR 0 = (x::bit)"
+ by (cases x, auto)+
+
+lemma bit_ops_comm:
+ "(x::bit) AND y = y AND x"
+ "(x::bit) OR y = y OR x"
+ "(x::bit) XOR y = y XOR x"
+ by (cases y, auto)+
+
+lemma bit_ops_same [simp]:
+ "(x::bit) AND x = x"
+ "(x::bit) OR x = x"
+ "(x::bit) XOR x = 0"
+ by (cases x, auto)+
+
+lemma bit_not_not [simp]: "NOT (NOT (x::bit)) = x"
+ by (cases x) auto
+
+lemma bit_or_def: "(b::bit) OR c = NOT (NOT b AND NOT c)"
+ by (induct b, simp_all)
+
+lemma bit_xor_def: "(b::bit) XOR c = (b AND NOT c) OR (NOT b AND c)"
+ by (induct b, simp_all)
+
+lemma bit_NOT_eq_1_iff [simp]: "NOT (b::bit) = 1 \<longleftrightarrow> b = 0"
+ by (induct b, simp_all)
+
+lemma bit_AND_eq_1_iff [simp]: "(a::bit) AND b = 1 \<longleftrightarrow> a = 1 \<and> b = 1"
+ by (induct a, simp_all)
+
+end
--- a/src/HOL/Word/Bit_Int.thy Thu Oct 31 11:44:20 2013 +0100
+++ b/src/HOL/Word/Bit_Int.thy Thu Oct 31 11:44:20 2013 +0100
@@ -9,7 +9,7 @@
header {* Bitwise Operations on Binary Integers *}
theory Bit_Int
-imports Bit_Representation Bit_Operations
+imports Bit_Representation Bit_Bit
begin
subsection {* Logical operations *}
--- a/src/HOL/Word/Bit_Operations.thy Thu Oct 31 11:44:20 2013 +0100
+++ b/src/HOL/Word/Bit_Operations.thy Thu Oct 31 11:44:20 2013 +0100
@@ -8,8 +8,6 @@
imports "~~/src/HOL/Library/Bit"
begin
-subsection {* Abstract syntactic bit operations *}
-
class bit =
fixes bitNOT :: "'a \<Rightarrow> 'a" ("NOT _" [70] 71)
and bitAND :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "AND" 64)
@@ -37,69 +35,5 @@
class bitss = bits +
fixes msb :: "'a \<Rightarrow> bool"
-
-subsection {* Bitwise operations on @{typ bit} *}
-
-instantiation bit :: bit
-begin
-
-primrec bitNOT_bit where
- "NOT 0 = (1::bit)"
- | "NOT 1 = (0::bit)"
-
-primrec bitAND_bit where
- "0 AND y = (0::bit)"
- | "1 AND y = (y::bit)"
-
-primrec bitOR_bit where
- "0 OR y = (y::bit)"
- | "1 OR y = (1::bit)"
-
-primrec bitXOR_bit where
- "0 XOR y = (y::bit)"
- | "1 XOR y = (NOT y :: bit)"
-
-instance ..
-
end
-lemmas bit_simps =
- bitNOT_bit.simps bitAND_bit.simps bitOR_bit.simps bitXOR_bit.simps
-
-lemma bit_extra_simps [simp]:
- "x AND 0 = (0::bit)"
- "x AND 1 = (x::bit)"
- "x OR 1 = (1::bit)"
- "x OR 0 = (x::bit)"
- "x XOR 1 = NOT (x::bit)"
- "x XOR 0 = (x::bit)"
- by (cases x, auto)+
-
-lemma bit_ops_comm:
- "(x::bit) AND y = y AND x"
- "(x::bit) OR y = y OR x"
- "(x::bit) XOR y = y XOR x"
- by (cases y, auto)+
-
-lemma bit_ops_same [simp]:
- "(x::bit) AND x = x"
- "(x::bit) OR x = x"
- "(x::bit) XOR x = 0"
- by (cases x, auto)+
-
-lemma bit_not_not [simp]: "NOT (NOT (x::bit)) = x"
- by (cases x) auto
-
-lemma bit_or_def: "(b::bit) OR c = NOT (NOT b AND NOT c)"
- by (induct b, simp_all)
-
-lemma bit_xor_def: "(b::bit) XOR c = (b AND NOT c) OR (NOT b AND c)"
- by (induct b, simp_all)
-
-lemma bit_NOT_eq_1_iff [simp]: "NOT (b::bit) = 1 \<longleftrightarrow> b = 0"
- by (induct b, simp_all)
-
-lemma bit_AND_eq_1_iff [simp]: "(a::bit) AND b = 1 \<longleftrightarrow> a = 1 \<and> b = 1"
- by (induct a, simp_all)
-
-end
--- a/src/HOL/Word/Word.thy Thu Oct 31 11:44:20 2013 +0100
+++ b/src/HOL/Word/Word.thy Thu Oct 31 11:44:20 2013 +0100
@@ -8,6 +8,7 @@
imports
Type_Length
"~~/src/HOL/Library/Boolean_Algebra"
+ Bit_Bit
Bool_List_Representation
Misc_Typedef
Word_Miscellaneous