Theory BitSyntax

Up to index of Isabelle/HOL/HOL-Word

theory BitSyntax
imports BinGeneral
begin

(* 
  ID:     $Id: BitSyntax.thy,v 1.8 2008/04/04 11:40:26 haftmann Exp $
  Author: Brian Huffman, PSU and Gerwin Klein, NICTA

  Syntactic class for bitwise operations.
*)


header {* Syntactic classes for bitwise operations *}

theory BitSyntax
imports BinGeneral
begin

class bit = type +
  fixes bitNOT :: "'a => 'a"       ("NOT _" [70] 71)
    and bitAND :: "'a => 'a => 'a" (infixr "AND" 64)
    and bitOR  :: "'a => 'a => 'a" (infixr "OR"  59)
    and bitXOR :: "'a => 'a => 'a" (infixr "XOR" 59)

text {*
  We want the bitwise operations to bind slightly weaker
  than @{text "+"} and @{text "-"}, but @{text "~~"} to 
  bind slightly stronger than @{text "*"}.
*}

text {*
  Testing and shifting operations.
*}

class bits = bit +
  fixes test_bit :: "'a => nat => bool" (infixl "!!" 100)
    and lsb      :: "'a => bool"
    and set_bit  :: "'a => nat => bool => 'a"
    and set_bits :: "(nat => bool) => 'a" (binder "BITS " 10)
    and shiftl   :: "'a => nat => 'a" (infixl "<<" 55)
    and shiftr   :: "'a => nat => 'a" (infixl ">>" 55)

class bitss = bits +
  fixes msb      :: "'a => bool"


subsection {* Bitwise operations on @{typ bit} *}

instantiation bit :: bit
begin

primrec bitNOT_bit where
  "NOT bit.B0 = bit.B1"
  | "NOT bit.B1 = bit.B0"

primrec bitAND_bit where
  "bit.B0 AND y = bit.B0"
  | "bit.B1 AND y = y"

primrec bitOR_bit where
  "bit.B0 OR y = y"
  | "bit.B1 OR y = bit.B1"

primrec bitXOR_bit where
  "bit.B0 XOR y = y"
  | "bit.B1 XOR y = NOT y"

instance  ..

end

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"
  "x AND bit.B1 = x"
  "x OR bit.B1 = bit.B1"
  "x OR bit.B0 = x"
  "x XOR bit.B1 = NOT x"
  "x XOR bit.B0 = x"
  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 = bit.B0"
  by (cases x, auto)+

lemma bit_not_not [simp]: "NOT (NOT (x::bit)) = x"
  by (cases x) auto

end

Bitwise operations on @{typ bit}

lemma bit_simps:

  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

lemma bit_extra_simps:

  x AND bit.B0 = bit.B0
  x AND bit.B1 = x
  x OR bit.B1 = bit.B1
  x OR bit.B0 = x
  x XOR bit.B1 = NOT x
  x XOR bit.B0 = x

lemma bit_ops_comm:

  x AND y = y AND x
  x OR y = y OR x
  x XOR y = y XOR x

lemma bit_ops_same:

  x AND x = x
  x OR x = x
  x XOR x = bit.B0

lemma bit_not_not:

  NOT NOT x = x