src/HOL/Word/BitSyntax.thy
author haftmann
Tue, 21 Aug 2007 13:30:36 +0200
changeset 24380 c215e256beca
parent 24366 08b116049547
child 25382 72cfe89f7b21
permissions -rw-r--r--
moved ordered_ab_semigroup_add to OrderedGroup.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24334
22863f364531 added header
kleing
parents: 24333
diff changeset
     1
(* 
22863f364531 added header
kleing
parents: 24333
diff changeset
     2
  ID:     $Id$
22863f364531 added header
kleing
parents: 24333
diff changeset
     3
  Author: Brian Huffman, PSU and Gerwin Klein, NICTA
22863f364531 added header
kleing
parents: 24333
diff changeset
     4
22863f364531 added header
kleing
parents: 24333
diff changeset
     5
  Syntactic class for bitwise operations.
22863f364531 added header
kleing
parents: 24333
diff changeset
     6
*)
22863f364531 added header
kleing
parents: 24333
diff changeset
     7
22863f364531 added header
kleing
parents: 24333
diff changeset
     8
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     9
header {* Syntactic class for bitwise operations *}
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    10
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    11
theory BitSyntax
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    12
imports Main
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    13
begin
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    14
24352
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    15
class bit = type +
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    16
  fixes bitNOT :: "'a \<Rightarrow> 'a"
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    17
    and bitAND :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    18
    and bitOR  :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    19
    and bitXOR :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    20
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    21
text {*
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    22
  We want the bitwise operations to bind slightly weaker
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    23
  than @{text "+"} and @{text "-"}, but @{text "~~"} to 
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    24
  bind slightly stronger than @{text "*"}.
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    25
*}
24352
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    26
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    27
notation
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    28
  bitNOT  ("NOT _" [70] 71)
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    29
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    30
notation
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    31
  bitAND  (infixr "AND" 64)
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    32
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    33
notation
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    34
  bitOR   (infixr "OR"  59)
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    35
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    36
notation
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    37
  bitXOR  (infixr "XOR" 59)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    38
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    39
text {*
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    40
  Testing and shifting operations.
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    41
*}
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    42
consts
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    43
  test_bit :: "'a::bit \<Rightarrow> nat \<Rightarrow> bool" (infixl "!!" 100)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    44
  lsb      :: "'a::bit \<Rightarrow> bool"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    45
  msb      :: "'a::bit \<Rightarrow> bool"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    46
  set_bit  :: "'a::bit \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> 'a"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    47
  set_bits :: "(nat \<Rightarrow> bool) \<Rightarrow> 'a::bit" (binder "BITS " 10)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    48
  shiftl   :: "'a::bit \<Rightarrow> nat \<Rightarrow> 'a" (infixl "<<" 55)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    49
  shiftr   :: "'a::bit \<Rightarrow> nat \<Rightarrow> 'a" (infixl ">>" 55)
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    50
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    51
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    52
subsection {* Bitwise operations on type @{typ bit} *}
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    53
24352
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    54
instance bit :: bit
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    55
  NOT_bit_def: "NOT x \<equiv> case x of bit.B0 \<Rightarrow> bit.B1 | bit.B1 \<Rightarrow> bit.B0"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    56
  AND_bit_def: "x AND y \<equiv> case x of bit.B0 \<Rightarrow> bit.B0 | bit.B1 \<Rightarrow> y"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    57
  OR_bit_def: "x OR y :: bit \<equiv> NOT (NOT x AND NOT y)"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    58
  XOR_bit_def: "x XOR y :: bit \<equiv> (x AND NOT y) OR (NOT x AND y)"
24352
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    59
  ..
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    60
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    61
lemma bit_simps [simp]:
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    62
  "NOT bit.B0 = bit.B1"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    63
  "NOT bit.B1 = bit.B0"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    64
  "bit.B0 AND y = bit.B0"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    65
  "bit.B1 AND y = y"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    66
  "bit.B0 OR y = y"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    67
  "bit.B1 OR y = bit.B1"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    68
  "bit.B0 XOR y = y"
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    69
  "bit.B1 XOR y = NOT y"
24334
22863f364531 added header
kleing
parents: 24333
diff changeset
    70
  by (simp_all add: NOT_bit_def AND_bit_def OR_bit_def
22863f364531 added header
kleing
parents: 24333
diff changeset
    71
                    XOR_bit_def split: bit.split)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    72
24366
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    73
lemma bit_extra_simps [simp]: 
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    74
  "x AND bit.B0 = bit.B0"
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    75
  "x AND bit.B1 = x"
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    76
  "x OR bit.B1 = bit.B1"
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    77
  "x OR bit.B0 = x"
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    78
  "x XOR bit.B1 = NOT x"
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    79
  "x XOR bit.B0 = x"
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    80
  by (cases x, auto)+
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    81
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    82
lemma bit_ops_comm: 
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    83
  "(x::bit) AND y = y AND x"
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    84
  "(x::bit) OR y = y OR x"
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    85
  "(x::bit) XOR y = y XOR x"
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    86
  by (cases y, auto)+
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    87
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    88
lemma bit_ops_same [simp]: 
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    89
  "(x::bit) AND x = x"
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    90
  "(x::bit) OR x = x"
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    91
  "(x::bit) XOR x = bit.B0"
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    92
  by (cases x, auto)+
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    93
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    94
lemma bit_not_not [simp]: "NOT (NOT (x::bit)) = x"
08b116049547 move bit simps from BinOperations to BitSyntax
huffman
parents: 24352
diff changeset
    95
  by (cases x) auto
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    96
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    97
end