src/HOL/Word/Bits_Bit.thy
author haftmann
Mon, 06 Feb 2017 20:56:34 +0100
changeset 64990 c6a7de505796
parent 61799 4cf66f21b764
child 65363 5eb619751b14
permissions -rw-r--r--
more explicit errors in pathological cases
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55210
d1e3b708d74b tuned headers;
wenzelm
parents: 54854
diff changeset
     1
(*  Title:      HOL/Word/Bits_Bit.thy
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
     2
    Author:     Author: Brian Huffman, PSU and Gerwin Klein, NICTA
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
     3
*)
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
     4
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 58874
diff changeset
     5
section \<open>Bit operations in $\cal Z_2$\<close>
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
     6
54854
3324a0078636 prefer "Bits" as theory name for abstract bit operations, similar to "Orderings", "Lattices", "Groups" etc.
haftmann
parents: 54394
diff changeset
     7
theory Bits_Bit
3324a0078636 prefer "Bits" as theory name for abstract bit operations, similar to "Orderings", "Lattices", "Groups" etc.
haftmann
parents: 54394
diff changeset
     8
imports Bits "~~/src/HOL/Library/Bit"
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
     9
begin
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    10
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    11
instantiation bit :: bit
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    12
begin
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    13
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    14
primrec bitNOT_bit where
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    15
  "NOT 0 = (1::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    16
  | "NOT 1 = (0::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    17
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    18
primrec bitAND_bit where
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    19
  "0 AND y = (0::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    20
  | "1 AND y = (y::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    21
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    22
primrec bitOR_bit where
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    23
  "0 OR y = (y::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    24
  | "1 OR y = (1::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    25
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    26
primrec bitXOR_bit where
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    27
  "0 XOR y = (y::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    28
  | "1 XOR y = (NOT y :: bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    29
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    30
instance  ..
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    31
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    32
end
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    33
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    34
lemmas bit_simps =
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    35
  bitNOT_bit.simps bitAND_bit.simps bitOR_bit.simps bitXOR_bit.simps
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    36
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    37
lemma bit_extra_simps [simp]: 
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    38
  "x AND 0 = (0::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    39
  "x AND 1 = (x::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    40
  "x OR 1 = (1::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    41
  "x OR 0 = (x::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    42
  "x XOR 1 = NOT (x::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    43
  "x XOR 0 = (x::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    44
  by (cases x, auto)+
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    45
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    46
lemma bit_ops_comm: 
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    47
  "(x::bit) AND y = y AND x"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    48
  "(x::bit) OR y = y OR x"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    49
  "(x::bit) XOR y = y XOR x"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    50
  by (cases y, auto)+
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    51
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    52
lemma bit_ops_same [simp]: 
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    53
  "(x::bit) AND x = x"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    54
  "(x::bit) OR x = x"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    55
  "(x::bit) XOR x = 0"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    56
  by (cases x, auto)+
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    57
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    58
lemma bit_not_not [simp]: "NOT (NOT (x::bit)) = x"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    59
  by (cases x) auto
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    60
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    61
lemma bit_or_def: "(b::bit) OR c = NOT (NOT b AND NOT c)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    62
  by (induct b, simp_all)
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    63
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    64
lemma bit_xor_def: "(b::bit) XOR c = (b AND NOT c) OR (NOT b AND c)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    65
  by (induct b, simp_all)
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    66
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    67
lemma bit_NOT_eq_1_iff [simp]: "NOT (b::bit) = 1 \<longleftrightarrow> b = 0"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    68
  by (induct b, simp_all)
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    69
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    70
lemma bit_AND_eq_1_iff [simp]: "(a::bit) AND b = 1 \<longleftrightarrow> a = 1 \<and> b = 1"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    71
  by (induct a, simp_all)
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    72
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    73
end