src/HOL/Word/Bits_Z2.thy
author wenzelm
Fri, 27 Mar 2020 12:46:56 +0100
changeset 71598 269dc4bf1f40
parent 70342 e4d626692640
permissions -rw-r--r--
clarified signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
70342
e4d626692640 clear separation of types for bits (False / True) and Z2 (0 / 1)
haftmann
parents: 70191
diff changeset
     1
(*  Title:      HOL/Word/Bits_Z2.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
70342
e4d626692640 clear separation of types for bits (False / True) and Z2 (0 / 1)
haftmann
parents: 70191
diff changeset
     7
theory Bits_Z2
e4d626692640 clear separation of types for bits (False / True) and Z2 (0 / 1)
haftmann
parents: 70191
diff changeset
     8
  imports Bits "HOL-Library.Z2"
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
70191
bdc835d934b7 no need to maintain two separate type classes
haftmann
parents: 70183
diff changeset
    11
instantiation bit :: bit_operations
54224
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
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    14
primrec bitNOT_bit
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    15
  where
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    16
    "NOT 0 = (1::bit)"
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    17
  | "NOT 1 = (0::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    18
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    19
primrec bitAND_bit
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    20
  where
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    21
    "0 AND y = (0::bit)"
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    22
  | "1 AND y = (y::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    23
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    24
primrec bitOR_bit
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    25
  where
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    26
    "0 OR y = (y::bit)"
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    27
  | "1 OR y = (1::bit)"
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    28
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    29
primrec bitXOR_bit
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    30
  where
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    31
    "0 XOR y = (y::bit)"
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    32
  | "1 XOR y = (NOT y :: bit)"
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
instance  ..
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    35
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    36
end
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    37
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    38
lemmas bit_simps =
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    39
  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
    40
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    41
lemma bit_extra_simps [simp]:
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    42
  "x AND 0 = 0"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    43
  "x AND 1 = x"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    44
  "x OR 1 = 1"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    45
  "x OR 0 = x"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    46
  "x XOR 1 = NOT x"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    47
  "x XOR 0 = x"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    48
  for x :: bit
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    49
  by (cases x; auto)+
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    50
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    51
lemma bit_ops_comm:
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    52
  "x AND y = y AND x"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    53
  "x OR y = y OR x"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    54
  "x XOR y = y XOR x"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    55
  for x :: bit
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    56
  by (cases y; auto)+
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    57
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    58
lemma bit_ops_same [simp]:
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    59
  "x AND x = x"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    60
  "x OR x = x"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    61
  "x XOR x = 0"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    62
  for x :: bit
67120
491fd7f0b5df misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    63
  by (cases x; auto)+
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    64
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    65
lemma bit_not_not [simp]: "NOT (NOT x) = x"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    66
  for x :: bit
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    67
  by (cases x) auto
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    68
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    69
lemma bit_or_def: "b OR c = NOT (NOT b AND NOT c)"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    70
  for b c :: bit
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    71
  by (induct b) simp_all
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    72
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    73
lemma bit_xor_def: "b XOR c = (b AND NOT c) OR (NOT b AND c)"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    74
  for b c :: bit
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    75
  by (induct b) simp_all
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    76
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    77
lemma bit_NOT_eq_1_iff [simp]: "NOT b = 1 \<longleftrightarrow> b = 0"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    78
  for b :: bit
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    79
  by (induct b) simp_all
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    80
65363
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    81
lemma bit_AND_eq_1_iff [simp]: "a AND b = 1 \<longleftrightarrow> a = 1 \<and> b = 1"
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    82
  for a b :: bit
5eb619751b14 misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    83
  by (induct a) simp_all
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    84
70183
3ea80c950023 incorporated various material from the AFP into the distribution
haftmann
parents: 67120
diff changeset
    85
lemma bit_nand_same [simp]: "x AND NOT x = 0"
3ea80c950023 incorporated various material from the AFP into the distribution
haftmann
parents: 67120
diff changeset
    86
  for x :: bit
3ea80c950023 incorporated various material from the AFP into the distribution
haftmann
parents: 67120
diff changeset
    87
  by (cases x) simp_all
3ea80c950023 incorporated various material from the AFP into the distribution
haftmann
parents: 67120
diff changeset
    88
54224
9fda41a04c32 separated bit operations on type bit from generic syntactic bit operations
haftmann
parents:
diff changeset
    89
end