src/HOL/ex/Bit_Lists.thy
author wenzelm
Tue, 05 Nov 2019 14:28:00 +0100
changeset 71047 87c132cf5860
parent 70926 b4564de51fa7
child 71042 400e9512f1d3
permissions -rw-r--r--
more options;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     1
(*  Author:  Florian Haftmann, TUM
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     2
*)
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     3
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
     4
section \<open>Proof(s) of concept for algebraically founded lists of bits\<close>
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     5
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
     6
theory Bit_Lists
70926
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
     7
  imports
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
     8
    Main
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
     9
    "HOL-Library.Boolean_Algebra"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
    10
begin
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
    11
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
    12
context ab_group_add
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    13
begin
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    14
70926
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
    15
lemma minus_diff_commute:
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
    16
  "- b - a = - a - b"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
    17
  by (simp only: diff_conv_add_uminus add.commute)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
    18
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
    19
end
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
    20
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
    21
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    22
subsection \<open>Bit representations\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    23
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    24
subsubsection \<open>Mere syntactic bit representation\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    25
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    26
class bit_representation =
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    27
  fixes bits_of :: "'a \<Rightarrow> bool list"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    28
    and of_bits :: "bool list \<Rightarrow> 'a"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    29
  assumes of_bits_of [simp]: "of_bits (bits_of a) = a"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    30
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    31
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    32
subsubsection \<open>Algebraic bit representation\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    33
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    34
context comm_semiring_1
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    35
begin
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    36
 
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    37
primrec radix_value :: "('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b list \<Rightarrow> 'a"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    38
  where "radix_value f b [] = 0"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    39
  | "radix_value f b (a # as) = f a + radix_value f b as * b"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    40
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    41
abbreviation (input) unsigned_of_bits :: "bool list \<Rightarrow> 'a"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    42
  where "unsigned_of_bits \<equiv> radix_value of_bool 2"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    43
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    44
lemma unsigned_of_bits_replicate_False [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    45
  "unsigned_of_bits (replicate n False) = 0"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    46
  by (induction n) simp_all
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    47
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    48
end
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    49
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    50
context unique_euclidean_semiring_with_nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    51
begin
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
    52
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    53
lemma unsigned_of_bits_append [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    54
  "unsigned_of_bits (bs @ cs) = unsigned_of_bits bs
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    55
    + push_bit (length bs) (unsigned_of_bits cs)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    56
  by (induction bs) (simp_all add: push_bit_double,
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    57
    simp_all add: algebra_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    58
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    59
lemma unsigned_of_bits_take [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    60
  "unsigned_of_bits (take n bs) = take_bit n (unsigned_of_bits bs)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    61
proof (induction bs arbitrary: n)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    62
  case Nil
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    63
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    64
    by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    65
next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    66
  case (Cons b bs)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    67
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    68
    by (cases n) (simp_all add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    69
qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    70
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    71
lemma unsigned_of_bits_drop [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    72
  "unsigned_of_bits (drop n bs) = drop_bit n (unsigned_of_bits bs)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    73
proof (induction bs arbitrary: n)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    74
  case Nil
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    75
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    76
    by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    77
next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    78
  case (Cons b bs)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    79
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    80
    by (cases n) simp_all
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    81
qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    82
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    83
end
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    84
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    85
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    86
subsubsection \<open>Instances\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    87
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    88
text \<open>Unclear whether a \<^typ>\<open>bool\<close> instantiation is needed or not\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    89
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    90
instantiation nat :: bit_representation
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    91
begin
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    92
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    93
fun bits_of_nat :: "nat \<Rightarrow> bool list"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    94
  where "bits_of (n::nat) =
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    95
    (if n = 0 then [] else odd n # bits_of (n div 2))"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    96
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    97
lemma bits_of_nat_simps [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    98
  "bits_of (0::nat) = []"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
    99
  "n > 0 \<Longrightarrow> bits_of n = odd n # bits_of (n div 2)" for n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   100
  by simp_all
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   101
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   102
declare bits_of_nat.simps [simp del]
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   103
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   104
definition of_bits_nat :: "bool list \<Rightarrow> nat"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   105
  where [simp]: "of_bits_nat = unsigned_of_bits"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   106
  \<comment> \<open>remove simp\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   107
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   108
instance proof
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   109
  show "of_bits (bits_of n) = n" for n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   110
    by (induction n rule: nat_bit_induct) simp_all
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   111
qed
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   112
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   113
end
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   114
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   115
lemma bits_of_Suc_0 [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   116
  "bits_of (Suc 0) = [True]"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   117
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   118
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   119
lemma bits_of_1_nat [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   120
  "bits_of (1 :: nat) = [True]"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   121
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   122
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   123
lemma bits_of_nat_numeral_simps [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   124
  "bits_of (numeral Num.One :: nat) = [True]" (is ?One)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   125
  "bits_of (numeral (Num.Bit0 n) :: nat) = False # bits_of (numeral n :: nat)" (is ?Bit0)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   126
  "bits_of (numeral (Num.Bit1 n) :: nat) = True # bits_of (numeral n :: nat)" (is ?Bit1)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   127
proof -
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   128
  show ?One
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   129
    by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   130
  define m :: nat where "m = numeral n"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   131
  then have "m > 0" and *: "numeral n = m" "numeral (Num.Bit0 n) = 2 * m" "numeral (Num.Bit1 n) = Suc (2 * m)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   132
    by simp_all
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   133
  from \<open>m > 0\<close> show ?Bit0 ?Bit1
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   134
    by (simp_all add: *)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   135
qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   136
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   137
lemma unsigned_of_bits_of_nat [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   138
  "unsigned_of_bits (bits_of n) = n" for n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   139
  using of_bits_of [of n] by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   140
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   141
instantiation int :: bit_representation
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   142
begin
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   143
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   144
fun bits_of_int :: "int \<Rightarrow> bool list"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   145
  where "bits_of_int k = odd k #
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   146
    (if k = 0 \<or> k = - 1 then [] else bits_of_int (k div 2))"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   147
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   148
lemma bits_of_int_simps [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   149
  "bits_of (0 :: int) = [False]"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   150
  "bits_of (- 1 :: int) = [True]"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   151
  "k \<noteq> 0 \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> bits_of k = odd k # bits_of (k div 2)" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   152
  by simp_all
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   153
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   154
lemma bits_of_not_Nil [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   155
  "bits_of k \<noteq> []" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   156
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   157
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   158
declare bits_of_int.simps [simp del]
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   159
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   160
definition of_bits_int :: "bool list \<Rightarrow> int"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   161
  where "of_bits_int bs = (if bs = [] \<or> \<not> last bs then unsigned_of_bits bs
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   162
    else unsigned_of_bits bs - 2 ^ length bs)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   163
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   164
lemma of_bits_int_simps [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   165
  "of_bits [] = (0 :: int)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   166
  "of_bits [False] = (0 :: int)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   167
  "of_bits [True] = (- 1 :: int)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   168
  "of_bits (bs @ [b]) = (unsigned_of_bits bs :: int) - (2 ^ length bs) * of_bool b"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   169
  "of_bits (False # bs) = 2 * (of_bits bs :: int)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   170
  "bs \<noteq> [] \<Longrightarrow> of_bits (True # bs) = 1 + 2 * (of_bits bs :: int)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   171
  by (simp_all add: of_bits_int_def push_bit_of_1)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   172
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   173
instance proof
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   174
  show "of_bits (bits_of k) = k" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   175
    by (induction k rule: int_bit_induct) simp_all
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   176
qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   177
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   178
lemma bits_of_1_int [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   179
  "bits_of (1 :: int) = [True, False]"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   180
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   181
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   182
lemma bits_of_int_numeral_simps [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   183
  "bits_of (numeral Num.One :: int) = [True, False]" (is ?One)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   184
  "bits_of (numeral (Num.Bit0 n) :: int) = False # bits_of (numeral n :: int)" (is ?Bit0)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   185
  "bits_of (numeral (Num.Bit1 n) :: int) = True # bits_of (numeral n :: int)" (is ?Bit1)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   186
  "bits_of (- numeral (Num.Bit0 n) :: int) = False # bits_of (- numeral n :: int)" (is ?nBit0)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   187
  "bits_of (- numeral (Num.Bit1 n) :: int) = True # bits_of (- numeral (Num.inc n) :: int)" (is ?nBit1)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   188
proof -
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   189
  show ?One
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   190
    by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   191
  define k :: int where "k = numeral n"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   192
  then have "k > 0" and *: "numeral n = k" "numeral (Num.Bit0 n) = 2 * k" "numeral (Num.Bit1 n) = 2 * k + 1"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   193
    "numeral (Num.inc n) = k + 1"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   194
    by (simp_all add: add_One)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   195
  have "- (2 * k) div 2 = - k" "(- (2 * k) - 1) div 2 = - k - 1"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   196
    by simp_all
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   197
  with \<open>k > 0\<close> show ?Bit0 ?Bit1 ?nBit0 ?nBit1
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   198
    by (simp_all add: *)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   199
qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   200
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   201
lemma of_bits_append [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   202
  "of_bits (bs @ cs) = of_bits bs + push_bit (length bs) (of_bits cs :: int)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   203
    if "bs \<noteq> []" "\<not> last bs"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   204
using that proof (induction bs rule: list_nonempty_induct)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   205
  case (single b)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   206
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   207
    by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   208
next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   209
  case (cons b bs)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   210
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   211
    by (cases b) (simp_all add: push_bit_double)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   212
qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   213
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   214
lemma of_bits_replicate_False [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   215
  "of_bits (replicate n False) = (0 :: int)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   216
  by (auto simp add: of_bits_int_def)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   217
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   218
lemma of_bits_drop [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   219
  "of_bits (drop n bs) = drop_bit n (of_bits bs :: int)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   220
    if "n < length bs"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   221
using that proof (induction bs arbitrary: n)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   222
  case Nil
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   223
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   224
    by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   225
next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   226
  case (Cons b bs)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   227
  show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   228
  proof (cases n)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   229
    case 0
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   230
    then show ?thesis
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   231
      by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   232
  next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   233
    case (Suc n)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   234
    with Cons.prems have "bs \<noteq> []"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   235
      by auto
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   236
    with Suc Cons.IH [of n] Cons.prems show ?thesis
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   237
      by (cases b) simp_all
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   238
  qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   239
qed
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   240
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   241
end
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   242
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   243
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   244
subsection \<open>Syntactic bit operations\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   245
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   246
class bit_operations = bit_representation +
70926
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   247
  fixes not :: "'a \<Rightarrow> 'a"  ("NOT")
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   248
    and "and" :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixr "AND" 64)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   249
    and or :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixr "OR"  59)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   250
    and xor :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixr "XOR" 59)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   251
    and shift_left :: "'a \<Rightarrow> nat \<Rightarrow> 'a"  (infixl "<<" 55)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   252
    and shift_right :: "'a \<Rightarrow> nat \<Rightarrow> 'a"  (infixl ">>" 55)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   253
  assumes not_eq: "not = of_bits \<circ> map Not \<circ> bits_of"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   254
    and and_eq: "length bs = length cs \<Longrightarrow>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   255
      of_bits bs AND of_bits cs = of_bits (map2 (\<and>) bs cs)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   256
    and semilattice_and: "semilattice (AND)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   257
    and or_eq: "length bs = length cs \<Longrightarrow>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   258
      of_bits bs OR of_bits cs = of_bits (map2 (\<or>) bs cs)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   259
    and semilattice_or: "semilattice (OR)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   260
    and xor_eq: "length bs = length cs \<Longrightarrow>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   261
      of_bits bs XOR of_bits cs = of_bits (map2 (\<noteq>) bs cs)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   262
    and abel_semigroup_xor: "abel_semigroup (XOR)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   263
    and shift_right_eq: "a << n = of_bits (replicate n False @ bits_of a)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   264
    and shift_left_eq: "n < length (bits_of a) \<Longrightarrow> a >> n = of_bits (drop n (bits_of a))"
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   265
begin
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   266
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   267
text \<open>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   268
  We want the bitwise operations to bind slightly weaker
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   269
  than \<open>+\<close> and \<open>-\<close>, but \<open>~~\<close> to
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   270
  bind slightly stronger than \<open>*\<close>.
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   271
\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   272
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   273
sublocale "and": semilattice "(AND)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   274
  by (fact semilattice_and)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   275
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   276
sublocale or: semilattice "(OR)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   277
  by (fact semilattice_or)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   278
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   279
sublocale xor: abel_semigroup "(XOR)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   280
  by (fact abel_semigroup_xor)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   281
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   282
end
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   283
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   284
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   285
subsubsection \<open>Instance \<^typ>\<open>nat\<close>\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   286
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   287
locale zip_nat = single: abel_semigroup f
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   288
    for f :: "bool \<Rightarrow> bool \<Rightarrow> bool"  (infixl "\<^bold>*" 70) +
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   289
  assumes end_of_bits: "\<not> False \<^bold>* False"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   290
begin
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   291
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   292
lemma False_P_imp:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   293
  "False \<^bold>* True \<and> P" if "False \<^bold>* P"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   294
  using that end_of_bits by (cases P) simp_all
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   295
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   296
function F :: "nat \<Rightarrow> nat \<Rightarrow> nat"  (infixl "\<^bold>\<times>" 70)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   297
  where "m \<^bold>\<times> n = (if m = 0 \<and> n = 0 then 0
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   298
    else of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\<times> (n div 2) * 2)"
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   299
  by auto
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   300
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   301
termination
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   302
  by (relation "measure (case_prod (+))") auto
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   303
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   304
lemma zero_left_eq:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   305
  "0 \<^bold>\<times> n = of_bool (False \<^bold>* True) * n"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   306
  by (induction n rule: nat_bit_induct) (simp_all add: end_of_bits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   307
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   308
lemma zero_right_eq:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   309
  "m \<^bold>\<times> 0 = of_bool (True \<^bold>* False) * m"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   310
  by (induction m rule: nat_bit_induct) (simp_all add: end_of_bits)
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   311
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   312
lemma simps [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   313
  "0 \<^bold>\<times> 0 = 0"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   314
  "0 \<^bold>\<times> n = of_bool (False \<^bold>* True) * n"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   315
  "m \<^bold>\<times> 0 = of_bool (True \<^bold>* False) * m"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   316
  "m > 0 \<Longrightarrow> n > 0 \<Longrightarrow> m \<^bold>\<times> n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\<times> (n div 2) * 2"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   317
  by (simp_all only: zero_left_eq zero_right_eq) simp
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   318
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   319
lemma rec:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   320
  "m \<^bold>\<times> n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\<times> (n div 2) * 2"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   321
  by (cases "m = 0 \<and> n = 0") (auto simp add: end_of_bits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   322
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   323
declare F.simps [simp del]
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   324
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   325
sublocale abel_semigroup F
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   326
proof
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   327
  show "m \<^bold>\<times> n \<^bold>\<times> q = m \<^bold>\<times> (n \<^bold>\<times> q)" for m n q :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   328
  proof (induction m arbitrary: n q rule: nat_bit_induct)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   329
    case zero
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   330
    show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   331
      by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   332
  next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   333
    case (even m)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   334
    with rec [of "2 * m"] rec [of _ q] show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   335
      by (cases "even n") (auto dest: False_P_imp)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   336
  next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   337
    case (odd m)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   338
    with rec [of "Suc (2 * m)"] rec [of _ q] show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   339
      by (cases "even n"; cases "even q")
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   340
        (auto dest: False_P_imp simp add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   341
  qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   342
  show "m \<^bold>\<times> n = n \<^bold>\<times> m" for m n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   343
  proof (induction m arbitrary: n rule: nat_bit_induct)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   344
    case zero
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   345
    show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   346
      by (simp add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   347
  next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   348
    case (even m)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   349
    with rec [of "2 * m" n] rec [of n "2 * m"] show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   350
      by (simp add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   351
  next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   352
    case (odd m)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   353
    with rec [of "Suc (2 * m)" n] rec [of n "Suc (2 * m)"] show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   354
      by (simp add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   355
  qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   356
qed
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   357
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   358
lemma self [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   359
  "n \<^bold>\<times> n = of_bool (True \<^bold>* True) * n"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   360
  by (induction n rule: nat_bit_induct) (simp_all add: end_of_bits)
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   361
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   362
lemma even_iff [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   363
  "even (m \<^bold>\<times> n) \<longleftrightarrow> \<not> (odd m \<^bold>* odd n)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   364
proof (induction m arbitrary: n rule: nat_bit_induct)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   365
  case zero
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   366
  show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   367
    by (cases "even n") (simp_all add: end_of_bits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   368
next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   369
  case (even m)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   370
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   371
    by (simp add: rec [of "2 * m"])
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   372
next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   373
  case (odd m)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   374
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   375
    by (simp add: rec [of "Suc (2 * m)"])
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   376
qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   377
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   378
lemma of_bits:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   379
  "of_bits bs \<^bold>\<times> of_bits cs = (of_bits (map2 (\<^bold>*) bs cs) :: nat)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   380
    if "length bs = length cs" for bs cs
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   381
using that proof (induction bs cs rule: list_induct2)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   382
  case Nil
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   383
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   384
    by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   385
next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   386
  case (Cons b bs c cs)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   387
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   388
    by (cases "of_bits bs = (0::nat) \<or> of_bits cs = (0::nat)")
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   389
      (auto simp add: ac_simps end_of_bits rec [of "Suc 0"])
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   390
qed
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   391
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   392
end
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   393
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   394
instantiation nat :: bit_operations
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   395
begin
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   396
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   397
definition not_nat :: "nat \<Rightarrow> nat"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   398
  where "not_nat = of_bits \<circ> map Not \<circ> bits_of"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   399
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   400
global_interpretation and_nat: zip_nat "(\<and>)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   401
  defines and_nat = and_nat.F
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   402
  by standard auto
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   403
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   404
global_interpretation and_nat: semilattice "(AND) :: nat \<Rightarrow> nat \<Rightarrow> nat"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   405
proof (rule semilattice.intro, fact and_nat.abel_semigroup_axioms, standard)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   406
  show "n AND n = n" for n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   407
    by (simp add: and_nat.self)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   408
qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   409
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   410
declare and_nat.simps [simp] \<comment> \<open>inconsistent declaration handling by
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   411
  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   412
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   413
lemma zero_nat_and_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   414
  "0 AND n = 0" for n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   415
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   416
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   417
lemma and_zero_nat_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   418
  "n AND 0 = 0" for n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   419
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   420
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   421
global_interpretation or_nat: zip_nat "(\<or>)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   422
  defines or_nat = or_nat.F
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   423
  by standard auto
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   424
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   425
global_interpretation or_nat: semilattice "(OR) :: nat \<Rightarrow> nat \<Rightarrow> nat"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   426
proof (rule semilattice.intro, fact or_nat.abel_semigroup_axioms, standard)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   427
  show "n OR n = n" for n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   428
    by (simp add: or_nat.self)
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   429
qed
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   430
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   431
declare or_nat.simps [simp] \<comment> \<open>inconsistent declaration handling by
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   432
  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   433
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   434
lemma zero_nat_or_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   435
  "0 OR n = n" for n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   436
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   437
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   438
lemma or_zero_nat_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   439
  "n OR 0 = n" for n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   440
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   441
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   442
global_interpretation xor_nat: zip_nat "(\<noteq>)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   443
  defines xor_nat = xor_nat.F
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   444
  by standard auto
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   445
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   446
declare xor_nat.simps [simp] \<comment> \<open>inconsistent declaration handling by
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   447
  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   448
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   449
lemma zero_nat_xor_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   450
  "0 XOR n = n" for n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   451
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   452
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   453
lemma xor_zero_nat_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   454
  "n XOR 0 = n" for n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   455
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   456
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   457
definition shift_left_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   458
  where [simp]: "m << n = push_bit n m" for m :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   459
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   460
definition shift_right_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   461
  where [simp]: "m >> n = drop_bit n m" for m :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   462
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   463
instance proof
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   464
  show "semilattice ((AND) :: nat \<Rightarrow> _)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   465
    by (fact and_nat.semilattice_axioms)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   466
  show "semilattice ((OR):: nat \<Rightarrow> _)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   467
    by (fact or_nat.semilattice_axioms)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   468
  show "abel_semigroup ((XOR):: nat \<Rightarrow> _)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   469
    by (fact xor_nat.abel_semigroup_axioms)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   470
  show "(not :: nat \<Rightarrow> _) = of_bits \<circ> map Not \<circ> bits_of"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   471
    by (fact not_nat_def)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   472
  show "of_bits bs AND of_bits cs = (of_bits (map2 (\<and>) bs cs) :: nat)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   473
    if "length bs = length cs" for bs cs
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   474
    using that by (fact and_nat.of_bits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   475
  show "of_bits bs OR of_bits cs = (of_bits (map2 (\<or>) bs cs) :: nat)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   476
    if "length bs = length cs" for bs cs
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   477
    using that by (fact or_nat.of_bits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   478
  show "of_bits bs XOR of_bits cs = (of_bits (map2 (\<noteq>) bs cs) :: nat)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   479
    if "length bs = length cs" for bs cs
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   480
    using that by (fact xor_nat.of_bits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   481
  show "m << n = of_bits (replicate n False @ bits_of m)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   482
    for m n :: nat
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   483
    by simp
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   484
  show "m >> n = of_bits (drop n (bits_of m))"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   485
    for m n :: nat
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   486
    by simp
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   487
qed
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   488
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   489
end
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   490
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   491
global_interpretation or_nat: semilattice_neutr "(OR)" "0 :: nat"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   492
  by standard simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   493
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   494
global_interpretation xor_nat: comm_monoid "(XOR)" "0 :: nat"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   495
  by standard simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   496
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   497
lemma not_nat_simps [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   498
  "NOT 0 = (0::nat)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   499
  "n > 0 \<Longrightarrow> NOT n = of_bool (even n) + 2 * NOT (n div 2)" for n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   500
  by (simp_all add: not_eq)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   501
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   502
lemma not_1_nat [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   503
  "NOT 1 = (0::nat)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   504
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   505
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   506
lemma not_Suc_0 [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   507
  "NOT (Suc 0) = (0::nat)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   508
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   509
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   510
lemma Suc_0_and_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   511
  "Suc 0 AND n = n mod 2"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   512
  by (cases n) auto
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   513
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   514
lemma and_Suc_0_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   515
  "n AND Suc 0 = n mod 2"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   516
  using Suc_0_and_eq [of n] by (simp add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   517
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   518
lemma Suc_0_or_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   519
  "Suc 0 OR n = n + of_bool (even n)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   520
  by (cases n) (simp_all add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   521
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   522
lemma or_Suc_0_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   523
  "n OR Suc 0 = n + of_bool (even n)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   524
  using Suc_0_or_eq [of n] by (simp add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   525
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   526
lemma Suc_0_xor_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   527
  "Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   528
  by (cases n) (simp_all add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   529
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   530
lemma xor_Suc_0_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   531
  "n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   532
  using Suc_0_xor_eq [of n] by (simp add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   533
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   534
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   535
subsubsection \<open>Instance \<^typ>\<open>int\<close>\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   536
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   537
abbreviation (input) complement :: "int \<Rightarrow> int"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   538
  where "complement k \<equiv> - k - 1"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   539
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   540
lemma complement_half:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   541
  "complement (k * 2) div 2 = complement k"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   542
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   543
70926
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   544
lemma complement_div_2:
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   545
  "complement (k div 2) = complement k div 2"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   546
  by linarith
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   547
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   548
locale zip_int = single: abel_semigroup f
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   549
  for f :: "bool \<Rightarrow> bool \<Rightarrow> bool"  (infixl "\<^bold>*" 70)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   550
begin
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   551
 
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   552
lemma False_False_imp_True_True:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   553
  "True \<^bold>* True" if "False \<^bold>* False"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   554
proof (rule ccontr)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   555
  assume "\<not> True \<^bold>* True"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   556
  with that show False
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   557
    using single.assoc [of False True True]
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   558
    by (cases "False \<^bold>* True") simp_all
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   559
qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   560
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   561
function F :: "int \<Rightarrow> int \<Rightarrow> int"  (infixl "\<^bold>\<times>" 70)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   562
  where "k \<^bold>\<times> l = (if k \<in> {0, - 1} \<and> l \<in> {0, - 1}
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   563
    then - of_bool (odd k \<^bold>* odd l)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   564
    else of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\<times> (l div 2) * 2)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   565
  by auto
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   566
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   567
termination
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   568
  by (relation "measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))") auto
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   569
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   570
lemma zero_left_eq:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   571
  "0 \<^bold>\<times> l = (case (False \<^bold>* False, False \<^bold>* True)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   572
    of (False, False) \<Rightarrow> 0
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   573
     | (False, True) \<Rightarrow> l
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   574
     | (True, False) \<Rightarrow> complement l
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   575
     | (True, True) \<Rightarrow> - 1)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   576
  by (induction l rule: int_bit_induct)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   577
   (simp_all split: bool.split) 
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   578
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   579
lemma minus_left_eq:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   580
  "- 1 \<^bold>\<times> l = (case (True \<^bold>* False, True \<^bold>* True)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   581
    of (False, False) \<Rightarrow> 0
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   582
     | (False, True) \<Rightarrow> l
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   583
     | (True, False) \<Rightarrow> complement l
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   584
     | (True, True) \<Rightarrow> - 1)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   585
  by (induction l rule: int_bit_induct)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   586
   (simp_all split: bool.split) 
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   587
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   588
lemma zero_right_eq:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   589
  "k \<^bold>\<times> 0 = (case (False \<^bold>* False, False \<^bold>* True)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   590
    of (False, False) \<Rightarrow> 0
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   591
     | (False, True) \<Rightarrow> k
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   592
     | (True, False) \<Rightarrow> complement k
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   593
     | (True, True) \<Rightarrow> - 1)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   594
  by (induction k rule: int_bit_induct)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   595
    (simp_all add: ac_simps split: bool.split)
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   596
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   597
lemma minus_right_eq:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   598
  "k \<^bold>\<times> - 1 = (case (True \<^bold>* False, True \<^bold>* True)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   599
    of (False, False) \<Rightarrow> 0
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   600
     | (False, True) \<Rightarrow> k
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   601
     | (True, False) \<Rightarrow> complement k
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   602
     | (True, True) \<Rightarrow> - 1)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   603
  by (induction k rule: int_bit_induct)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   604
    (simp_all add: ac_simps split: bool.split)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   605
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   606
lemma simps [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   607
  "0 \<^bold>\<times> 0 = - of_bool (False \<^bold>* False)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   608
  "- 1 \<^bold>\<times> 0 = - of_bool (True \<^bold>* False)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   609
  "0 \<^bold>\<times> - 1 = - of_bool (False \<^bold>* True)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   610
  "- 1 \<^bold>\<times> - 1 = - of_bool (True \<^bold>* True)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   611
  "0 \<^bold>\<times> l = (case (False \<^bold>* False, False \<^bold>* True)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   612
    of (False, False) \<Rightarrow> 0
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   613
     | (False, True) \<Rightarrow> l
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   614
     | (True, False) \<Rightarrow> complement l
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   615
     | (True, True) \<Rightarrow> - 1)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   616
  "- 1 \<^bold>\<times> l = (case (True \<^bold>* False, True \<^bold>* True)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   617
    of (False, False) \<Rightarrow> 0
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   618
     | (False, True) \<Rightarrow> l
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   619
     | (True, False) \<Rightarrow> complement l
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   620
     | (True, True) \<Rightarrow> - 1)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   621
  "k \<^bold>\<times> 0 = (case (False \<^bold>* False, False \<^bold>* True)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   622
    of (False, False) \<Rightarrow> 0
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   623
     | (False, True) \<Rightarrow> k
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   624
     | (True, False) \<Rightarrow> complement k
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   625
     | (True, True) \<Rightarrow> - 1)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   626
  "k \<^bold>\<times> - 1 = (case (True \<^bold>* False, True \<^bold>* True)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   627
    of (False, False) \<Rightarrow> 0
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   628
     | (False, True) \<Rightarrow> k
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   629
     | (True, False) \<Rightarrow> complement k
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   630
     | (True, True) \<Rightarrow> - 1)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   631
  "k \<noteq> 0 \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> l \<noteq> - 1
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   632
    \<Longrightarrow> k \<^bold>\<times> l = of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\<times> (l div 2) * 2"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   633
  by simp_all[4] (simp_all only: zero_left_eq minus_left_eq zero_right_eq minus_right_eq, simp)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   634
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   635
declare F.simps [simp del]
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   636
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   637
lemma rec:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   638
  "k \<^bold>\<times> l = of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\<times> (l div 2) * 2"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   639
  by (cases "k \<in> {0, - 1} \<and> l \<in> {0, - 1}") (auto simp add: ac_simps F.simps [of k l] split: bool.split)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   640
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   641
sublocale abel_semigroup F
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   642
proof
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   643
  show "k \<^bold>\<times> l \<^bold>\<times> r = k \<^bold>\<times> (l \<^bold>\<times> r)" for k l r :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   644
  proof (induction k arbitrary: l r rule: int_bit_induct)
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   645
    case zero
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   646
    have "complement l \<^bold>\<times> r = complement (l \<^bold>\<times> r)" if "False \<^bold>* False" "\<not> False \<^bold>* True"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   647
    proof (induction l arbitrary: r rule: int_bit_induct)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   648
      case zero
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   649
      from that show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   650
        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   651
    next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   652
      case minus
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   653
      from that show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   654
        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   655
    next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   656
      case (even l)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   657
      with that rec [of _ r] show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   658
        by (cases "even r")
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   659
          (auto simp add: complement_half ac_simps False_False_imp_True_True split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   660
    next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   661
      case (odd l)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   662
      moreover have "- l - 1 = - 1 - l"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   663
        by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   664
      ultimately show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   665
        using that rec [of _ r]
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   666
        by (cases "even r")
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   667
          (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   668
    qed
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   669
    then show ?case
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   670
      by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   671
  next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   672
    case minus
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   673
    have "complement l \<^bold>\<times> r = complement (l \<^bold>\<times> r)" if "\<not> True \<^bold>* True" "False \<^bold>* True"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   674
    proof (induction l arbitrary: r rule: int_bit_induct)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   675
      case zero
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   676
      from that show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   677
        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   678
    next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   679
      case minus
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   680
      from that show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   681
        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   682
    next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   683
      case (even l)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   684
      with that rec [of _ r] show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   685
        by (cases "even r")
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   686
          (auto simp add: complement_half ac_simps False_False_imp_True_True split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   687
    next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   688
      case (odd l)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   689
      moreover have "- l - 1 = - 1 - l"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   690
        by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   691
      ultimately show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   692
        using that rec [of _ r]
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   693
        by (cases "even r")
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   694
          (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   695
    qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   696
    then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   697
      by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   698
  next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   699
    case (even k)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   700
    with rec [of "k * 2"] rec [of _ r] show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   701
      by (cases "even r"; cases "even l") (auto simp add: ac_simps False_False_imp_True_True)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   702
  next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   703
    case (odd k)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   704
    with rec [of "1 + k * 2"] rec [of _ r] show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   705
      by (cases "even r"; cases "even l") (auto simp add: ac_simps False_False_imp_True_True)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   706
  qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   707
  show "k \<^bold>\<times> l = l \<^bold>\<times> k" for k l :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   708
  proof (induction k arbitrary: l rule: int_bit_induct)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   709
    case zero
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   710
    show ?case
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   711
      by simp
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   712
  next
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   713
    case minus
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   714
    show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   715
      by simp
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   716
  next
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   717
    case (even k)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   718
    with rec [of "k * 2" l] rec [of l "k * 2"] show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   719
      by (simp add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   720
  next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   721
    case (odd k)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   722
    with rec [of "k * 2 + 1" l] rec [of l "k * 2 + 1"] show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   723
      by (simp add: ac_simps)
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   724
  qed
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   725
qed
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   726
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   727
lemma self [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   728
  "k \<^bold>\<times> k = (case (False \<^bold>* False, True \<^bold>* True)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   729
    of (False, False) \<Rightarrow> 0
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   730
     | (False, True) \<Rightarrow> k
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   731
     | (True, True) \<Rightarrow> - 1)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   732
  by (induction k rule: int_bit_induct) (auto simp add: False_False_imp_True_True split: bool.split)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   733
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   734
lemma even_iff [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   735
  "even (k \<^bold>\<times> l) \<longleftrightarrow> \<not> (odd k \<^bold>* odd l)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   736
proof (induction k arbitrary: l rule: int_bit_induct)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   737
  case zero
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   738
  show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   739
    by (cases "even l") (simp_all split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   740
next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   741
  case minus
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   742
  show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   743
    by (cases "even l") (simp_all split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   744
next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   745
  case (even k)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   746
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   747
    by (simp add: rec [of "k * 2"])
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   748
next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   749
  case (odd k)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   750
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   751
    by (simp add: rec [of "1 + k * 2"])
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   752
qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   753
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   754
lemma of_bits:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   755
  "of_bits bs \<^bold>\<times> of_bits cs = (of_bits (map2 (\<^bold>*) bs cs) :: int)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   756
    if "length bs = length cs" and "\<not> False \<^bold>* False" for bs cs
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   757
using \<open>length bs = length cs\<close> proof (induction bs cs rule: list_induct2)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   758
  case Nil
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   759
  then show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   760
    using \<open>\<not> False \<^bold>* False\<close> by (auto simp add: False_False_imp_True_True split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   761
next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   762
  case (Cons b bs c cs)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   763
  show ?case
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   764
  proof (cases "bs = []")
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   765
    case True
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   766
    with Cons show ?thesis
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   767
      using \<open>\<not> False \<^bold>* False\<close> by (cases b; cases c)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   768
        (auto simp add: ac_simps split: bool.splits)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   769
  next
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   770
    case False
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   771
    with Cons.hyps have "cs \<noteq> []"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   772
      by auto
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   773
    with \<open>bs \<noteq> []\<close> have "map2 (\<^bold>*) bs cs \<noteq> []"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   774
      by (simp add: zip_eq_Nil_iff)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   775
    from \<open>bs \<noteq> []\<close> \<open>cs \<noteq> []\<close> \<open>map2 (\<^bold>*) bs cs \<noteq> []\<close> Cons show ?thesis
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   776
      by (cases b; cases c)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   777
        (auto simp add: \<open>\<not> False \<^bold>* False\<close> ac_simps
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   778
          rec [of "of_bits bs * 2"] rec [of "of_bits cs * 2"]
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   779
          rec [of "1 + of_bits bs * 2"])
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   780
  qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   781
qed
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   782
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   783
end
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   784
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   785
instantiation int :: bit_operations
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   786
begin
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   787
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   788
definition not_int :: "int \<Rightarrow> int"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   789
  where "not_int = complement"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   790
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   791
global_interpretation and_int: zip_int "(\<and>)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   792
  defines and_int = and_int.F
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   793
  by standard
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   794
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   795
declare and_int.simps [simp] \<comment> \<open>inconsistent declaration handling by
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   796
  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   797
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   798
global_interpretation and_int: semilattice "(AND) :: int \<Rightarrow> int \<Rightarrow> int"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   799
proof (rule semilattice.intro, fact and_int.abel_semigroup_axioms, standard)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   800
  show "k AND k = k" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   801
    by (simp add: and_int.self)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   802
qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   803
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   804
lemma zero_int_and_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   805
  "0 AND k = 0" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   806
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   807
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   808
lemma and_zero_int_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   809
  "k AND 0 = 0" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   810
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   811
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   812
lemma minus_int_and_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   813
  "- 1 AND k = k" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   814
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   815
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   816
lemma and_minus_int_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   817
  "k AND - 1 = k" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   818
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   819
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   820
global_interpretation or_int: zip_int "(\<or>)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   821
  defines or_int = or_int.F
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   822
  by standard
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   823
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   824
declare or_int.simps [simp] \<comment> \<open>inconsistent declaration handling by
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   825
  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   826
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   827
global_interpretation or_int: semilattice "(OR) :: int \<Rightarrow> int \<Rightarrow> int"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   828
proof (rule semilattice.intro, fact or_int.abel_semigroup_axioms, standard)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   829
  show "k OR k = k" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   830
    by (simp add: or_int.self)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   831
qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   832
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   833
lemma zero_int_or_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   834
  "0 OR k = k" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   835
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   836
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   837
lemma and_zero_or_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   838
  "k OR 0 = k" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   839
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   840
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   841
lemma minus_int_or_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   842
  "- 1 OR k = - 1" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   843
  by simp
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   844
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   845
lemma or_minus_int_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   846
  "k OR - 1 = - 1" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   847
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   848
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   849
global_interpretation xor_int: zip_int "(\<noteq>)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   850
  defines xor_int = xor_int.F
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   851
  by standard
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   852
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   853
declare xor_int.simps [simp] \<comment> \<open>inconsistent declaration handling by
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   854
  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   855
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   856
lemma zero_int_xor_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   857
  "0 XOR k = k" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   858
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   859
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   860
lemma and_zero_xor_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   861
  "k XOR 0 = k" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   862
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   863
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   864
lemma minus_int_xor_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   865
  "- 1 XOR k = complement k" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   866
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   867
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   868
lemma xor_minus_int_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   869
  "k XOR - 1 = complement k" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   870
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   871
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   872
definition shift_left_int :: "int \<Rightarrow> nat \<Rightarrow> int"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   873
  where [simp]: "k << n = push_bit n k" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   874
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   875
definition shift_right_int :: "int \<Rightarrow> nat \<Rightarrow> int"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   876
  where [simp]: "k >> n = drop_bit n k" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   877
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   878
instance proof
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   879
  show "semilattice ((AND) :: int \<Rightarrow> _)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   880
    by (fact and_int.semilattice_axioms)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   881
  show "semilattice ((OR) :: int \<Rightarrow> _)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   882
    by (fact or_int.semilattice_axioms)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   883
  show "abel_semigroup ((XOR) :: int \<Rightarrow> _)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   884
    by (fact xor_int.abel_semigroup_axioms)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   885
  show "(not :: int \<Rightarrow> _) = of_bits \<circ> map Not \<circ> bits_of"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   886
  proof (rule sym, rule ext)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   887
    fix k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   888
    show "(of_bits \<circ> map Not \<circ> bits_of) k = NOT k"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   889
      by (induction k rule: int_bit_induct) (simp_all add: not_int_def)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   890
  qed
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   891
  show "of_bits bs AND of_bits cs = (of_bits (map2 (\<and>) bs cs) :: int)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   892
    if "length bs = length cs" for bs cs
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   893
    using that by (rule and_int.of_bits) simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   894
  show "of_bits bs OR of_bits cs = (of_bits (map2 (\<or>) bs cs) :: int)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   895
    if "length bs = length cs" for bs cs
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   896
    using that by (rule or_int.of_bits) simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   897
  show "of_bits bs XOR of_bits cs = (of_bits (map2 (\<noteq>) bs cs) :: int)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   898
    if "length bs = length cs" for bs cs
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   899
    using that by (rule xor_int.of_bits) simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   900
  show "k << n = of_bits (replicate n False @ bits_of k)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   901
    for k :: int and n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   902
    by (cases "n = 0") simp_all
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   903
  show "k >> n = of_bits (drop n (bits_of k))"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   904
    if "n < length (bits_of k)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   905
    for k :: int and n :: nat
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   906
    using that by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   907
qed
67909
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   908
f55b07f4d1ee proof of concept for algebraically founded bit lists
haftmann
parents:
diff changeset
   909
end
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   910
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   911
global_interpretation and_int: semilattice_neutr "(AND)" "- 1 :: int"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   912
  by standard simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   913
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   914
global_interpretation or_int: semilattice_neutr "(OR)" "0 :: int"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   915
  by standard simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   916
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   917
global_interpretation xor_int: comm_monoid "(XOR)" "0 :: int"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   918
  by standard simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   919
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   920
lemma not_int_simps [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   921
  "NOT 0 = (- 1 :: int)"
70926
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   922
  "NOT (- 1) = (0 :: int)"
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   923
  "k \<noteq> 0 \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> NOT k = of_bool (even k) + 2 * NOT (k div 2)" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   924
  by (auto simp add: not_int_def elim: oddE)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   925
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   926
lemma not_one_int [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   927
  "NOT 1 = (- 2 :: int)"
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   928
  by simp
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   929
70926
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   930
lemma even_not_iff [simp]:
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   931
  "even (NOT k) \<longleftrightarrow> odd k"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   932
  for k :: int
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   933
  by (simp add: not_int_def)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   934
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   935
lemma not_div_2:
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   936
  "NOT k div 2 = NOT (k div 2)"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   937
  for k :: int
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   938
  by (simp add: complement_div_2 not_int_def)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   939
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   940
lemma one_and_int_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   941
  "1 AND k = k mod 2" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   942
  using and_int.rec [of 1] by (simp add: mod2_eq_if)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   943
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   944
lemma and_one_int_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   945
  "k AND 1 = k mod 2" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   946
  using one_and_int_eq [of 1] by (simp add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   947
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   948
lemma one_or_int_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   949
  "1 OR k = k + of_bool (even k)" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   950
  using or_int.rec [of 1] by (auto elim: oddE)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   951
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   952
lemma or_one_int_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   953
  "k OR 1 = k + of_bool (even k)" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   954
  using one_or_int_eq [of k] by (simp add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   955
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   956
lemma one_xor_int_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   957
  "1 XOR k = k + of_bool (even k) - of_bool (odd k)" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   958
  using xor_int.rec [of 1] by (auto elim: oddE)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   959
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   960
lemma xor_one_int_eq [simp]:
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   961
  "k XOR 1 = k + of_bool (even k) - of_bool (odd k)" for k :: int
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   962
  using one_xor_int_eq [of k] by (simp add: ac_simps)
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
   963
70926
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   964
global_interpretation bit_int: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: int"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   965
  rewrites "bit_int.xor = ((XOR) :: int \<Rightarrow> _)"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   966
proof -
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   967
  interpret bit_int: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: int"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   968
  proof
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   969
    show "k AND (l OR r) = k AND l OR k AND r"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   970
      for k l r :: int
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   971
    proof (induction k arbitrary: l r rule: int_bit_induct)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   972
      case zero
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   973
      show ?case
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   974
        by simp
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   975
    next
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   976
      case minus
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   977
      show ?case
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   978
        by simp
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   979
    next
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   980
      case (even k)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   981
      then show ?case by (simp add: and_int.rec [of "k * 2"]
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   982
        or_int.rec [of "(k AND l div 2) * 2"] or_int.rec [of l])
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   983
    next
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   984
      case (odd k)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   985
      then show ?case by (simp add: and_int.rec [of "1 + k * 2"]
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   986
        or_int.rec [of "(k AND l div 2) * 2"] or_int.rec [of "1 + (k AND l div 2) * 2"] or_int.rec [of l])
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   987
    qed
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   988
    show "k OR l AND r = (k OR l) AND (k OR r)"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   989
      for k l r :: int
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   990
    proof (induction k arbitrary: l r rule: int_bit_induct)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   991
      case zero
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   992
      then show ?case
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   993
        by simp
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   994
    next
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   995
      case minus
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   996
      then show ?case
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   997
        by simp
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   998
    next
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
   999
      case (even k)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1000
      then show ?case by (simp add: or_int.rec [of "k * 2"]
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1001
        and_int.rec [of "(k OR l div 2) * 2"] and_int.rec [of "1 + (k OR l div 2) * 2"] and_int.rec [of l])
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1002
    next
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1003
      case (odd k)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1004
      then show ?case by (simp add: or_int.rec [of "1 + k * 2"]
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1005
        and_int.rec [of "1 + (k OR l div 2) * 2"] and_int.rec [of l])
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1006
    qed
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1007
    show "k AND NOT k = 0" for k :: int
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1008
      by (induction k rule: int_bit_induct)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1009
        (simp_all add: not_int_def complement_half minus_diff_commute [of 1])
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1010
    show "k OR NOT k = - 1" for k :: int
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1011
      by (induction k rule: int_bit_induct)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1012
        (simp_all add: not_int_def complement_half minus_diff_commute [of 1])
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1013
  qed simp_all
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1014
  show "boolean_algebra (AND) (OR) NOT 0 (- 1 :: int)"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1015
    by (fact bit_int.boolean_algebra_axioms)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1016
  show "bit_int.xor = ((XOR) :: int \<Rightarrow> _)"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1017
  proof (rule ext)+
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1018
    fix k l :: int
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1019
    have "k XOR l = k AND NOT l OR NOT k AND l"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1020
    proof (induction k arbitrary: l rule: int_bit_induct)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1021
      case zero
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1022
      show ?case
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1023
        by simp
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1024
    next
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1025
      case minus
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1026
      show ?case
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1027
        by (simp add: not_int_def)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1028
    next
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1029
      case (even k)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1030
      then show ?case
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1031
        by (simp add: xor_int.rec [of "k * 2"] and_int.rec [of "k * 2"]
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1032
          or_int.rec [of _ "1 + 2 * NOT k AND l"] not_div_2)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1033
          (simp add: and_int.rec [of _ l])
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1034
    next
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1035
      case (odd k)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1036
      then show ?case
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1037
        by (simp add: xor_int.rec [of "1 + k * 2"] and_int.rec [of "1 + k * 2"]
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1038
          or_int.rec [of _ "2 * NOT k AND l"] not_div_2)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1039
          (simp add: and_int.rec [of _ l])
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1040
    qed
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1041
    then show "bit_int.xor k l = k XOR l"
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1042
      by (simp add: bit_int.xor_def)
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1043
  qed
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1044
qed
b4564de51fa7 bit operations form an boolean algebra
haftmann
parents: 70912
diff changeset
  1045
70912
8c2bef3df488 refined proof of concept for bit operations
haftmann
parents: 70353
diff changeset
  1046
end