src/HOL/Word/Bits.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.thy
37655
f4d616d41a59 more speaking theory names
haftmann
parents: 29631
diff changeset
     2
    Author:     Author: Brian Huffman, PSU and Gerwin Klein, NICTA
24334
22863f364531 added header
kleing
parents: 24333
diff changeset
     3
*)
22863f364531 added header
kleing
parents: 24333
diff changeset
     4
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 58874
diff changeset
     5
section \<open>Syntactic classes for bitwise operations\<close>
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     6
54854
3324a0078636 prefer "Bits" as theory name for abstract bit operations, similar to "Orderings", "Lattices", "Groups" etc.
haftmann
parents: 54847
diff changeset
     7
theory Bits
54847
d6cf9a5b9be9 prefer plain bool over dedicated type for binary digits
haftmann
parents: 54224
diff changeset
     8
imports Main
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
     9
begin
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    10
29608
564ea783ace8 no base sort in class import
haftmann
parents: 26559
diff changeset
    11
class bit =
26559
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    12
  fixes bitNOT :: "'a \<Rightarrow> 'a"       ("NOT _" [70] 71)
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    13
    and bitAND :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "AND" 64)
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    14
    and bitOR  :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "OR"  59)
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    15
    and bitXOR :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "XOR" 59)
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    16
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 58874
diff changeset
    17
text \<open>
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    18
  We want the bitwise operations to bind slightly weaker
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 58874
diff changeset
    19
  than \<open>+\<close> and \<open>-\<close>, but \<open>~~\<close> to 
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 58874
diff changeset
    20
  bind slightly stronger than \<open>*\<close>.
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 58874
diff changeset
    21
\<close>
24352
eda777a2785b use class instead of axclass
huffman
parents: 24334
diff changeset
    22
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 58874
diff changeset
    23
text \<open>
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    24
  Testing and shifting operations.
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 58874
diff changeset
    25
\<close>
26559
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    26
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    27
class bits = bit +
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    28
  fixes test_bit :: "'a \<Rightarrow> nat \<Rightarrow> bool" (infixl "!!" 100)
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    29
    and lsb      :: "'a \<Rightarrow> bool"
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    30
    and set_bit  :: "'a \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> 'a"
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    31
    and set_bits :: "(nat \<Rightarrow> bool) \<Rightarrow> 'a" (binder "BITS " 10)
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    32
    and shiftl   :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infixl "<<" 55)
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    33
    and shiftr   :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infixl ">>" 55)
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    34
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    35
class bitss = bits +
799983936aad syntactic classes for bit operations
haftmann
parents: 26086
diff changeset
    36
  fixes msb      :: "'a \<Rightarrow> bool"
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    37
25762
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 25382
diff changeset
    38
end
24333
e77ea0ea7f2c * HOL-Word:
kleing
parents:
diff changeset
    39